Library mcertikos.mcslock.MCSLockIntroGenDef


This file provide the contextual refinement proof between MBoot layer and MALInit 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 LoadStoreSem1.
Require Export AsmImplLemma.
Require Export GenSem.
Require Export RefinementTactic.
Require Export PrimSemantics.

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 LayerCalculusLemma.

Require Export AbstractDataType.
Require Import DeviceStateDataType.

Require Export CalMCSLock.
Require Export GlobalOracle.
Require Export MMCSLockIntro.

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

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

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

Definition of the refinement relation


Ltac sp_with_arg H Hval HFB :=
  specialize (Mem.load_inject _ _ _ _ _ _ _ _ _ H Hval HFB).

Ltac sp_with_arg´ H Hval HFB :=
  specialize (Mem.valid_access_inject _ _ _ _ _ _ _ _ _ HFB H Hval).

Ltac mcslock_destruct HMCSLock :=
  destruct HMCSLock as (la_v & bs0_v & ne0_v & bs1_v & ne1_v & bs2_v & ne2_v & bs3_v & ne3_v &
                        bs4_v & ne4_v & bs5_v & ne5_v & bs6_v & ne6_v & bs7_v & ne7_v &
                        HL__la & HL__bs0 & HL__ne0 & HL__bs1 & HL__ne1 & HL__bs2 & HL__ne2 & HL__bs3 & HL__ne3 &
                        HL__bs4 & HL__ne4 & HL__bs5 & HL__ne5 & HL__bs6 & HL__ne6 & HL__bs7 & HL__ne7 &
                        HV__la & HV__bs0 & HV__ne0 & HV__bs1 & HV__ne1 & HV__bs2 & HV__ne2 & HV__bs3 & HV__ne3 &
                        HV__bs4 & HV__ne4 & HV__bs5 & HV__ne5 & HV__bs6 & HV__ne6 & HV__bs7 & HV__ne7 & Hm).

Ltac mcslock_destruct´ HMCSLock :=
  destruct HMCSLock as (la_v´ & bs0_v´ & ne0_v´ & bs1_v´ & ne1_v´ & bs2_v´ & ne2_v´ & bs3_v´ & ne3_v´ &
                        bs4_v´ & ne4_v´ & bs5_v´ & ne5_v´ & bs6_v´ & ne6_v´ & bs7_v´ & ne7_v´ &
                        HL__la´ & HL__bs0´ & HL__ne0´ & HL__bs1´ & HL__ne1´ & HL__bs2´ & HL__ne2´ & HL__bs3´ & HL__ne3´ &
                        HL__bs4´ & HL__ne4´ & HL__bs5´ & HL__ne5´ & HL__bs6´ & HL__ne6´ & HL__bs7´ & HL__ne7´ &
                        HV__la´ & HV__bs0´ & HV__ne0´ & HV__bs1´ & HV__ne1´ & HV__bs2´ & HV__ne2´ & HV__bs3´ & HV__ne3´ &
                        HV__bs4´ & HV__ne4´ & HV__bs5´ & HV__ne5´ & HV__bs6´ & HV__ne6´ & HV__bs7´ & HV__ne7´ & Hm´).

Ltac ex_vint val := (Vint (Int.repr val)).
Ltac ex_vbool val := (Vint (Int.repr (BooltoZ val))).

Lemma destruct_CPU_ID_cases:
   i, 0 Int.unsigned i < 8 →
             Int.unsigned i = 0 Int.unsigned i = 1 Int.unsigned i = 2 Int.unsigned i = 3
             Int.unsigned i = 4 Int.unsigned i = 5 Int.unsigned i = 6 Int.unsigned i = 7.
Proof. intros; omega. Qed.

Section Refinement.

  Section WITHMEM.

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

