Library mcertikos.proc.EPTIntroGenDef


This file provide the contextual refinement proof between MAL layer and MPTIntro layer
Require Export LoadStoreSem3.
Require Export Coqlib.
Require Export Errors.
Require Export AST.
Require Export Integers.
Require Export Floats.
Require Export Op.
Require Export Asm.
Require Export Events.
Require Export Globalenvs.
Require Export Smallstep.
Require Export Values.
Require Export Memory.
Require Export Maps.
Require Export CommonTactic.
Require Export AuxLemma.
Require Export FlatMemory.
Require Export AuxStateDataType.
Require Export Constant.
Require Export GlobIdent.
Require Export RealParams.
Require Export AsmImplLemma.
Require Export LAsm.
Require Export RefinementTactic.
Require Export PrimSemantics.

Require Export liblayers.logic.PTreeModules.
Require Export liblayers.logic.LayerLogicImpl.
Require Export liblayers.compcertx.ClightModules.
Require Export liblayers.compat.CompatLayers.
Require Export liblayers.compat.CompatGenSem.
Require Export liblayers.compat.CompatClightSem.
Require Export LayerCalculusLemma.
Require Export GenSem.

Require Export AbstractDataType.
Require Export VEPTIntro.
Require Export XOmega.

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

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

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

Definition of the refinement relation

