summaryrefslogtreecommitdiff
path: root/cfrontend/Cshmgen.v
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2008-09-27 08:57:55 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2008-09-27 08:57:55 +0000
commitec6d00d94bcb1a0adc5c698367634b5e2f370c6e (patch)
tree2e66a3c9211e2952cdfb50374c76baea6fa68eec /cfrontend/Cshmgen.v
parent2cf153d684a48ed7ab910c77d9d98b4c9da3fe2e (diff)
Clight: ajout Econdition, suppression Eindex.
caml/PrintCsyntax.ml: afficher les formes a[b] et a->fld. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@789 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'cfrontend/Cshmgen.v')
-rw-r--r--cfrontend/Cshmgen.v10
1 files changed, 3 insertions, 7 deletions
diff --git a/cfrontend/Cshmgen.v b/cfrontend/Cshmgen.v
index 20eb17f..64a58ad 100644
--- a/cfrontend/Cshmgen.v
+++ b/cfrontend/Cshmgen.v
@@ -383,11 +383,11 @@ Fixpoint transl_expr (a: Csyntax.expr) {struct a} : res expr :=
| Expr (Csyntax.Ecast ty b) _ =>
do tb <- transl_expr b;
OK (make_cast (typeof b) ty tb)
- | Expr (Csyntax.Eindex b c) ty =>
+ | Expr (Csyntax.Econdition b c d) _ =>
do tb <- transl_expr b;
do tc <- transl_expr c;
- do ts <- make_add tb (typeof b) tc (typeof c);
- make_load ts ty
+ do td <- transl_expr d;
+ OK(Econdition (make_boolean tb (typeof b)) tc td)
| Expr (Csyntax.Eandbool b c) _ =>
do tb <- transl_expr b;
do tc <- transl_expr c;
@@ -425,10 +425,6 @@ with transl_lvalue (a: Csyntax.expr) {struct a} : res expr :=
OK (Eaddrof id)
| Expr (Csyntax.Ederef b) _ =>
transl_expr b
- | Expr (Csyntax.Eindex b c) _ =>
- do tb <- transl_expr b;
- do tc <- transl_expr c;
- make_add tb (typeof b) tc (typeof c)
| Expr (Csyntax.Efield b i) ty =>
match typeof b with
| Tstruct _ fld =>