Definition the refinement relation: relate_RData + match_RData

    Section REFINEMENT_REL.

      Inductive match_MCSLOCK_info: MultiLogTypeval
                                    (val × val) → (val × val) → (val × val) → (val × val) →
                                    (val × val) → (val × val) → (val × val) → (val × val) → Prop :=
      | MATCH_MCSLOCK_UNDEF: la bs0 ne0 bs1 ne1 bs2 ne2 bs3 ne3 bs4 ne4 bs5 ne5 bs6 ne6 bs7 ne7,
          match_MCSLOCK_info MultiUndef la (bs0, ne0) (bs1, ne1) (bs2, ne2) (bs3, ne3) (bs4, ne4)
                             (bs5, ne5) (bs6, ne6) (bs7, ne7)
      | MATCH_MCSLOCK_VALID: l la np tq bs0 ne0 bs1 ne1 bs2 ne2 bs3 ne3 bs4 ne4 bs5 ne5 bs6 ne6 bs7 ne7
                                    (Hcal: CalMCSLock l = Some (MCSLOCK la np tq))
                                    (Hnc0: ZMap.get 0 np = (bs0, ne0)) (Hnc1: ZMap.get 1 np = (bs1, ne1))
                                    (Hnc2: ZMap.get 2 np = (bs2, ne2)) (Hnc3: ZMap.get 3 np = (bs3, ne3))
                                    (Hnc4: ZMap.get 4 np = (bs4, ne4)) (Hnc5: ZMap.get 5 np = (bs5, ne5))
                                    (Hnc6: ZMap.get 6 np = (bs6, ne6)) (Hnc7: ZMap.get 7 np = (bs7, ne7)),
          match_MCSLOCK_info (MultiDef l) (Vint (Int.repr la))
                             ((Vint (Int.repr (BooltoZ bs0))), (Vint (Int.repr ne0)))
                             ((Vint (Int.repr (BooltoZ bs1))), (Vint (Int.repr ne1)))
                             ((Vint (Int.repr (BooltoZ bs2))), (Vint (Int.repr ne2)))
                             ((Vint (Int.repr (BooltoZ bs3))), (Vint (Int.repr ne3)))
                             ((Vint (Int.repr (BooltoZ bs4))), (Vint (Int.repr ne4)))
                             ((Vint (Int.repr (BooltoZ bs5))), (Vint (Int.repr ne5)))
                             ((Vint (Int.repr (BooltoZ bs6))), (Vint (Int.repr ne6)))
                             ((Vint (Int.repr (BooltoZ bs7))), (Vint (Int.repr ne7))).

