Library mcertikos.ticketlog.TicketLockOpGen


This file provide the contextual refinement proof between MALInit layer and MALOp 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 LAsm.
Require Import RefinementTactic.
Require Import PrimSemantics.

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 MTicketLockOp.
Require Import TicketLockOpGenSpec.
Require Import DeviceStateDataType.

Notation 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 `{multi_oracle_prop: MultiOracleProp}.

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

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

  Section WITHMEM.

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

Definition the refinement relation: relate_RData + match_RData

    Record relate_RData (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)
        }.

    Inductive match_RData: stencilHDATAmemmeminjProp :=
    | MATCH_RDATA: habd m f s, match_RData s habd m f.

    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 := nil
      }.

Properties of relations

    Section Rel_Property.

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; intros; simpl; trivial.
        eapply relate_incr; eauto.
      Qed.

    End Rel_Property.

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

    Section OneStep_Forward_Relation.

      Section FRESH_PRIM.

        Lemma ticket_lock_init_kern_mode:
           i d ,
            ticket_lock_init_spec i d = Some
            → kernel_mode d.
        Proof.
          unfold ticket_lock_init_spec. simpl; intros.
          subdestruct. auto.
        Qed.

        Lemma ticket_lock_init_exist:
           habd habd´ labd n f,
            ticket_lock_init_spec n habd = Some habd´
            → relate_RData f habd labd
            → labd´, ticket_lock_init_spec n labd = Some labd´ relate_RData f habd´ labd´.
        Proof.
          unfold ticket_lock_init_spec; intros until f; exist_simpl.
        Qed.

        Lemma ticket_lock_init_spec_ref:
          compatsim (crel HDATA LDATA) (gensem ticket_lock_init_spec) ticket_lock_init_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit ticket_lock_init_exist; eauto 1.
          intros [labd´ [HP HM]].
          refine_split; try econstructor; eauto.
          - eapply ticket_lock_init_kern_mode; eauto.
          - constructor.
        Qed.

        Lemma pass_lock_kern_mode:
           i ofs d ,
            pass_lock_spec i ofs d = Some
            → kernel_mode d.
        Proof.
          unfold pass_lock_spec. simpl; intros.
          subdestruct. auto.
        Qed.

        Lemma pass_lock_exist:
           habd habd´ labd n ofs f,
            pass_lock_spec n ofs habd = Some habd´
            → relate_RData f habd labd
            → labd´, pass_lock_spec n ofs labd = Some labd´ relate_RData f habd´ labd´.
        Proof.
          unfold pass_lock_spec; intros until f; exist_simpl.
        Qed.

        Lemma pass_lock_spec_ref:
          compatsim (crel HDATA LDATA) (gensem pass_lock_spec) pass_lock_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit pass_lock_exist; eauto 1.
          intros [labd´ [HP HM]].
          refine_split; try econstructor; eauto.
          - eapply pass_lock_kern_mode; eauto.
          - constructor.
        Qed.

        Lemma wait_lock_kern_mode:
           bound i ofs d ,
            wait_lock_spec_wraparound bound i ofs d = Some
            → kernel_mode d.
        Proof.
          unfold wait_lock_spec_wraparound. simpl; intros.
          subdestruct; auto.
        Qed.

        Lemma CalTicketLockExist:
           l t n,
            CalTicketLock l = Some (t, n, )
            CalTicketLockWraparound l = Some (Int.unsigned (Int.repr t), Int.unsigned (Int.repr n), ).
        Proof.
          Local Transparent CalTicketLock CalTicketLockWraparound.
          induction l.
          simpl.
          intros.
          inv H.
          reflexivity.
          simpl.
          intros.
          assert (zeq: z, Int.unsigned (Int.repr (Int.unsigned (Int.repr z) + 1)) = Int.unsigned (Int.repr (z + 1))).
          {
            intros.
            repeat rewrite Int.unsigned_repr_eq.
            unfold Int.modulus, two_power_nat, shift_nat.
            simpl.
            Require Import XOmega.
            repeat match goal with | [H: _ |- _] ⇒ clear H end.
            xomega.
          }
          subdestruct; inv H;
          erewrite IHl; eauto;
          rewrite Int.repr_unsigned;
          try rewrite Int.repr_unsigned;
          try rewrite zeq;
          try reflexivity.
        Qed.

        Lemma CalTicketLock_range:
           q l t t0 n n0 l0
                 (Hcal: CalTicketLock l = Some (t, n, ))
                 (Hcal0: CalTicketLock (q ++ l) = Some (t0, n0, l0)),
            t t0 n n0.
        Proof.
          induction q; simpl; intros.
          - rewrite Hcal in Hcal0. inv Hcal0.
            split; reflexivity.
          - subdestruct; contra_inv; inv Hcal0; eauto.
            + exploit (IHq l); eauto.
              intros (? & ?). split; omega.
            + exploit (IHq l); eauto.
              intros (? & ?). split; omega.
            + exploit (IHq l); eauto.
              intros (? & ?). split; omega.
        Qed.

        Lemma CalTicketWait_range:
           n i t l to t0 n0 l0
                 (Hwait: CalTicketWait n i t l to = Some )
                 (Hcal: CalTicketLock l = Some (t0, n0, l0)),
            n0 t.
        Proof.
          Local Transparent CalTicketWait.
          Opaque CalTicketLock.
          induction n; simpl; intros; try congruence.
          subdestruct; subst.
          - rewrite app_comm_cons in Hdestruct.
            eapply CalTicketLock_range; eauto.
          - eapply IHn in Hwait; try eapply Hdestruct.
            rewrite app_comm_cons in Hdestruct.
            exploit CalTicketLock_range; eauto.
            intros (_ & ?). omega.
        Qed.

        Lemma CalTicketWaitExist:
           n l i t to t0 n0 l0,
            CalTicketLock l = Some (t0, n0, l0)
            t < n0 + TOTAL_CPU
            CalTicketWait n i t l to = Some
            CalTicketWaitWraparound n i (Int.unsigned (Int.repr t)) l to = Some .
        Proof.
          Local Transparent CalTicketWait CalTicketWaitWraparound.
          Opaque CalTicketLock CalTicketLockWraparound.
          induction n.
          intros.
          simpl in ×.
          auto.
          intros.
          simpl in ×.
          subdestruct; inv H.
          - eapply CalTicketLockExist in Hdestruct; eauto.
            rewrite Hdestruct.
            rewrite zeq_true. eauto.
          - pose proof Hdestruct as Ht.
            eapply CalTicketLockExist in Hdestruct; eauto.
            rewrite Hdestruct.
            assert (Ht_range: z0 < t < z0 + TOTAL_CPU).
            {
              pose proof H1 as Hwait.
              eapply CalTicketWait_range in Hwait; eauto.
              split; [omega|].
              rewrite app_comm_cons in Ht.
              eapply CalTicketLock_range in H3; eauto.
              destruct H3 as (_ & ?). omega.
            }
            rewrite zeq_false.
            eapply IHn; eauto.
            + omega.
            + revert Ht_range. clear; intros.
              repeat rewrite Int.unsigned_repr_eq.
              unfold Int.modulus, two_power_nat, shift_nat.
              simpl.
              assert (4294967296 > 0) by omega.
              repeat rewrite Zmod_eq; trivial.
              assert (Ht: a, t = z0 + a 0 < a < 8).
              {
                 (t - z0).
                split; omega.
              }
              destruct Ht as (a & Ht & Ha). subst. clear Ht_range.
              assert (Hz0: c d, z0 = c × 4294967296 + d 0 d < 4294967296).
              {
                pose proof (Z_div_mod_eq z0 _ H) as Hz0.
                replace (4294967296 × (z0 / 4294967296)) with ((z0 / 4294967296) × 4294967296) in Hz0 by omega.
                 (z0 / 4294967296), (z0 mod 4294967296).
                split; trivial.
                eapply Z_mod_lt; eauto.
              }
              destruct Hz0 as (c & d & Hz0 & Hd).
              rewrite Hz0.
              replace (c × 4294967296 + d + a) with (c × 4294967296 + (d + a)) by omega.
              assert (4294967296 0) by omega.
              repeat rewrite Z_div_plus_full_l; trivial.
              repeat rewrite Z.mul_add_distr_r.
              repeat rewrite Zminus_plus_simpl_l.
              assert (Hdv: d / 4294967296 = 0) by xomega.
              rewrite Hdv. simpl.
              replace (d - 0) with d by omega.
              assert (Hr: 0 (d + a) < 4294967296 4294967296 (d + a) < 4294967296 + TOTAL_CPU).
              {
                destruct (zlt (d + a) 4294967296).
                - left. omega.
                - right. omega.
              }
              destruct Hr.
              {
                assert (Hdv´: (d + a) / 4294967296 = 0) by xomega.
                rewrite Hdv´. simpl. omega.
              }
              {
                assert (Hdv´: (d + a) / 4294967296 = 1) by xomega.
                rewrite Hdv´. simpl. omega.
              }
        Qed.

        Lemma wait_lock_exist:
           habd habd´ labd bound n ofs s f,
            wait_lock_spec bound n ofs habd = Some habd´
            → relate_AbData s f habd labd
            → labd´, wait_lock_spec_wraparound bound n ofs labd = Some labd´ relate_AbData s f habd´ labd´.
        Proof.
          Opaque CalTicketLock CalTicketLockWraparound CalTicketWait CalTicketWaitWraparound.
          unfold wait_lock_spec, wait_lock_spec_wraparound.
          intros.
          exploit relate_impl_multi_oracle_eq; eauto; inversion 1.
          exploit relate_impl_multi_log_eq; eauto; inversion 1.
          exploit relate_impl_CPU_ID_eq; eauto; inversion 1.
          exploit relate_impl_iflags_eq; eauto; inversion 1.
          subdestruct.
          rewrite <- ikern_eq, <- ihost_eq.
          match goal with
            | [H: CalTicketWait _ _ _ _ _ = Some _ |- _] ⇒
              eapply CalTicketWaitExist in H; eauto
          end.
          match goal with
            | [H: CalTicketLock _ = Some (_, _, _) |- _] ⇒
              eapply CalTicketLockExist in H; eauto
          end.
          subrewrite´.
          assert(z0eq: (Int.unsigned (Int.repr (Int.unsigned (Int.repr z0) - 1))) = (Int.unsigned (Int.repr (z0 - 1)))).
          {
            rewrite Int.unsigned_repr_eq.
            rewrite Int.unsigned_repr_eq.
            rewrite Int.unsigned_repr_eq.
            Require Import XOmega.
            unfold Int.modulus, two_power_nat, shift_nat; simpl.
            repeat match goal with | [H: _ |- _] ⇒ clear H end.
            xomega.
          }
          rewrite z0eq.
          exploit relate_impl_multi_oracle_eq; eauto; inversion 1.
          exploit relate_impl_multi_log_eq; eauto; inversion 1.
          exploit relate_impl_CPU_ID_eq; eauto; inversion 1.
          subrewrite´.
          refine_split.
          reflexivity.
          inv H. inv H0.
          subrewrite´.
          econstructor; simpl; eauto.
          omega.
        Qed.

        Lemma wait_lock_spec_ref:
          compatsim (crel HDATA LDATA) (gensem wait_lock_spec) wait_lock_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit wait_lock_exist; eauto 1.
          intros [labd´ [HP HM]].
          refine_split; try econstructor; eauto.
          - eapply wait_lock_kern_mode; eauto.
          - constructor.
        Qed.

      End FRESH_PRIM.

      Section PASSTHROUGH_RPIM.

        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) mticketlockop_passthrough mticketlockintro.
        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 release_shared0_sim.
          - apply (acquire_shared0_sim (valid_id_args:= Shared2ID_valid0)).
            intros. inv H.

          - apply get_curid_sim.
          - apply set_curid_sim.
          - apply set_curid_init_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.
          - 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.
          - layer_sim_simpl.
            + eapply load_correct1.
            + eapply store_correct1.
        Qed.

      End PASSTHROUGH_RPIM.

    End OneStep_Forward_Relation.

  End WITHMEM.

End Refinement.