summaryrefslogtreecommitdiff
path: root/backend/RTLtyping.v
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2011-06-13 18:11:19 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2011-06-13 18:11:19 +0000
commita5ffc59246b09a389e5f8cbc2f217e323e76990f (patch)
treee1bc7cc54518aad7c20645f187cee8110de8cff9 /backend/RTLtyping.v
parent4daccd62b92b23016d3f343d5691f9c164a8a951 (diff)
Revised handling of annotation statements, and more generally built-in functions, and more generally external functions
git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1672 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'backend/RTLtyping.v')
-rw-r--r--backend/RTLtyping.v5
1 files changed, 3 insertions, 2 deletions
diff --git a/backend/RTLtyping.v b/backend/RTLtyping.v
index 49f339d..02359b9 100644
--- a/backend/RTLtyping.v
+++ b/backend/RTLtyping.v
@@ -110,7 +110,7 @@ Inductive wt_instr : instruction -> Prop :=
forall ef args res s,
List.map env args = (ef_sig ef).(sig_args) ->
env res = proj_sig_res (ef_sig ef) ->
- arity_ok (ef_sig ef).(sig_args) = true ->
+ arity_ok (ef_sig ef).(sig_args) = true \/ ef_reloads ef = false ->
valid_successor s ->
wt_instr (Ibuiltin ef args res s)
| wt_Icond:
@@ -236,7 +236,7 @@ Definition check_instr (i: instruction) : bool :=
| Ibuiltin ef args res s =>
check_regs args (ef_sig ef).(sig_args)
&& check_reg res (proj_sig_res (ef_sig ef))
- && arity_ok (ef_sig ef).(sig_args)
+ && (if ef_reloads ef then arity_ok (ef_sig ef).(sig_args) else true)
&& check_successor s
| Icond cond args s1 s2 =>
check_regs args (type_of_condition cond)
@@ -349,6 +349,7 @@ Proof.
apply check_regs_correct; auto.
apply check_reg_correct; auto.
auto.
+ destruct (ef_reloads e); auto.
apply check_successor_correct; auto.
(* cond *)
constructor. apply check_regs_correct; auto.