Relation between the allocation table and the underline memory
      Inductive match_MCSLOCK: stencilMultiLogPoolmemProp :=
      | MATCH_MCSLOCK: log m b s
                              (HMCSLock:
                                 ( lock_index
                                         (Hvalid: 0 lock_index < lock_range),
                                     ( la bs0 ne0 bs1 ne1 bs2 ne2 bs3 ne3 bs4 ne4 bs5 ne5 bs6 ne6 bs7 ne7,
                                         Mem.load Mint32 m b (tail_pos lock_index) = Some la
                                         Mem.load Mint32 m b (busy_pos lock_index 0) = Some bs0
                                         Mem.load Mint32 m b (next_pos lock_index 0) = Some ne0
                                         Mem.load Mint32 m b (busy_pos lock_index 1) = Some bs1
                                         Mem.load Mint32 m b (next_pos lock_index 1) = Some ne1
                                         Mem.load Mint32 m b (busy_pos lock_index 2) = Some bs2
                                         Mem.load Mint32 m b (next_pos lock_index 2) = Some ne2
                                         Mem.load Mint32 m b (busy_pos lock_index 3) = Some bs3
                                         Mem.load Mint32 m b (next_pos lock_index 3) = Some ne3
                                         Mem.load Mint32 m b (busy_pos lock_index 4) = Some bs4
                                         Mem.load Mint32 m b (next_pos lock_index 4) = Some ne4
                                         Mem.load Mint32 m b (busy_pos lock_index 5) = Some bs5
                                         Mem.load Mint32 m b (next_pos lock_index 5) = Some ne5
                                         Mem.load Mint32 m b (busy_pos lock_index 6) = Some bs6
                                         Mem.load Mint32 m b (next_pos lock_index 6) = Some ne6
                                         Mem.load Mint32 m b (busy_pos lock_index 7) = Some bs7
                                         Mem.load Mint32 m b (next_pos lock_index 7) = Some ne7

                                         Mem.valid_access m Mint32 b (tail_pos lock_index) Writable
                                         Mem.valid_access m Mint32 b (busy_pos lock_index 0) Writable
                                         Mem.valid_access m Mint32 b (next_pos lock_index 0) Writable
                                         Mem.valid_access m Mint32 b (busy_pos lock_index 1) Writable
                                         Mem.valid_access m Mint32 b (next_pos lock_index 1) Writable
                                         Mem.valid_access m Mint32 b (busy_pos lock_index 2) Writable
                                         Mem.valid_access m Mint32 b (next_pos lock_index 2) Writable
                                         Mem.valid_access m Mint32 b (busy_pos lock_index 3) Writable
                                         Mem.valid_access m Mint32 b (next_pos lock_index 3) Writable
                                         Mem.valid_access m Mint32 b (busy_pos lock_index 4) Writable
                                         Mem.valid_access m Mint32 b (next_pos lock_index 4) Writable
                                         Mem.valid_access m Mint32 b (busy_pos lock_index 5) Writable
                                         Mem.valid_access m Mint32 b (next_pos lock_index 5) Writable
                                         Mem.valid_access m Mint32 b (busy_pos lock_index 6) Writable
                                         Mem.valid_access m Mint32 b (next_pos lock_index 6) Writable
                                         Mem.valid_access m Mint32 b (busy_pos lock_index 7) Writable
                                         Mem.valid_access m Mint32 b (next_pos lock_index 7) Writable
                                         
                                         match_MCSLOCK_info (ZMap.get lock_index log) la
                                                            (bs0, ne0) (bs1, ne1) (bs2, ne2) (bs3, ne3)
                                                            (bs4, ne4) (bs5, ne5) (bs6, ne6) (bs7, ne7))))
                              (Hsymbol: find_symbol s MCS_LOCK_LOC = Some b),
          match_MCSLOCK s log 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
                            (Hlog: match_MCSLOCK s (multi_log hadt) m),
          match_RData s hadt m f.

Relation between the shared 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);
            MM_re: MM hadt = MM ladt;
            MMSize_re: MMSize hadt = MMSize 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;
            ti_fst_re: (fst (ti hadt)) = (fst (ti ladt));
            ti_snd_re: val_inject f (snd (ti hadt)) (snd (ti ladt));
            init_re: init hadt = init ladt;

            buffer_re: buffer hadt = buffer 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;
            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)

          }.

      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 := MCS_LOCK_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; intros.
        inv Hlog.
        assert (HFB0: j b = Some (b, 0)).
        { eapply stencil_find_symbol_inject´; eauto. }
        econstructor; eauto; intros.
        - specialize (HMCSLock _ Hvalid).
          mcslock_destruct HMCSLock.
          sp_with_arg H0 HL__la HFB0;
            sp_with_arg H0 HL__bs0 HFB0; sp_with_arg H0 HL__ne0 HFB0; sp_with_arg H0 HL__bs1 HFB0; sp_with_arg H0 HL__ne1 HFB0;
            sp_with_arg H0 HL__bs2 HFB0; sp_with_arg H0 HL__ne2 HFB0; sp_with_arg H0 HL__bs3 HFB0; sp_with_arg H0 HL__ne3 HFB0;
            sp_with_arg H0 HL__bs4 HFB0; sp_with_arg H0 HL__ne4 HFB0; sp_with_arg H0 HL__bs5 HFB0; sp_with_arg H0 HL__ne5 HFB0;
            sp_with_arg H0 HL__bs6 HFB0; sp_with_arg H0 HL__ne6 HFB0; sp_with_arg H0 HL__bs7 HFB0; sp_with_arg H0 HL__ne7 HFB0.
          repeat rewrite Z.add_0_r;
            intros [ne7´ [HLD__ne7´ HV__ne7´]] [bs7´ [HLD__bs7´ HV__bs7´]] [ne6´ [HLD__ne6´ HV__ne6´]] [bs6´ [HLD__bs6´ HV__bs6´]]
