aboutsummaryrefslogtreecommitdiffstats
path: root/backend/Cminor.v
diff options
context:
space:
mode:
Diffstat (limited to 'backend/Cminor.v')
-rw-r--r--backend/Cminor.v20
1 files changed, 16 insertions, 4 deletions
diff --git a/backend/Cminor.v b/backend/Cminor.v
index ca01ad50..91a4c104 100644
--- a/backend/Cminor.v
+++ b/backend/Cminor.v
@@ -676,12 +676,24 @@ Definition outcome_block (out: outcome) : outcome :=
| out => out
end.
+(*
Definition outcome_result_value
- (out: outcome) (retsig: option typ) (vres: val) : Prop :=
+ (out: outcome) (retsig: rettype) (vres: val) : Prop :=
match out with
| Out_normal => vres = Vundef
| Out_return None => vres = Vundef
- | Out_return (Some v) => retsig <> None /\ vres = v
+ | Out_return (Some v) => retsig <> Tvoid /\ vres = v
+ | Out_tailcall_return v => vres = v
+ | _ => False
+ end.
+*)
+
+Definition outcome_result_value
+ (out: outcome) (vres: val) : Prop :=
+ match out with
+ | Out_normal => vres = Vundef
+ | Out_return None => vres = Vundef
+ | Out_return (Some v) => vres = v
| Out_tailcall_return v => vres = v
| _ => False
end.
@@ -711,7 +723,7 @@ Inductive eval_funcall:
Mem.alloc m 0 f.(fn_stackspace) = (m1, sp) ->
set_locals f.(fn_vars) (set_params vargs f.(fn_params)) = e ->
exec_stmt f (Vptr sp Ptrofs.zero) e m1 f.(fn_body) t e2 m2 out ->
- outcome_result_value out f.(fn_sig).(sig_res) vres ->
+ outcome_result_value out vres ->
outcome_free_mem out m2 sp f.(fn_stackspace) m3 ->
eval_funcall m (Internal f) vargs t m3 vres
| eval_funcall_external:
@@ -995,7 +1007,7 @@ Proof.
subst vres. replace k with (call_cont k') by congruence.
apply star_one. apply step_return_0; auto.
(* Out_return Some *)
- destruct H3. subst vres.
+ subst vres.
replace k with (call_cont k') by congruence.
apply star_one. eapply step_return_1; eauto.
(* Out_tailcall_return *)