Library mcertikos.devdrivers.HandlerAsmGen


This file provide the contextual refinement proof between PThreadInit layer and PQueueIntro layer
Require Import Coqlib.
Require Import Errors.
Require Import AST.
Require Import Integers.
Require Import Floats.
Require Import Op.
Require Import Asm.
Require Import Events.
Require Import Globalenvs.
Require Import Smallstep.
Require Import Values.
Require Import Memory.
Require Import Maps.
Require Import CommonTactic.
Require Import AuxLemma.
Require Import FlatMemory.
Require Import AuxStateDataType.
Require Import Constant.
Require Import GlobIdent.
Require Import RealParams.
Require Import LoadStoreSem1.
Require Import AsmImplLemma.
Require Import GenSem.
Require Import RefinementTactic.
Require Import DevicePrimSemantics.
Require Import XOmega.

Require Import liblayers.logic.PTreeModules.
Require Import liblayers.logic.LayerLogicImpl.
Require Import liblayers.compcertx.Stencil.
Require Import liblayers.compcertx.MakeProgram.
Require Import liblayers.compat.CompatLayers.
Require Import liblayers.compat.CompatGenSem.
Require Import compcert.cfrontend.Ctypes.
Require Import LayerCalculusLemma.
Require Import AbstractDataType.

Require Import DHandlerAsm.
Require Import HandlerAsmGenSpec.
Require Import DeviceStateDataType.
Require Import ObjX86.
Require Import ListLemma.
Require Import FutureTactic.
Require Import ObjX86Proof.

Definition of the refinement relation

