summaryrefslogtreecommitdiff
path: root/backend/Cminor.v
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2014-08-17 07:52:12 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2014-08-17 07:52:12 +0000
commit17f519651feb4a09aa90c89c949469e8a5ab0e88 (patch)
treec7bda5e43a2d1f950180521a1b854ac9592eea73 /backend/Cminor.v
parent88613c0f5415a0d3f2e0e0e9ff74bd32b6b4685e (diff)
- Support "switch" statements over 64-bit integers
(in CompCert C to Cminor, included) - Translation of "switch" to decision trees or jumptables made generic over the sizes of integers and moved to the Cminor->CminorSel pass instead of CminorSel->RTL as before. - CminorSel: add "exitexpr" to support the above. - ValueDomain: more precise analysis of comparisons against an integer literal. E.g. "x >=u 0" is always true. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@2565 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'backend/Cminor.v')
-rw-r--r--backend/Cminor.v19
1 files changed, 10 insertions, 9 deletions
diff --git a/backend/Cminor.v b/backend/Cminor.v
index bf20de2..7ac23bf 100644
--- a/backend/Cminor.v
+++ b/backend/Cminor.v
@@ -148,7 +148,7 @@ Inductive stmt : Type :=
| Sloop: stmt -> stmt
| Sblock: stmt -> stmt
| Sexit: nat -> stmt
- | Sswitch: expr -> list (int * nat) -> nat -> stmt
+ | Sswitch: bool -> expr -> list (Z * nat) -> nat -> stmt
| Sreturn: option expr -> stmt
| Slabel: label -> stmt -> stmt
| Sgoto: label -> stmt.
@@ -510,9 +510,10 @@ Inductive step: state -> trace -> state -> Prop :=
step (State f (Sexit (S n)) (Kblock k) sp e m)
E0 (State f (Sexit n) k sp e m)
- | step_switch: forall f a cases default k sp e m n,
- eval_expr sp e m a (Vint n) ->
- step (State f (Sswitch a cases default) k sp e m)
+ | step_switch: forall f islong a cases default k sp e m v n,
+ eval_expr sp e m a v ->
+ switch_argument islong v n ->
+ step (State f (Sswitch islong a cases default) k sp e m)
E0 (State f (Sexit (switch_target n default cases)) k sp e m)
| step_return_0: forall f k sp e m m',
@@ -733,9 +734,10 @@ with exec_stmt:
forall f sp e m n,
exec_stmt f sp e m (Sexit n) E0 e m (Out_exit n)
| exec_Sswitch:
- forall f sp e m a cases default n,
- eval_expr ge sp e m a (Vint n) ->
- exec_stmt f sp e m (Sswitch a cases default)
+ forall f sp e m islong a cases default v n,
+ eval_expr ge sp e m a v ->
+ switch_argument islong v n ->
+ exec_stmt f sp e m (Sswitch islong a cases default)
E0 e m (Out_exit (switch_target n default cases))
| exec_Sreturn_none:
forall f sp e m,
@@ -1051,8 +1053,7 @@ Proof.
(* switch *)
econstructor; split.
- apply star_one. econstructor; eauto.
- constructor.
+ apply star_one. econstructor; eauto. constructor.
(* return none *)
econstructor; split. apply star_refl. constructor; auto.