aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2021-02-01 15:07:46 +0100
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2021-02-01 15:07:46 +0100
commitd2dc422b91ed628f0f8d6286a23f6f4fb9869248 (patch)
tree41d37d0c2201ddc30dc571946408aa84d08410e6
parent6eabac7e12de09dc4b2a32fa2458e3b91fb34471 (diff)
downloadcompcert-kvx-d2dc422b91ed628f0f8d6286a23f6f4fb9869248.tar.gz
compcert-kvx-d2dc422b91ed628f0f8d6286a23f6f4fb9869248.zip
Obits_of_single etc
-rw-r--r--riscV/NeedOp.v2
-rw-r--r--riscV/Op.v19
-rw-r--r--riscV/ValueAOp.v28
3 files changed, 46 insertions, 3 deletions
diff --git a/riscV/NeedOp.v b/riscV/NeedOp.v
index 117bbcb4..9a406b6f 100644
--- a/riscV/NeedOp.v
+++ b/riscV/NeedOp.v
@@ -87,6 +87,8 @@ Definition needs_of_operation (op: operation) (nv: nval): list nval :=
| Ointofsingle | Ointuofsingle | Osingleofint | Osingleofintu => op1 (default nv)
| Olongofsingle | Olonguofsingle | Osingleoflong | Osingleoflongu => op1 (default nv)
| Ocmp c => needs_of_condition c
+ | Obits_of_single => op1 (default nv)
+ | Obits_of_float => op1 (default nv)
end.
Definition operation_is_redundant (op: operation) (nv: nval): bool :=
diff --git a/riscV/Op.v b/riscV/Op.v
index 2271ecd2..160e4412 100644
--- a/riscV/Op.v
+++ b/riscV/Op.v
@@ -32,6 +32,7 @@
Require Import BoolEqual Coqlib.
Require Import AST Integers Floats.
Require Import Values Memory Globalenvs Events.
+Require ExtValues.
Set Implicit Arguments.
@@ -152,7 +153,9 @@ Inductive operation : Type :=
| Osingleoflong (**r [rd = float32_of_signed_long(r1)] *)
| Osingleoflongu (**r [rd = float32_of_unsigned_int(r1)] *)
(*c Boolean tests: *)
- | Ocmp (cond: condition). (**r [rd = 1] if condition holds, [rd = 0] otherwise. *)
+ | Ocmp (cond: condition) (**r [rd = 1] if condition holds, [rd = 0] otherwise. *)
+ | Obits_of_single
+ | Obits_of_float.
(** Addressing modes. [r1], [r2], etc, are the arguments to the
addressing. *)
@@ -317,6 +320,8 @@ Definition eval_operation
| Olonguofsingle, v1::nil => Some (Val.maketotal (Val.longuofsingle v1))
| Osingleoflong, v1::nil => Some (Val.maketotal (Val.singleoflong v1))
| Osingleoflongu, v1::nil => Some (Val.maketotal (Val.singleoflongu v1))
+ | Obits_of_single, v1::nil => Some (ExtValues.bits_of_single v1)
+ | Obits_of_float, v1::nil => Some (ExtValues.bits_of_float v1)
| Ocmp c, _ => Some (Val.of_optbool (eval_condition c vl m))
| _, _ => None
end.
@@ -348,9 +353,9 @@ Qed.
Ltac FuncInv :=
match goal with
| H: (match ?x with nil => _ | _ :: _ => _ end = Some _) |- _ =>
- destruct x; simpl in H; FuncInv
+ destruct x; cbn in H; FuncInv
| H: (match ?v with Vundef => _ | Vint _ => _ | Vfloat _ => _ | Vptr _ _ => _ end = Some _) |- _ =>
- destruct v; simpl in H; FuncInv
+ destruct v; cbn in H; FuncInv
| H: (if Archi.ptr64 then _ else _) = Some _ |- _ =>
destruct Archi.ptr64 eqn:?; FuncInv
| H: (Some _ = Some _) |- _ =>
@@ -474,6 +479,8 @@ Definition type_of_operation (op: operation) : list typ * typ :=
| Osingleoflong => (Tlong :: nil, Tsingle)
| Osingleoflongu => (Tlong :: nil, Tsingle)
| Ocmp c => (type_of_condition c, Tint)
+ | Obits_of_single => (Tsingle :: nil, Tint)
+ | Obits_of_float => (Tfloat :: nil, Tlong)
end.
Definition type_of_addressing (addr: addressing) : list typ :=
@@ -680,6 +687,9 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
- destruct v0; cbn; trivial.
(* cmp *)
- destruct (eval_condition cond vl m)... destruct b...
+ (* Bits_of_single, float *)
+ - destruct v0; cbn; trivial.
+ - destruct v0; cbn; trivial.
Qed.
(* This should not be simplified to "false" because it breaks proofs elsewhere. *)
@@ -1246,6 +1256,9 @@ Proof.
exploit eval_condition_inj; eauto. intros EQ; rewrite EQ.
destruct b; simpl; constructor.
simpl; constructor.
+ (* Bits_of_single, double *)
+ - inv H4; simpl; auto.
+ - inv H4; simpl; auto.
Qed.
Lemma eval_addressing_inj:
diff --git a/riscV/ValueAOp.v b/riscV/ValueAOp.v
index f4b7b4d6..e1ba878e 100644
--- a/riscV/ValueAOp.v
+++ b/riscV/ValueAOp.v
@@ -42,6 +42,18 @@ Definition eval_static_addressing (addr: addressing) (vl: list aval): aval :=
| _, _ => Vbot
end.
+Definition bits_of_single (v : aval) : aval :=
+ match v with
+ | FS f => I (Float32.to_bits f)
+ | _ => ntop1 v
+ end.
+
+Definition bits_of_float (v : aval) : aval :=
+ match v with
+ | F f => L (Float.to_bits f)
+ | _ => ntop1 v
+ end.
+
Definition eval_static_operation (op: operation) (vl: list aval): aval :=
match op, vl with
| Omove, v1::nil => v1
@@ -137,6 +149,8 @@ Definition eval_static_operation (op: operation) (vl: list aval): aval :=
| Osingleoflong, v1::nil => singleoflong v1
| Osingleoflongu, v1::nil => singleoflongu v1
| Ocmp c, _ => of_optbool (eval_static_condition c vl)
+ | Obits_of_single, v1::nil => bits_of_single v1
+ | Obits_of_float, v1::nil => bits_of_float v1
| _, _ => Vbot
end.
@@ -148,6 +162,20 @@ Hypothesis GENV: genv_match bc ge.
Variable sp: block.
Hypothesis STACK: bc sp = BCstack.
+Lemma bits_of_single_sound:
+ forall v x, vmatch bc v x -> vmatch bc (ExtValues.bits_of_single v) (bits_of_single x).
+Proof.
+ unfold ExtValues.bits_of_single; intros. inv H; cbn; constructor.
+Qed.
+
+Lemma bits_of_float_sound:
+ forall v x, vmatch bc v x -> vmatch bc (ExtValues.bits_of_float v) (bits_of_float x).
+Proof.
+ unfold ExtValues.bits_of_float; intros. inv H; cbn; constructor.
+Qed.
+
+Hint Resolve bits_of_single_sound bits_of_float_sound : va.
+
Theorem eval_static_condition_sound:
forall cond vargs m aargs,
list_forall2 (vmatch bc) vargs aargs ->