Require Import Coqlib Maps. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import RTL Op Registers OptionMonad BTL. Require Export BTLmatchRTL. Require Import Errors Linking. (** External oracle *) Axiom btl2rtl: BTL.function -> RTL.code * node * (PTree.t node). Extract Constant btl2rtl => "BTLtoRTLaux.btl2rtl". Local Open Scope error_monad_scope. Definition transf_function (f: BTL.function) : res RTL.function := let (tcte, dupmap) := btl2rtl f in let (tc, te) := tcte in let f' := RTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in do u <- verify_function dupmap f f'; OK f'. Definition transf_fundef (f: fundef) : res RTL.fundef := transf_partial_fundef transf_function f. Definition transf_program (p: program) : res RTL.program := transform_partial_program transf_fundef p.