aboutsummaryrefslogtreecommitdiffstats
path: root/exportclight/Clightdefs.v
diff options
context:
space:
mode:
Diffstat (limited to 'exportclight/Clightdefs.v')
-rw-r--r--exportclight/Clightdefs.v42
1 files changed, 26 insertions, 16 deletions
diff --git a/exportclight/Clightdefs.v b/exportclight/Clightdefs.v
index fda5bb55..83d82d88 100644
--- a/exportclight/Clightdefs.v
+++ b/exportclight/Clightdefs.v
@@ -15,17 +15,8 @@
(** All imports and definitions used by .v Clight files generated by clightgen *)
-Require Export String.
-Require Export List.
-Require Export ZArith.
-Require Export Integers.
-Require Export Floats.
-Require Export AST.
-Require Export Ctypes.
-Require Export Cop.
-Require Export Clight.
-Require Import Maps.
-Require Import Errors.
+From Coq Require Import String List ZArith.
+From compcert Require Import Integers Floats Maps Errors AST Ctypes Cop Clight.
Definition tvoid := Tvoid.
Definition tschar := Tint I8 Signed noattr.
@@ -65,8 +56,27 @@ Definition talignas (n: N) (ty: type) :=
Definition tvolatile_alignas (n: N) (ty: type) :=
tattr {| attr_volatile := true; attr_alignas := Some n |} ty.
-Definition make_composite_env (comps: list composite_definition): composite_env :=
- match build_composite_env comps with
- | OK e => e
- | Error _ => PTree.empty _
- end.
+Definition wf_composites (types: list composite_definition) : Prop :=
+ match build_composite_env types with OK _ => True | Error _ => False end.
+
+Definition build_composite_env' (types: list composite_definition)
+ (WF: wf_composites types)
+ : { ce | build_composite_env types = OK ce }.
+Proof.
+ revert WF. unfold wf_composites. case (build_composite_env types); intros.
+- exists c; reflexivity.
+- contradiction.
+Defined.
+
+Definition mkprogram (types: list composite_definition)
+ (defs: list (ident * globdef fundef type))
+ (public: list ident)
+ (main: ident)
+ (WF: wf_composites types) : Clight.program :=
+ let (ce, EQ) := build_composite_env' types WF in
+ {| prog_defs := defs;
+ prog_public := public;
+ prog_main := main;
+ prog_types := types;
+ prog_comp_env := ce;
+ prog_comp_env_eq := EQ |}.