Section Refinement.

  Section WITHMEM.

    Context `{Hstencil: Stencil}.
    Context `{Hmem: Mem.MemoryModelX}.
    Context `{Hmwd: UseMemWithData mem}.
    Context `{multi_oracle_prop: MultiOracleProp}.
    Context `{real_params: RealParams}.

Definition the refinement relation: relate_RData + match_RData

    Section REFINEMENT_REL.

      Lemma EPT_PML4_INDEX_max:
        EPT_PML4_INDEX Int.max_unsigned = 0.
      Proof.
        rewrite_omega.
        reflexivity.
      Qed.

      Lemma EPT_PDPT_INDEX_max:
        EPT_PDPT_INDEX Int.max_unsigned = 3.
      Proof.
        rewrite_omega.
        reflexivity.
      Qed.

      Lemma EPT_PDIR_INDEX_max:
        EPT_PDIR_INDEX Int.max_unsigned = 511.
      Proof.
        rewrite_omega.
        reflexivity.
      Qed.

      Lemma EPT_PTAB_INDEX_max:
        EPT_PTAB_INDEX Int.max_unsigned = 511.
      Proof.
        rewrite_omega.
        reflexivity.
      Qed.

      Lemma EPT_PML4_INDEX_unsigned:
         i,
          EPT_PML4_INDEX (Int.unsigned i) = 0.
      Proof.
        unfold EPT_PML4_INDEX. intros.
        replace (Int.unsigned i / (4096 × 512 × 512 × 512)) with 0.
        - reflexivity.
        - pose proof (Int.unsigned_range i) as Hr.
          rewrite Zdiv_small; trivial.
          change Int.modulus with 4294967296 in Hr.
          omega.
      Qed.

      Lemma EPT_PDPT_INDEX_range:
         i,
          0 EPT_PDPT_INDEX (Int.unsigned i) EPT_PDPT_INDEX Int.max_unsigned.
      Proof.
        intros. rewrite EPT_PDPT_INDEX_max.
        pose proof (Int.unsigned_range i) as Hr.
        change Int.modulus with 4294967296 in Hr.
        unfold EPT_PDPT_INDEX.
        assert (HR: 0 Int.unsigned i / (4096 × 512 × 512) 3).
        {
          xomega.
        }
        rewrite Zmod_small; trivial.
        omega.
      Qed.

      Lemma EPT_PDIR_INDEX_range:
         i,
          0 EPT_PDIR_INDEX (Int.unsigned i) EPT_PDIR_INDEX Int.max_unsigned.
      Proof.
        intros. rewrite EPT_PDIR_INDEX_max.
        pose proof (Int.unsigned_range i) as Hr.
        change Int.modulus with 4294967296 in Hr.
        unfold EPT_PDIR_INDEX.
        xomega.
      Qed.

      Lemma EPT_PTAB_INDEX_range:
         i,
          0 EPT_PTAB_INDEX (Int.unsigned i) EPT_PTAB_INDEX Int.max_unsigned.
      Proof.
        intros. rewrite EPT_PTAB_INDEX_max.
        pose proof (Int.unsigned_range i) as Hr.
        change Int.modulus with 4294967296 in Hr.
        unfold EPT_PTAB_INDEX.
        xomega.
      Qed.

      Inductive match_EPTE: stencilZEPTEZZZmemblockProp :=
      | MATCH_EPTE_UNDEF:
           s m i j k b vmid,
            match_EPTE s vmid EPTEUndef i j k m b
      | MATCH_EPTE_VALID:
           hpa i j k m b s vmid
                 (HLD: Mem.load Mint64 m b (8413184 × vmid + (6 + i × 512 + j) × PgSize + k × 8) = Some (Vlong (Int64.repr (Int.unsigned hpa)))),
            match_EPTE s vmid (EPTEValid (Int.unsigned hpa)) i j k m b.

      Inductive match_EPDTE: stencilZEPDTEZZmemblockProp :=
      | MATCH_EPDTE_UNDEF:
           s i j m b vmid,
            match_EPDTE s vmid EPDTEUndef i j m b
      | MATCH_EPDTE_VALID:
           epte m b s ofs i j vmid
                 (HMat: ( k,
                           0 k EPT_PTAB_INDEX Int.max_unsigned
                           match_EPTE s vmid (ZMap.get k epte) i j k m b))
                 (Hofs: Int.unsigned ofs = 8413184 × vmid + (6 + i × 512 + j) × PgSize + EPTEPERM)
                 (HLD: Mem.load Mint32 m b (8413184 × vmid + (i + 2) × PgSize + j × 8) = Some (Vptr b ofs)),
            match_EPDTE s vmid (EPDTEValid epte) i j m b.

      Inductive match_EPDPTE: stencilZEPDPTEZmemblockProp :=
      | MATCH_EPDPTE_UNDEF:
           s i m b vmid,
            match_EPDPTE s vmid EPDPTEUndef i m b
      | MATCH_EPDPTE_VALID:
           epdte m b s ofs i vmid
                 (HMat: ( j,
                           0 j EPT_PDIR_INDEX Int.max_unsigned
                           match_EPDTE s vmid (ZMap.get j epdte) i j m b))
                 (Hofs: Int.unsigned ofs = 8413184 × vmid + (i + 2) × PgSize + EPTEPERM)
                 (HLD: Mem.load Mint32 m b (8413184 × vmid + PgSize + i × 8) = Some (Vptr b ofs)),
            match_EPDPTE s vmid (EPDPTEValid epdte) i m b.

      Inductive match_EPML4E: stencilZEPML4EmemblockProp :=
      | MATCH_EPML4E_UNDEF:
           s m b vmid,
            match_EPML4E s vmid EPML4EUndef m b
      | MATCH_EPML4E_VALID:
           epdpt m b s ofs vmid
                 (HMat: ( j,
                           0 j EPT_PDPT_INDEX Int.max_unsigned
                           match_EPDPTE s vmid (ZMap.get j epdpt) j m b))
                 (Hofs: Int.unsigned ofs = 8413184 × vmid + PgSize + EPTEPERM)
                 (HLD: Mem.load Mint32 m b (8413184 × vmid) = Some (Vptr b ofs)),
            match_EPML4E s vmid (EPML4EValid epdpt) m b.

      Inductive match_EPT: stencilZEPTmemProp :=
      | MATCH_EPT: ept m b s vmid
                          (HMat: match_EPML4E s vmid (ZMap.get 0 ept) m b)
                          (Hsymbol: find_symbol s EPT_LOC = Some b)
                          (HVA: Mem.valid_access m Mint32 b (8413184 × vmid) Writable)
                          (HVA´: Mem.valid_access m Mint32 b (8413184 × vmid + 4) Writable)
                          (HVA1: ( i,
                                    0 i EPT_PDPT_INDEX Int.max_unsigned
                                    Mem.valid_access m Mint32 b (8413184 × vmid + PgSize + i × 8) Writable
                                     Mem.valid_access m Mint32 b (8413184 × vmid + PgSize + i × 8 + 4) Writable
                                     ( j,
                                          0 j EPT_PDIR_INDEX Int.max_unsigned
                                          Mem.valid_access m Mint32 b (8413184 × vmid + (i + 2) × PgSize + j × 8) Writable
                                           Mem.valid_access m Mint32 b (8413184 × vmid + (i + 2) × PgSize + j × 8 + 4) Writable
                                           ( k,
                                                0 k EPT_PTAB_INDEX Int.max_unsigned
                                                Mem.valid_access m Mint64 b (8413184 × vmid + (6 + i × 512 + j) × PgSize + k × 8) Writable)))),
                     match_EPT s vmid ept m.

Relation between the new raw data at the higher layer with the mememory at lower layer
      Inductive match_RData: stencilHDATAmemmeminjProp :=
      | MATCH_RDATA: hadt m s f eptati,
                       eptati = ZMap.get (CPU_ID hadt) (ept hadt)
                       → match_EPT s (CPU_ID hadt) eptati m
                       → match_RData s hadt m f.

Relation between the shared 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;
            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));
            LAT_re: LAT hadt = LAT ladt;
            nps_re: nps hadt = nps ladt;
            init_re: init hadt = init ladt;

            pperm_re: pperm hadt = pperm ladt;
            PT_re: PT hadt = PT ladt;
            ptp_re: ptpool hadt = ptpool ladt;
            idpde_re: idpde hadt = idpde ladt;
            ipt_re: ipt hadt = ipt ladt;
            smspool_re: smspool hadt = smspool 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;

            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;
            in_intr_re: in_intr ladt = in_intr hadt;
            drv_serial_re: drv_serial hadt = drv_serial ladt;

            kctxt_re: kctxt_inj f num_proc (kctxt hadt) (kctxt ladt);
            uctxt_re: uctxt_inj f (uctxt hadt) (uctxt ladt);
            
            syncchpool_re: syncchpool hadt = syncchpool 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 := EPT_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.
        econstructor; eauto.
        - inv H1.
          assert (HFB0: j b = Some (b, 0)).
          {
            eapply stencil_find_symbol_inject´; eauto.
          }
          econstructor; eauto; intros.
          inv HMat; try constructor.
          specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H0 HLD HFB0).
          rewrite Z.add_0_r.
          intros [v1´[HLD1´ HVA1´]].
          econstructor; eauto.
          {
            intros. specialize (HMat0 _ H1).
            inv HMat0; try constructor.
            specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H0 HLD0 HFB0).
            rewrite Z.add_0_r.
            intros [v2´[HLD2´ HVA2´]].
            econstructor; eauto.
            {
              intros. specialize (HMat _ H3).
              inv HMat; try constructor.
              specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H0 HLD1 HFB0).
              rewrite Z.add_0_r.
              intros [v3´[HLD3´ HVA3´]].
              econstructor; eauto.
              {
                intros. specialize (HMat0 _ H4).
                inv HMat0; try constructor.
                specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H0 HLD2 HFB0).
                rewrite Z.add_0_r.
                intros [v4´[HLD4´ HVA4´]].
                inv HVA4´. trivial.
              }
              inv HVA3´. rewrite HFB0 in H8. inv H8.
              rewrite Int.add_zero in HLD3´. trivial.
            }
            inv HVA2´. rewrite HFB0 in H6. inv H6.
            rewrite Int.add_zero in HLD2´. trivial.
          }
          + inv HVA1´. rewrite HFB0 in H4. inv H4.
            rewrite Int.add_zero in HLD1´. trivial.
          + specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H0 HVA).
            rewrite Z.add_0_r; trivial.
          + specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H0 HVA´).
            rewrite Z.add_0_r; trivial.
          + specialize (HVA1 _ H1).
            destruct HVA1 as (HVA11 & HVA12 & HVA13).
            refine_split´.
            × specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H0 HVA11).
              rewrite Z.add_0_r; trivial.
            × specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H0 HVA12).
              rewrite Z.add_0_r; trivial.
            × intros. specialize (HVA13 _ H3).
              destruct HVA13 as (HVA131 & HVA132 & HVA133).
              refine_split´.
              {
                specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H0 HVA131).
                rewrite Z.add_0_r; trivial.
              }
              {
                specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H0 HVA132).
                rewrite Z.add_0_r; trivial.
              }
              {
                intros. specialize (HVA133 _ H4).
                specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H0 HVA133).
                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. econstructor; eauto.
        inv H3. pose proof Hsymbol as Hsymbol´.
        eapply H0 in Hsymbol; simpl; eauto.
        econstructor; eauto;
        try eapply Mem.store_valid_access_1; eauto; intros.
        - inv HMat; try econstructor;
          try rewrite (Mem.load_store_other _ _ _ _ _ _ H1);
            eauto; intros.

          specialize (HMat0 _ H).
          inv HMat0; try econstructor;
          try rewrite (Mem.load_store_other _ _ _ _ _ _ H1);
          eauto; intros.

          specialize (HMat _ H2).
          inv HMat; try econstructor;
          try rewrite (Mem.load_store_other _ _ _ _ _ _ H1);
          eauto; intros.

          specialize (HMat0 _ H3).
          inv HMat0; try econstructor;
          try rewrite (Mem.load_store_other _ _ _ _ _ _ H1); eauto.

        - specialize (HVA1 _ H).
          destruct HVA1 as (HVA11 & HVA12 & HVA13).
          refine_split´; try eapply Mem.store_valid_access_1; eauto; intros.
          specialize (HVA13 _ H2).
          destruct HVA13 as (HVA131 & HVA132 & HVA133).
          refine_split´; try eapply Mem.store_valid_access_1; eauto; intros.
          specialize (HVA133 _ H3).
          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. econstructor; eauto.
        inv H3. pose proof Hsymbol as Hsymbol´.
        eapply H0 in Hsymbol; simpl; eauto.
        econstructor; eauto;
        try eapply Mem.storebytes_valid_access_1; eauto; intros.
        - inv HMat; try econstructor;
          try rewrite (Mem.load_storebytes_other _ _ _ _ _ H1);
          eauto; intros.

          specialize (HMat0 _ H).
          inv HMat0; try econstructor;
          try rewrite (Mem.load_storebytes_other _ _ _ _ _ H1);
          eauto; intros.

          specialize (HMat _ H2).
          inv HMat; try econstructor;
          try rewrite (Mem.load_storebytes_other _ _ _ _ _ H1);
          eauto; intros.

          specialize (HMat0 _ H3).
          inv HMat0; try econstructor;
          try rewrite (Mem.load_storebytes_other _ _ _ _ _ H1); eauto.

        - specialize (HVA1 _ H).
          destruct HVA1 as (HVA11 & HVA12 & HVA13).
          refine_split´; try eapply Mem.storebytes_valid_access_1; eauto; intros.
          specialize (HVA13 _ H2).
          destruct HVA13 as (HVA131 & HVA132 & HVA133).
          refine_split´; try eapply Mem.storebytes_valid_access_1; eauto; intros.
          specialize (HVA133 _ H3).
          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. econstructor; eauto.
        inv H3. pose proof Hsymbol as Hsymbol´.
        eapply H0 in Hsymbol; simpl; eauto.
        econstructor; eauto;
        try eapply Mem.valid_access_free_1; eauto; intros.
        - inv HMat; try econstructor;
          try rewrite (Mem.load_free _ _ _ _ _ H1);
          eauto; intros.

          specialize (HMat0 _ H).
          inv HMat0; try econstructor;
          try rewrite (Mem.load_free _ _ _ _ _ H1);
          eauto; intros.

          specialize (HMat _ H2).
          inv HMat; try econstructor;
          try rewrite (Mem.load_free _ _ _ _ _ H1);
          eauto; intros.

          specialize (HMat0 _ H3).
          inv HMat0; try econstructor;
          try rewrite (Mem.load_free _ _ _ _ _ H1); eauto.

        - specialize (HVA1 _ H).
          destruct HVA1 as (HVA11 & HVA12 & HVA13).
          refine_split´; try eapply Mem.valid_access_free_1; eauto; intros.
          specialize (HVA13 _ H2).
          destruct HVA13 as (HVA131 & HVA132 & HVA133).
          refine_split´; try eapply Mem.valid_access_free_1; eauto; intros.
          specialize (HVA133 _ H3).
          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. econstructor; eauto.
        inv H2.
        econstructor; eauto;
        try eapply Mem.valid_access_alloc_other; eauto; intros.
        - inv HMat; try econstructor;
          try apply (Mem.load_alloc_other _ _ _ _ _ H0);
          eauto; intros.

          specialize (HMat0 _ H).
          inv HMat0; try econstructor;
          try apply (Mem.load_alloc_other _ _ _ _ _ H0);
          eauto; intros.

          specialize (HMat _ H1).
          inv HMat; try econstructor;
          try apply (Mem.load_alloc_other _ _ _ _ _ H0);
          eauto; intros.

          specialize (HMat0 _ H2).
          inv HMat0; try econstructor;
          try apply (Mem.load_alloc_other _ _ _ _ _ H0); eauto.

        - specialize (HVA1 _ H).
          destruct HVA1 as (HVA11 & HVA12 & HVA13).
          refine_split´; try eapply Mem.valid_access_alloc_other; eauto; intros.
          specialize (HVA13 _ H1).
          destruct HVA13 as (HVA131 & HVA132 & HVA133).
          refine_split´; try eapply Mem.valid_access_alloc_other; eauto; intros.
          specialize (HVA133 _ H2).
          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 kctxt_inj_incr; eauto.
        - eapply uctxt_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.

    Global Instance: (LoadStoreProp (hflatmem_store:= flatmem_store) (lflatmem_store:= flatmem_store)).
    Proof.
      accessor_prop_tac.
      - exploit flatmem_store_exists; eauto.
      - exploit flatmem_store_match; eauto.
    Qed.

  End WITHMEM.

End Refinement.