aboutsummaryrefslogtreecommitdiffstats
path: root/backend/Allocation.v
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@college-de-france.fr>2020-02-18 16:57:17 +0100
committerXavier Leroy <xavierleroy@users.noreply.github.com>2020-02-21 13:29:39 +0100
commitbe0b1872bf2ad36df9b0c7a0ffa63b9e77fa769b (patch)
tree0d54cce547d12567d7e9e9f2c4d650e5a1b94b39 /backend/Allocation.v
parenta9eaf4897c825093aba2137ff76e56bfbf1e72d5 (diff)
downloadcompcert-kvx-be0b1872bf2ad36df9b0c7a0ffa63b9e77fa769b.tar.gz
compcert-kvx-be0b1872bf2ad36df9b0c7a0ffa63b9e77fa769b.zip
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.
Diffstat (limited to 'backend/Allocation.v')
-rw-r--r--backend/Allocation.v10
1 files changed, 5 insertions, 5 deletions
diff --git a/backend/Allocation.v b/backend/Allocation.v
index 13e14530..08e0a4f4 100644
--- a/backend/Allocation.v
+++ b/backend/Allocation.v
@@ -734,11 +734,11 @@ Function add_equations_args (rl: list reg) (tyl: list typ) (ll: list (rpair loc)
(** [add_equations_res] is similar but is specialized to the case where
there is only one pseudo-register. *)
-Function add_equations_res (r: reg) (oty: option typ) (p: rpair mreg) (e: eqs) : option eqs :=
- match p, oty with
+Function add_equations_res (r: reg) (ty: typ) (p: rpair mreg) (e: eqs) : option eqs :=
+ match p, ty with
| One mr, _ =>
Some (add_equation (Eq Full r (R mr)) e)
- | Twolong mr1 mr2, Some Tlong =>
+ | Twolong mr1 mr2, Tlong =>
if Archi.ptr64 then None else
Some (add_equation (Eq Low r (R mr2)) (add_equation (Eq High r (R mr1)) e))
| _, _ =>
@@ -1084,7 +1084,7 @@ Definition transfer_aux (f: RTL.function) (env: regenv)
| BStailcall sg ros args mv1 ros' =>
let args' := loc_arguments sg in
assertion (tailcall_is_possible sg);
- assertion (opt_typ_eq sg.(sig_res) f.(RTL.fn_sig).(sig_res));
+ assertion (rettype_eq sg.(sig_res) f.(RTL.fn_sig).(sig_res));
assertion (ros_compatible_tailcall ros');
do e1 <- add_equation_ros ros ros' empty_eqs;
do e2 <- add_equations_args args (sig_args sg) args' e1;
@@ -1114,7 +1114,7 @@ Definition transfer_aux (f: RTL.function) (env: regenv)
track_moves env mv empty_eqs
| BSreturn (Some arg) mv =>
let arg' := loc_result (RTL.fn_sig f) in
- do e1 <- add_equations_res arg (sig_res (RTL.fn_sig f)) arg' empty_eqs;
+ do e1 <- add_equations_res arg (proj_sig_res (RTL.fn_sig f)) arg' empty_eqs;
track_moves env mv e1
end.