aboutsummaryrefslogtreecommitdiffstats
path: root/backend/Duplicate.v
diff options
context:
space:
mode:
Diffstat (limited to 'backend/Duplicate.v')
-rw-r--r--backend/Duplicate.v26
1 files changed, 21 insertions, 5 deletions
diff --git a/backend/Duplicate.v b/backend/Duplicate.v
index 0e04b07d..6ae5475d 100644
--- a/backend/Duplicate.v
+++ b/backend/Duplicate.v
@@ -18,14 +18,28 @@
Require Import AST RTL Maps Globalenvs.
Require Import Coqlib Errors Op.
-Local Open Scope error_monad_scope.
-Local Open Scope positive_scope.
-(** External oracle returning the new RTL code (entry point unchanged),
+
+Module Type DuplicateParam.
+
+ (** External oracle returning the new RTL code (entry point unchanged),
along with the new entrypoint, and a mapping of new nodes to old nodes *)
-Axiom duplicate_aux: function -> code * node * (PTree.t node).
+ Parameter duplicate_aux: function -> code * node * (PTree.t node).
+
+End DuplicateParam.
+
+
-Extract Constant duplicate_aux => "Duplicateaux.duplicate_aux".
+Module Duplicate (D: DuplicateParam).
+
+Export D.
+
+Definition duplicate_aux := duplicate_aux.
+
+(* Extract Constant duplicate_aux => "Duplicateaux.duplicate_aux". *)
+
+Local Open Scope error_monad_scope.
+Local Open Scope positive_scope.
(** * Verification of node duplications *)
@@ -215,3 +229,5 @@ Definition transf_fundef (f: fundef) : res fundef :=
Definition transf_program (p: program) : res program :=
transform_partial_program transf_fundef p.
+
+End Duplicate.