aboutsummaryrefslogtreecommitdiffstats
path: root/backend/RTLgenspec.v
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@inria.fr>2015-09-15 14:26:42 +0200
committerXavier Leroy <xavier.leroy@inria.fr>2015-09-15 14:26:42 +0200
commitde40fce9c16ced8d23389cbcfc55ef6d99466fe8 (patch)
tree569187ad4a38e7a8e15d2d151706867a5955db15 /backend/RTLgenspec.v
parent700e9b41253e04ac3a0a16edfa1f41a2d4084462 (diff)
downloadcompcert-kvx-de40fce9c16ced8d23389cbcfc55ef6d99466fe8.tar.gz
compcert-kvx-de40fce9c16ced8d23389cbcfc55ef6d99466fe8.zip
Issue with ignoring the result of non-void builtin functions.
In RTL and beyond, the result of a builtin function that has return type different from "void" must be BR, never BR_none. Otherwise, we get compile-time fatal errors, either in Asmexpand or in Lineartyping.
Diffstat (limited to 'backend/RTLgenspec.v')
-rw-r--r--backend/RTLgenspec.v18
1 files changed, 12 insertions, 6 deletions
diff --git a/backend/RTLgenspec.v b/backend/RTLgenspec.v
index 41b5016f..1e665002 100644
--- a/backend/RTLgenspec.v
+++ b/backend/RTLgenspec.v
@@ -814,7 +814,10 @@ Inductive tr_builtin_res: mapping -> builtin_res ident -> builtin_res reg -> Pro
map.(map_vars)!id = Some r ->
tr_builtin_res map (BR id) (BR r)
| tr_builtin_res_none: forall map,
- tr_builtin_res map BR_none BR_none.
+ tr_builtin_res map BR_none BR_none
+ | tr_builtin_res_fresh: forall map r,
+ ~reg_in_map map r ->
+ tr_builtin_res map BR_none (BR r).
(** [tr_stmt c map stmt ns ncont nexits nret rret] holds if the graph [c],
starting at node [ns], contains instructions that perform the Cminor
@@ -1214,14 +1217,17 @@ Proof.
Qed.
Lemma convert_builtin_res_charact:
- forall map res s res' s' INCR
- (TR: convert_builtin_res map res s = OK res' s' INCR)
+ forall map oty res s res' s' INCR
+ (TR: convert_builtin_res map oty res s = OK res' s' INCR)
(WF: map_valid map s),
tr_builtin_res map res res'.
Proof.
- destruct res; simpl; intros; monadInv TR.
-- constructor. unfold find_var in EQ. destruct (map_vars map)!x; inv EQ; auto.
-- constructor.
+ destruct res; simpl; intros.
+- monadInv TR. constructor. unfold find_var in EQ. destruct (map_vars map)!x; inv EQ; auto.
+- destruct oty; monadInv TR.
++ constructor. eauto with rtlg.
++ constructor.
+- monadInv TR.
Qed.
Lemma transl_stmt_charact: