aboutsummaryrefslogtreecommitdiffstats
path: root/cfrontend/Csharpminor.v
diff options
context:
space:
mode:
authorxleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2006-09-05 12:32:59 +0000
committerxleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2006-09-05 12:32:59 +0000
commit6c0511a03c8c970435d8b97e600312ac45340801 (patch)
treed42b16c8f77d1aa34e570647eb6728a4a9f4e72b /cfrontend/Csharpminor.v
parent81afbd38d1597cefc03dd699fd115c4261c6877f (diff)
downloadcompcert-kvx-6c0511a03c8c970435d8b97e600312ac45340801.tar.gz
compcert-kvx-6c0511a03c8c970435d8b97e600312ac45340801.zip
Revu traitement des variables globales dans AST.program et dans Globalenvs.
Adaptation frontend, backend en consequence. Integration passe C -> C#minor dans common/Main.v. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@77 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'cfrontend/Csharpminor.v')
-rw-r--r--cfrontend/Csharpminor.v20
1 files changed, 5 insertions, 15 deletions
diff --git a/cfrontend/Csharpminor.v b/cfrontend/Csharpminor.v
index 246ebf53..016fab4d 100644
--- a/cfrontend/Csharpminor.v
+++ b/cfrontend/Csharpminor.v
@@ -122,11 +122,7 @@ Record function : Set := mkfunction {
Definition fundef := AST.fundef function.
-Record program : Set := mkprogram {
- prog_funct: list (ident * fundef);
- prog_main: ident;
- prog_vars: list (ident * var_kind * list init_data)
-}.
+Definition program : Set := AST.program fundef var_kind.
Definition funsig (fd: fundef) :=
match fd with
@@ -187,12 +183,6 @@ Definition sizeof (lv: var_kind) : Z :=
| Varray sz => Zmax 0 sz
end.
-Definition program_of_program (p: program): AST.program fundef :=
- AST.mkprogram
- p.(prog_funct)
- p.(prog_main)
- (List.map (fun x => match x with (id, k, init) => (id, init) end) p.(prog_vars)).
-
Definition fn_variables (f: function) :=
List.map
(fun id_chunk => (fst id_chunk, Vscalar (snd id_chunk))) f.(fn_params)
@@ -206,7 +196,7 @@ Definition fn_vars_names (f: function) :=
Definition global_var_env (p: program): gvarenv :=
List.fold_left
- (fun gve x => match x with (id, k, init) => PTree.set id k gve end)
+ (fun gve x => match x with (id, init, k) => PTree.set id k gve end)
p.(prog_vars) (PTree.empty var_kind).
(** Evaluation of operator applications. *)
@@ -316,7 +306,7 @@ Inductive bind_parameters: env ->
Section RELSEM.
Variable prg: program.
-Let ge := Genv.globalenv (program_of_program prg).
+Let ge := Genv.globalenv prg.
(* Evaluation of the address of a variable:
[eval_var_addr prg ge e id b] states that variable [id]
@@ -554,8 +544,8 @@ End RELSEM.
in the initial memory state for [p] eventually returns value [r]. *)
Definition exec_program (p: program) (t: trace) (r: val) : Prop :=
- let ge := Genv.globalenv (program_of_program p) in
- let m0 := Genv.init_mem (program_of_program p) in
+ let ge := Genv.globalenv p in
+ let m0 := Genv.init_mem p in
exists b, exists f, exists m,
Genv.find_symbol ge p.(prog_main) = Some b /\
Genv.find_funct_ptr ge b = Some f /\