Section Refinement.

  Local Open Scope string_scope.
  Local Open Scope error_monad_scope.
  Local Open Scope Z_scope.

  Context `{real_params: RealParams}.
  Context `{oracle_prop: MultiOracleProp}.

  Notation HDATA := RData.
  Notation LDATA := RData.

  Notation HDATAOps := (cdata (cdata_ops := dhandlerintro_data_ops) HDATA).
  Notation LDATAOps := (cdata (cdata_ops := dhandlerintro_data_ops) LDATA).

  Section WITHMEM.

    Context `{Hstencil: Stencil}.
    Context `{Hmem: Mem.MemoryModelX}.
    Context `{Hmwd: UseMemWithData mem}.

Definition the refinement relation: relate_RData + match_RData

    Section REFINEMENT_REL.

      Inductive match_set_tf: stencilvalmemmeminjProp :=
      | MATCH_SET_TF:
           v m b f s,
            Mem.load Mint32 m b 0 = Some (Vint v)
            → Mem.valid_access m Mint32 b 0 Writable
            → find_symbol s set_tf = Some b
            → match_set_tf s (Vint v) m f.

Relation between the new raw data at the higher layer with the mememory at lower layer
      Inductive match_RData: stencilHDATAmemmeminjProp :=
      | MATCH_RDATA:
           hadt m f s,
            match_set_tf s (tf_reg hadt) m f
            → match_RData s hadt m f.

Relation between raw data at two layers
      Record relate_RData (f: meminj) (hadt: HDATA) (ladt: LDATA) :=
        mkrelate_RData {
            flatmem_re: FlatMem.flatmem_inj (HP hadt) (HP ladt);
            vmxinfo_re: vmxinfo hadt = vmxinfo ladt;
            CR3_re: CR3 hadt = CR3 ladt;
            ikern_re: ikern hadt = ikern ladt;
            MM_re: MM hadt = MM ladt;
            MMSize_re: MMSize hadt = MMSize ladt;
            pg_re: pg hadt = pg ladt;
            ihost_re: ihost hadt = ihost ladt;
            AC_re: AC hadt = AC ladt;
            ti_fst_re: (fst (ti hadt)) = (fst (ti ladt));
            ti_snd_re: val_inject f (snd (ti hadt)) (snd (ti ladt));
            AT_re: AT hadt = AT ladt;
            nps_re: nps hadt = nps ladt;
            init_re: init hadt = init ladt;

            buffer_re: buffer hadt = buffer ladt;

            com1_re: com1 hadt = com1 ladt;
            console_re: console hadt = console ladt;
            console_concrete_re: console_concrete hadt = console_concrete ladt;
            ioapic_re: ioapic ladt = ioapic hadt;
            lapic_re: lapic ladt = lapic hadt;
            intr_flag_re: intr_flag ladt = intr_flag hadt;
            curr_intr_num_re: curr_intr_num ladt = curr_intr_num hadt;
            drv_serial_re: drv_serial hadt = drv_serial ladt;
            in_intr_re: in_intr hadt = in_intr ladt;
            tf_re: tfs_inj f (tf hadt) (tf ladt);
            tf_reg_re: tf_reg hadt = tf_reg ladt;

            CPU_ID_re: CPU_ID hadt = CPU_ID ladt;
            cid_re: cid hadt = cid ladt;
            multi_oracle_re: multi_oracle hadt = multi_oracle ladt;
            multi_log_re: multi_log hadt = multi_log ladt;
            lock_re: lock hadt = lock ladt
          }.

      Local Hint Resolve MATCH_RDATA.

      Global Instance rel_ops: CompatRelOps HDATAOps LDATAOps :=
        {
          relate_AbData s f d1 d2 := relate_RData f d1 d2;
          match_AbData s d1 m f := match_RData s d1 m f;
          new_glbl := set_tf :: nil
        }.

    End REFINEMENT_REL.

Properties of relations

    Section Rel_Property.

      Lemma inject_match_correct:
         s d1 m2 f m2´ j,
          match_RData s d1 m2 f
          Mem.inject j m2 m2´
          inject_incr (Mem.flat_inj (genv_next s)) j
          match_RData s d1 m2´ (compose_meminj f j).
      Proof.
        inversion 1; subst; intros.
        inv H0.
        assert (HFB0: j b = Some (b, 0)).
        {
          eapply stencil_find_symbol_inject´; eauto.
        }
        econstructor; eauto; intros.
        specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H1 H4 HFB0).
        repeat rewrite Z.add_0_r.
        intros [v1´[HLD1´ HV1´]].
        inversion HV1´; subst.
        rewrite <- H3.
        econstructor; eauto; intros.
        specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H1 H5).
        rewrite Z.add_0_r; trivial.
      Qed.

      Lemma store_match_correct:
         s abd m0 m0´ f b2 v chunk,
          match_RData s abd m0 f
          ( i b,
             In i new_glbl
             find_symbol s i = Some bb b2) →
          Mem.store chunk m0 b2 v = Some m0´
          match_RData s abd m0´ f.
      Proof.
        intros. inv H. inv H2.
        econstructor; eauto.
        rewrite <- H.
        econstructor; eauto.
        eapply H0 in H6; simpl; eauto.
        rewrite (Mem.load_store_other _ _ _ _ _ _ H1); auto.
        eapply Mem.store_valid_access_1; eauto.
      Qed.

      Lemma storebytes_match_correct:
         s abd m0 m0´ f b2 v ,
          match_RData s abd m0 f
          ( i b,
             In i new_glbl
             find_symbol s i = Some bb b2) →
          Mem.storebytes m0 b2 v = Some m0´
          match_RData s abd m0´ f.
      Proof.
        intros. inv H. inv H2.
        econstructor; eauto.
        rewrite <- H.
        econstructor; eauto.
        eapply H0 in H6; simpl; eauto.
        repeat rewrite (Mem.load_storebytes_other _ _ _ _ _ H1); eauto.
        eapply Mem.storebytes_valid_access_1; eauto.
      Qed.

      Lemma free_match_correct:
         s abd m0 m0´ f ofs sz b2,
          match_RData s abd m0 f
          ( i b,
             In i new_glbl
             find_symbol s i = Some bb b2) →
          Mem.free m0 b2 ofs sz = Some m0´
          match_RData s abd m0´ f.
      Proof.
        intros; inv H; inv H2.
        econstructor; eauto.
        rewrite <- H.
        econstructor; eauto.
        eapply H0 in H6; simpl; eauto.
        repeat rewrite (Mem.load_free _ _ _ _ _ H1); auto.
        eapply H0 in H6; simpl; eauto.
        eapply Mem.valid_access_free_1; eauto.
      Qed.

      Lemma alloc_match_correct:
         s abd m´0 m´1 f ofs sz b0 b´1,
          match_RData s abd m´0 f
          Mem.alloc m´0 ofs sz = (m´1, b´1)
           b0 = Some (b´1, 0%Z)
          ( b : block, b b0 b = f b) →
          inject_incr f
          ( i b,
             In i new_glbl
             find_symbol s i = Some bb b0) →
          match_RData s abd m´1 .
      Proof.
        intros. rename H1 into HF1, H2 into HB. inv H; inv H1.
        econstructor; eauto.
        rewrite <- H.
        econstructor; eauto;
        try (apply (Mem.load_alloc_other _ _ _ _ _ H0));
        try (eapply Mem.valid_access_alloc_other); eauto.
      Qed.

Prove that after taking one step, the refinement relation still holds
      Lemma relate_incr:
         abd abd´ f ,
          relate_RData f abd abd´
          → inject_incr f
          → relate_RData abd abd´.
      Proof.
        inversion 1; subst; intros; inv H; constructor; eauto.
        eapply tfs_inj_incr; eauto.
      Qed.

      Global Instance rel_prf: CompatRel HDATAOps LDATAOps.
      Proof.
        constructor.
        - apply inject_match_correct.
        - apply store_match_correct.
        - apply alloc_match_correct.
        - apply free_match_correct.
        - apply storebytes_match_correct.
        - intros. eapply relate_incr; eauto.
      Qed.

    End Rel_Property.

Proofs the one-step forward simulations for the low level specifications

    Section OneStep_Forward_Relation.

      Ltac pattern2_refinement_simpl:=
        pattern2_refinement_simpl´ (@relate_AbData).

      Section FRESH_PRIM.

        Lemma serial_intr_handler_asm_exist:
           habd habd´ labd rs rs´ rs´´ s f,
            serial_intr_handler_asm_spec rs habd = Some (habd´, rs´)
            → relate_AbData s f habd labd
            → high_level_invariant habd
            → ( r, val_inject f (Pregmap.get r rs) (Pregmap.get r rs´´))
            → labd´ rs´´´, serial_intr_handler_asm_spec rs´´ labd = Some (labd´, rs´´´)
                                 ( r, val_inject f (Pregmap.get r rs´) (Pregmap.get r rs´´´))
                                  relate_AbData s f habd´ labd´.
        Proof.
          unfold serial_intr_handler_asm_spec; intros until f.
          intros HP HR HINV HINJ; pose proof HR as HR´; inv HR; revert HP.
          simpl; subrewrite´. intros HQ.
          subdestruct. inv HQ.

          exploit (lapic_eoi_exist s _ _ (((save_context_spec rs´´ labd) { ioapic
                                                                                : (ioapic habd) { s / IoApicEnables [4]:false}}) { intr_flag
                                                                                                                                      : true}) f Hdestruct3); eauto.
          apply relate_impl_intr_flag_update; eauto.
          apply relate_impl_ioapic_update; eauto.
          unfold save_context_spec.
          rewrite curr_intr_num_re0.
          apply relate_impl_tf_update; eauto.
          eapply tfs_inj_rev; eauto.
          constructor; simpl; try split; eauto.

          intros (labd´ & HP´ & HM´).
          rewrite HP´.
          exploit (cons_intr_aux_exist s _ _ labd´ f Hdestruct4); eauto.
          intros (labd´´ & HP´´ & HM´´).
          rewrite HP´´.
          inversion HM´.
          inversion HM´´.
          subst.
          functional inversion Hdestruct5; subst.
          unfold restore_context_spec.
          simpl in ×.
          rewrite H0 in tf_re2.
          simpl in tf_re2.
          subdestruct.
          esplit. esplit.
          split.
          reflexivity.
          generalize relate_impl_ioapic_update; intro.
          simpl in H.
          rewrite ioapic_re2.
          split.
          intros.
          unfold Pregmap.get.
          change (Pregmap.init Vundef r1) with Vundef.
          constructor.
          eapply H; eauto.
          inv tf_re2.

          unfold restore_context_spec.
          simpl in H0.
          rewrite H0 in tf_re2.
          simpl in tf_re2.
          subdestruct.
          inv tf_re2.
          rewrite H0.
          rewrite Hdestruct6.
          esplit. esplit.
          split.
          reflexivity.
          split.
          destruct tf_re2.
          generalize H, H1; repeat match goal with | [H : _ |- _] ⇒ clear H end.
          generalize dependent t.
          generalize dependent t0.
          generalize dependent _x.
          induction _x0.
          intros.
          simpl in H1.
          subdestruct.
          unfold tf_inj in H.
          specialize (H r).
          destruct H; auto.
          inv H1.
          intros.
          simpl in H1.
          destruct t0.
          inv H1.
          change (last (_x :: a :: _x0)) with (last (a :: _x0)).
          change (last (t :: t0 :: t1)) with (last (t0 :: t1)).
          destruct H1.
          eapply IH_x0; eauto.

          generalize relate_impl_tf_update; intro.
          simpl in H.
          eapply H; eauto.
          generalize relate_impl_ioapic_update; intro.
          simpl in H1.
          rewrite ioapic_re2.
          eapply H1; eauto.
          eapply tfs_inj_removelast.
          destruct tf_re2.
          econstructor; eauto.
          intro contra; inv contra.
        Qed.

        Lemma lapic_eoi_spec_preserves_tf:
           d ,
            lapic_eoi_spec d = Some
            tf d = tf .
        Proof.
          intros.
          unfold lapic_eoi_spec.
          functional inversion H; simpl in *; subst; simpl in ×.
          reflexivity.
        Qed.

        Lemma cons_intr_aux_preserves_tf:
           d ,
            cons_intr_aux d = Some
            tf d = tf .
        Proof.
          intros.
          unfold cons_intr_aux.
          functional inversion H; simpl in *; subst; simpl in *;
          reflexivity.
        Qed.

        Lemma lapic_eoi_spec_preserves_curr_intr_num:
           d ,
            lapic_eoi_spec d = Some
            curr_intr_num d = curr_intr_num .
        Proof.
          intros.
          unfold lapic_eoi_spec.
          functional inversion H; simpl in *; subst; simpl in ×.
          reflexivity.
        Qed.

        Lemma cons_intr_aux_preserves_curr_intr_num:
           d ,
            cons_intr_aux d = Some
            curr_intr_num d = curr_intr_num .
        Proof.
          intros.
          unfold cons_intr_aux.
          functional inversion H; simpl in *; subst; simpl in *;
          reflexivity.
        Qed.

        Lemma lapic_eoi_spec_independent_tf:
           b d ,
            lapic_eoi_spec d = Some
            lapic_eoi_spec d {tf: b} = Some ( {tf: b}).
        Proof.
          intros.
          unfold lapic_eoi_spec.
          functional inversion H; simpl in *; subst; simpl in ×.
          rewrite H1, H2, H3, H4;
            reflexivity.
        Qed.

        Lemma cons_intr_aux_independent_tf:
           b d ,
            cons_intr_aux d = Some
            cons_intr_aux d {tf: b} = Some ( {tf: b}).
        Proof.
          intros.
          unfold cons_intr_aux.
          functional inversion H; simpl in *; subst; simpl in *;
          rewrite H1, H2, H3, H4, H5, H6, H7, H8, H9;
          reflexivity.
        Qed.

        Lemma serial_intr_handler_asm_spec_inv´:
           rs rs´ d ,
            serial_intr_handler_asm_spec rs d = Some (, rs´)
            (ikern d = true
             ihost d = true
             init d = true
             in_intr d = true
              d1 d2 d3 d4 d5,
               save_context_spec rs d = d1
               d1 {ioapic / s / IoApicEnables [4%nat]: false} {intr_flag: true} = d2
               lapic_eoi_spec d2 = Some d3
               cons_intr_aux d3 = Some d4
               d4 {ioapic / s / IoApicEnables [4%nat]: true} = d5
               restore_context_spec d5 = (rs´, )).
        Proof.
          intros rs rs´ d .
          unfold serial_intr_handler_asm_spec.
          intros H.
          split; [subdestruct; auto|].
          split; [subdestruct; auto|].
          split; [subdestruct; auto|].
          split; [subdestruct; auto|].
          subdestruct.
          repeat esplit.
          eassumption.
          eassumption.
          inv H.
          assumption.
        Qed.

        Lemma serial_intr_handler_asm_spec_inv´´:
           rs rs´ d ,
            serial_intr_handler_asm_spec rs d = Some (, rs´)
            (ikern d = true
             ihost d = true
             init d = true
             in_intr d = true
              d1 d2,
               save_context_spec rs d = d1
               serial_intr_handler_sw_spec d1 = Some d2
               restore_context_spec d2 = (rs´, )).
        Proof.
          intros rs rs´ d Hspec.
          apply serial_intr_handler_asm_spec_inv´ in Hspec.
          blast Hspec.
          split; [assumption|].
          split; [assumption|].
          split; [assumption|].
          split; [assumption|].
           e. e3.
          split; [assumption|].
          split; [|reflexivity].
          unfold serial_intr_handler_sw_spec.
          if_true. if_true. if_true. if_true.
          rewrite <- Hex2 in Hex1. change (Z.to_nat 4) with 4%nat. rewrite Hex1.
          rewrite Hex3. rewrite Hex4. reflexivity.
          functional inversion Hex0; simpl; trivial.
          functional inversion Hex0; simpl; trivial.
          functional inversion Hex0; simpl; trivial.
          functional inversion Hex0; simpl; trivial.
        Qed.

        Lemma save_restore_basic_transparent´:
           (rs rs´: regset) (d d1 d2 : RData)
            (spec: RDataoption (RData)),
            ( dx dy, spec dx = Some (dy) → tf dy = tf dx) →
            ( dx dy x, spec dx = Some (dy) → spec (dx {tf: x}) = Some (dy {tf: x})) →
            save_context_spec rs d = d1
            spec d1 = Some (d2) →
            restore_context_spec d2 = (rs´, )
            spec d = Some ().
        Proof.
          intros rs rs´ d d1 d2 spec Hisotf Hnonint Hsave Hspec Hres.
          functional inversion Hsave. functional inversion Hres; subst.
          eapply Hisotf in Hspec.
          simpl in Hspec.
          rewrite H1 in Hspec.
          symmetry in Hspec.
          eapply app_eq_nil in Hspec.
          destruct Hspec.
          inv H0.
          generalize (Hisotf _ _ Hspec). intros Htf21.
          generalize (Hnonint _ _ (removelast (tf d2)) Hspec).
          intros Hd12.
          rewrite <- Hd12.
          rewrite Htf21.
          simpl.
          rewrite removelast_app.
          simpl.
          rewrite app_nil_r.
          destruct d; simpl; reflexivity.
          intro contra; inv contra.
        Qed.

        Lemma serial_intr_handler_asm_to_sw:
           d rs rs´,
            serial_intr_handler_asm_spec rs d = Some (, rs´)
            serial_intr_handler_sw_spec d = Some .
        Proof.
          intros d rs rs´ Hspec.
          apply serial_intr_handler_asm_spec_inv´´ in Hspec.
          blast Hspec. clear H0 H1.
          eapply save_restore_basic_transparent´.
          apply serial_intr_handler_sw_spec_does_not_change_tf.
          apply serial_intr_handler_sw_spec_independent_from_tf.
          eapply Hex1.
          eapply Hex0.
          eapply Hex3.
        Qed.

        Lemma serial_intr_handler_asm_rs_independent:
           d1 d2 rs1 rs2 rs3,
            serial_intr_handler_asm_spec rs1 d1 = Some (d2, rs2)
            serial_intr_handler_asm_spec rs3 d1 = Some (d2, rs3).
        Proof.
          intros d1 d2 rs1 rs2 rs3 Hspec.
          apply serial_intr_handler_asm_to_sw in Hspec.
          apply serial_intr_handler_sw_to_asm. assumption.
        Qed.

        Lemma restore_context_preserves_tf_reg:
           d1 d2 rs,
            restore_context_spec d1 = (rs, d2)
            tf_reg d1 = tf_reg d2.
        Proof.
          intros.
          functional inversion H; reflexivity.
        Qed.

        Lemma lapic_eoi_spec_preserves_tf_reg:
           d ,
            lapic_eoi_spec d = Some
            tf_reg d = tf_reg .
        Proof.
          intros.
          unfold lapic_eoi_spec.
          functional inversion H; simpl in *; subst; simpl in ×.
          reflexivity.
        Qed.

        Lemma cons_intr_aux_preserves_tf_reg:
           d ,
            cons_intr_aux d = Some
            tf_reg d = tf_reg .
        Proof.
          intros.
          unfold cons_intr_aux.
          functional inversion H; simpl in *; subst; simpl in *;
          reflexivity.
        Qed.


        Lemma serial_intr_handler_asm_spec_ref:
          compatsim (crel HDATA LDATA)
                    (primcall_serial_intr_handler_asm_compatsem serial_intr_handler_asm_spec)
                    serial_intr_handler_asm_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          intros.
          inv H4.
          inv match_extcall_states.
          exploit serial_intr_handler_asm_exist; eauto 1.
          intros [labd´ [r´0[HP [HM HReg]]]].
          generalize (serial_intr_handler_asm_rs_independent _ _ _ _ rs2 HP); intro.
          generalize match_match; intro match´.
          inv match_match.
          inv H1.
          specialize (Mem.valid_access_store _ _ _ _ (rs2 EDX) H4); intros [ HST].
          assert (Hstore´: m´´, Mem.store Mint32 b0 0 (Vint Int.zero) = Some m´´).
          {
            apply (Mem.valid_access_store); auto.
            intros; eapply Mem.store_valid_access_1; eauto.
          }
          destruct Hstore´.
          refine_split; try econstructor; eauto.
          eapply reg_symbol_inject; eassumption.
          eapply Mem.load_store_same in HST; eauto.
          inv ASM_INV.
          unfold AsmX.wt_regset in inv_reg_wt.
          generalize (inv_reg_wt EDX); intro.
          simpl in H1.
          simpl in HST.
          destruct (rs2 EDX); simpl in *; try assumption; try inversion H5.
          econstructor; eauto.

          inv Hhigh´.
          rewrite valid_tf_reg in H2.
          inv H2.
          pattern2_refinement_simpl.
          constructor; eauto.
          unfold serial_intr_handler_asm_spec in H0.
          subdestruct.
          inv H0.
          eapply restore_context_preserves_tf_reg in Hdestruct5.
          eapply cons_intr_aux_preserves_tf_reg in Hdestruct4.
          eapply lapic_eoi_spec_preserves_tf_reg in Hdestruct3.
          simpl in ×.
          rewrite <- Hdestruct5, <- Hdestruct4, <- Hdestruct3.
          rewrite <- H2.
          econstructor; eauto.
          eapply Mem.load_store_same in H1; eauto.
          simpl in H1.
          inv Hhigh´.
          rewrite H2, valid_tf_reg.
          assumption.
          eapply Mem.store_valid_access_1; eauto.
          eapply Mem.store_valid_access_1; eauto.
          rewrite (Mem.nextblock_store _ _ _ _ _ _ H1).
          rewrite (Mem.nextblock_store _ _ _ _ _ _ HST).
          assumption.
          val_inject_simpl; eapply HReg;
          apply PregToZ_correct; reflexivity.
        Qed.

      End FRESH_PRIM.

      Section PASSTHROUGH_PRIM.

        Global Instance: (LoadStoreProp (hflatmem_store:= flatmem_store´) (lflatmem_store:= flatmem_store´)).
        Proof.
          accessor_prop_tac.
          - eapply flatmem_store´_exists; eauto.
        Qed.

        Lemma passthrough_correct:
          sim (crel HDATA LDATA) dhandlerasm_passthrough dhandlersw.
        Proof.
          sim_oplus.
          - apply fload´_sim.
          - apply fstore´_sim.
          - apply page_copy´_sim.
          - apply page_copy_back´_sim.

          - apply vmxinfo_get_sim.
          - apply setPG_sim.
          - apply setCR3_sim.
          - apply get_size_sim.
          - apply is_mm_usable_sim.
          - apply get_mm_s_sim.
          - apply get_mm_l_sim.
          - apply get_CPU_ID_sim.
          - apply get_curid_sim.
          - apply set_curid_sim.
          - apply set_curid_init_sim.

          - apply (release_lock_sim (valid_arg_imply:= Shared2ID0_imply)).
          -
            eapply acquire_lock_sim0; eauto.
            intros. inv H; trivial; try inv H0.
          - apply ticket_lock_init0_sim.
          - apply serial_irq_check_sim.
          - apply iret_sim.
          - apply cli_sim.
          - apply sti_sim.
          - apply serial_irq_current_sim.
          - apply ic_intr_sim.
          - apply ioapic_mask_sim.
          - apply ioapic_unmask_sim.
          - apply serial_putc_sim.
          - apply serial_hw_intr_sim.
          - apply cons_buf_read_sim.
          - apply trapin_sim.
          - apply trapout´_sim.
          - apply hostin_sim.
          - apply hostout´_sim.
          - apply proc_create_postinit_sim.
          - apply trap_info_get_sim.
          - apply trap_info_ret_sim.
          - layer_sim_simpl.
            + eapply load_correct1.
            + eapply store_correct1.
        Qed.

      End PASSTHROUGH_PRIM.

    End OneStep_Forward_Relation.

  End WITHMEM.

End Refinement.