Library mcertikos.virt.intel.VVMCSInit

***********************************************************************
*                                                                     *
*            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 Maps.
Require Import ASTExtra.
Require Import Integers.
Require Import Floats.
Require Import Values.
Require Import Memory.
Require Import Events.
Require Import Stacklayout.
Require Import Globalenvs.
Require Import AsmX.
Require Import Smallstep.
Require Import AuxStateDataType.
Require Import Constant.
Require Import GlobIdent.
Require Import FlatMemory.
Require Import CommonTactic.
Require Import AuxLemma.
Require Import RealParams.
Require Import PrimSemantics.
Require Import LAsm.
Require Import XOmega.

Require Import liblayers.logic.PTreeModules.
Require Import liblayers.logic.LayerLogicImpl.
Require Import liblayers.compat.CompatLayers.
Require Import liblayers.compat.CompatGenSem.

Require Import CalRealPTPool.
Require Import CalRealPT.
Require Import CalRealIDPDE.
Require Import CalRealInitPTE.
Require Import CalRealPTB.
Require Import CalRealSMSPool.
Require Import CalRealProcModule.

Require Import INVLemmaContainer.
Require Import INVLemmaMemory.
Require Import INVLemmaThread.
Require Import INVLemmaProc.
Require Import INVLemmaIntel.

Require Import AbstractDataType.

Require Export VVMCSIntro.
Require Import LoadStoreSem3.


Abstract Data and Primitives at this layer