[ne5´ [HLD__ne5´ HV__ne5´]] [bs5´ [HLD__bs5´ HV__bs5´]] [ne4´ [HLD__ne4´ HV__ne4´]] [bs4´ [HLD__bs4´ HV__bs4´]]
[ne3´ [HLD__ne3´ HV__ne3´]] [bs3´ [HLD__bs3´ HV__bs3´]] [ne2´ [HLD__ne2´ HV__ne2´]] [bs2´ [HLD__bs2´ HV__bs2´]]
[ne1´ [HLD__ne1´ HV__ne1´]] [bs1´ [HLD__bs1´ HV__bs1´]] [ne0´ [HLD__ne0´ HV__ne0´]] [bs0´ [HLD__bs0´ HV__bs0´]]
[la´ [HLD__la´ HV__la´]].
          refine_split´; eauto;
          [ sp_with_arg´ H0 HV__la HFB0 |
            sp_with_arg´ H0 HV__bs0 HFB0 | sp_with_arg´ H0 HV__ne0 HFB0 | sp_with_arg´ H0 HV__bs1 HFB0 | sp_with_arg´ H0 HV__ne1 HFB0 |
            sp_with_arg´ H0 HV__bs2 HFB0 | sp_with_arg´ H0 HV__ne2 HFB0 | sp_with_arg´ H0 HV__bs3 HFB0 | sp_with_arg´ H0 HV__ne3 HFB0 |
            sp_with_arg´ H0 HV__bs4 HFB0 | sp_with_arg´ H0 HV__ne4 HFB0 | sp_with_arg´ H0 HV__bs5 HFB0 | sp_with_arg´ H0 HV__ne5 HFB0 |
            sp_with_arg´ H0 HV__bs6 HFB0 | sp_with_arg´ H0 HV__ne6 HFB0 | sp_with_arg´ H0 HV__bs7 HFB0 | sp_with_arg´ H0 HV__ne7 HFB0 | ];
          try (rewrite Z.add_0_r; trivial; fail).
          + inv Hm; try constructor.
            inv HV__la´; inv HV__bs0´; inv HV__ne0´; inv HV__bs1´; inv HV__ne1´; inv HV__bs2´; inv HV__ne2´; inv HV__bs3´; inv HV__ne3´;
            inv HV__bs4´; inv HV__ne4´; inv HV__bs5´; inv HV__ne5´; inv HV__bs6´; inv HV__ne6´; inv HV__bs7´; inv HV__ne7´; econstructor; eauto.
      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 Hlog.
        econstructor; eauto.
        econstructor; eauto.
        intros.
        specialize (HMCSLock _ Hvalid).
        mcslock_destruct HMCSLock.
        eapply H0 in Hsymbol; simpl; eauto.
        repeat rewrite (Mem.load_store_other _ _ _ _ _ _ H1); auto.
        refine_split´; eauto; try (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 Hlog.
        econstructor; eauto.
        econstructor; eauto.
        intros.
        specialize (HMCSLock _ Hvalid).
        mcslock_destruct HMCSLock.
        eapply H0 in Hsymbol; simpl; eauto.
        repeat rewrite (Mem.load_storebytes_other _ _ _ _ _ H1); eauto.
        refine_split´; eauto; try (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 Hlog.
        econstructor; eauto.
        econstructor; eauto.
        intros.
        specialize (HMCSLock _ Hvalid).
        mcslock_destruct HMCSLock.
        eapply H0 in Hsymbol; simpl; eauto.
        repeat rewrite (Mem.load_free _ _ _ _ _ H1); auto.
        refine_split´; eauto; try (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 Hlog.
        econstructor; eauto.
        econstructor; eauto.
        intros.
        specialize (HMCSLock _ Hvalid).
        mcslock_destruct HMCSLock.
        refine_split´; eauto; try (apply (Mem.load_alloc_other _ _ _ _ _ H0); auto; fail); try (eapply Mem.valid_access_alloc_other; 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 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.

  End WITHMEM.

End Refinement.