Library mcertikos.devdrivers.ConsoleBuffIntroGen


This file provide the contextual refinement proof between MPTInit layer and MPTBit 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 AsmImplLemma.
Require Import GenSem.
Require Import LoadStoreSem1.
Require Import RefinementTactic.
Require Import PrimSemantics.
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 AbstractDataType.
Require Import DeviceStateDataType.
Require Import CLemmas.

Require Import ConsoleBuffIntroGenSpec.
Require Import DConsoleBuffIntro.
Require Import LayerCalculusLemma.

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 := mhticketlockop_data_ops) HDATA).
  Notation LDATAOps := (cdata (cdata_ops := mhticketlockop_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.

Relation between each entry of the console buffer and the underline memory
        Inductive match_Console_Buff: ConsoleBufValvalProp :=
        | MATCH_CONSBUFFUNDEF:
             v, match_Console_Buff BufValUndef v
        | MATCH_CONSBUFF:
             bv v,
              bv = BufVal v
              → match_Console_Buff bv (Vint (Int.repr v)).

        Inductive match_Console_Buff_rpos: ZvalProp :=
        | MATCH_RPOS:
             rpos v,
              rpos = v
              → 0 v < CONSOLE_BUFFER_SIZE
              → match_Console_Buff_rpos rpos (Vint (Int.repr v)).

        Inductive match_Console_Buff_wpos: ZvalProp :=
        | MATCH_WPOS:
             wpos v,
              wpos = v
              → 0 v < CONSOLE_BUFFER_SIZE
              → match_Console_Buff_wpos wpos (Vint (Int.repr v)).

Relation between console buffer and the underline memory
        Inductive match_Console: stencilConsoleDriverConcretememmeminjProp :=
        | MATCH_CONSOLE_BUFF: console m b f s,
                       (( ofs, 0 ofs < CONSOLE_BUFFER_SIZE
                                    ( v, Mem.load Mint32 m b (ofs×4) = Some v
                                               Mem.valid_access m Mint32 b (ofs×4) Writable
                                               match_Console_Buff (ZMap.get ofs (cons_buf_concrete console)) v))
                                      
                                    ( v, Mem.load Mint32 m b (CONSOLE_BUFFER_SIZE×4) = Some v
                                               Mem.valid_access m Mint32 b (CONSOLE_BUFFER_SIZE×4) Writable
                                               match_Console_Buff_rpos (rpos_concrete console) v)
                                      
                                    ( v, Mem.load Mint32 m b ((CONSOLE_BUFFER_SIZE+1)*4) = Some v
                                               Mem.valid_access m Mint32 b ((CONSOLE_BUFFER_SIZE+1)*4) Writable
                                               match_Console_Buff_wpos (wpos_concrete console) v)
                       )
                       → find_symbol s cons_buf_LOC = Some b
                       → match_Console s console 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_Console s (console_concrete 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;
              MM_re: MM hadt = MM ladt;
              MMSize_re: MMSize hadt = MMSize ladt;
              CR3_re: CR3 hadt = CR3 ladt;
              ikern_re: ikern ladt = ikern hadt;
              pg_re: pg ladt = pg hadt;
              ihost_re: ihost ladt = ihost hadt;
              AC_re: AC ladt = AC hadt;
              ti_fst_re: (fst (ti ladt)) = (fst (ti hadt));
              ti_snd_re: val_inject f (snd (ti hadt)) (snd (ti ladt));
              LAT_re: LAT ladt = LAT hadt;
              nps_re: nps ladt = nps hadt;
              PT_re: PT ladt = PT hadt;
              ipt_re: ipt ladt = ipt hadt;
              init_re: init ladt = init hadt;

              buffer_re: buffer hadt = buffer ladt;

              com1_re: com1 ladt = com1 hadt;
              ioapic_re: ioapic ladt = ioapic hadt;
              lapic_re: lapic ladt = lapic hadt;
              intr_flag_re: intr_flag ladt = intr_flag hadt;
              saved_intr_flags_re: saved_intr_flags ladt = saved_intr_flags hadt;
              curr_intr_num_re: curr_intr_num ladt = curr_intr_num hadt;
              in_intr_re: in_intr hadt = in_intr ladt;
              tf_re: tfs_inj f (tf hadt) (tf 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
            }.

        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 := cons_buf_LOC :: 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.
        econstructor; eauto; intros.
        destruct H3.
        destruct H3.
        split.
        {
          intros.
          specialize (H0 _ H6).
          destruct H0 as [v[HL[HV HM]]].
          specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H1 HL HFB0).
          repeat rewrite Z.add_0_r; intros [[HLD´ HV´]].
          refine_split´; eauto.
          specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H1 HV).
          rewrite Z.add_0_r; trivial.
          inv HM. constructor.
          inv HV´. constructor; auto.
        }
        split.
        {
          destruct H3 as [v[HL[HV HM]]].
          specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H1 HL HFB0).
          repeat rewrite Z.add_0_r; intros [[HLD´ HV´]].
          refine_split´; eauto.
          specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H1 HV).
          rewrite Z.add_0_r; trivial.
          inv HM.
          inv HV´. constructor; auto.
        }
        {
          destruct H5 as [v[HL[HV HM]]].
          specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H1 HL HFB0).
          repeat rewrite Z.add_0_r; intros [[HLD´ HV´]].
          refine_split´; eauto.
          specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H1 HV).
          rewrite Z.add_0_r; trivial.
          inv HM.
          inv HV´. constructor; auto.
        }
      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.
        econstructor; eauto.
        intros.
        destruct H.
        destruct H2.
        split.
        {
          intros.
          specialize (H _ H5).
          destruct H as [v0[HL[HV HM]]].
          eapply H0 in H3; simpl; eauto.
          repeat rewrite (Mem.load_store_other _ _ _ _ _ _ H1); auto.
          refine_split´; eauto.
          eapply Mem.store_valid_access_1; eauto.
        }
        split.
        {
          destruct H2 as [v0[HL[HV HM]]].
          eapply H0 in H3; simpl; eauto.
          repeat rewrite (Mem.load_store_other _ _ _ _ _ _ H1); auto.
          refine_split´; eauto.
          eapply Mem.store_valid_access_1; eauto.
        }
        {
          destruct H4 as [v0[HL[HV HM]]].
          eapply H0 in H3; simpl; eauto.
          repeat rewrite (Mem.load_store_other _ _ _ _ _ _ H1); auto.
          refine_split´; eauto.
          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.
        econstructor; eauto.
        destruct H.
        destruct H2.
        split.
        {
          intros.
          specialize (H _ H5).
          destruct H as [v0[HL[HV HM]]].
          eapply H0 in H3; simpl; eauto.
          repeat rewrite (Mem.load_storebytes_other _ _ _ _ _ H1); eauto.
          refine_split´; eauto.
          eapply Mem.storebytes_valid_access_1; eauto.
        }
        split.
        {
          destruct H2 as [v0[HL[HV HM]]].
          eapply H0 in H3; simpl; eauto.
          repeat rewrite (Mem.load_storebytes_other _ _ _ _ _ H1); eauto.
          refine_split´; eauto.
          eapply Mem.storebytes_valid_access_1; eauto.
        }
        {
          destruct H4 as [v0[HL[HV HM]]].
          eapply H0 in H3; simpl; eauto.
          repeat rewrite (Mem.load_storebytes_other _ _ _ _ _ H1); eauto.
          refine_split´; 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.
        econstructor; eauto.
        destruct H.
        destruct H2.
        split.
        {
          intros.
          specialize (H _ H5).
          destruct H as [v1[HL1 [HV1 HM]]].
          eapply H0 in H3; simpl; eauto.
          repeat rewrite (Mem.load_free _ _ _ _ _ H1); auto.
          refine_split´; eauto.
          eapply Mem.valid_access_free_1; eauto.
        }
        split.
        {
          destruct H2 as [v1[HL1 [HV1 HM]]].
          simpl in ×.
          eapply H0 in H3; simpl; eauto.
          repeat rewrite (Mem.load_free _ _ _ _ _ H1); auto.
          refine_split´; eauto.
          eapply Mem.valid_access_free_1; eauto.
        }
        {
          destruct H4 as [v1[HL1 [HV1 HM]]].
          simpl in ×.
          eapply H0 in H3; simpl; eauto.
          repeat rewrite (Mem.load_free _ _ _ _ _ H1); auto.
          refine_split´; 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.
        econstructor; eauto.
        destruct H.
        destruct H1.
        split.
        {
          intros.
          specialize (H _ H6).
          destruct H as [v1[HL1 [HV1 HM]]].
          refine_split´; eauto.
          apply (Mem.load_alloc_other _ _ _ _ _ H0); auto.
          eapply Mem.valid_access_alloc_other; eauto.
        }
        split.
        {
          destruct H1 as [v1[HL1 [HV1 HM]]].
          refine_split´; eauto.
          apply (Mem.load_alloc_other _ _ _ _ _ H0); auto.
          eapply Mem.valid_access_alloc_other; eauto.
        }
        {
          destruct H5 as [v1[HL1 [HV1 HM]]].
          refine_split´; eauto.
          apply (Mem.load_alloc_other _ _ _ _ _ H0); auto.
          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 cons_buf_init_concrete_spec_ref:
          compatsim (crel HDATA LDATA) (gensem cons_buf_init_concrete_spec) cons_buf_init_concrete_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          assert (Hkern: kernel_mode d2).
          {
            inv match_related. functional inversion H1; subst.
            repeat (split; eauto); try congruence.
          }
          destruct Hkern as [Hkern HOS].
          inv H.
          destruct H0 as [HMB H0].
          destruct H0 as [HRPOS HWPOS].
          destruct HRPOS as [vr[HLR[HVR HMR]]].
          destruct HWPOS as [vw[HLW[HVW HMW]]].
          specialize (Mem.valid_access_store _ _ _ _ (Vint (Int.zero)) HVR); intros [m´R HSTR].
          assert(Mem.valid_access m´R Mint32 b ((512 + 1) × 4) Writable).
          {
              eapply Mem.store_valid_access_1; eauto.
          }
          specialize (Mem.valid_access_store _ _ _ _ (Vint (Int.zero)) H); intros [m´W HSTW].
          refine_split.
          - econstructor; eauto.
            lift_unfold. split; eauto.
            instantiate (1:= (m´R, d2)).
            eauto.
            eauto.
            lift_unfold. split; eauto.
            simpl.
            eauto.
          - constructor.
          - pose proof H1 as Hspec.
            functional inversion Hspec; subst.
            split; eauto; pattern2_refinement_simpl.
            econstructor; simpl; eauto.
            econstructor; eauto; intros.
            split.
            {
              intros.
              specialize (HMB _ H0).
              destruct HMB as [v1[HL1 [HV1 HM]]].
              refine_split´; eauto.
              assert(Mem.load Mint32 m´W b (ofs × 4) = Mem.load Mint32 m´R b (ofs × 4)).
              {
                eapply Mem.load_store_other; eauto.
                right.
                left.
                simpl.
                omega.
              }
              rewrite H6.
              rewrite <- HL1.
              eapply Mem.load_store_other; eauto.
              right.
              left.
              simpl.
              omega.
              eapply Mem.store_valid_access_1; eauto.
              eapply Mem.store_valid_access_1; eauto.
            }
            split.
            {
              refine_split´; eauto.
              assert(Mem.load Mint32 m´W b (512 × 4) = Mem.load Mint32 m´R b (512 × 4)).
              {
                eapply Mem.load_store_other; eauto.
                right.
                left.
                simpl.
                omega.
              }
              rewrite H0.
              eapply Mem.load_store_same; eauto.
              eapply Mem.store_valid_access_1; eauto.
              eapply Mem.store_valid_access_1; eauto.
              inversion HMR; constructor; eauto.
              omega.
            }
            {
              refine_split´; eauto.
              eapply Mem.load_store_same; eauto.
              eapply Mem.store_valid_access_1; eauto.
              inversion HMW; constructor; eauto.
              omega.
            }
          - apply inject_incr_refl.
        Qed.

        Lemma cons_buf_write_concrete_spec_ref:
          compatsim (crel HDATA LDATA) (gensem cons_buf_write_concrete_spec) cons_buf_write_concrete_spec_low.
        Proof.
          generalize max_unsigned_val; intro muval.
          compatsim_simpl (@match_AbData).
          assert (Hkern: kernel_mode d2).
          {
            inv match_related. functional inversion H1; subst;
            repeat (split; eauto); try congruence.
          }
          inv H.
          destruct H0 as [HMB H0].
          destruct H0 as [HRPOS HWPOS].
          destruct HRPOS as [vr[HLR[HVR HMR]]].
          destruct HWPOS as [vw[HLW[HVW HMW]]].
          functional inversion H1; subst.
          {
            inv HMW.
            generalize (HMB _ H6); intro tmpH.
            destruct tmpH as [[HL´[HV´ HM´]]].
            specialize (Mem.valid_access_store _ _ _ _ (Vint i) HV´); intros [m´´ HST´].
            assert(Mem.valid_access m´´ Mint32 b ((512 + 1) × 4) Writable).
            {
              eapply Mem.store_valid_access_1; eauto.
            }
            specialize (Mem.valid_access_store _ _ _ _ (Vint (Int.repr ((wpos + 1) mod 512))) H); intros [m´W HSTW].
            assert(Mem.valid_access m´W Mint32 b (512 × 4) Writable).
            {
              eapply Mem.store_valid_access_1; eauto.
              eapply Mem.store_valid_access_1; eauto.
            }
            specialize (Mem.valid_access_store _ _ _ _ (Vint (Int.repr ((rpos_concrete (console_concrete d1) + 1) mod 512))) H7); intros [m´W´ HSTW´].
            refine_split.
            - unfold wpos in ×.
              eapply cons_buf_write_concrete_spec_low_intro2; eauto.
              rewrite Int.unsigned_repr; try omega.
              lift_unfold.
              split; eauto.
              instantiate (1:= (m´´, d2)).
              rewrite Int.unsigned_repr; try omega.
              eauto.
              eauto.
              lift_unfold.
              rewrite Int.unsigned_repr; try omega.
              split; eauto.
              instantiate (1:= (m´W, d2)).
              auto.
              auto.
              rewrite Int.unsigned_repr; try omega.
              inv HMR.
              lift_unfold.
              rewrite _x.
              eauto.
              rewrite Int.unsigned_repr; try omega.
              rewrite Int.unsigned_repr; try omega.
              lift_unfold.
              split; eauto.
              rewrite _x.
              eauto.
              rewrite Int.unsigned_repr; try omega.
              xomega.
            - constructor.
            - pose proof H1 as Hspec.
              split; eauto; pattern2_refinement_simpl.
              econstructor; simpl; eauto.
              econstructor; eauto; intros.
              split.
              {
                intros.
                specialize (HMB _ H8).
                destruct HMB as [v1[HL1 [HV1 HM]]].
                assert(ofscase: ofs = wpos ofs wpos) by omega.
                Caseeq ofscase.
                {
                  assert(ofscase´: ofs = CONSOLE_BUFFER_SIZE ofs CONSOLE_BUFFER_SIZE) by omega.
                  intro; subst.
                  Caseeq ofscase´.
                  {
                    intro; subst.
                    omega.
                  }
                  {
                    intro; subst.
                    refine_split´; eauto.
                    unfold wpos in ×.
                    assert(tmprw: Mem.load Mint32 m´W´ b (wpos_concrete (console_concrete d1) × 4) = Mem.load Mint32 m´W b (wpos_concrete (console_concrete d1) × 4)).
                    {
                      eapply Mem.load_store_other; eauto.
                      right.
                      left.
                      simpl.
                      omega.
                    }
                    rewrite tmprw; clear tmprw.
                    assert(tmprw: Mem.load Mint32 m´W b (wpos_concrete (console_concrete d1) × 4) = Mem.load Mint32 m´´ b (wpos_concrete (console_concrete d1) × 4)).
                    {
                      eapply Mem.load_store_other; eauto.
                      right.
                      left.
                      simpl.
                      omega.
                    }
                    rewrite tmprw; clear tmprw.
                    eapply Mem.load_store_same; eauto.
                    eapply Mem.store_valid_access_1; eauto.
                    eapply Mem.store_valid_access_1; eauto.
                    eapply Mem.store_valid_access_1; eauto.
                    simpl.
                    rewrite ZMap.gss.
                    rewrite <- Int.repr_unsigned with i.
                    constructor; eauto.
                    rewrite Int.repr_unsigned.
                    reflexivity.
                  }
                }
                {
                  intro.
                  refine_split´; eauto.
                  instantiate (1:= v1).
                  rewrite <- HL1.
                  assert(tmprw: Mem.load Mint32 m´W´ b (ofs × 4) = Mem.load Mint32 m´W b (ofs × 4)).
                  {
                    eapply Mem.load_store_other; eauto.
                    right.
                    left.
                    simpl.
                    omega.
                  }
                  rewrite tmprw; clear tmprw.
                  assert(tmprw: Mem.load Mint32 m´W b (ofs × 4) = Mem.load Mint32 m´´ b (ofs × 4)).
                  {
                    eapply Mem.load_store_other; eauto.
                    right.
                    left.
                    simpl.
                    omega.
                  }
                  rewrite tmprw; clear tmprw.
                  eapply Mem.load_store_other; eauto.
                  right.
                  simpl.
                  unfold wpos in ×.
                  assert(tmpcase: ofs < wpos_concrete (console_concrete d1) wpos_concrete (console_concrete d1) < ofs) by omega.
                  Caseeq tmpcase.
                  {
                    intro.
                    left.
                    omega.
                  }
                  {
                    intro.
                    right.
                    omega.
                  }
                  eapply Mem.store_valid_access_1; eauto.
                  eapply Mem.store_valid_access_1; eauto.
                  eapply Mem.store_valid_access_1; eauto.
                  simpl.
                  rewrite ZMap.gso; try assumption.
                }
              }
              split.
              {
                refine_split´; eauto.
                eapply Mem.load_store_same; eauto.
                eapply Mem.store_valid_access_1; eauto.
                simpl.
                constructor; eauto.
                xomega.
              }
              {
                refine_split´; eauto.
                assert(tmprw: Mem.load Mint32 m´W´ b ((512 + 1) × 4) = Mem.load Mint32 m´W b ((512 + 1) × 4)).
                {
                  eapply Mem.load_store_other; eauto.
                  right.
                  right.
                  simpl.
                  omega.
                }
                rewrite tmprw; clear tmprw.
                eapply Mem.load_store_same; eauto.
                eapply Mem.store_valid_access_1; eauto.
                eapply Mem.store_valid_access_1; eauto.
                simpl.
                constructor; eauto.
                xomega.
              }
            - apply inject_incr_refl.
          }
          {
            inv HMW.
            generalize (HMB _ H6); intro tmpH.
            destruct tmpH as [[HL´[HV´ HM´]]].
            specialize (Mem.valid_access_store _ _ _ _ (Vint i) HV´); intros [m´´ HST´].
            assert(Mem.valid_access m´´ Mint32 b ((512 + 1) × 4) Writable).
            {
              eapply Mem.store_valid_access_1; eauto.
            }
            specialize (Mem.valid_access_store _ _ _ _ (Vint (Int.repr ((wpos + 1) mod 512))) H); intros [m´W HSTW].
            refine_split.
            - unfold wpos in ×.
              econstructor; eauto.
              rewrite Int.unsigned_repr; try omega.
              lift_unfold.
              split; eauto.
              instantiate (1:= (m´´, d2)).
              rewrite Int.unsigned_repr; try omega.
              eauto.
              eauto.
              lift_unfold.
              inv HMR.
              eauto.
              rewrite Int.unsigned_repr; try omega.
              rewrite Int.unsigned_repr; try omega.
              unfold rpos in ×.
              rewrite Int.unsigned_repr; try omega.
              inv HMR.
              omega.
              rewrite Int.unsigned_repr; try omega.
              xomega.
              rewrite Int.unsigned_repr; try omega.
              lift_unfold.
              split; eauto.
            - constructor.
            - pose proof H1 as Hspec.
              split; eauto; pattern2_refinement_simpl.
              econstructor; simpl; eauto.
              econstructor; eauto; intros.
              inv match_match.
              inv H7.
              destruct H8.
              destruct H8.
              split.
              {
                intros.
                specialize (HMB _ H11).
                destruct HMB as [v1[HL1 [HV1 HM]]].
                assert(tmprw: Mem.load Mint32 m´W b (ofs × 4) = Mem.load Mint32 m´´ b (ofs × 4)).
                {
                  eapply Mem.load_store_other; eauto.
                  right.
                  left.
                  simpl.
                  omega.
                }
                rewrite tmprw; clear tmprw.
                assert(ofscase: ofs = wpos ofs wpos) by omega.
                Caseeq ofscase.
                {
                  intro; subst.
                  refine_split´; eauto.
                  eapply Mem.load_store_same; eauto.
                  eapply Mem.store_valid_access_1; eauto.
                  eapply Mem.store_valid_access_1; eauto.
                  simpl.
                  rewrite ZMap.gss.
                  rewrite <- Int.repr_unsigned with i.
                  constructor; eauto.
                  rewrite Int.repr_unsigned.
                  reflexivity.
                }
                {
                  intro.
                  simpl.
                  refine_split´; eauto.
                  instantiate (1:= v1).
                  rewrite <- HL1.
                  eapply Mem.load_store_other; eauto.
                  right.
                  simpl.
                  unfold wpos in ×.
                  assert(tmpcase: ofs < wpos_concrete (console_concrete d1) wpos_concrete (console_concrete d1) < ofs) by omega.
                  Caseeq tmpcase.
                  {
                    intro.
                    left.
                    omega.
                  }
                  {
                    intro.
                    right.
                    omega.
                  }
                  eapply Mem.store_valid_access_1; eauto.
                  eapply Mem.store_valid_access_1; eauto.
                  simpl.
                  rewrite ZMap.gso; try assumption.
                }
              }
              split.
              {
                refine_split´; eauto.
                assert(tmprw: Mem.load Mint32 m´W b (512 × 4) = Mem.load Mint32 m´´ b (512 × 4)).
                {
                  eapply Mem.load_store_other; eauto.
                  right.
                  left.
                  simpl.
                  omega.
                }
                rewrite tmprw; clear tmprw.
                rewrite <- HLR.
                eapply Mem.load_store_other; eauto.
                right.
                right.
                simpl.
                xomega.
                eapply Mem.store_valid_access_1; eauto.
                eapply Mem.store_valid_access_1; eauto.
              }
              {
                refine_split´; eauto.
                eapply Mem.load_store_same; eauto.
                eapply Mem.store_valid_access_1; eauto.
                constructor; eauto.
                xomega.
              }
            - apply inject_incr_refl.
          }
        Qed.

        Lemma cons_buf_wpos_concrete_spec_ref:
          compatsim (crel HDATA LDATA) (gensem cons_buf_wpos_concrete_spec) cons_buf_wpos_concrete_spec_low.
        Proof.
          compatsim_simpl (@match_AbData). inv H.
          assert(HKern: kernel_mode d2).
          {
            simpl; inv match_related.
            functional inversion H2; repeat (split; trivial); congruence.
          }
          pose proof H0 as HBit.
          destruct HBit.
          destruct H3.
          refine_split; eauto.
          econstructor; eauto.
          destruct H4 as [v[HL[_ HM]]].
          lift_unfold.
          assert (HP: v = Vint z).
          {
            functional inversion H2; subst; rewrite H4 in HM; inv HM.
            rewrite Int.repr_unsigned with z; reflexivity.
          }
          rewrite <- HP.
          assumption.
        Qed.

        Lemma cons_buf_read_concrete_spec_ref:
          compatsim (crel HDATA LDATA) (gensem cons_buf_read_concrete_spec) cons_buf_read_concrete_spec_low.
        Proof.
          generalize max_unsigned_val; intro muval.
          compatsim_simpl (@match_AbData).
          assert (Hkern: kernel_mode d2).
          {
            inv match_related. functional inversion H1; unfold rpos, wpos, s0, buf in *; subst;
            repeat (split; eauto); try congruence.
          }
          inv H.
          destruct H0 as [HMB H0].
          destruct H0 as [HRPOS HWPOS].
          destruct HRPOS as [vr[HLR[HVR HMR]]].
          destruct HWPOS as [vw[HLW[HVW HMW]]].
          functional inversion H1; unfold rpos, wpos, s0, buf in *; subst.
          {
            refine_split.
            - destruct HMR; destruct HMW; subst.
              econstructor; eauto.
              subst.
              reflexivity.
            - rewrite <- Int.repr_unsigned with z.
              rewrite <- H0.
              constructor.
            - pose proof H1 as Hspec.
              split; eauto; pattern2_refinement_simpl.
            - apply inject_incr_refl.
          }
          {
            generalize HMR; intro tmpd.
            destruct tmpd; subst; eauto.
            generalize (HMB _ H0); intro tmpH.
            destruct tmpH as [[HL´[HV´ HM´]]].
            specialize (Mem.valid_access_store _ _ _ _ (Vint (Int.repr ((v + 1) mod CONSOLE_BUFFER_SIZE))) HVR); intros [m´´ HST´].
            inversion HM´; subst.
            unfold s0 in *; rewrite H7 in H8; discriminate H8.
            refine_split.
            - eapply cons_buf_read_concrete_spec_low_intro2; eauto.
              lift_unfold.
              inv HMW.
              eassumption.
              rewrite Int.unsigned_repr; try omega.
              rewrite Int.unsigned_repr; try omega.
              inv HMW.
              omega.
              rewrite Int.unsigned_repr; try omega.
              lift_unfold.
              eassumption.
              rewrite Int.unsigned_repr; try omega.
              lift_unfold. split; eauto.
              rewrite Int.unsigned_repr; try omega.
            - unfold s0 in ×.
              rewrite H7 in H.
              injection H; intros; subst.
              rewrite Int.repr_unsigned.
              constructor.
            - pose proof H1 as Hspec.
              split; eauto; pattern2_refinement_simpl.
              econstructor; simpl; eauto.
              econstructor; eauto; intros.
              split.
              {
                intros.
                specialize (HMB _ H8).
                destruct HMB as [v1[HL1 [HV1 HM]]].
                assert(tmprw: Mem.load Mint32 m´´ b (ofs × 4) = Mem.load Mint32 m2 b (ofs × 4)).
                {
                  eapply Mem.load_store_other; eauto.
                  right.
                  left.
                  simpl.
                  omega.
                }
                rewrite tmprw; clear tmprw.
                refine_split´; eauto.
                eapply Mem.store_valid_access_1; eauto.
              }
              split.
              {
                refine_split´; eauto.
                eapply Mem.load_store_same; eauto.
                eapply Mem.store_valid_access_1; eauto.
                inversion HMR; constructor; eauto.
                xomega.
              }
              {
                refine_split´; eauto.
                assert(tmprw: Mem.load Mint32 m´´ b ((512 + 1) × 4) = Mem.load Mint32 m2 b ((512 + 1) × 4)).
                {
                  eapply Mem.load_store_other; eauto.
                  right.
                  right.
                  simpl.
                  omega.
                }
                rewrite tmprw; clear tmprw.
                assumption.
                eapply Mem.store_valid_access_1; eauto.
              }
            - apply inject_incr_refl.
          }
        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) dconsolebuffintro_passthrough mqticketlock.
        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. 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 save_context_sim.
          - apply restore_context_sim.
          - apply local_irq_save_sim.
          - apply local_irq_restore_sim.
          - apply serial_in_sim.
          - apply serial_out_sim.
          - apply serial_hw_intr_sim.
          - apply ioapic_read_sim.
          - apply ioapic_write_sim.
          - apply lapic_read_sim.
          - apply lapic_write_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.