Skip to content
Snippets Groups Projects
Commit 0d2f4a41 authored by xleroy's avatar xleroy
Browse files

Relaxation de la regle d'evaluation Ecast

git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@420 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
parent 8a07279b
No related branches found
No related tags found
No related merge requests found
...@@ -562,10 +562,10 @@ Inductive eval_expr: expr -> val -> Prop := ...@@ -562,10 +562,10 @@ Inductive eval_expr: expr -> val -> Prop :=
eval_expr a2 v2 -> eval_expr a2 v2 ->
bool_of_val v2 (typeof a2) v -> bool_of_val v2 (typeof a2) v ->
eval_expr (Expr (Eandbool a1 a2) ty) v eval_expr (Expr (Eandbool a1 a2) ty) v
| eval_Ecast: forall a ty v1 v, | eval_Ecast: forall a ty ty' v1 v,
eval_expr a v1 -> eval_expr a v1 ->
cast v1 (typeof a) ty v -> cast v1 (typeof a) ty v ->
eval_expr (Expr (Ecast ty a) ty) v eval_expr (Expr (Ecast ty a) ty') v
(** [eval_lvalue ge e m a b ofs] defines the evaluation of expression [a] (** [eval_lvalue ge e m a b ofs] defines the evaluation of expression [a]
in l-value position. The result is the memory location [b, ofs] in l-value position. The result is the memory location [b, ofs]
......
...@@ -628,10 +628,10 @@ Proof. ...@@ -628,10 +628,10 @@ Proof.
Qed. Qed.
Lemma transl_Ecast_correct: Lemma transl_Ecast_correct:
forall (a : Csyntax.expr) (ty : type) (v1 v : val), forall (a : Csyntax.expr) (ty ty': type) (v1 v : val),
Csem.eval_expr ge e m a v1 -> Csem.eval_expr ge e m a v1 ->
eval_expr_prop a v1 -> eval_expr_prop a v1 ->
cast v1 (typeof a) ty v -> eval_expr_prop (Expr (Ecast ty a) ty) v. cast v1 (typeof a) ty v -> eval_expr_prop (Expr (Ecast ty a) ty') v.
Proof. Proof.
intros; red; intros. inversion WT; clear WT; subst. monadInv TR. intros; red; intros. inversion WT; clear WT; subst. monadInv TR.
eapply make_cast_correct; eauto. eapply make_cast_correct; eauto.
......
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Finish editing this message first!
Please register or to comment