Library mcertikos.virt.intel.VMCSInitGen

***********************************************************************
*                                                                     *
*            The CertiKOS Certified Kit Operating System              *
*                                                                     *
*                   The FLINT Group, Yale University                  *
*                                                                     *
*  Copyright The FLINT Group, Yale University.  All rights reserved.  *
*  This file is distributed under the terms of the Yale University    *
*  Non-Commercial License Agreement.                                  *
*                                                                     *
*********************************************************************** 


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 VVMCSInit.
Require Import VMCSInitGenSpec.
Require Import LoadStoreSem3.

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

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

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

  Section WITHMEM.

    Context `{Hstencil: Stencil}.
    Context `{Hmem: Mem.MemoryModel}.
    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;
          smspool_re: smspool hadt = smspool ladt;

          pb_re: pb hadt = pb ladt;
          kctxt_re: kctxt_inj f num_proc (kctxt hadt) (kctxt ladt);
          abtcb_re: abtcb hadt = abtcb ladt;
          abq_re: abq hadt = abq ladt;
          cid_re: cid hadt = cid ladt;
          chpool_re: syncchpool hadt = syncchpool ladt;
          uctxt_re: uctxt_inj f (uctxt hadt) (uctxt ladt);
          ept_re: ept hadt = ept ladt;
          vmcs_re: VMCS_inj f (vmcs hadt) (vmcs 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 hadt = in_intr ladt;
          drv_serial_re: drv_serial hadt = drv_serial 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.
        - eapply VMCS_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.

      Section FRESH_PRIM.

        Lemma vmx_set_intercept_intwin_kernel_mode:
           d2 d2´ v1,
            vmx_set_intercept_intwin_spec v1 d2 = Some d2´
            → kernel_mode d2.
        Proof.
          intros. functional inversion H; subst;
          simpl; auto.
        Qed.

        Lemma vmx_set_intercept_intwin_spec_ref:
          compatsim (crel HDATA LDATA) (gensem vmx_set_intercept_intwin_spec) vmx_set_intercept_intwin_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit vmx_set_intercept_intwin_exist; eauto 1.
          intros (labd´ & HP & HM).
          refine_split; try econstructor; eauto.
          - eapply vmx_set_intercept_intwin_kernel_mode; eauto.
          - constructor.
        Qed.

        Lemma vmx_set_desc_kernel_mode:
           d2 d2´ v1 v2 v3 v4 v5,
            vmx_set_desc_spec v1 v2 v3 v4 v5 d2 = Some d2´
            → kernel_mode d2.
        Proof.
          intros. functional inversion H; subst;
          simpl; auto.
        Qed.

        Lemma vmx_set_desc_spec_ref:
          compatsim (crel HDATA LDATA) (gensem vmx_set_desc_spec) vmx_set_desc_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit vmx_set_desc_exist; eauto 1.
          intros (labd´ & HP & HM).
          refine_split; try econstructor; eauto.
          - eapply vmx_set_desc_kernel_mode; eauto.
          - constructor.
        Qed.

        Lemma vmx_inject_event_kernel_mode:
           d2 d2´ v1 v2 v3 v4,
            vmx_inject_event_spec v1 v2 v3 v4 d2 = Some d2´
            → kernel_mode d2.
        Proof.
          intros. functional inversion H; subst;
          simpl; auto.
        Qed.

        Lemma vmx_inject_event_spec_ref:
          compatsim (crel HDATA LDATA) (gensem vmx_inject_event_spec) vmx_inject_event_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit vmx_inject_event_exist; eauto 1.
          intros (labd´ & HP & HM).
          refine_split; try econstructor; eauto.
          - eapply vmx_inject_event_kernel_mode; eauto.
          - constructor.
        Qed.

        Lemma vmx_set_tsc_offset_kernel_mode:
           d2 d2´ v1,
            vmx_set_tsc_offset_spec v1 d2 = Some d2´
            → kernel_mode d2.
        Proof.
          intros. functional inversion H; subst;
          simpl; auto.
        Qed.

        Lemma vmx_set_tsc_offset_spec_ref:
          compatsim (crel HDATA LDATA) (gensem vmx_set_tsc_offset_spec) vmx_set_tsc_offset_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit vmx_set_tsc_offset_exist; eauto 1.
          intros (labd´ & HP & HM).
          refine_split; try econstructor; eauto.
          - eapply vmx_set_tsc_offset_kernel_mode; eauto.
          - constructor.
        Qed.

        Lemma vmx_get_tsc_offset_kernel_mode:
           d2 d2´,
            vmx_get_tsc_offset_spec d2 = Some d2´
            → kernel_mode d2.
        Proof.
          intros. functional inversion H; subst;
          simpl; auto.
        Qed.

        Lemma vmx_get_tsc_offset_spec_ref:
          compatsim (crel HDATA LDATA) (gensem vmx_get_tsc_offset_spec) vmx_get_tsc_offset_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit vmx_get_tsc_offset_exists; eauto 1.
          refine_split; try econstructor; eauto.
          eapply vmx_get_tsc_offset_kernel_mode; eauto.
        Qed.

        Lemma vmx_get_exit_reason_kernel_mode:
           d2 d2´,
            vmx_get_exit_reason_spec d2 = Some d2´
            → kernel_mode d2.
        Proof.
          intros. functional inversion H; subst;
          simpl; auto.
        Qed.

        Lemma vmx_get_exit_reason_spec_ref:
          compatsim (crel HDATA LDATA) (gensem vmx_get_exit_reason_spec) vmx_get_exit_reason_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit vmx_get_exit_reason_exists; eauto 1.
          refine_split; try econstructor; eauto.
          eapply vmx_get_exit_reason_kernel_mode; eauto.
        Qed.

        Lemma vmx_get_exit_fault_addr_kernel_mode:
           d2 d2´,
            vmx_get_exit_fault_addr_spec d2 = Some d2´
            → kernel_mode d2.
        Proof.
          intros. functional inversion H; subst;
          simpl; auto.
        Qed.

        Lemma vmx_get_exit_fault_addr_spec_ref:
          compatsim (crel HDATA LDATA) (gensem vmx_get_exit_fault_addr_spec) vmx_get_exit_fault_addr_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit vmx_get_exit_fault_addr_exists; eauto 1.
          refine_split; try econstructor; eauto.
          eapply vmx_get_exit_fault_addr_kernel_mode; eauto.
        Qed.

        Lemma vmx_check_pending_event_kernel_mode:
           d2 d2´,
            vmx_check_pending_event_spec d2 = Some d2´
            → kernel_mode d2.
        Proof.
          intros. functional inversion H; subst;
          simpl; auto.
        Qed.

        Lemma vmx_check_pending_event_spec_ref:
          compatsim (crel HDATA LDATA) (gensem vmx_check_pending_event_spec) vmx_check_pending_event_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit vmx_check_pending_event_exists; eauto 1.
          refine_split; try econstructor; eauto.
          eapply vmx_check_pending_event_kernel_mode; eauto.
        Qed.

        Lemma vmx_check_int_shadow_kernel_mode:
           d2 d2´,
            vmx_check_int_shadow_spec d2 = Some d2´
            → kernel_mode d2.
        Proof.
          intros. functional inversion H; subst;
          simpl; auto.
        Qed.

        Lemma vmx_check_int_shadow_spec_ref:
          compatsim (crel HDATA LDATA) (gensem vmx_check_int_shadow_spec) vmx_check_int_shadow_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit vmx_check_int_shadow_exists; eauto 1.
          refine_split; try econstructor; eauto.
          eapply vmx_check_int_shadow_kernel_mode; eauto.
        Qed.

        Lemma vmcs_set_defaults_kernel_mode:
           d2 d2´ v1 pml4ept_b tss_b gdt_b idt_b msr_bitmap_b io_bitmap_b host_rip_b ,
            vmcs_set_defaults_spec v1 pml4ept_b tss_b gdt_b idt_b msr_bitmap_b io_bitmap_b host_rip_b d2 = Some d2´
            → kernel_mode d2.
        Proof.
          intros. functional inversion H; subst;
          simpl; auto.
        Qed.

        Lemma vmcs_set_defaults_spec_ref:
          compatsim (crel HDATA LDATA) (vmcs_set_defaults_compatsem vmcs_set_defaults_spec) vmcs_set_defaults_spec_low.
        Proof.
          compatsim_simpl (@match_AbData).
          exploit vmcs_set_defaults_exist; eauto 1; intros.
          - exploit (stencil_find_symbol_inject´ s ι EPT_LOC pml4ept_b); eauto.
            intros HFB.
            econstructor; eauto. rewrite Int.add_zero; trivial.
          - exploit (stencil_find_symbol_inject´ s ι tss_LOC tss_b); eauto.
            intros HFB.
            econstructor; eauto. rewrite Int.add_zero; trivial.
          - exploit (stencil_find_symbol_inject´ s ι gdt_LOC gdt_b); eauto.
            intros HFB.
            econstructor; eauto. rewrite Int.add_zero; trivial.
          - exploit (stencil_find_symbol_inject´ s ι idt_LOC idt_b); eauto.
            intros HFB.
            econstructor; eauto. rewrite Int.add_zero; trivial.
          - exploit (stencil_find_symbol_inject´ s ι msr_bitmap_LOC msr_bitmap_b); eauto.
            intros HFB.
            econstructor; eauto. rewrite Int.add_zero; trivial.
          - exploit (stencil_find_symbol_inject´ s ι io_bitmap_LOC io_bitmap_b); eauto.
            intros HFB.
            econstructor; eauto. rewrite Int.add_zero; trivial.
          - exploit (stencil_find_symbol_inject´ s ι vmx_return_from_guest host_rip_b); eauto.
            intros HFB.
            econstructor; eauto. rewrite Int.add_zero; trivial.
          - destruct H as (labd´ & HP & HM).
            refine_split; try econstructor; eauto.
            + eapply vmcs_set_defaults_kernel_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) vmcsinit_passthrough vmcsintro.
        Proof.
          sim_oplus.
          - apply fload_sim.
          - apply fstore_sim.
          - apply palloc_sim.
          - apply ptRead_sim.
          - apply ptResv_sim.
          - apply shared_mem_status_sim.
          - apply offer_shared_mem_sim.
          - apply get_curid_sim.
          - apply thread_wakeup_sim.
          - apply syncreceive_chan_sim.
          - apply syncsendto_chan_pre_sim.
          - apply syncsendto_chan_post_sim.
          - apply uctx_get_sim.
          - apply uctx_set_sim.
          - apply proc_create_sim.
          - apply rdmsr_sim.
          - apply wrmsr_sim.
          - apply cli_sim.
          - apply sti_sim.
          - apply cons_buf_read_sim.
          - apply serial_putc_sim.
          - apply serial_intr_disable_sim.
          - apply serial_intr_enable_sim.
          - apply vmcs_readz_sim.
          - apply vmcs_writez_sim.
          - apply vmcs_readz64_sim.
          - apply vmcs_writez64_sim.
          - apply vmptrld_sim.
          - apply ept_add_mapping_sim.
          - apply ept_invalidate_mappings_sim.
          - apply container_get_quota_sim.
          - apply container_get_usage_sim.
          - apply container_can_consume_sim.
          - apply hostin_sim.
          - apply hostout_sim.
          - apply trap_info_get_sim.
          - apply trap_info_ret_sim.
          - apply thread_yield_sim.
          - apply thread_sleep_sim.
          - apply proc_start_user_sim.
            intros. eapply valid_curid; eauto.
          - apply proc_exit_user_sim.
          - layer_sim_simpl.
            + eapply load_correct3.
            + eapply store_correct3.
        Qed.

      End PASSTHROUGH_PRIM.

    End OneStep_Forward_Relation.

  End WITHMEM.

End Refinement.