Module BTL_SEimpl_SE


Require Import Coqlib AST Integers.
Require Import Op Registers.
Require Import BTL.
Require Import BTL_SEsimuref BTL_SEtheory OptionMonad.
Require Import BTL_SEsimplify BTL_SEsimplifyproof BTL_SEsimplifyMem BTL_SEimpl_prelude.
Require Import Values FunInd.

Import Notations.
Import HConsing.
Import SvalNotations.

Import ListNotations.
Local Open Scope list_scope.
Local Open Scope lazy_bool_scope.
Local Open Scope option_monad_scope.
Local Open Scope impure.

Section SymbolicExecution.

  Context `{HCF : HashConsingFcts}.
  Context `{HC : HashConsingHyps HCF}.
  Context `{RRULES: rrules_set}.

Additional okclauses



Definition adp_op_okset (op : operation) (sargs : list_sval) (iinfo : inst_info)
  : ?? OKset.t :=
  match is_promotable_op iinfo.(inst_promotion) with
  | Some prom =>
      DO ok <~ hOKpromotableOp op prom sargs;;
      RET (OKset.singleton (`ok))
  | _ =>
      RET OKset.empty
  end.

Lemma adp_op_okset_correct ctx op sargs iinfo
  (ALIVE : alive (eval_list_sval ctx sargs)):
  WHEN adp_op_okset op sargs iinfo ~> ok THEN
  OKset.eval ctx ok <-> proj1_sig (adp_op_okpred op sargs iinfo) ctx.
Proof.
  simpl; unfold adp_op_okset, adp_op, proj1_sig.
  apply elim_alive in ALIVE as [? ARGS]; rewrite ARGS.
  destruct is_promotable_op; simpl; wlp_seq.
  intros [ok OK] _.
  rewrite OKset.eval_singleton, (okclause_eqP OK); simpl.
  rewrite ARGS, <- ex_Some_iff.
  all:intuition auto using OKset.eval_empty.
Qed.

Global Opaque adp_op_okset.


