Require Import Coqlib Maps. Require Import AST Integers Values Events Memory Globalenvs Smallstep. Require Import RTL Op Registers OptionMonad BTL. Require Import Errors Linking BTL_SEtheory BTL_SEimpl. (** External oracle *) Axiom scheduler: BTL.function -> BTL.code. Extract Constant scheduler => "BTLScheduleraux.btl_scheduler". Definition equiv_input_regs (f1 f2: BTL.function): Prop := (forall pc, (fn_code f1)!pc = None <-> (fn_code f2)!pc = None) /\ (forall pc ib1 ib2, (fn_code f1)!pc = Some ib1 -> (fn_code f2)!pc = Some ib2 -> ib1.(input_regs) = ib2.(input_regs)). Lemma equiv_input_regs_union f1 f2: equiv_input_regs f1 f2 -> forall tbl, union_inputs f1 tbl = union_inputs f2 tbl. Proof. intros (EQNone & EQSome). induction tbl as [|pc l']; simpl; auto. generalize (EQNone pc) (EQSome pc); clear EQNone EQSome; intros EQN EQS. do 2 autodestruct; intuition; try_simplify_someHyps. intros; exploit EQS; eauto; clear EQS. congruence. Qed. Lemma equiv_input_regs_pre f1 f2 tbl or: equiv_input_regs f1 f2 -> pre_inputs f1 tbl or = pre_inputs f2 tbl or. Proof. intros; unfold pre_inputs; erewrite equiv_input_regs_union; auto. Qed. Lemma equiv_input_regs_tr_inputs f1 f2 l oreg rs: equiv_input_regs f1 f2 -> tr_inputs f1 l oreg rs = tr_inputs f2 l oreg rs. Proof. intros; unfold tr_inputs; erewrite equiv_input_regs_pre; eauto. Qed. (* a specification of the verification to do on each function *) Record match_function (f1 f2: BTL.function): Prop := { preserv_fnsig: fn_sig f1 = fn_sig f2; preserv_fnparams: fn_params f1 = fn_params f2; preserv_fnstacksize: fn_stacksize f1 = fn_stacksize f2; preserv_entrypoint: fn_entrypoint f1 = fn_entrypoint f2; preserv_inputs: equiv_input_regs f1 f2; symbolic_simu_ok: forall pc ib1, (fn_code f1)!pc = Some ib1 -> exists ib2, (fn_code f2)!pc = Some ib2 /\ symbolic_simu f1 f2 (entry ib1) (entry ib2); }. Local Open Scope error_monad_scope. Fixpoint check_symbolic_simu_rec (f1 f2: BTL.function) (lpc: list node): res unit := match lpc with | nil => OK tt | pc :: lpc' => match (fn_code f1)!pc, (fn_code f2)!pc with | Some ibf1, Some ibf2 => do _ <- simu_check f1 f2 ibf1 ibf2; check_symbolic_simu_rec f1 f2 lpc' | _, _ => Error (msg "check_symbolic_simu_rec: code tree mismatch") end end. Definition check_symbolic_simu (f1 f2: BTL.function): res unit := check_symbolic_simu_rec f1 f2 (List.map (fun elt => fst elt) (PTree.elements (fn_code f1))). Lemma check_symbolic_simu_rec_input_equiv lpc: forall f1 f2 x, check_symbolic_simu_rec f1 f2 lpc = OK x -> lpc = (List.map (fun elt => fst elt) (PTree.elements (fn_code f1))) -> equiv_input_regs f1 f2. Proof. Admitted. (* unfold equiv_input_regs; induction lpc; simpl; intros f1 f2 x X EQL1. - destruct (map _ _) eqn:EQL2; inv X; intuition; symmetry in EQL1; apply map_eq_nil in EQL1; apply map_eq_nil in EQL2. + destruct (fn_code f1), (fn_code f2); simpl in *; auto. inv EQL2. unfold PTree.elements in EQL1. + admit. + destruct ((fn_code f1) ! a), ((fn_code f2) ! a); monadInv X. eapply IHlpc; eauto. Qed. *) Lemma check_symbolic_simu_input_equiv x f1 f2: check_symbolic_simu f1 f2 = OK x -> equiv_input_regs f1 f2. Proof. unfold check_symbolic_simu; simpl; intros X. eapply check_symbolic_simu_rec_input_equiv; eauto. Qed. Lemma check_symbolic_simu_rec_correct lpc: forall f1 f2 x, check_symbolic_simu_rec f1 f2 lpc = OK x -> forall pc ib1, (fn_code f1)!pc = Some ib1 -> exists ib2, (fn_code f2)!pc = Some ib2 /\ symbolic_simu f1 f2 (entry ib1) (entry ib2). Proof. Admitted. (* induction lpc; simpl; intros f1 f2 x X pc ib1 PC; try (inv X; fail). destruct ((fn_code f1) ! pc), ((fn_code f2) ! pc); inv X. eapply IHlpc; eauto. Qed. *) Lemma check_symbolic_simu_correct x f1 f2: check_symbolic_simu f1 f2 = OK x -> forall pc ib1, (fn_code f1)!pc = Some ib1 -> exists ib2, (fn_code f2)!pc = Some ib2 /\ symbolic_simu f1 f2 (entry ib1) (entry ib2). Proof. unfold check_symbolic_simu; intros X pc ib1 PC. eapply check_symbolic_simu_rec_correct; eauto. Qed. Definition transf_function (f: BTL.function) := let tf := BTL.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) (scheduler f) (fn_entrypoint f) in do _ <- check_symbolic_simu f tf; OK tf. Definition transf_fundef (f: fundef) : res fundef := transf_partial_fundef (fun f => transf_function f) f. Definition transf_program (p: program) : res program := transform_partial_program transf_fundef p.