aboutsummaryrefslogtreecommitdiffstats
path: root/backend/Mach.v
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@inria.fr>2015-03-27 08:55:05 +0100
committerXavier Leroy <xavier.leroy@inria.fr>2015-03-27 08:55:05 +0100
commit4622f49fd089ae47d0c853343cb0a05f986c962a (patch)
treebd045e1ef59d57f8e4b5f5734470cae56a4e68b7 /backend/Mach.v
parent8d75ab2d38fa20dc7d8e3839967015cc276cd642 (diff)
downloadcompcert-kvx-4622f49fd089ae47d0c853343cb0a05f986c962a.tar.gz
compcert-kvx-4622f49fd089ae47d0c853343cb0a05f986c962a.zip
Extend annotations so that they can keep track of global variables and local variables whose address is taken.
- CminorSel, RTL: add "annot" instructions. - CminorSel to Asm: use type "annot_arg" for arguments of "annot" instructions. - AST, Events: simplify EF_annot because constants are now part of the arguments. Implementation is not complete yet.
Diffstat (limited to 'backend/Mach.v')
-rw-r--r--backend/Mach.v25
1 files changed, 4 insertions, 21 deletions
diff --git a/backend/Mach.v b/backend/Mach.v
index ff72a828..fe00d42d 100644
--- a/backend/Mach.v
+++ b/backend/Mach.v
@@ -61,16 +61,12 @@ Inductive instruction: Type :=
| Mcall: signature -> mreg + ident -> instruction
| Mtailcall: signature -> mreg + ident -> instruction
| Mbuiltin: external_function -> list mreg -> list mreg -> instruction
- | Mannot: external_function -> list annot_param -> instruction
+ | Mannot: external_function -> list (annot_arg mreg) -> instruction
| Mlabel: label -> instruction
| Mgoto: label -> instruction
| Mcond: condition -> list mreg -> label -> instruction
| Mjumptable: mreg -> list label -> instruction
- | Mreturn: instruction
-
-with annot_param: Type :=
- | APreg: mreg -> annot_param
- | APstack: memory_chunk -> Z -> annot_param.
+ | Mreturn: instruction.
Definition code := list instruction.
@@ -225,19 +221,6 @@ Definition extcall_arguments
(rs: regset) (m: mem) (sp: val) (sg: signature) (args: list val) : Prop :=
list_forall2 (extcall_arg rs m sp) (loc_arguments sg) args.
-(** Extract the values of the arguments to an annotation. *)
-
-Inductive annot_arg: regset -> mem -> val -> annot_param -> val -> Prop :=
- | annot_arg_reg: forall rs m sp r,
- annot_arg rs m sp (APreg r) (rs r)
- | annot_arg_stack: forall rs m stk base chunk ofs v,
- Mem.load chunk m stk (Int.unsigned base + ofs) = Some v ->
- annot_arg rs m (Vptr stk base) (APstack chunk ofs) v.
-
-Definition annot_arguments
- (rs: regset) (m: mem) (sp: val) (params: list annot_param) (args: list val) : Prop :=
- list_forall2 (annot_arg rs m sp) params args.
-
(** Mach execution states. *)
(** Mach execution states. *)
@@ -352,8 +335,8 @@ Inductive step: state -> trace -> state -> Prop :=
t (State s f sp b rs' m')
| exec_Mannot:
forall s f sp rs m ef args b vargs t v m',
- annot_arguments rs m sp args vargs ->
- external_call' ef ge vargs m t v m' ->
+ eval_annot_args ge rs sp m args vargs ->
+ external_call ef ge vargs m t v m' ->
step (State s f sp (Mannot ef args :: b) rs m)
t (State s f sp b rs m')
| exec_Mgoto: