summaryrefslogtreecommitdiff
path: root/powerpc/Op.v
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2013-05-19 09:54:40 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2013-05-19 09:54:40 +0000
commitbe4d6e42dfa287b93b1a35ec820ab2a5aaf8c7ec (patch)
treec51b66e9154bc64cf4fd4191251f29d102928841 /powerpc/Op.v
parent60e1fd71c7e97b2214daf574e0f41b55a3e0bceb (diff)
Merge of the float32 branch:
- added RTL type "Tsingle" - ABI-compatible passing of single-precision floats on ARM and x86 git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@2260 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'powerpc/Op.v')
-rw-r--r--powerpc/Op.v21
1 files changed, 4 insertions, 17 deletions
diff --git a/powerpc/Op.v b/powerpc/Op.v
index 5835717..3963c6b 100644
--- a/powerpc/Op.v
+++ b/powerpc/Op.v
@@ -265,7 +265,7 @@ Definition type_of_operation (op: operation) : list typ * typ :=
match op with
| Omove => (nil, Tint) (* treated specially *)
| Ointconst _ => (nil, Tint)
- | Ofloatconst _ => (nil, Tfloat)
+ | Ofloatconst f => (nil, if Float.is_single_dec f then Tsingle else Tfloat)
| Oaddrsymbol _ _ => (nil, Tint)
| Oaddrstack _ => (nil, Tint)
| Ocast8signed => (Tint :: nil, Tint)
@@ -303,7 +303,7 @@ Definition type_of_operation (op: operation) : list typ * typ :=
| Osubf => (Tfloat :: Tfloat :: nil, Tfloat)
| Omulf => (Tfloat :: Tfloat :: nil, Tfloat)
| Odivf => (Tfloat :: Tfloat :: nil, Tfloat)
- | Osingleoffloat => (Tfloat :: nil, Tfloat)
+ | Osingleoffloat => (Tfloat :: nil, Tsingle)
| Ointoffloat => (Tfloat :: nil, Tint)
| Ofloatofwords => (Tint :: Tint :: nil, Tfloat)
| Omakelong => (Tint :: Tint :: nil, Tlong)
@@ -340,7 +340,7 @@ Proof with (try exact I).
destruct op; simpl in H0; FuncInv; subst; simpl.
congruence.
exact I.
- exact I.
+ destruct (Float.is_single_dec f); auto.
unfold symbol_address. destruct (Genv.find_symbol genv i)...
destruct sp...
destruct v0...
@@ -380,7 +380,7 @@ Proof with (try exact I).
destruct v0; destruct v1...
destruct v0; destruct v1...
destruct v0; destruct v1...
- destruct v0...
+ destruct v0... simpl. apply Float.singleoffloat_is_single.
destruct v0; simpl in H0; inv H0. destruct (Float.intoffloat f); inv H2...
destruct v0; destruct v1...
destruct v0; destruct v1...
@@ -389,19 +389,6 @@ Proof with (try exact I).
destruct (eval_condition c vl m); simpl... destruct b...
Qed.
-Lemma type_of_chunk_correct:
- forall chunk m addr v,
- Mem.loadv chunk m addr = Some v ->
- Val.has_type v (type_of_chunk chunk).
-Proof.
- intro chunk.
- assert (forall v, Val.has_type (Val.load_result chunk v) (type_of_chunk chunk)).
- destruct v; destruct chunk; exact I.
- intros until v. unfold Mem.loadv.
- destruct addr; intros; try discriminate.
- eapply Mem.load_type; eauto.
-Qed.
-
End SOUNDNESS.
(** * Manipulating and transforming operations *)