Definition adp_cond_okset (cond : condition) (sargs : list_sval) (iinfo : inst_info)
  : ?? OKset.t :=
  match is_promotable_cond iinfo.(inst_promotion) with
  | Some prom =>
      DO ok <~ hOKpromotableCond cond prom sargs;;
      RET (OKset.singleton (`ok))
  | _ =>
      RET OKset.empty
  end.

Lemma adp_cond_okset_correct ctx cond sargs iinfo
  (ALIVE : alive (eval_list_sval ctx sargs)):
  WHEN adp_cond_okset cond sargs iinfo ~> ok THEN
  OKset.eval ctx ok <-> proj1_sig (adp_cond_okpred cond sargs iinfo) ctx.
Proof.
  simpl; unfold adp_cond_okset, adp_cond, proj1_sig.
  apply elim_alive in ALIVE as [? ARGS]; rewrite ARGS.
  destruct is_promotable_cond; simpl; wlp_seq.
  intros [ok OK] _.
  rewrite OKset.eval_singleton, (okclause_eqP OK); simpl.
  rewrite ARGS, <- ex_Some_iff.
  all:intuition auto using OKset.eval_empty.
Qed.

Global Opaque adp_cond_okset.

Normalization and rewritings


Definition cbranch_expanse (cond: condition) (args: list_sval) (iinfo: inst_info)
  : ?? (condition * list_sval * OKset.t) :=
  match rewrite_cbranches RRULES cond (list_of_lsv args) iinfo with
  | Some ((cond', vargs), pre) =>
      DO vargs' <~ @fsval_list_proj HCF HC vargs;;
      DO pre' <~ imp_map (@fokclause_proj HCF HC) pre;;
      RET (cond', vargs', OKset.of_list pre')
  | None =>
      RET (cond, args, OKset.empty)
  end.

Lemma cbranch_expanse_spec ctx c l iinfo:
  WHEN cbranch_expanse c l iinfo ~> r THEN
  let '(c', l', pre) := r in
  forall (PRE : OKset.eval ctx pre)
         (ALIVE : alive (eval_list_sval ctx l)),
  eval_scondition ctx c' l' = eval_scondition ctx c l.
Proof.
  unfold cbranch_expanse.
  case rewrite_cbranches as [((cond', vargs), pre)|] eqn:REW; wlp_seq.
  - intros vargs' PJ0 pre' PJ1; intros.
    rewrite OKset.eval_of_list in PRE.
    apply fsval_list_proj_correct in PJ0. eapply okclause_proj_list_correct in PJ1.
    eapply rewrite_cbranches_correct in REW.
    + unfold eval_scondition in *.
      rewrite <- PJ0, REW.
      rewrite !eval_list_sval_eq, l_lsv_l.
      reflexivity.
    + apply PJ1. apply PRE.
    + rewrite eval_list_sval_eq in ALIVE.
      apply map_opt_alive in ALIVE; auto.
  - reflexivity.
Qed.
Global Opaque cbranch_expanse.

Lemma bits_of_float_would_work:
  forall (f: Floats.float),
  Int64.cmpu Cne (Floats.Float.to_bits f) (Floats.Float.to_bits f) = false.
Proof.
  intros f. simpl.
  rewrite Int64.eq_true. auto.
Qed.

Assignment of memory


Definition hrexec_store m chk addr args src iinfo hrs: ?? ristate :=
  DO hargs <~ @hrs_lr_get HCF hrs args;;
  DO hsrc <~ @hrs_sreg_get HCF hrs src;;
  @hrs_set_store HCF HC RRULES m (ris_smem hrs) chk addr hargs hsrc (BTL.store_info_of_iinfo iinfo) hrs.

Lemma hrexec_store_spec ctx sis m chunk addr args src iinfo hrs
  (RR: ris_refines_ok ctx hrs sis):
  WHEN hrexec_store m chunk addr args src iinfo hrs ~> hrs' THEN
  exists sis',
    sexec_store chunk addr args src iinfo sis = Some sis' /\
    ris_refines m ctx hrs' sis'.
Proof.
  unfold hrexec_store, sexec_store; wlp_seq.
  intros hargs GARGS hsrc GSRC hrs' STORE.
  eapply hrs_lr_get_spec_alive in GARGS as (sargs & -> & EQA & AARGS); eauto.
  eapply hrs_sreg_get_spec_alive in GSRC as (ssrc & -> & EQS & ASRC ); eauto.
  eexists; split. reflexivity.
  eapply hrs_set_store_spec in STORE; eauto using ris_refines_ok_rmem.
  refine (ris_refines_morph _ STORE).
  apply set_smem_morph. reflexivity.
  simpl; rewrite EQA, EQS; apply MEM_EQ in RR as ->; reflexivity.
Qed.

Lemma hrexec_store_rel m chunk addr args aaddr src hrs:
  WHEN hrexec_store m chunk addr args aaddr src hrs ~> rst THEN
  ris_rel m hrs rst.
Proof.
  unfold hrexec_store; wlp_seq; intros.
  apply hrs_set_store_rel.
Qed.
Global Opaque hrexec_store.

Evaluation of builtins


Fixpoint hbuiltin_arg (hrs: ristate) (arg : builtin_arg reg): ?? builtin_arg sval :=
  match arg with
  | BA r =>
         DO v <~ @hrs_sreg_get HCF hrs r;;
         RET (BA v)
  | BA_int n => RET (BA_int n)
  | BA_long n => RET (BA_long n)
  | BA_float f0 => RET (BA_float f0)
  | BA_single s => RET (BA_single s)
  | BA_loadstack chunk ptr => RET (BA_loadstack chunk ptr)
  | BA_addrstack ptr => RET (BA_addrstack ptr)
  | BA_loadglobal chunk id ptr => RET (BA_loadglobal chunk id ptr)
  | BA_addrglobal id ptr => RET (BA_addrglobal id ptr)
  | BA_splitlong ba1 ba2 =>
    DO v1 <~ hbuiltin_arg hrs ba1;;
    DO v2 <~ hbuiltin_arg hrs ba2;;
    RET (BA_splitlong v1 v2)
  | BA_addptr ba1 ba2 =>
    DO v1 <~ hbuiltin_arg hrs ba1;;
    DO v2 <~ hbuiltin_arg hrs ba2;;
    RET (BA_addptr v1 v2)
  end.

Lemma hbuiltin_arg_spec hrs arg:
  WHEN hbuiltin_arg hrs arg ~> hargs THEN
  forall ctx sis
  (RR: ris_refines_ok ctx hrs sis),
  exists sarg,
    map_builtin_arg_opt sis arg = Some sarg /\
    eval_builtin_sval ctx hargs = eval_builtin_sval ctx sarg.
Proof.
  unfold eval_builtin_sval.
  induction arg; try solve [wlp_simplify; inv SARG; eauto]; simpl; wlp_seq; simpl.
  1 :intros v V; intros. eapply hrs_sreg_get_spec in V as (? & -> & ->); eauto.
  all:intros v1 V1 v2 V2; intros;
      eapply IHarg1 in V1 as (? & -> & ->); eauto;
      eapply IHarg2 in V2 as (? & -> & ->); eauto.
Qed.
Global Opaque hbuiltin_arg.


Definition hbuiltin_args hrs : list (builtin_arg reg) -> ?? list (builtin_arg sval) :=
  imp_map (hbuiltin_arg hrs).

Lemma hbuiltin_args_spec hrs args:
  WHEN hbuiltin_args hrs args ~> hargs THEN
  forall ctx sis
  (RR: ris_refines_ok ctx hrs sis),
  exists sargs,
    map_opt (map_builtin_arg_opt sis) args = Some sargs /\
    bargs_simu ctx hargs sargs.
Proof.
  unfold hbuiltin_args, bargs_simu, eval_list_builtin_sval.
  induction args; simpl; wlp_seq; eauto.
  intros ha HA hl HL; intros; simpl.
  eapply hbuiltin_arg_spec in HA as (? & -> & ->); eauto.
  eapply IHargs in HL as (? & -> & ->); eauto.
Qed.
Global Opaque hbuiltin_args.


Evaluation of final instructions


Definition hsum_left_optmap {A B C} (f: A -> ?? B) (x: A + C): ?? (B + C) :=
  match x with
  | inl r => DO hr <~ f r;; RET (inl hr)
  | inr s => RET (inr s)
  end.

Lemma hsum_left_optmap_spec hrs (ros: reg + qualident):
  WHEN hsum_left_optmap (@hrs_sreg_get HCF hrs) ros ~> hsi THEN forall ctx (sis: sistate)
  (RR: ris_refines_ok ctx hrs sis),
  exists svi,
    map_sum_left_opt sis ros = Some svi /\
    svident_simu ctx hsi svi.
Proof.
  unfold svident_simu.
  case ros as [r|]; simpl; wlp_seq.
  - intros hr HR; intros; simpl.
    eapply hrs_sreg_get_spec in HR as (? & -> & ->); eauto.
  - repeat econstructor.
Qed.
Global Opaque hsum_left_optmap.


Definition hrexec_final_sfv (i: final) hrs: ?? rfval :=
  match i with
  | Bgoto pc => RET (Sgoto pc)
  | Bcall sig ros args res pc =>
      DO svos <~ hsum_left_optmap (@hrs_sreg_get HCF hrs) ros;;
      DO sargs <~ @hrs_lr_get HCF hrs args;;
      RET (Scall sig svos sargs res pc)
  | Btailcall sig ros args =>
      DO svos <~ hsum_left_optmap (@hrs_sreg_get HCF hrs) ros;;
      DO sargs <~ @hrs_lr_get HCF hrs args;;
      RET (Stailcall sig svos sargs)
  | Bbuiltin ef args res pc =>
      DO sargs <~ hbuiltin_args hrs args;;
      RET (Sbuiltin ef sargs res pc)
  | Breturn or =>
      match or with
      | Some r => DO hr <~ @hrs_sreg_get HCF hrs r;; RET (Sreturn (Some hr))
      | None => RET (Sreturn None)
      end
  | Bjumptable reg tbl =>
      DO sv <~ @hrs_sreg_get HCF hrs reg;;
      RET (Sjumptable sv tbl)
  end.

Lemma hrexec_final_sfv_spec ctx sis fi hrs
  (RR: ris_refines_ok ctx hrs sis):
  WHEN hrexec_final_sfv fi hrs ~> rfv THEN
  exists sfv,
    sexec_final_sfv fi sis = Some sfv /\
    rfv_refines ctx rfv sfv.
Proof.
  unfold rfv_refines.
  destruct fi; simpl. 2:destruct res. all:wlp_seq; simpl.
  2,7:intros sv SV; intros;
        eapply hrs_sreg_get_spec in SV as (? & -> & ?); eauto.
  5,6:intros svos SVOS sargs SARGS; intros;
        eapply hsum_left_optmap_spec in SVOS as (? & -> & ?); eauto;
        eapply hrs_lr_get_spec in SARGS as (? & SARGS & SARGS_EQ); eauto;
        rewrite !eval_list_sval_eq in SARGS_EQ;
        revert SARGS; rewrite lmap_sv_eq; autodestruct; try_simplify_someHyps;
        intros; rewrite l_lsv_l in SARGS_EQ.
  7: intros sargs SARGS; intros;
        eapply hbuiltin_args_spec in SARGS as (? & -> & ?); eauto.
  all:repeat econstructor; auto.
Qed.
Global Opaque hrexec_final_sfv.


Symbolic execution of the whole block


Fixpoint hrexec_rec (m : se_mode) ib hrs (k: ristate -> ?? rstate): ?? rstate :=
  match ib with
  | BF fin iinfo =>
      DO sfv <~ hrexec_final_sfv fin hrs;;
      RET (Sfinal hrs sfv iinfo.(meminv_annot))
  | Bnop _ => k hrs
  | Bop op args dst iinfo =>
      DO lsv <~ @hrs_lr_get HCF hrs args;;
      DO ok <~ adp_op_okset op lsv iinfo;;
      DO hrs1 <~ hrs_add_okset m false hrs ok;;
      DO next <~ @hrs_rsv_set HCF HC RRULES m dst (Rop op) lsv (Some iinfo) hrs1.(ris_smem) hrs1;;
      k next
  | Bload trap chunk addr args dst iinfo =>
      DO lsv <~ @hrs_lr_get HCF hrs args;;
      DO next <~ @hrs_rsv_set HCF HC RRULES m dst
                    (Rload trap chunk addr iinfo.(addr_aval)) lsv (Some iinfo) hrs.(ris_smem) hrs;;
      k next
  | Bstore chunk addr args src iinfo =>
      DO next <~ hrexec_store m chunk addr args src iinfo hrs;;
      k next
  | Bseq ib1 ib2 =>
      hrexec_rec m ib1 hrs (fun hrs2 => hrexec_rec m ib2 hrs2 k)
  | Bcond cond args ifso ifnot iinfo =>
      DO lsv <~ @hrs_lr_get HCF hrs args;;
      DO prune <~ branch_pruning cond lsv;;
      match prune with
      | Some b =>
          if b then hrexec_rec m ifso hrs k
          else (
            if is_Some (is_promotable_cond iinfo.(inst_promotion)) then
              FAILWITH "hrexec_rec: promotable cond found while pruning (this should never happen)"
            else hrexec_rec m ifnot hrs k)
      | None =>
          DO ok <~ adp_cond_okset cond lsv iinfo;;
          DO hrs1 <~ hrs_add_okset m false hrs ok;;
          DO res <~ cbranch_expanse cond lsv iinfo;;
          let '(cond, vargs, pre) := res in
          DO hrs2 <~ hrs_add_okset m true hrs1 pre;;
          DO ifso <~ hrexec_rec m ifso hrs2 k;;
          DO ifnot <~ hrexec_rec m ifnot hrs2 k;;
          RET (Scond cond vargs ifso ifnot)
      end
  end.

Definition hrexec m ib hrinit :=
  hrexec_rec m ib hrinit (fun _ => RET error_rstate).


Lemma hrexec_rec_okpreserv m ctx hrsf rfv ib
  (ROK: ris_ok ctx hrsf):
  forall k
  (CONT: forall hrs,
    WHEN k hrs ~> rst THEN
    get_routcome ctx rst = sout hrsf rfv ->
    ris_ok ctx hrs)
  hrs,
  WHEN hrexec_rec m ib hrs k ~> rst THEN forall
  (ROUT: get_routcome ctx rst = sout hrsf rfv),
  ris_ok ctx hrs.
Proof.
  induction ib; simpl; wlp_simplify.
  - (* BF*)
    injection ROUT as ?; subst; auto.
  - (* Bop *)
    exploit @hrs_rsv_set_rel; eauto.
    exploit hrs_add_okset_rel; eauto.
    intros <- REL. apply REL; auto.
  - (* Bload *)
    exploit @hrs_rsv_set_rel; eauto using RRL_OK.
  - (* Bstore *)
    exploit hrexec_store_rel; eauto using RRL_OK.
  - (* Bcond *)
    apply hrs_add_okset_rel in Hexta2, Hexta4;
    destruct eval_scondition as [[|]|] in ROUT;
      [ rename H into HB | rename H0 into HB
      | injection ROUT as ?; subst; apply ris_error_not_ok in ROK as [ ]];
    apply Hexta4 in HB; eauto; apply Hexta2 in HB; eauto.
Qed.

Lemma add_pre_adp_cond_noprom iinfo cond lsv ctx
  (NOPROM: is_Some (is_promotable_cond (inst_promotion iinfo)) = false)
  (ALIVE: alive (eval_list_sval ctx lsv))
  : ` (adp_cond_okpred cond lsv iinfo) ctx.
Proof.
  unfold adp_cond_okpred, adp_cond.
  simpl. unfold ex_Some; autodestruct.
  intros _. destruct (is_promotable_cond _); try discriminate NOPROM.
  simpl; trivial.
Qed.

Lemma hrexec_rec_spec m ctx ib hrs sis:
  forall rk k
  (STRICT_R: forall hrs hrsf rfv,
    ris_ok ctx hrsf ->
    WHEN rk hrs ~> rst THEN
    get_routcome ctx rst = sout hrsf rfv ->
    ris_ok ctx hrs)
  (STRICT_S: forall sis sisf sfv,
    get_soutcome ctx (k sis) = sout sisf sfv ->
    sis_rel sis sisf)
  (CONT: forall hrs sis,
    ris_refines m ctx hrs sis ->
    WHEN rk hrs ~> rst THEN
    rst_refines m ctx rst (k sis))
  (RR: ris_refines m ctx hrs sis),
  WHEN hrexec_rec m ib hrs rk ~> rst THEN
  rst_refines m ctx rst (sexec_rec ib sis k).
Proof.
  pattern ib, hrs, sis. match goal with |- ?G ib hrs sis => set (Goal := G) end; revert ib hrs sis.
  assert (INTRO: forall ib hrs sis, (ris_refines_ok ctx hrs sis -> Goal ib hrs sis) -> Goal ib hrs sis). {
    (* We handle here the case when [ris_ok ctx hrs] does not hold using [ris_refines_step]. *)
    intros ib hrs sis C_OK rk k ? ? ? RR ? EXEC.
    apply (ris_refines_step_rst RR).
    + destruct get_routcome eqn:GET; simpl.
      intros; eapply hrexec_rec_okpreserv in EXEC; eauto.
    + destruct get_soutcome eqn:GET; simpl.
      apply sexec_rec_rel in GET; eauto.
    + intro. eapply C_OK; eauto.
  }

  induction ib;
      intros hrs sis; apply INTRO; clear INTRO;
      intros RR rk k ? ? ? _;
      simpl; try wlp_seq; auto.
  - (* BF *)
    intros sfv SFV.
    eapply hrexec_final_sfv_spec in SFV as (sfv' & -> & SIMU); eauto.
    apply wrst_ref; constructor; simpl; auto.
  - (* Bop *)
    intros lsv LSV ok OK hrs1 ADD next SET rst K; unfold sexec_op.
    eapply hrs_lr_get_spec_alive in LSV as (lsv' & -> & LSV_EQ & AARG); eauto.
    eapply CONT; eauto.
    (* add_okpred *)
    set (sis1 := add_pre _ sis).
    assert (RR1 : ris_refines m ctx hrs1 sis1). {
      eapply hrs_add_okset_pre_correct in ADD.
      refine (ris_refines_morph _ ADD).
      all:eauto; cycle 1.
      - apply adp_op_okset_correct; eauto.
      - apply add_pre_morph; simpl; rewrite ?LSV_EQ; reflexivity.
    }
    (* set_sreg *)
    apply (ris_refines_step RR1).
      { apply hrs_rsv_set_rel in SET. apply SET. }
      { apply set_sreg_rel. }
    clear RR1; intro RR1.
    eapply hrs_rsv_set_spec in SET as RR2; eauto.
      2:{ eapply ris_refines_ok_rmem, RR1. }
    eapply ris_refines_morph. 2:exact RR2.
    apply set_sreg_morph. reflexivity.
    simpl; rewrite LSV_EQ; reflexivity.
  - (* Bload *)
    intros lsv LSV next SET rst K; unfold sexec_load.
    eapply hrs_lr_get_spec_alive in LSV as (lsv' & -> & LSV_EQ & AARG); eauto.
    eapply CONT; eauto.
    eapply hrs_rsv_set_spec in SET as RR1; cycle 1; eauto.
      { eapply ris_refines_ok_rmem, RR. }
    eapply ris_refines_morph. 2:exact RR1.
    apply set_sreg_morph. reflexivity.
    apply root_apply_morph_equiv.
    + apply LSV_EQ.
    + apply RR.
  - (* Bstore *)
    intros next STORE rst K.
    eapply hrexec_store_spec in STORE as (sis1 & -> & RR1); eauto.
    eapply CONT; eauto.
  - (* Bseq *)
    eapply IHib1; eauto.
    + intros. apply hrexec_rec_okpreserv; auto.
    + intros sis1 ? ?; apply sexec_rec_rel; eauto.
    + intros; apply IHib2; auto.
  - (* Bcond *)
    intros lsv LSV b PRUNE; destruct b as [[ ]|];
    eapply hrs_lr_get_spec_alive in LSV as (lsv' & -> & LSV_EQ & AARG); eauto.
    + eapply branch_pruning_no_true in PRUNE; eauto.
      contradiction.
    + eapply branch_pruning_some_false in PRUNE; eauto.
      unfold rst_refines; subst; simpl.
      replace (eval_scondition _ _ lsv') with (eval_scondition ctx cond lsv)
        by (unfold eval_scondition; rewrite LSV_EQ; reflexivity).
      rewrite PRUNE; auto.
      wlp_simplify.
      eapply IHib2; eauto.
      apply wref; inv RR; constructor; auto.
      apply ok_add_pre; split; auto.
      unfold adp_cond_okpred; simpl.
      apply add_pre_adp_cond_noprom; auto.
      rewrite <- LSV_EQ; exact AARG.
    + wlp_seq. intros ok OK hrs1 ADD1 ((cond0, vargs), pre) BRANCH. wlp_seq.
      intros hrs2 ADD2 ifso0 RECso ifnot0 RECnot.
      set (sis1 := add_pre _ sis).
      (* add pre *)
      assert (RR1: ris_refines m ctx hrs1 sis1). {
        eapply hrs_add_okset_pre_correct in ADD1.
        refine (ris_refines_morph _ ADD1).
        all:eauto; cycle 1.
        - apply adp_cond_okset_correct; eauto.
        - apply add_pre_morph; simpl; rewrite ?LSV_EQ; reflexivity.
      }
      clear RR.
      (* add hyp *)
      apply hrs_add_okset_impl with (ctx := ctx) in ADD2 as PRE.
      assert (RR2 : ris_refines m ctx hrs2 sis1). {
        apply (ris_refines_step RR1).
          { apply hrs_add_okset_rel in ADD2; apply ADD2. }
          { reflexivity. }
        clear RR1; intro RR1.
        eapply hrs_add_okset_rhy_correct in ADD2; eauto.
      }
      (* condition *)
      apply (ris_refines_step_rst RR2).
        { simpl. case eval_scondition as [[|]|]; intro ROK;
          [ rename RECso into RECb | rename RECnot into RECb
          | apply ris_error_not_ok in ROK as [ ] ];
          destruct get_routcome eqn:GET in ROK;
          eapply hrexec_rec_okpreserv in RECb; eauto. }
        { simpl. case eval_scondition as [[|]|]. 3:apply error_sis_rel.
          all:destruct get_soutcome eqn:GET; eapply sexec_rec_rel; eauto. }
      clear RR1 RR2; intro RR2.
      eapply cbranch_expanse_spec in BRANCH; eauto.
      unfold rst_refines; simpl; rewrite BRANCH; auto.
        2:{ apply PRE; apply RR2. }
      replace (eval_scondition _ _ lsv') with (eval_scondition ctx cond lsv)
        by (unfold eval_scondition; rewrite LSV_EQ; reflexivity).
      case eval_scondition as [[|]|] in |-*.
      * eapply IHib1; eauto.
      * eapply IHib2; eauto.
      * apply rout_refines_error. apply RR2.
Qed.
Global Opaque hrexec_rec.

Theorem hrexec_spec m ctx ib hrs sis
  (RR: ris_refines m ctx hrs sis):
  WHEN hrexec m ib hrs ~> rst THEN
  rst_refines m ctx rst (sexec ib sis).
Proof.
  apply hrexec_rec_spec; auto; simpl; intros; try wlp_seq.
  - injection 1 as ?; subst.
    apply ris_error_not_ok in H as [ ].
  - injection H as ?; subst.
    apply error_sis_rel.
  - apply (ris_refines_step_rst H).
    + intro OK. apply ris_error_not_ok in OK as [ ].
    + apply error_sis_rel.
    + intro RR1. apply rout_refines_error. apply RR1.
Qed.
Global Opaque hrexec.

End SymbolicExecution.