Library mcertikos.proc.CVIntroGenDef


This file provide the contextual refinement proof between PKContext layer and PKCtxtNew layer
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 LoadStoreSem2.
Require Export AsmImplLemma.
Require Export GenSem.
Require Export RefinementTactic.
Require Export PrimSemantics.
Require Export XOmega.

Require Export liblayers.logic.PTreeModules.
Require Export liblayers.logic.LayerLogicImpl.
Require Export liblayers.compcertx.Stencil.
Require Export liblayers.compcertx.MakeProgram.
Require Export liblayers.compat.CompatLayers.
Require Export liblayers.compat.CompatGenSem.
Require Export compcert.cfrontend.Ctypes.

Require Export FutureTactic.

Require Export AbstractDataType.
Require Export LayerCalculusLemma.

Require Export PAbQueueAtomic.
Require Export PCVIntro.

Require Export CVIntroGenSpec.

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

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

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

Definition of the refinement relation

Section Refinement.

  Section WITHMEM.

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

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

Definition the refinement relation: relate_RData + match_RData

    Section REFINEMENT_REL.




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
                 (HSCP: match_ChanPool s (syncchpool hadt) m),
            match_RData s hadt m f.

Relation between raw data at two layers
      Record relate_RData (s: stencil) (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;
            
            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;

            abq_re: abq ladt = abq hadt;
            
            abtcb_re: abtcb ladt = abtcb hadt;
            sleeper_re: sleeper ladt = sleeper hadt;

            kctxt_re: kctxt_inj f num_proc (kctxt hadt) (kctxt ladt);

            multi_oracle_re: relate_SC_Oracle_Pool s (multi_oracle hadt) (multi_oracle ladt);
            multi_log_re: relate_SC_Log_Pool s (multi_log hadt) (multi_log ladt)

          }.

      Local Hint Resolve MATCH_RDATA.

      Global Instance rel_ops: CompatRelOps HDATAOps LDATAOps :=
        {
          relate_AbData s f d1 d2 := relate_RData s f d1 d2;
          match_AbData s d1 m f := match_RData s d1 m f;
          new_glbl := SYNCCHPOOL_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.
        {
          inv HSCP.
          assert (HFB0: j b = Some (b, 0)).
          {
            eapply stencil_find_symbol_inject´; eauto.
          }
          econstructor; eauto.
          intros id Hid. specialize (H2 id Hid).
          destruct H2 as (v1 & v2 & v3 & v4 & HL1 & HV1 & HL2 & HV2 & HL3 & HV3 & HL4 & HV4 & HM).
          specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H0 HL1 HFB0).
          specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H0 HL2 HFB0).
          specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H0 HL3 HFB0).
          specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H0 HL4 HFB0).
          repeat rewrite Z.add_0_r.
          intros (v4´ & HLD4´ & HV4´).
          intros (v3´ & HLD3´ & HV3´).
          intros (v2´ & HLD2´ & HV2´).
          intros (v1´ & HLD1´ & HV1´).
          refine_split´; eauto.
          specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H0 HV1).
          rewrite Z.add_0_r; trivial.
          specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H0 HV2).
          rewrite Z.add_0_r; trivial.
          specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H0 HV3).
          rewrite Z.add_0_r; trivial.
          specialize(Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB0 H0 HV4).
          rewrite Z.add_0_r; trivial.
          inv HM.
          + inv HV1´. inv HV2´. inv HV3´. inv HV4´. constructor; auto.
          + inv HV1´. inv HV2´. inv HV3´. inv HV4´. 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.
        econstructor.
        {
          inv HSCP.
          - econstructor; eauto.
            eapply H0 in H2.
            intros id Hid. specialize (H id Hid).
            blast H. refine_split´.
            + rewrite (Mem.load_store_other _ _ _ _ _ _ H1). eassumption.
              left. assumption.
            + eapply Mem.store_valid_access_1; eauto.
            + rewrite (Mem.load_store_other _ _ _ _ _ _ H1). eassumption.
              left. assumption.
            + eapply Mem.store_valid_access_1; eauto.
            + rewrite (Mem.load_store_other _ _ _ _ _ _ H1). eassumption.
              left. assumption.
            + eapply Mem.store_valid_access_1; eauto.
            + rewrite (Mem.load_store_other _ _ _ _ _ _ H1). eassumption.
              left. assumption.
            + eapply Mem.store_valid_access_1; eauto.
            + eassumption.
            + econstructor. reflexivity.
        }
        
      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.
        {
          inv HSCP.
          - econstructor; eauto.
            intros id Hid. specialize (H id Hid).
            blast H. eapply H0 in H2.
            refine_split´.
            + repeat rewrite (Mem.load_storebytes_other _ _ _ _ _ H1); eauto.
            + eapply Mem.storebytes_valid_access_1; eauto.
            + repeat rewrite (Mem.load_storebytes_other _ _ _ _ _ H1); eauto.
            + eapply Mem.storebytes_valid_access_1; eauto.
            + repeat rewrite (Mem.load_storebytes_other _ _ _ _ _ H1); eauto.
            + eapply Mem.storebytes_valid_access_1; eauto.
            + repeat rewrite (Mem.load_storebytes_other _ _ _ _ _ H1); eauto.
            + eapply Mem.storebytes_valid_access_1; eauto.
            + assumption.
            + econstructor; reflexivity.
        }
        
      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.
        {
          inv HSCP.
          - econstructor; eauto.
            generalize H2; intros Hx.
            eapply H0 in H2.
            intros id Hid. specialize (H id Hid).
            blast H. refine_split´.
            + repeat rewrite (Mem.load_free _ _ _ _ _ H1); eauto.
            + eapply Mem.valid_access_free_1; eauto.
            + repeat rewrite (Mem.load_free _ _ _ _ _ H1); eauto.
            + eapply Mem.valid_access_free_1; eauto.
            + repeat rewrite (Mem.load_free _ _ _ _ _ H1); eauto.
            + eapply Mem.valid_access_free_1; eauto.
            + repeat rewrite (Mem.load_free _ _ _ _ _ H1); eauto.
            + eapply Mem.valid_access_free_1; eauto.
            + assumption.
            + econstructor. reflexivity.
        }
        
      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.
        {
          inv HSCP.
          - econstructor; eauto.
            intros id Hid. specialize (H id Hid).
            blast H. refine_split´.
            + apply (Mem.load_alloc_other _ _ _ _ _ H0); eauto.
            + eapply Mem.valid_access_alloc_other; eauto.
            + apply (Mem.load_alloc_other _ _ _ _ _ H0); eauto.
            + eapply Mem.valid_access_alloc_other; eauto.
            + apply (Mem.load_alloc_other _ _ _ _ _ H0); eauto.
            + eapply Mem.valid_access_alloc_other; eauto.
            + apply (Mem.load_alloc_other _ _ _ _ _ H0); eauto.
            + eapply Mem.valid_access_alloc_other; eauto.
            + eauto.
        }
        
      Qed.

Prove that after taking one step, the refinement relation still holds
      Lemma relate_incr:
         s abd abd´ f ,
          relate_RData s f abd abd´
          → inject_incr f
          → relate_RData s abd abd´.
      Proof.
        inversion 1; subst; intros; inv H; constructor; eauto.
        eapply kctxt_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.
  End WITHMEM.
End Refinement.