Library mcertikos.trap.TrapArgGen


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

Require Import TTrapArg.
Require Import TrapArgGenSpec.

Require Import GlobalOracleProp.
Require Import SingleOracle.

Definition of the refinement relation

Section Refinement.

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

  Context `{single_oracle_prop: SingleOracleProp}.
  Context `{real_params : RealParams}.
  Context `{multi_oracle_prop: MultiOracleProp}.

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

  Notation HDATAOps := (cdata (cdata_ops := pipcintro_data_ops) HDATA).
  Notation LDATAOps := (cdata (cdata_ops := pipcintro_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;
          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;
          

          CPU_ID_re: CPU_ID hadt = CPU_ID ladt;
          cid_re: cid hadt = cid ladt;
          big_log_re: big_log hadt = big_log ladt;
          big_oracle_re: big_oracle hadt = big_oracle 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;
          

          

          ept_re: ept hadt = ept ladt;
          vmcs_re: VMCSPool_inj f (vmcs hadt) (vmcs ladt);
          vmx_re: VMXPool_inj f (vmx hadt) (vmx 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 kctxt_inj_incr; eauto.
        - eapply uctxt_inj_incr; eauto.
        - unfold VMCSPool_inj; intros.
          eapply VMCS_inj_incr; eauto.
        - unfold VMXPool_inj; intros.
          eapply VMX_inj_incr; eauto.
      Qed.

    End Rel_Property.

    Global Instance rel_prf: CompatRel HDATAOps LDATAOps.
    Proof.
      constructor; intros; simpl; trivial.
      eapply relate_incr; eauto.
    Qed.

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

    Section OneStep_Forward_Relation.

The low level specifications exist

      Section Exists.

        Lemma uctx_argn_exist´:
           n s habd z labd f,
            uctx_argn_spec n habd = Some z
            0 n < UCTXT_SIZE
            0 ZMap.get (CPU_ID habd) (cid habd) < num_proc
            relate_AbData s f habd labd
            uctx_argn_spec n labd = Some z
             kernel_mode labd.
        Proof.
          intros. split.
          - eapply uctx_argn_exist; eauto.
          - unfold uctx_argn_spec in ×.
            inv H2. revert H; subrewrite.
            simpl; subdestruct. eauto.
        Qed.

        Lemma uctx_set_regk_exist´:
           k s habd habd´ labd n f,
            uctx_set_regk_spec k n habd = Some habd´
            → relate_AbData s f habd labd
            → 0 ZMap.get (CPU_ID habd) (cid habd) < num_proc
            → labd´, uctx_set_regk_spec k n labd = Some labd´
                              relate_AbData s f habd´ labd´
                              kernel_mode labd.
        Proof.
          intros. exploit uctx_set_regk_exist; eauto.
          intros (labd´ & He & HR).
          refine_split´; eauto.
          unfold uctx_set_regk_spec in He.
          simpl; subdestruct. eauto.
        Qed.

      End Exists.

      Section FRESH_PRIM.

        Lemma uctx_arg1_spec_ref:
          compatsim (crel HDATA LDATA) (gensem uctx_arg1_spec) uctx_arg1_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit uctx_argn_exist´; eauto 1.
          - omega.
          - eapply valid_curid; eauto.
          - intros [HP Hkern].
            refine_split; try econstructor; eauto.
        Qed.

        Lemma uctx_arg2_spec_ref:
          compatsim (crel HDATA LDATA) (gensem uctx_arg2_spec) uctx_arg2_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit uctx_argn_exist´; eauto 1.
          - omega.
          - eapply valid_curid; eauto.
          - intros [HP Hkern].
            refine_split; try econstructor; eauto.
        Qed.

        Lemma uctx_arg3_spec_ref:
          compatsim (crel HDATA LDATA) (gensem uctx_arg3_spec) uctx_arg3_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit uctx_argn_exist´; eauto 1.
          - omega.
          - eapply valid_curid; eauto.
          - intros [HP Hkern].
            refine_split; try econstructor; eauto.
        Qed.

        Lemma uctx_arg4_spec_ref:
          compatsim (crel HDATA LDATA) (gensem uctx_arg4_spec) uctx_arg4_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          inv Hhigh´.
          exploit uctx_argn_exist´; eauto 1.
          - intros [HP Hkern].
            refine_split; try econstructor; eauto.
        Qed.

        Lemma uctx_set_errno_spec_ref:
          compatsim (crel HDATA LDATA) (gensem uctx_set_errno_spec) uctx_set_errno_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit uctx_set_regk_exist´; eauto 1.
          - eapply valid_curid; eauto.
          - intros [labd´ [HP [HM Hkern]]].
            refine_split; try econstructor; eauto.
            constructor.
        Qed.

        Lemma uctx_set_retval1_spec_ref:
          compatsim (crel HDATA LDATA) (gensem uctx_set_retval1_spec) uctx_set_retval1_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit uctx_set_regk_exist´; eauto 1.
          - eapply valid_curid; eauto.
          - intros [labd´ [HP [HM Hkern]]].
            refine_split; try econstructor; eauto.
            constructor.
        Qed.

        Lemma uctx_set_retval2_spec_ref:
          compatsim (crel HDATA LDATA) (gensem uctx_set_retval2_spec) uctx_set_retval2_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit uctx_set_regk_exist´; eauto 1.
          - eapply valid_curid; eauto.
          - intros [labd´ [HP [HM Hkern]]].
            refine_split; try econstructor; eauto.
            constructor.
        Qed.

        Lemma uctx_set_retval3_spec_ref:
          compatsim (crel HDATA LDATA) (gensem uctx_set_retval3_spec) uctx_set_retval3_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit uctx_set_regk_exist´; eauto 1.
          - eapply valid_curid; eauto.
          - intros [labd´ [HP [HM Hkern]]].
            refine_split; try econstructor; eauto.
            constructor.
        Qed.

      End FRESH_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) ttraparg_passthrough pipc.
      Proof.
        sim_oplus.
        - apply thread_vmxinfo_get_sim.
        - apply big2_palloc_sim.
        - apply thread_ptRead_sim.
        - apply big2_ptResv_sim.         - apply big2_thread_yield_sim.

        - apply thread_uctx_get_sim.
        - apply thread_uctx_set_sim.
        - apply big2_proc_create_sim.

        - apply thread_container_get_quota_sim.
        - apply thread_container_get_usage_sim.
        - apply thread_container_can_consume_sim.

        - apply thread_get_CPU_ID_sim.
        - apply thread_get_curid_sim.

        - apply thread_syncsendto_chan_sim.
        - apply thread_syncreceive_chan_sim.

        - apply thread_rdmsr_sim.
        - apply thread_wrmsr_sim.

        - apply thread_vmx_set_intercept_intwin_sim.
        - apply thread_vmx_set_desc1_sim.
        - apply thread_vmx_set_desc2_sim.
        - apply thread_vmx_inject_event_sim.
        - apply thread_vmx_set_tsc_offset_sim.
        - apply thread_vmx_get_tsc_offset_sim.
        - apply thread_vmx_get_exit_reason_sim.
        - apply thread_vmx_get_exit_fault_addr_sim.
        - apply thread_vmx_get_exit_qualification_sim.
        - apply thread_vmx_check_pending_event_sim.
        - apply thread_vmx_check_int_shadow_sim.
        - apply thread_vmx_get_reg_sim.
        - apply thread_vmx_set_reg_sim.

        - apply thread_vmx_get_next_eip_sim.
        - apply thread_vmx_get_io_width_sim.
        - apply thread_vmx_get_io_write_sim.
        - apply thread_vmx_get_exit_io_rep_sim.
        - apply thread_vmx_get_exit_io_str_sim.
        - apply thread_vmx_get_exit_io_port_sim.
        - apply thread_vmx_set_mmap_sim.
        - apply thread_vm_run_sim, VMX_INJECT.
        - apply thread_vmx_return_from_guest_sim.
        - apply thread_vmx_init_sim.

        - apply thread_cli_sim.
        - apply thread_sti_sim.
        - apply thread_serial_intr_disable_sim.
        - apply thread_serial_intr_enable_sim.
        - apply thread_serial_putc_sim.
        - apply thread_cons_buf_read_sim.

        - apply thread_proc_create_postinit_sim.
        - apply thread_trap_info_get_sim.
        - apply thread_trap_info_ret_sim.
        - apply thread_proc_start_user_sim.
          intros. eapply valid_curid; eauto.
        - apply thread_proc_exit_user_sim.
        - apply thread_proc_start_user_sim2.
          intros. eapply valid_curid; eauto.
        - apply thread_proc_exit_user_sim2.
        - layer_sim_simpl.
          + eapply load_correct3high.
          + eapply store_correct3high.
      Qed.

    End OneStep_Forward_Relation.

  End WITHMEM.

End Refinement.