Section WITHMEM.

  Local Open Scope Z_scope.

  Context `{real_params: RealParams}.

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

Proofs that the primitives satisfies the invariants at this layer

  Section INV.

    Section INTERCEPT_INTWIN.

      Lemma vmx_set_intercept_intwin_high_inv:
         d i,
          vmx_set_intercept_intwin_spec i d = Some
          high_level_invariant d
          high_level_invariant .
      Proof.
        intros. inv H0; functional inversion H; subst; simpl;
        constructor; simpl; intros; try congruence; eauto 2.
      Qed.

      Lemma vmx_set_intercept_intwin_low_inv:
         d n i,
          vmx_set_intercept_intwin_spec i d = Some
          low_level_invariant n d
          low_level_invariant n .
      Proof.
        intros. inv H0; functional inversion H; subst; simpl;
        constructor; simpl; intros; try congruence; eauto 2;
        eapply VMCS_inject_neutral_gss_Vint; eauto.
      Qed.

      Lemma vmx_set_intercept_intwin_kernel_mode:
         d i,
          vmx_set_intercept_intwin_spec i d = Some
          kernel_mode d
          kernel_mode .
      Proof.
        intros. functional inversion H; subst; constructor; trivial.
      Qed.

      Global Instance vmx_set_intercept_intwin_inv: PreservesInvariants vmx_set_intercept_intwin_spec.
      Proof.
        preserves_invariants_simpl´.
        - eapply vmx_set_intercept_intwin_low_inv; eauto.
        - eapply vmx_set_intercept_intwin_high_inv; eauto.
        - eapply vmx_set_intercept_intwin_kernel_mode; eauto.
      Qed.

    End INTERCEPT_INTWIN.

    Section INJECT_EVENT.

      Lemma vmx_inject_event_high_inv:
         d i i0 i1 i2,
          vmx_inject_event_spec i i0 i1 i2 d = Some
          high_level_invariant d
          high_level_invariant .
      Proof.
        intros. inv H0; functional inversion H; subst; simpl;
        constructor; simpl; intros; try congruence; eauto 2.
      Qed.

      Lemma vmx_inject_event_low_inv:
         d n i i0 i1 i2,
          vmx_inject_event_spec i i0 i1 i2 d = Some
          low_level_invariant n d
          low_level_invariant n .
      Proof.
        intros. inv H0; functional inversion H; subst; simpl;
        constructor; simpl; intros; try congruence; eauto 2;
        rewrite H4 in *; repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
      Qed.

      Lemma vmx_inject_event_kernel_mode:
         d i i0 i1 i2,
          vmx_inject_event_spec i i0 i1 i2 d = Some
          kernel_mode d
          kernel_mode .
      Proof.
        intros. functional inversion H; subst; constructor; trivial.
      Qed.

      Global Instance vmx_inject_event_inv: PreservesInvariants vmx_inject_event_spec.
      Proof.
        preserves_invariants_simpl´.
        - eapply vmx_inject_event_low_inv; eauto.
        - eapply vmx_inject_event_high_inv; eauto.
        - eapply vmx_inject_event_kernel_mode; eauto.
      Qed.

    End INJECT_EVENT.

    Section SET_DESC.

      Lemma vmx_set_desc_high_inv:
         d i i0 i1 i2 i3,
          vmx_set_desc_spec i i0 i1 i2 i3 d = Some
          high_level_invariant d
          high_level_invariant .
      Proof.
        intros. inv H0; functional inversion H; subst; simpl;
        constructor; simpl; intros; try congruence; eauto 2.
      Qed.

      Lemma vmx_set_desc_low_inv:
         d n i i0 i1 i2 i3,
          vmx_set_desc_spec i i0 i1 i2 i3 d = Some
          low_level_invariant n d
          low_level_invariant n .
      Proof.
        intros. inv H0; functional inversion H; subst; simpl;
        constructor; simpl; intros; try congruence; eauto 2.
        functional inversion H5;
        rewrite H4 in *; repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
      Qed.

      Lemma vmx_set_desc_kernel_mode:
         d i i0 i1 i2 i3,
          vmx_set_desc_spec i i0 i1 i2 i3 d = Some
          kernel_mode d
          kernel_mode .
      Proof.
        intros. functional inversion H; subst; constructor; trivial.
      Qed.

      Global Instance vmx_set_desc_inv: PreservesInvariants vmx_set_desc_spec.
      Proof.
        preserves_invariants_simpl´.
        - eapply vmx_set_desc_low_inv; eauto.
        - eapply vmx_set_desc_high_inv; eauto.
        - eapply vmx_set_desc_kernel_mode; eauto.
      Qed.

    End SET_DESC.

    Section SET_TSC_OFFSET.

      Lemma vmx_set_tsc_offset_high_inv:
         d i,
          vmx_set_tsc_offset_spec i d = Some
          high_level_invariant d
          high_level_invariant .
      Proof.
        intros. inv H0; functional inversion H; subst; simpl;
        constructor; simpl; intros; try congruence; eauto 2.
      Qed.

      Lemma vmx_set_tsc_offset_low_inv:
         d n i,
          vmx_set_tsc_offset_spec i d = Some
          low_level_invariant n d
          low_level_invariant n .
      Proof.
        intros. inv H0; functional inversion H; subst; simpl;
        constructor; simpl; intros; try congruence; eauto 2.
        rewrite H4 in *; repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
      Qed.

      Lemma vmx_set_tsc_offset_kernel_mode:
         d i,
          vmx_set_tsc_offset_spec i d = Some
          kernel_mode d
          kernel_mode .
      Proof.
        intros. functional inversion H; subst; constructor; trivial.
      Qed.

      Global Instance vmx_set_tsc_offset_inv: PreservesInvariants vmx_set_tsc_offset_spec.
      Proof.
        preserves_invariants_simpl´.
        - eapply vmx_set_tsc_offset_low_inv; eauto.
        - eapply vmx_set_tsc_offset_high_inv; eauto.
        - eapply vmx_set_tsc_offset_kernel_mode; eauto.
      Qed.

    End SET_TSC_OFFSET.

    Global Instance vmcs_set_defaults_inv: VMCSSetDefaultsInvariants vmcs_set_defaults_spec.
    Proof.
      constructor; intros; functional inversion H; simpl; auto.
      -
        inv H7. constructor; trivial; intros; simpl in ×.
        eapply real_vmcs_inject_neutral; eauto.
      -
        inv H0. constructor; simpl; try congruence; intros.
        + apply real_nps_range.
        + apply real_lat_kern_valid.
        + apply real_lat_usr_valid.
        + apply AC_init_container_valid.
        + rewrite init_pperm; try assumption.
          apply Lreal_pperm_valid.
        + eapply real_pt_PMap_valid; eauto.
        + apply real_pt_PMap_kern.
        + omega.
        + assumption.
        + apply real_idpde_init.
        + apply real_pt_consistent_pmap.
        + apply real_pt_consistent_pmap_domain.
        + eapply Lreal_at_consistent_lat_domain; eauto.
        + apply real_PTB_0_valid.
        + apply real_PTB_valid.
        + eapply real_PTB_AC_valid; eauto.
        + apply real_abtcb_strong_range´; eauto.
        + apply real_abq_range; auto.
        + eapply real_abtcb_pb_notInQ´; eauto.
        + eapply real_abtcb_abq_QCount´; eauto.
        + eapply real_abq_tcb_inQ; eauto.
        + omega.
        + eapply real_abtcb_PTB_CurIDValid; eauto.
        + eapply real_abtcb_SingleRun; eauto.
        + eapply real_syncchpool_valid´; eauto.
    Qed.

  End INV.

These should be put into mcertikos.layerlib.GlobIdent
  Definition vmcsinit_fresh : compatlayer (cdata RData) :=
    vmx_set_intercept_intwin gensem vmx_set_intercept_intwin_spec
                              vmx_set_desc gensem vmx_set_desc_spec
                              vmx_inject_event gensem vmx_inject_event_spec
                              vmx_set_tsc_offset gensem vmx_set_tsc_offset_spec
                              vmx_get_tsc_offset gensem vmx_get_tsc_offset_spec
                              vmx_get_exit_reason gensem vmx_get_exit_reason_spec
                              vmx_get_exit_fault_addr gensem vmx_get_exit_fault_addr_spec
                              vmx_check_pending_event gensem vmx_check_pending_event_spec
                              vmx_check_int_shadow gensem vmx_check_int_shadow_spec
                              vmcs_set_defaults vmcs_set_defaults_compatsem vmcs_set_defaults_spec.

  Definition vmcsinit_passthrough : compatlayer (cdata RData) :=
    fload gensem fload_spec
           fstore gensem fstore_spec
           palloc gensem palloc_spec
          
           pt_read gensem ptRead_spec
           pt_resv gensem ptResv_spec
           shared_mem_status gensem shared_mem_status_spec
           offer_shared_mem gensem offer_shared_mem_spec
           get_curid gensem get_curid_spec
           thread_wakeup gensem thread_wakeup_spec

          
           syncreceive_chan gensem syncreceive_chan_spec
           syncsendto_chan_pre gensem syncsendto_chan_pre_spec
           syncsendto_chan_post gensem syncsendto_chan_post_spec

           uctx_get gensem uctx_get_spec
           uctx_set gensem uctx_set_spec
           proc_create proc_create_compatsem proc_create_spec
           rdmsr gensem rdmsr_spec
           wrmsr gensem wrmsr_spec

           cli gensem cli_spec
           sti gensem sti_spec
           cons_buf_read gensem cons_buf_read_spec
           serial_putc gensem serial_putc_spec
           serial_intr_disable gensem serial_intr_disable_spec
           serial_intr_enable gensem serial_intr_enable_spec

           vmcs_readz gensem vmcs_readz_spec
           vmcs_writez gensem vmcs_writez_spec
           vmcs_readz64 gensem vmcs_readz64_spec
           vmcs_writez64 gensem vmcs_writez64_spec
           vmptrld gensem vmptrld_spec

           ept_add_mapping gensem ept_add_mapping_spec
           ept_invalidate_mappings gensem ept_invalidate_mappings_spec

           container_get_quota gensem container_get_quota_spec
           container_get_usage gensem container_get_usage_spec
           container_can_consume gensem container_can_consume_spec
           host_in primcall_general_compatsem hostin_spec
           host_out primcall_general_compatsem hostout_spec
           trap_get primcall_trap_info_get_compatsem trap_info_get_spec
           trap_set primcall_trap_info_ret_compatsem trap_info_ret_spec
           thread_yield primcall_thread_schedule_compatsem thread_yield_spec (prim_ident:= thread_yield)
           thread_sleep primcall_thread_transfer_compatsem thread_sleep_spec
           proc_start_user primcall_start_user_compatsem proc_start_user_spec
           proc_exit_user primcall_exit_user_compatsem proc_exit_user_spec
           accessors {| exec_load := (@exec_loadex _ _ Hmwd);
                           exec_store := (@exec_storeex _ _ Hmwd) |}.

  Definition vmcsinit : compatlayer (cdata RData) := vmcsinit_fresh vmcsinit_passthrough.

End WITHMEM.