From be0b1872bf2ad36df9b0c7a0ffa63b9e77fa769b Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Tue, 18 Feb 2020 16:57:17 +0100 Subject: Refine the type of function results in AST.signature Before it was "option typ". Now it is a proper inductive type that can also express small integer types (8/16-bit unsigned/signed integers). One benefit is that external functions get more precise types that control better their return values. As a consequence, the CompCert C type preservation property now holds unconditionally, without extra typing hypotheses on external functions. --- arm/Conventions1.v | 26 ++++++++++++-------------- 1 file changed, 12 insertions(+), 14 deletions(-) (limited to 'arm/Conventions1.v') diff --git a/arm/Conventions1.v b/arm/Conventions1.v index c5277e8d..45008bff 100644 --- a/arm/Conventions1.v +++ b/arm/Conventions1.v @@ -104,13 +104,12 @@ Definition is_float_reg (r: mreg): bool := representation with a single LDM instruction. *) Definition loc_result (s: signature) : rpair mreg := - match s.(sig_res) with - | None => One R0 - | Some (Tint | Tany32) => One R0 - | Some (Tfloat | Tsingle | Tany64) => One F0 - | Some Tlong => if Archi.big_endian - then Twolong R0 R1 - else Twolong R1 R0 + match proj_sig_res s with + | Tint | Tany32 => One R0 + | Tfloat | Tsingle | Tany64 => One F0 + | Tlong => if Archi.big_endian + then Twolong R0 R1 + else Twolong R1 R0 end. (** The result registers have types compatible with that given in the signature. *) @@ -119,7 +118,7 @@ Lemma loc_result_type: forall sig, subtype (proj_sig_res sig) (typ_rpair mreg_type (loc_result sig)) = true. Proof. - intros. unfold proj_sig_res, loc_result. destruct (sig_res sig) as [[]|]; destruct Archi.big_endian; auto. + intros. unfold loc_result. destruct (proj_sig_res sig); destruct Archi.big_endian; auto. Qed. (** The result locations are caller-save registers *) @@ -129,7 +128,7 @@ Lemma loc_result_caller_save: forall_rpair (fun r => is_callee_save r = false) (loc_result s). Proof. intros. - unfold loc_result. destruct (sig_res s) as [[]|]; destruct Archi.big_endian; simpl; auto. + unfold loc_result. destruct (proj_sig_res s); destruct Archi.big_endian; simpl; auto. Qed. (** If the result is in a pair of registers, those registers are distinct and have type [Tint] at least. *) @@ -139,14 +138,13 @@ Lemma loc_result_pair: match loc_result sg with | One _ => True | Twolong r1 r2 => - r1 <> r2 /\ sg.(sig_res) = Some Tlong + r1 <> r2 /\ proj_sig_res sg = Tlong /\ subtype Tint (mreg_type r1) = true /\ subtype Tint (mreg_type r2) = true /\ Archi.ptr64 = false end. Proof. - intros; unfold loc_result; destruct (sig_res sg) as [[]|]; destruct Archi.big_endian; auto. - intuition congruence. - intuition congruence. + intros; unfold loc_result; destruct (proj_sig_res sg); auto. + destruct Archi.big_endian; intuition congruence. Qed. (** The location of the result depends only on the result part of the signature *) @@ -154,7 +152,7 @@ Qed. Lemma loc_result_exten: forall s1 s2, s1.(sig_res) = s2.(sig_res) -> loc_result s1 = loc_result s2. Proof. - intros. unfold loc_result. rewrite H; auto. + intros. unfold loc_result, proj_sig_res. rewrite H; auto. Qed. (** ** Location of function arguments *) -- cgit