Library mcertikos.devdrivers.HandlerSwGen


This file provide the contextual refinement proof between PThreadInit layer and PQueueIntro 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 GenSem.
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 LayerCalculusLemma.
Require Import AbstractDataType.

Require Import DHandlerSw.
Require Import HandlerSwGenSpec.
Require Import DeviceStateDataType.

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

  Section WITHMEM.

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

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;
          CR3_re: CR3 hadt = CR3 ladt;
          ikern_re: ikern hadt = ikern ladt;
          MM_re: MM hadt = MM ladt;
          MMSize_re: MMSize hadt = MMSize 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));
          AT_re: AT hadt = AT ladt;
          nps_re: nps hadt = nps ladt;
          init_re: init hadt = init ladt;

          buffer_re: buffer hadt = buffer 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;
          drv_serial_re: drv_serial hadt = drv_serial ladt;
          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
        }.

    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 serial_intr_handler_sw_kern_mode:
           d ,
            serial_intr_handler_sw_spec d = Some
            → kernel_mode d.
        Proof.
          unfold serial_intr_handler_sw_spec. simpl; intros.
          subdestruct; auto.
        Qed.

        Lemma serial_intr_handler_sw_exist:
           s habd habd´ labd f,
            serial_intr_handler_sw_spec habd = Some habd´
            → relate_AbData s f habd labd
            → labd´, serial_intr_handler_sw_spec labd = Some labd´
                              relate_AbData s f habd´ labd´.
        Proof.
          unfold serial_intr_handler_sw_spec; intros.
          exploit relate_impl_iflags_eq; eauto; inversion 1.
          exploit relate_impl_ioapic_eq; eauto; inversion 1.
          exploit relate_impl_init_eq; eauto; inversion 1.
          exploit relate_impl_in_intr_eq; eauto; inversion 1.
          exploit relate_impl_intr_flag_eq; eauto; inversion 1.
          revert H. subrewrite.
          subdestruct.
          exploit (lapic_eoi_exist s _ _ ((labd { ioapic / s / IoApicEnables [Pos.to_nat 4]:false}) {
                                              intr_flag : true}) f Hdestruct3); eauto.
          apply relate_impl_intr_flag_update; eauto.
          apply relate_impl_ioapic_update; eauto.
          intros (labd´ & HP´ & HM´).
          exploit (cons_intr_aux_exist s _ _ labd´ f Hdestruct4); eauto.
          intros (labd´´ & HP´´ & HM´´).
          subdestruct;
            inv HQ; refine_split´; trivial.
          rewrite HP´.
          rewrite HP´´.
          reflexivity.
          inv HM´´.
          econstructor; simpl; eauto.
          rewrite ioapic_re0.
          reflexivity.
        Qed.

        Lemma serial_intr_handler_sw_spec_ref:
          compatsim (crel HDATA LDATA) (gensem serial_intr_handler_sw_spec) serial_intr_handler_sw_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit serial_intr_handler_sw_exist; eauto 1.
          intros [labd´ [HP HM]].
          refine_split; try econstructor; eauto.
          - eapply serial_intr_handler_sw_kern_mode; eauto.
          - constructor.
        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) dhandlersw_passthrough dhandlerintro.
        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; try 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 ioapic_mask_sim.
          - apply ioapic_unmask_sim.
          - apply serial_putc_sim.
          - apply serial_hw_intr_sim.
          - apply cons_buf_read_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.