Library mcertikos.proc.PBThread
This file defines the abstract data and the primitives for the PThread layer,
which will introduce the primtives of thread
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 LoadStoreSem3.
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 CalRealSMSPool.
Require Import CalRealProcModule.
Require Import CalRealIntelModule.
Require Import INVLemmaContainer.
Require Import INVLemmaMemory.
Require Import INVLemmaThread.
Require Import INVLemmaIntel.
Require Import INVLemmaProc.
Require Import AbstractDataType.
Require Export ObjCPU.
Require Export ObjFlatMem.
Require Export ObjContainer.
Require Export ObjVMM.
Require Export ObjLMM.
Require Export ObjShareMem.
Require Export ObjThread.
Require Export ObjMultiprocessor.
Require Import CalTicketLock.
Require Import INVLemmaQLock.
Require Import INVLemmaInterrupt.
Require Import INVLemmaDriver.
Require Import DeviceStateDataType.
Require Import FutureTactic.
Require Export ObjQLock.
Require Export ObjInterruptManagement.
Require Export ObjInterruptController.
Require Export ObjConsole.
Require Export ObjSerialDriver.
Require Export ObjInterruptDriver.
Require Export ObjQueue.
Require Export ObjCV.
Require Export ObjBigThread.
Require Export ObjEPT.
Require Export ObjVMCS.
Require Export ObjVMX.
Require Import ObjProc.
Require Import IntelPrimSemantics.
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 LoadStoreSem3.
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 CalRealSMSPool.
Require Import CalRealProcModule.
Require Import CalRealIntelModule.
Require Import INVLemmaContainer.
Require Import INVLemmaMemory.
Require Import INVLemmaThread.
Require Import INVLemmaIntel.
Require Import INVLemmaProc.
Require Import AbstractDataType.
Require Export ObjCPU.
Require Export ObjFlatMem.
Require Export ObjContainer.
Require Export ObjVMM.
Require Export ObjLMM.
Require Export ObjShareMem.
Require Export ObjThread.
Require Export ObjMultiprocessor.
Require Import CalTicketLock.
Require Import INVLemmaQLock.
Require Import INVLemmaInterrupt.
Require Import INVLemmaDriver.
Require Import DeviceStateDataType.
Require Import FutureTactic.
Require Export ObjQLock.
Require Export ObjInterruptManagement.
Require Export ObjInterruptController.
Require Export ObjConsole.
Require Export ObjSerialDriver.
Require Export ObjInterruptDriver.
Require Export ObjQueue.
Require Export ObjCV.
Require Export ObjBigThread.
Require Export ObjEPT.
Require Export ObjVMCS.
Require Export ObjVMX.
Require Import ObjProc.
Require Import IntelPrimSemantics.
Section WITHMEM.
Local Open Scope Z_scope.
Context `{real_params: RealParams}.
Context `{multi_oracle_prop: MultiOracleProp}.
Local Open Scope Z_scope.
Context `{real_params: RealParams}.
Context `{multi_oracle_prop: MultiOracleProp}.
**Definition of the invariants at MPTNew layer 0th page map is reserved for the kernel thread
Record high_level_invariant (abd: RData) :=
mkInvariant {
valid_nps: pg abd = true → kern_low ≤ nps abd ≤ maxpage;
valid_kern: ipt abd = false → pg abd = true;
valid_iptt: ipt abd = true → ikern abd = true;
valid_iptf: ikern abd = false → ipt abd = false;
valid_ihost: ihost abd = false → pg abd = true ∧ ikern abd = true;
valid_container: Container_valid (AC abd);
init_pperm: pg abd = false → (pperm abd) = ZMap.init PGUndef;
valid_PMap: pg abd = true →
(∀ i, 0≤ i < num_proc →
PMap_valid (ZMap.get i (ptpool abd)));
valid_PT_kern: pg abd = true → ipt abd = true → (PT abd) = 0;
valid_PMap_kern: pg abd = true → PMap_kern (ZMap.get 0 (ptpool abd));
valid_PT: pg abd = true → 0≤ PT abd < num_proc;
valid_dirty: dirty_ppage´ (pperm abd) (HP abd);
valid_idpde: pg abd = true → IDPDE_init (idpde abd);
valid_root: ∀ i , 0 ≤ i ≤ TOTAL_CPU → (pg abd) = true → cused (ZMap.get i (AC abd)) = true;
valid_cons_buf_rpos: 0 ≤ rpos (console abd) < CONSOLE_BUFFER_SIZE;
valid_cons_buf_length: 0 ≤ Zlength (cons_buf (console abd)) < CONSOLE_BUFFER_SIZE;
CPU_ID_range: 0 ≤ (CPU_ID abd) < TOTAL_CPU;
valid_curid: 0 ≤ ZMap.get (CPU_ID abd) (cid abd) < num_proc;
valid_CPU_ID_eq: CPU_ID abd = current_CPU_ID;
valid_big_oracle_inv: valid_big_oracle (big_oracle abd)
}.
mkInvariant {
valid_nps: pg abd = true → kern_low ≤ nps abd ≤ maxpage;
valid_kern: ipt abd = false → pg abd = true;
valid_iptt: ipt abd = true → ikern abd = true;
valid_iptf: ikern abd = false → ipt abd = false;
valid_ihost: ihost abd = false → pg abd = true ∧ ikern abd = true;
valid_container: Container_valid (AC abd);
init_pperm: pg abd = false → (pperm abd) = ZMap.init PGUndef;
valid_PMap: pg abd = true →
(∀ i, 0≤ i < num_proc →
PMap_valid (ZMap.get i (ptpool abd)));
valid_PT_kern: pg abd = true → ipt abd = true → (PT abd) = 0;
valid_PMap_kern: pg abd = true → PMap_kern (ZMap.get 0 (ptpool abd));
valid_PT: pg abd = true → 0≤ PT abd < num_proc;
valid_dirty: dirty_ppage´ (pperm abd) (HP abd);
valid_idpde: pg abd = true → IDPDE_init (idpde abd);
valid_root: ∀ i , 0 ≤ i ≤ TOTAL_CPU → (pg abd) = true → cused (ZMap.get i (AC abd)) = true;
valid_cons_buf_rpos: 0 ≤ rpos (console abd) < CONSOLE_BUFFER_SIZE;
valid_cons_buf_length: 0 ≤ Zlength (cons_buf (console abd)) < CONSOLE_BUFFER_SIZE;
CPU_ID_range: 0 ≤ (CPU_ID abd) < TOTAL_CPU;
valid_curid: 0 ≤ ZMap.get (CPU_ID abd) (cid abd) < num_proc;
valid_CPU_ID_eq: CPU_ID abd = current_CPU_ID;
valid_big_oracle_inv: valid_big_oracle (big_oracle abd)
}.
Global Instance pbthread_data_ops : CompatDataOps RData :=
{
empty_data := init_adt multi_oracle_init7;
high_level_invariant := high_level_invariant;
low_level_invariant := low_level_invariant;
kernel_mode adt := ikern adt = true ∧ ihost adt = true
}.
{
empty_data := init_adt multi_oracle_init7;
high_level_invariant := high_level_invariant;
low_level_invariant := low_level_invariant;
kernel_mode adt := ikern adt = true ∧ ihost adt = true
}.
Section Property_Abstract_Data.
Lemma empty_data_high_level_invariant:
high_level_invariant (init_adt multi_oracle_init7).
Proof.
constructor; simpl; intros; auto; try inv H.
- apply empty_container_valid.
- eapply dirty_ppage´_init.
- inv H0.
- destruct oracle_ops; auto.
- rewrite ZMap.gi; intuition.
- apply valid_big_oracle0.
Qed.
Lemma empty_data_high_level_invariant:
high_level_invariant (init_adt multi_oracle_init7).
Proof.
constructor; simpl; intros; auto; try inv H.
- apply empty_container_valid.
- eapply dirty_ppage´_init.
- inv H0.
- destruct oracle_ops; auto.
- rewrite ZMap.gi; intuition.
- apply valid_big_oracle0.
Qed.
Global Instance pbthread_data_prf : CompatData RData.
Proof.
constructor.
- apply low_level_invariant_incr.
- apply empty_data_low_level_invariant.
- apply empty_data_high_level_invariant.
Qed.
End Property_Abstract_Data.
Context `{Hstencil: Stencil}.
Context `{Hmem: Mem.MemoryModelX}.
Context `{Hmwd: UseMemWithData mem}.
Proof.
constructor.
- apply low_level_invariant_incr.
- apply empty_data_low_level_invariant.
- apply empty_data_high_level_invariant.
Qed.
End Property_Abstract_Data.
Context `{Hstencil: Stencil}.
Context `{Hmem: Mem.MemoryModelX}.
Context `{Hmwd: UseMemWithData mem}.
Section INV.
Section RECEIVE.
Lemma page_copy_back_high_level_inv:
∀ d d´ chid cound addr,
big_page_copy_back_spec chid cound addr d = Some d´
→ high_level_invariant d
→ high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
inv H0.
constructor; simpl; intros;
try eapply dirty_ppage´_gss_page_copy_back; eauto.
Qed.
Lemma page_copy_back_low_level_inv:
∀ d d´ chid cound addr n,
big_page_copy_back_spec chid cound addr d = Some d´
→ low_level_invariant n d
→ low_level_invariant n d´.
Proof.
intros. functional inversion H; subst; eauto.
inv H0. constructor; eauto 2.
Qed.
Lemma page_copy_back_kernel_mode:
∀ d d´ chid cound addr,
big_page_copy_back_spec chid cound addr d = Some d´
→ kernel_mode d
→ kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
Qed.
Lemma ipc_receive_body_high_level_inv:
∀ fromid vaddr count d d´ n,
big_ipc_receive_body_spec fromid vaddr count d = Some (d´, n) →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
exploit page_copy_back_high_level_inv; eauto.
intros Hh. inv Hh. constructor; eauto 2; simpl; intros.
Qed.
Lemma ipc_receive_body_low_level_inv:
∀ fromid vaddr count d d´ n n´,
big_ipc_receive_body_spec fromid vaddr count d = Some (d´, n) →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
intros. functional inversion H; subst; eauto.
exploit page_copy_back_low_level_inv; eauto.
intros Hh. inv Hh. constructor; eauto 2.
Qed.
Lemma ipc_receive_body_kernel_mode:
∀ fromid vaddr count d d´ n,
big_ipc_receive_body_spec fromid vaddr count d = Some (d´, n) →
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
exploit page_copy_back_kernel_mode; eauto.
Qed.
Global Instance ipc_receive_body_inv: PreservesInvariants big_ipc_receive_body_spec.
Proof.
preserves_invariants_simpl´.
- eapply ipc_receive_body_low_level_inv; eassumption.
- eapply ipc_receive_body_high_level_inv; eassumption.
- eapply ipc_receive_body_kernel_mode; eassumption.
Qed.
End RECEIVE.
Section SEND.
Lemma page_copy_high_level_inv:
∀ d d´ chid cound addr,
big_page_copy_spec chid cound addr d = Some d´
→ high_level_invariant d
→ high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
inv H0.
constructor; simpl; intros; eauto.
Qed.
Lemma page_copy_low_level_inv:
∀ d d´ chid cound addr n,
big_page_copy_spec chid cound addr d = Some d´
→ low_level_invariant n d
→ low_level_invariant n d´.
Proof.
intros. functional inversion H; subst; eauto.
inv H0. constructor; eauto 2.
Qed.
Lemma page_copy_kernel_mode:
∀ d d´ chid cound addr,
big_page_copy_spec chid cound addr d = Some d´
→ kernel_mode d
→ kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
Qed.
Lemma ipc_send_body_high_level_inv:
∀ fromid vaddr count d d´ n,
big_ipc_send_body_spec fromid vaddr count d = Some (d´, n) →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
exploit page_copy_high_level_inv; eauto.
intros Hh. inv Hh. constructor; eauto 2; simpl; intros.
Qed.
Lemma ipc_send_body_low_level_inv:
∀ fromid vaddr count d d´ n n´,
big_ipc_send_body_spec fromid vaddr count d = Some (d´, n) →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
intros. functional inversion H; subst; eauto.
exploit page_copy_low_level_inv; eauto.
intros Hh. inv Hh. constructor; eauto 2.
Qed.
Lemma ipc_send_body_kernel_mode:
∀ fromid vaddr count d d´ n,
big_ipc_send_body_spec fromid vaddr count d = Some (d´, n) →
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
exploit page_copy_kernel_mode; eauto.
Qed.
Global Instance ipc_send_body_inv: PreservesInvariants big_ipc_send_body_spec.
Proof.
preserves_invariants_simpl´.
- eapply ipc_send_body_low_level_inv; eassumption.
- eapply ipc_send_body_high_level_inv; eassumption.
- eapply ipc_send_body_kernel_mode; eassumption.
Qed.
End SEND.
Global Instance set_sync_chan_busy_inv: PreservesInvariants set_sync_chan_busy_spec.
Proof.
preserves_invariants_simpl_auto.
Qed.
Global Instance cli_inv: PreservesInvariants cli_spec.
Proof.
preserves_invariants_direct low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance sti_inv: PreservesInvariants sti_spec.
Proof.
preserves_invariants_direct low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance cons_buf_read_inv:
PreservesInvariants cons_buf_read_spec.
Proof.
preserves_invariants_nested low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance serial_putc_inv:
PreservesInvariants serial_putc_spec.
Proof.
preserves_invariants_simpl_auto.
Qed.
Global Instance serial_intr_disable_inv: PreservesInvariants serial_intr_disable_spec.
Proof.
constructor; simpl; intros; inv_generic_sem H.
- inversion H0; econstructor; eauto 2 with serial_intr_disable_invariantdb.
generalize (serial_intr_disable_preserves_tf _ _ H2); intro tmprw; rewrite <- tmprw; assumption.
- inversion H0; econstructor; eauto 2 with serial_intr_disable_invariantdb; rest.
- eauto 2 with serial_intr_disable_invariantdb.
Qed.
Global Instance serial_intr_enable_inv:
PreservesInvariants serial_intr_enable_spec.
Proof.
constructor; simpl; intros; inv_generic_sem H.
- inversion H0; econstructor; eauto 2 with serial_intr_enable_invariantdb.
generalize (serial_intr_enable_preserves_tf _ _ H2); intro tmprw; rewrite <- tmprw; assumption.
- inversion H0; econstructor; eauto 2 with serial_intr_enable_invariantdb; rest.
- eauto 2 with serial_intr_enable_invariantdb.
Qed.
Section SC_lock.
Lemma acquire_lock_SC_high_level_inv:
∀ d d´ i,
big_acquire_lock_SC_spec i d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
unfold big_acquire_lock_SC_spec; intros.
subdestruct; inv H; eauto.
- inv H0. constructor; simpl; eauto; intros.
- inv H0. constructor; simpl; eauto; intros.
Qed.
Lemma acquire_lock_SC_low_level_inv:
∀ d d´ i n,
big_acquire_lock_SC_spec i d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
unfold big_acquire_lock_SC_spec; intros.
subdestruct; inv H; eauto.
- inv H0. constructor; simpl; eauto.
- inv H0. constructor; simpl; eauto.
Qed.
Lemma acquire_lock_SC_kernel_mode:
∀ d d´ i,
big_acquire_lock_SC_spec i d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
unfold big_acquire_lock_SC_spec; intros.
subdestruct; inv H; eauto.
Qed.
Global Instance acquire_lock_SC_inv: PreservesInvariants big_acquire_lock_SC_spec.
Proof.
preserves_invariants_simpl´.
- eapply acquire_lock_SC_low_level_inv; eassumption.
- eapply acquire_lock_SC_high_level_inv; eassumption.
- eapply acquire_lock_SC_kernel_mode; eassumption.
Qed.
Lemma release_lock_SC_high_level_inv:
∀ d d´ i,
big_release_lock_SC_spec i d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. unfold big_release_lock_SC_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto; intros.
Qed.
Lemma release_lock_SC_low_level_inv:
∀ d d´ i n,
big_release_lock_SC_spec i d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. unfold big_release_lock_SC_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto.
Qed.
Lemma release_lock_SC_kernel_mode:
∀ d d´ i,
big_release_lock_SC_spec i d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
intros. unfold big_release_lock_SC_spec in *; subdestruct; inv H.
- inv H0; eauto.
Qed.
Global Instance release_lock_SC_inv: PreservesInvariants big_release_lock_SC_spec.
Proof.
preserves_invariants_simpl´.
- eapply release_lock_SC_low_level_inv; eassumption.
- eapply release_lock_SC_high_level_inv; eassumption.
- eapply release_lock_SC_kernel_mode; eassumption.
Qed.
End SC_lock.
Section PALLOC.
Lemma palloc_high_level_inv:
∀ d d´ i n,
big_palloc_spec i d = Some (d´, n) →
high_level_invariant d →
high_level_invariant d´.
Proof.
unfold big_palloc_spec; intros.
subdestruct; inv H; subst; eauto;
inv H0; constructor; simpl; eauto; intros.
- rewrite <- Hdestruct3.
eapply alloc_container_valid´; eauto.
- subst; simpl;
intros; congruence.
- eapply dirty_ppage´_gso_alloc; eauto.
- intros.
eapply (valid_root0 i0 H) in H0.
destruct (zeq i0 i); subst.
+ rewrite ZMap.gss; trivial.
+ rewrite ZMap.gso; auto.
Qed.
Lemma palloc_low_level_inv:
∀ d d´ i n n´,
big_palloc_spec i d = Some (d´, n) →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
unfold big_palloc_spec; intros.
subdestruct; inv H; subst; eauto;
inv H0; constructor; eauto.
Qed.
Lemma palloc_kernel_mode:
∀ d d´ i n,
big_palloc_spec i d = Some (d´, n) →
kernel_mode d →
kernel_mode d´.
Proof.
unfold big_palloc_spec; intros.
subdestruct; inv H; simpl; eauto.
Qed.
Global Instance palloc_inv: PreservesInvariants big_palloc_spec.
Proof.
preserves_invariants_simpl´.
- eapply palloc_low_level_inv; eassumption.
- eapply palloc_high_level_inv; eassumption.
- eapply palloc_kernel_mode; eassumption.
Qed.
End PALLOC.
Global Instance hostin_inv: PrimInvariants hostin_spec.
Proof.
PrimInvariants_simpl_auto.
Qed.
Global Instance hostout_inv: PrimInvariants hostout_spec.
Proof.
PrimInvariants_simpl_auto.
Qed.
Global Instance setPT_inv: PreservesInvariants setPT_spec.
Proof.
preserves_invariants_simpl_auto.
Qed.
Section PTINSERT.
Section PTINSERT_PTE.
Lemma ptInsertPTE_high_level_inv:
∀ d d´ n vadr padr p,
ptInsertPTE0_spec n vadr padr p d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
inv H0; constructor_gso_simpl_tac; intros.
- eapply PMap_valid_gso_valid; eauto.
- functional inversion H2. functional inversion H1.
eapply PMap_kern_gso; eauto.
Qed.
Lemma ptInsertPTE_low_level_inv:
∀ d d´ n vadr padr p n´,
ptInsertPTE0_spec n vadr padr p d = Some d´ →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
intros. functional inversion H; subst; eauto.
inv H0. constructor; eauto.
Qed.
Lemma ptInsertPTE_kernel_mode:
∀ d d´ n vadr padr p,
ptInsertPTE0_spec n vadr padr p d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
Qed.
End PTINSERT_PTE.
Section PTPALLOCPDE.
Lemma big_palloc_inv_prop:
∀ n d d´ v,
big_palloc_spec n d = Some (d´, v) →
ptpool d´ = ptpool d
∧ (v ≠ 0 →
ZMap.get v (pperm d´) = PGAlloc
∧ ZMap.get v (LAT d´) = LATCValid nil
∧ 0 < v < nps d´)
∧ pg d´ = true.
Proof.
unfold big_palloc_spec. intros.
subdestruct; inv H; simpl; subst.
- repeat rewrite ZMap.gss. refine_split´; trivial.
destruct a0 as (? & ? & ?). intros.
refine_split´; trivial.
- refine_split´; trivial.
intros HF; elim HF. trivial.
- refine_split´; trivial.
intros HF; elim HF. trivial.
Qed.
Lemma ptAllocPDE_high_level_inv:
∀ d d´ n vadr v,
big_ptAllocPDE_spec n vadr d = Some (d´, v) →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
- eapply palloc_high_level_inv; eauto.
- exploit palloc_high_level_inv; eauto.
intros.
exploit big_palloc_inv_prop; eauto. intros (HPT & Halloc & Hpg).
clear H11.
rewrite <- HPT in ×.
inv H1; constructor_gso_simpl_tac; try (intros; congruence); intros.
+ eapply PMap_valid_gso_pde_unp; eauto.
eapply real_init_PTE_defined.
+ functional inversion H3.
eapply PMap_kern_gso; eauto.
+ eapply dirty_ppage´_gss; eauto.
Qed.
Lemma ptAllocPDE_low_level_inv:
∀ d d´ n vadr v n´,
big_ptAllocPDE_spec n vadr d = Some (d´, v) →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
intros. functional inversion H; subst; eauto.
- eapply palloc_low_level_inv; eauto.
- exploit palloc_low_level_inv; eauto.
intros. inv H1. constructor; eauto.
Qed.
Lemma ptAllocPDE_kernel_mode:
∀ d d´ n vadr v,
big_ptAllocPDE_spec n vadr d = Some (d´, v) →
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
- eapply palloc_kernel_mode; eauto.
- exploit palloc_kernel_mode; eauto.
Qed.
End PTPALLOCPDE.
Lemma ptInsert_high_level_inv:
∀ d d´ n vadr padr p v,
big_ptInsert_spec n vadr padr p d = Some (d´, v) →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
- eapply ptInsertPTE_high_level_inv; eassumption.
- eapply ptAllocPDE_high_level_inv; eassumption.
- eapply ptInsertPTE_high_level_inv; try eassumption.
eapply ptAllocPDE_high_level_inv; eassumption.
Qed.
Lemma ptInsert_low_level_inv:
∀ d d´ n vadr padr p n´ v,
big_ptInsert_spec n vadr padr p d = Some (d´, v) →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
intros. functional inversion H; subst; eauto.
- eapply ptInsertPTE_low_level_inv; eassumption.
- eapply ptAllocPDE_low_level_inv; eassumption.
- eapply ptInsertPTE_low_level_inv; try eassumption.
eapply ptAllocPDE_low_level_inv; eassumption.
Qed.
Lemma ptInsert_kernel_mode:
∀ d d´ n vadr padr p v,
big_ptInsert_spec n vadr padr p d = Some (d´, v) →
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
- eapply ptInsertPTE_kernel_mode; eassumption.
- eapply ptAllocPDE_kernel_mode; eassumption.
- eapply ptInsertPTE_kernel_mode; try eassumption.
eapply ptAllocPDE_kernel_mode; eassumption.
Qed.
End PTINSERT.
Section PTRESV.
Lemma ptResv_high_level_inv:
∀ d d´ n vadr p v,
big_ptResv_spec n vadr p d = Some (d´, v) →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto; clear H.
- eapply palloc_high_level_inv; eassumption.
- eapply ptInsert_high_level_inv; try eassumption.
eapply palloc_high_level_inv; eassumption.
Qed.
Lemma ptResv_low_level_inv:
∀ d d´ n vadr p n´ v,
big_ptResv_spec n vadr p d = Some (d´, v) →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
intros. functional inversion H; subst; eauto.
eapply palloc_low_level_inv; eassumption.
eapply ptInsert_low_level_inv; try eassumption.
eapply palloc_low_level_inv; eassumption.
Qed.
Lemma ptResv_kernel_mode:
∀ d d´ n vadr p v,
big_ptResv_spec n vadr p d = Some (d´, v) →
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
eapply palloc_kernel_mode; eassumption.
eapply ptInsert_kernel_mode; try eassumption.
eapply palloc_kernel_mode; eassumption.
Qed.
Global Instance ptResv_inv: PreservesInvariants big_ptResv_spec.
Proof.
preserves_invariants_simpl´.
- eapply ptResv_low_level_inv; eassumption.
- eapply ptResv_high_level_inv; eassumption.
- eapply ptResv_kernel_mode; eassumption.
Qed.
End PTRESV.
Global Instance thread_wakeup_inv: PreservesInvariants big_thread_wakeup_spec.
Proof.
preserves_invariants_simpl´.
- intros. functional inversion H2; try subst; eauto.
+ inv H0. econstructor; eauto.
+ subst me l´ l1 to. inv H0. econstructor; eauto.
+ inv H0. econstructor; eauto.
- intros. functional inversion H2; try subst; eauto.
+ inv H0. econstructor; simpl; eauto.
+ subst me l´ l1 to. inv H0. econstructor; eauto.
+ inv H0. econstructor; simpl; eauto.
- intros. functional inversion H2; try subst; eauto; simpl.
Qed.
Global Instance big_proc_create_inv: PCreateInvariants big_proc_create_spec.
Proof.
constructor; intros; inv H0;
unfold big_proc_create_spec in *;
subdestruct; inv H; simpl; auto.
-
constructor; trivial; intros; simpl in ×.
+ eapply kctxt_inject_neutral_gss_flatinj´; eauto.
eapply kctxt_inject_neutral_gss_flatinj; eauto.
+ simpl.
eapply uctxt_inject_neutral_gss; eauto.
repeat eapply uctxt_inject_neutral0_gss;
try eapply uctxt_inject_neutral0_Vint.
× eapply uctxt_inject_neutral_impl; eauto. omega.
× eapply uctxt_inject_neutral0_Vptr_flat; eauto.
-
assert (Hrdyq: 0 ≤ rdy_q_id (CPU_ID d) < num_chan).
{
unfold rdy_q_id. omega.
}
constructor; simpl; eauto 2; try congruence; intros.
+ exploit split_container_valid; eauto.
simpl. omega.
rewrite Hdestruct3.
auto.
+ generalize (valid_root0 i H); intros Htemp.
exploit Htemp; auto; intros valid_idle_cid; clear Htemp.
destruct (zeq i ((ZMap.get (CPU_ID d) (cid d)) × 8 + 1 +
Z.of_nat (length (cchildren (ZMap.get (ZMap.get (CPU_ID d) (cid d)) (AC d)))))); subst.
rewrite ZMap.gss; simpl; split; auto.
rewrite ZMap.gso; auto.
destruct (zeq i (ZMap.get (CPU_ID d) (cid d))); subst.
× rewrite ZMap.gss; simpl; auto.
× rewrite ZMap.gso; auto.
Qed.
Global Instance sched_init_inv: PreservesInvariants big_sched_init_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant.
- apply real_nps_range.
- apply AC_init_container_valid.
- eapply real_pt_PMap_valid; eauto.
- apply real_pt_PMap_kern.
- omega.
- assumption.
- apply real_idpde_init.
- unfold AC_init.
unfold init_real_container.
simpl.
case_eq (zeq i0 8); intros; subst; [rewrite ZMap.gss; auto | ].
rewrite ZMap.gso; auto.
case_eq (zeq i0 7); intros; subst; [rewrite ZMap.gss; auto | ].
rewrite ZMap.gso; auto.
case_eq (zeq i0 6); intros; subst; [rewrite ZMap.gss; auto | ].
rewrite ZMap.gso; auto.
case_eq (zeq i0 5); intros; subst; [rewrite ZMap.gss; auto | ].
rewrite ZMap.gso; auto.
case_eq (zeq i0 4); intros; subst; [rewrite ZMap.gss; auto | ].
rewrite ZMap.gso; auto.
case_eq (zeq i0 3); intros; subst; [rewrite ZMap.gss; auto | ].
rewrite ZMap.gso; auto.
case_eq (zeq i0 2); intros; subst; [rewrite ZMap.gss; auto | ].
rewrite ZMap.gso; auto.
case_eq (zeq i0 1); intros; subst; [rewrite ZMap.gss; auto | ].
rewrite ZMap.gso; auto.
case_eq (zeq i0 0); intros; subst; [rewrite ZMap.gss; auto | ].
omega.
- erewrite real_valid_cidpool; eauto.
omega.
Qed.
Local Opaque remove.
Section YIELD.
Global Instance thread_yield_inv: ThreadScheduleInvariants big_thread_yield_spec.
Proof.
constructor; intros; functional inversion H.
- inv H1.
constructor; trivial; simpl.
eapply kctxt_inject_neutral_gss_mem; eauto.
- inv H0.
constructor; auto; simpl in *; intros; try congruence; auto.
rewrite ZMap.gss. omega.
Qed.
End YIELD.
Global Instance thread_sleep_inv: ThreadTransferInvariants big_thread_sleep_spec.
Proof.
constructor; intros; functional inversion H.
- inv H1.
constructor; trivial.
eapply kctxt_inject_neutral_gss_mem; eauto.
- inv H0.
constructor; auto; simpl in *; intros; try congruence; auto.
rewrite ZMap.gss. omega.
Qed.
Global Instance uctx_set_inv: PreservesInvariants uctx_set_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
eapply uctxt_inject_neutral_gss; eauto.
eapply uctxt_inject_neutral0_gss; eauto.
eapply uctxt_inject_neutral0_Vint.
Qed.
Global Instance proc_start_user_inv: StartUserInvariants proc_start_user_spec.
Proof.
constructor; intros; functional inversion H; inv H0.
- constructor; trivial.
- subst. constructor; auto; simpl in *; intros; congruence.
Qed.
Global Instance proc_exit_user_inv: ExitUserInvariants proc_exit_user_spec.
Proof.
constructor; intros; functional inversion H; inv H0.
- constructor; trivial.
eapply uctxt_inject_neutral_gss; eauto.
repeat eapply uctxt_inject_neutral0_gss;
try eapply uctxt_inject_neutral0_Vint.
eapply uctxt_inject_neutral0_init.
- constructor; auto; simpl in *; intros.
Qed.
Global Instance proc_start_user_inv2: StartUserInvariants2 proc_start_user_spec.
Proof.
constructor; intros; functional inversion H; inv H0.
- constructor; trivial.
- subst. constructor; auto; simpl in *; intros; congruence.
Qed.
Global Instance proc_exit_user_inv2: ExitUserInvariants2 proc_exit_user_spec.
Proof.
constructor; intros; functional inversion H; inv H0.
- constructor; trivial.
eapply uctxt_inject_neutral_gss; eauto.
repeat eapply uctxt_inject_neutral0_gss;
try eapply uctxt_inject_neutral0_Vint.
eapply uctxt_inject_neutral0_init.
- constructor; auto; simpl in *; intros.
Qed.
Global Instance proc_create_postinit_inv:
PreservesInvariants proc_create_postinit_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Section INTELVMSUPPORT.
Section INTERCEPT_INTWIN.
Lemma vmx_set_intercept_intwin_high_inv:
∀ d d´ i,
vmx_set_intercept_intwin_spec i d = Some d´→
high_level_invariant d →
high_level_invariant d´.
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 d´ n i,
vmx_set_intercept_intwin_spec i d = Some d´→
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
eapply VMCS_inject_neutral_gss_Vint; eauto.
+ rewrite ZMap.gso; eauto.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
eapply VMCS_inject_neutral_gss_Vint; eauto.
+ rewrite ZMap.gso; eauto.
Qed.
Lemma vmx_set_intercept_intwin_kernel_mode:
∀ d d´ i,
vmx_set_intercept_intwin_spec i d = Some d´→
kernel_mode d →
kernel_mode d´.
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 SET_DESC1.
Lemma vmx_set_desc1_high_inv:
∀ d d´ i i0 i1,
vmx_set_desc1_spec i i0 i1 d = Some d´→
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
Qed.
Lemma vmx_set_desc1_low_inv:
∀ d d´ n i i0 i1,
vmx_set_desc1_spec i i0 i1 d = Some d´→
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
functional inversion H6.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
+ rewrite ZMap.gso; eauto.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
+ rewrite ZMap.gso; eauto.
Qed.
Lemma vmx_set_desc1_kernel_mode:
∀ d d´ i i0 i1,
vmx_set_desc1_spec i i0 i1 d = Some d´→
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; constructor; trivial.
Qed.
Global Instance vmx_set_desc1_inv: PreservesInvariants vmx_set_desc1_spec.
Proof.
preserves_invariants_simpl´.
- eapply vmx_set_desc1_low_inv; eauto.
- eapply vmx_set_desc1_high_inv; eauto.
- eapply vmx_set_desc1_kernel_mode; eauto.
Qed.
End SET_DESC1.
Section SET_DESC2.
Lemma vmx_set_desc2_high_inv:
∀ d d´ i i0 i1,
vmx_set_desc2_spec i i0 i1 d = Some d´→
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
Qed.
Lemma vmx_set_desc2_low_inv:
∀ d d´ n i i0 i1,
vmx_set_desc2_spec i i0 i1 d = Some d´→
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
functional inversion H6.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
+ rewrite ZMap.gso; eauto.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
+ rewrite ZMap.gso; eauto.
Qed.
Lemma vmx_set_desc2_kernel_mode:
∀ d d´ i i0 i1,
vmx_set_desc2_spec i i0 i1 d = Some d´→
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; constructor; trivial.
Qed.
Global Instance vmx_set_desc2_inv: PreservesInvariants vmx_set_desc2_spec.
Proof.
preserves_invariants_simpl´.
- eapply vmx_set_desc2_low_inv; eauto.
- eapply vmx_set_desc2_high_inv; eauto.
- eapply vmx_set_desc2_kernel_mode; eauto.
Qed.
End SET_DESC2.
Section INJECT_EVENT.
Lemma vmx_inject_event_high_inv:
∀ d d´ i i0 i1 i2,
vmx_inject_event_spec i i0 i1 i2 d = Some d´→
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
Qed.
Lemma vmx_inject_event_low_inv:
∀ d d´ n i i0 i1 i2,
vmx_inject_event_spec i i0 i1 i2 d = Some d´→
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
+ rewrite ZMap.gso; eauto.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
+ rewrite ZMap.gso; eauto.
Qed.
Lemma vmx_inject_event_kernel_mode:
∀ d d´ i i0 i1 i2,
vmx_inject_event_spec i i0 i1 i2 d = Some d´→
kernel_mode d →
kernel_mode d´.
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_TSC_OFFSET.
Lemma vmx_set_tsc_offset_high_inv:
∀ d d´ i,
vmx_set_tsc_offset_spec i d = Some d´→
high_level_invariant d →
high_level_invariant d´.
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 d´ n i,
vmx_set_tsc_offset_spec i d = Some d´→
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
+ rewrite ZMap.gso; eauto.
Qed.
Lemma vmx_set_tsc_offset_kernel_mode:
∀ d d´ i,
vmx_set_tsc_offset_spec i d = Some d´→
kernel_mode d →
kernel_mode d´.
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.
Section SET_REG.
Lemma vmx_set_reg_high_inv:
∀ d d´ i i0,
vmx_set_reg_spec i i0 d = Some d´→
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
Qed.
Lemma vmx_set_reg_low_inv:
∀ d d´ n i i0,
vmx_set_reg_spec i i0 d = Some d´→
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
- unfold VMXPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss; eapply VMX_inject_neutral_gss_Vint; eauto.
+ rewrite ZMap.gso; eauto.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss; eapply VMCS_inject_neutral_gss_Vint; eauto.
+ rewrite ZMap.gso; eauto.
Qed.
Lemma vmx_set_reg_kernel_mode:
∀ d d´ i i0,
vmx_set_reg_spec i i0 d = Some d´→
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; constructor; trivial.
Qed.
Global Instance vmx_set_reg_inv: PreservesInvariants vmx_set_reg_spec.
Proof.
preserves_invariants_simpl´.
- eapply vmx_set_reg_low_inv; eauto.
- eapply vmx_set_reg_high_inv; eauto.
- eapply vmx_set_reg_kernel_mode; eauto.
Qed.
End SET_REG.
Section SET_MMAP.
Lemma vmx_set_mmap_high_inv:
∀ d d´ i i0 i1 v,
vmx_set_mmap_spec i i0 i1 d = Some (d´, v)→
high_level_invariant d →
high_level_invariant d´.
Proof.
intros.
inv H0; functional inversion H; functional inversion H2; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
Qed.
Lemma vmx_set_mmap_low_inv:
∀ d d´ n i i0 i1 v,
vmx_set_mmap_spec i i0 i1 d = Some (d´, v)→
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros.
inv H0; functional inversion H; functional inversion H2; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
Qed.
Lemma vmx_set_mmap_kernel_mode:
∀ d d´ i i0 i1 v,
vmx_set_mmap_spec i i0 i1 d = Some (d´, v)→
kernel_mode d →
kernel_mode d´.
Proof.
intros; functional inversion H;
functional inversion H3; subst; constructor; trivial.
Qed.
Global Instance vmx_set_mmap_inv: PreservesInvariants vmx_set_mmap_spec.
Proof.
preserves_invariants_simpl´.
- eapply vmx_set_mmap_low_inv; eauto.
- eapply vmx_set_mmap_high_inv; eauto.
- eapply vmx_set_mmap_kernel_mode; eauto.
Qed.
End SET_MMAP.
Global Instance vmx_run_inv: VMXEnterInvariants vm_run_spec.
Proof.
constructor; intros;
unfold vm_run_spec in *;
subdestruct; inv H; simpl; auto.
-
inv H1. constructor; trivial; intros; simpl in ×.
+ unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
× rewrite ZMap.gss.
eapply VMCS_inject_neutral_gss_Vint; eauto.
× rewrite ZMap.gso; eauto.
+ unfold VMXPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
× rewrite ZMap.gss.
eapply VMX_inject_neutral_gss_enter; eauto.
× rewrite ZMap.gso; eauto.
-
inv H0. constructor; simpl; eauto 2; try congruence; intros.
Qed.
Global Instance vmx_return_from_guest_inv: VMXReturnInvariants vmx_return_from_guest_spec.
Proof.
constructor; intros;
unfold vmx_return_from_guest_spec in *;
subdestruct; inv H; simpl; auto.
-
inv H0. constructor; trivial; intros; simpl in ×.
unfold VMXPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss; repeat eapply VMX_inject_neutral_gss_Vint; eauto.
+ rewrite ZMap.gso; eauto.
-
inv H0. constructor; simpl; eauto 2; try congruence; intros.
Qed.
Global Instance vmx_init_inv: VMCSSetDefaultsInvariants vmx_init_spec.
Proof.
constructor; intros; functional inversion H; simpl; auto.
-
inv H6. constructor; trivial; intros; simpl in ×.
+ eapply real_vmcs_inject_neutral; eauto.
+ eapply real_vmx_inject_neutral; eauto.
-
inv H0. constructor; simpl; try congruence; intros; eauto.
Qed.
End INTELVMSUPPORT.
End INV.
Definition exec_loadex {F V} := exec_loadex3 (F := F) (V := V).
Definition exec_storeex {F V} := exec_storeex3 (flatmem_store:= flatmem_store) (F := F) (V := V).
Global Instance flatmem_store_inv: FlatmemStoreInvariant (flatmem_store:= flatmem_store).
Proof.
split; inversion 1; intros.
- functional inversion H0. split; trivial.
- functional inversion H1.
split; simpl; try (eapply dirty_ppage´_store_unmapped´; try reflexivity; try eassumption); trivial.
Qed.
Global Instance trapinfo_set_inv: TrapinfoSetInvariant.
Proof.
split; inversion 1; intros; constructor; auto.
Qed.
Definition pbthread_passthrough : compatlayer (cdata RData) :=
vmxinfo_get ↦ gensem vmxinfo_get_spec
⊕ palloc ↦ gensem big_palloc_spec
⊕ set_pt ↦ gensem setPT_spec
⊕ pt_read ↦ gensem ptRead_spec
⊕ pt_resv ↦ gensem big_ptResv_spec
⊕ thread_wakeup ↦ gensem big_thread_wakeup_spec
⊕ thread_yield ↦ primcall_thread_schedule_compatsem big_thread_yield_spec (prim_ident:= thread_yield)
⊕ thread_sleep ↦ primcall_thread_transfer_compatsem big_thread_sleep_spec
⊕ sched_init ↦ gensem big_sched_init_spec
⊕ uctx_get ↦ gensem uctx_get_spec
⊕ uctx_set ↦ gensem uctx_set_spec
⊕ proc_create ↦ proc_create_compatsem big_proc_create_spec
⊕ container_get_nchildren ↦ gensem container_get_nchildren_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
⊕ get_CPU_ID ↦ gensem get_CPU_ID_spec
⊕ get_curid ↦ gensem get_curid_spec
⊕ acquire_lock_CHAN ↦ gensem big_acquire_lock_SC_spec
⊕ release_lock_CHAN ↦ gensem big_release_lock_SC_spec
⊕ get_sync_chan_busy ↦ gensem get_sync_chan_busy_spec
⊕ set_sync_chan_busy ↦ gensem set_sync_chan_busy_spec
⊕ ipc_send_body ↦ gensem big_ipc_send_body_spec
⊕ ipc_receive_body ↦ gensem big_ipc_receive_body_spec
⊕ rdmsr ↦ gensem rdmsr_spec
⊕ wrmsr ↦ gensem wrmsr_spec
⊕ vmx_set_intercept_intwin ↦ gensem vmx_set_intercept_intwin_spec
⊕ vmx_set_desc1 ↦ gensem vmx_set_desc1_spec
⊕ vmx_set_desc2 ↦ gensem vmx_set_desc2_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_get_exit_qualification ↦ gensem vmx_get_exit_qualification_spec
⊕ vmx_check_pending_event ↦ gensem vmx_check_pending_event_spec
⊕ vmx_check_int_shadow ↦ gensem vmx_check_int_shadow_spec
⊕ vmx_get_reg ↦ gensem vmx_get_reg_spec
⊕ vmx_set_reg ↦ gensem vmx_set_reg_spec
⊕ vmx_get_next_eip ↦ gensem vmx_get_next_eip_spec
⊕ vmx_get_io_width ↦ gensem vmx_get_io_width_spec
⊕ vmx_get_io_write ↦ gensem vmx_get_io_write_spec
⊕ vmx_get_exit_io_rep ↦ gensem vmx_get_exit_io_rep_spec
⊕ vmx_get_exit_io_str ↦ gensem vmx_get_exit_io_str_spec
⊕ vmx_get_exit_io_port ↦ gensem vmx_get_exit_io_port_spec
⊕ vmx_set_mmap ↦ gensem vmx_set_mmap_spec
⊕ vmx_run_vm ↦ primcall_vmx_enter_compatsem vm_run_spec vmx_run_vm
⊕ vmx_return_from_guest ↦ primcall_vmx_return_compatsem vmx_return_from_guest_spec
⊕ vmx_init ↦ vmcs_set_defaults_compatsem vmx_init_spec
⊕ cli ↦ gensem cli_spec
⊕ sti ↦ gensem sti_spec
⊕ serial_intr_disable ↦ gensem serial_intr_disable_spec
⊕ serial_intr_enable ↦ gensem serial_intr_enable_spec
⊕ serial_putc ↦ gensem serial_putc_spec
⊕ cons_buf_read ↦ gensem cons_buf_read_spec
⊕ host_in ↦ primcall_general_compatsem hostin_spec
⊕ host_out ↦ primcall_general_compatsem hostout_spec
⊕ proc_create_postinit ↦ gensem proc_create_postinit_spec
⊕ trap_get ↦ primcall_trap_info_get_compatsem trap_info_get_spec
⊕ trap_set ↦ primcall_trap_info_ret_compatsem trap_info_ret_spec
⊕ proc_start_user ↦ primcall_start_user_compatsem proc_start_user_spec
⊕ proc_exit_user ↦ primcall_exit_user_compatsem proc_exit_user_spec
⊕ proc_start_user2 ↦ primcall_start_user_compatsem2 proc_start_user_spec
⊕ proc_exit_user2 ↦ primcall_exit_user_compatsem2 proc_exit_user_spec
⊕ accessors ↦ {| exec_load := @exec_loadex; exec_store := @exec_storeex |}.
Definition pbthread_fresh : compatlayer (cdata RData) :=
∅ ⊕ ∅.
Definition pbthread :=
pbthread_fresh ⊕ pbthread_passthrough.
End WITHMEM.
Section RECEIVE.
Lemma page_copy_back_high_level_inv:
∀ d d´ chid cound addr,
big_page_copy_back_spec chid cound addr d = Some d´
→ high_level_invariant d
→ high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
inv H0.
constructor; simpl; intros;
try eapply dirty_ppage´_gss_page_copy_back; eauto.
Qed.
Lemma page_copy_back_low_level_inv:
∀ d d´ chid cound addr n,
big_page_copy_back_spec chid cound addr d = Some d´
→ low_level_invariant n d
→ low_level_invariant n d´.
Proof.
intros. functional inversion H; subst; eauto.
inv H0. constructor; eauto 2.
Qed.
Lemma page_copy_back_kernel_mode:
∀ d d´ chid cound addr,
big_page_copy_back_spec chid cound addr d = Some d´
→ kernel_mode d
→ kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
Qed.
Lemma ipc_receive_body_high_level_inv:
∀ fromid vaddr count d d´ n,
big_ipc_receive_body_spec fromid vaddr count d = Some (d´, n) →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
exploit page_copy_back_high_level_inv; eauto.
intros Hh. inv Hh. constructor; eauto 2; simpl; intros.
Qed.
Lemma ipc_receive_body_low_level_inv:
∀ fromid vaddr count d d´ n n´,
big_ipc_receive_body_spec fromid vaddr count d = Some (d´, n) →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
intros. functional inversion H; subst; eauto.
exploit page_copy_back_low_level_inv; eauto.
intros Hh. inv Hh. constructor; eauto 2.
Qed.
Lemma ipc_receive_body_kernel_mode:
∀ fromid vaddr count d d´ n,
big_ipc_receive_body_spec fromid vaddr count d = Some (d´, n) →
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
exploit page_copy_back_kernel_mode; eauto.
Qed.
Global Instance ipc_receive_body_inv: PreservesInvariants big_ipc_receive_body_spec.
Proof.
preserves_invariants_simpl´.
- eapply ipc_receive_body_low_level_inv; eassumption.
- eapply ipc_receive_body_high_level_inv; eassumption.
- eapply ipc_receive_body_kernel_mode; eassumption.
Qed.
End RECEIVE.
Section SEND.
Lemma page_copy_high_level_inv:
∀ d d´ chid cound addr,
big_page_copy_spec chid cound addr d = Some d´
→ high_level_invariant d
→ high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
inv H0.
constructor; simpl; intros; eauto.
Qed.
Lemma page_copy_low_level_inv:
∀ d d´ chid cound addr n,
big_page_copy_spec chid cound addr d = Some d´
→ low_level_invariant n d
→ low_level_invariant n d´.
Proof.
intros. functional inversion H; subst; eauto.
inv H0. constructor; eauto 2.
Qed.
Lemma page_copy_kernel_mode:
∀ d d´ chid cound addr,
big_page_copy_spec chid cound addr d = Some d´
→ kernel_mode d
→ kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
Qed.
Lemma ipc_send_body_high_level_inv:
∀ fromid vaddr count d d´ n,
big_ipc_send_body_spec fromid vaddr count d = Some (d´, n) →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
exploit page_copy_high_level_inv; eauto.
intros Hh. inv Hh. constructor; eauto 2; simpl; intros.
Qed.
Lemma ipc_send_body_low_level_inv:
∀ fromid vaddr count d d´ n n´,
big_ipc_send_body_spec fromid vaddr count d = Some (d´, n) →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
intros. functional inversion H; subst; eauto.
exploit page_copy_low_level_inv; eauto.
intros Hh. inv Hh. constructor; eauto 2.
Qed.
Lemma ipc_send_body_kernel_mode:
∀ fromid vaddr count d d´ n,
big_ipc_send_body_spec fromid vaddr count d = Some (d´, n) →
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
exploit page_copy_kernel_mode; eauto.
Qed.
Global Instance ipc_send_body_inv: PreservesInvariants big_ipc_send_body_spec.
Proof.
preserves_invariants_simpl´.
- eapply ipc_send_body_low_level_inv; eassumption.
- eapply ipc_send_body_high_level_inv; eassumption.
- eapply ipc_send_body_kernel_mode; eassumption.
Qed.
End SEND.
Global Instance set_sync_chan_busy_inv: PreservesInvariants set_sync_chan_busy_spec.
Proof.
preserves_invariants_simpl_auto.
Qed.
Global Instance cli_inv: PreservesInvariants cli_spec.
Proof.
preserves_invariants_direct low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance sti_inv: PreservesInvariants sti_spec.
Proof.
preserves_invariants_direct low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance cons_buf_read_inv:
PreservesInvariants cons_buf_read_spec.
Proof.
preserves_invariants_nested low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance serial_putc_inv:
PreservesInvariants serial_putc_spec.
Proof.
preserves_invariants_simpl_auto.
Qed.
Global Instance serial_intr_disable_inv: PreservesInvariants serial_intr_disable_spec.
Proof.
constructor; simpl; intros; inv_generic_sem H.
- inversion H0; econstructor; eauto 2 with serial_intr_disable_invariantdb.
generalize (serial_intr_disable_preserves_tf _ _ H2); intro tmprw; rewrite <- tmprw; assumption.
- inversion H0; econstructor; eauto 2 with serial_intr_disable_invariantdb; rest.
- eauto 2 with serial_intr_disable_invariantdb.
Qed.
Global Instance serial_intr_enable_inv:
PreservesInvariants serial_intr_enable_spec.
Proof.
constructor; simpl; intros; inv_generic_sem H.
- inversion H0; econstructor; eauto 2 with serial_intr_enable_invariantdb.
generalize (serial_intr_enable_preserves_tf _ _ H2); intro tmprw; rewrite <- tmprw; assumption.
- inversion H0; econstructor; eauto 2 with serial_intr_enable_invariantdb; rest.
- eauto 2 with serial_intr_enable_invariantdb.
Qed.
Section SC_lock.
Lemma acquire_lock_SC_high_level_inv:
∀ d d´ i,
big_acquire_lock_SC_spec i d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
unfold big_acquire_lock_SC_spec; intros.
subdestruct; inv H; eauto.
- inv H0. constructor; simpl; eauto; intros.
- inv H0. constructor; simpl; eauto; intros.
Qed.
Lemma acquire_lock_SC_low_level_inv:
∀ d d´ i n,
big_acquire_lock_SC_spec i d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
unfold big_acquire_lock_SC_spec; intros.
subdestruct; inv H; eauto.
- inv H0. constructor; simpl; eauto.
- inv H0. constructor; simpl; eauto.
Qed.
Lemma acquire_lock_SC_kernel_mode:
∀ d d´ i,
big_acquire_lock_SC_spec i d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
unfold big_acquire_lock_SC_spec; intros.
subdestruct; inv H; eauto.
Qed.
Global Instance acquire_lock_SC_inv: PreservesInvariants big_acquire_lock_SC_spec.
Proof.
preserves_invariants_simpl´.
- eapply acquire_lock_SC_low_level_inv; eassumption.
- eapply acquire_lock_SC_high_level_inv; eassumption.
- eapply acquire_lock_SC_kernel_mode; eassumption.
Qed.
Lemma release_lock_SC_high_level_inv:
∀ d d´ i,
big_release_lock_SC_spec i d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. unfold big_release_lock_SC_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto; intros.
Qed.
Lemma release_lock_SC_low_level_inv:
∀ d d´ i n,
big_release_lock_SC_spec i d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. unfold big_release_lock_SC_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto.
Qed.
Lemma release_lock_SC_kernel_mode:
∀ d d´ i,
big_release_lock_SC_spec i d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
intros. unfold big_release_lock_SC_spec in *; subdestruct; inv H.
- inv H0; eauto.
Qed.
Global Instance release_lock_SC_inv: PreservesInvariants big_release_lock_SC_spec.
Proof.
preserves_invariants_simpl´.
- eapply release_lock_SC_low_level_inv; eassumption.
- eapply release_lock_SC_high_level_inv; eassumption.
- eapply release_lock_SC_kernel_mode; eassumption.
Qed.
End SC_lock.
Section PALLOC.
Lemma palloc_high_level_inv:
∀ d d´ i n,
big_palloc_spec i d = Some (d´, n) →
high_level_invariant d →
high_level_invariant d´.
Proof.
unfold big_palloc_spec; intros.
subdestruct; inv H; subst; eauto;
inv H0; constructor; simpl; eauto; intros.
- rewrite <- Hdestruct3.
eapply alloc_container_valid´; eauto.
- subst; simpl;
intros; congruence.
- eapply dirty_ppage´_gso_alloc; eauto.
- intros.
eapply (valid_root0 i0 H) in H0.
destruct (zeq i0 i); subst.
+ rewrite ZMap.gss; trivial.
+ rewrite ZMap.gso; auto.
Qed.
Lemma palloc_low_level_inv:
∀ d d´ i n n´,
big_palloc_spec i d = Some (d´, n) →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
unfold big_palloc_spec; intros.
subdestruct; inv H; subst; eauto;
inv H0; constructor; eauto.
Qed.
Lemma palloc_kernel_mode:
∀ d d´ i n,
big_palloc_spec i d = Some (d´, n) →
kernel_mode d →
kernel_mode d´.
Proof.
unfold big_palloc_spec; intros.
subdestruct; inv H; simpl; eauto.
Qed.
Global Instance palloc_inv: PreservesInvariants big_palloc_spec.
Proof.
preserves_invariants_simpl´.
- eapply palloc_low_level_inv; eassumption.
- eapply palloc_high_level_inv; eassumption.
- eapply palloc_kernel_mode; eassumption.
Qed.
End PALLOC.
Global Instance hostin_inv: PrimInvariants hostin_spec.
Proof.
PrimInvariants_simpl_auto.
Qed.
Global Instance hostout_inv: PrimInvariants hostout_spec.
Proof.
PrimInvariants_simpl_auto.
Qed.
Global Instance setPT_inv: PreservesInvariants setPT_spec.
Proof.
preserves_invariants_simpl_auto.
Qed.
Section PTINSERT.
Section PTINSERT_PTE.
Lemma ptInsertPTE_high_level_inv:
∀ d d´ n vadr padr p,
ptInsertPTE0_spec n vadr padr p d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
inv H0; constructor_gso_simpl_tac; intros.
- eapply PMap_valid_gso_valid; eauto.
- functional inversion H2. functional inversion H1.
eapply PMap_kern_gso; eauto.
Qed.
Lemma ptInsertPTE_low_level_inv:
∀ d d´ n vadr padr p n´,
ptInsertPTE0_spec n vadr padr p d = Some d´ →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
intros. functional inversion H; subst; eauto.
inv H0. constructor; eauto.
Qed.
Lemma ptInsertPTE_kernel_mode:
∀ d d´ n vadr padr p,
ptInsertPTE0_spec n vadr padr p d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
Qed.
End PTINSERT_PTE.
Section PTPALLOCPDE.
Lemma big_palloc_inv_prop:
∀ n d d´ v,
big_palloc_spec n d = Some (d´, v) →
ptpool d´ = ptpool d
∧ (v ≠ 0 →
ZMap.get v (pperm d´) = PGAlloc
∧ ZMap.get v (LAT d´) = LATCValid nil
∧ 0 < v < nps d´)
∧ pg d´ = true.
Proof.
unfold big_palloc_spec. intros.
subdestruct; inv H; simpl; subst.
- repeat rewrite ZMap.gss. refine_split´; trivial.
destruct a0 as (? & ? & ?). intros.
refine_split´; trivial.
- refine_split´; trivial.
intros HF; elim HF. trivial.
- refine_split´; trivial.
intros HF; elim HF. trivial.
Qed.
Lemma ptAllocPDE_high_level_inv:
∀ d d´ n vadr v,
big_ptAllocPDE_spec n vadr d = Some (d´, v) →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
- eapply palloc_high_level_inv; eauto.
- exploit palloc_high_level_inv; eauto.
intros.
exploit big_palloc_inv_prop; eauto. intros (HPT & Halloc & Hpg).
clear H11.
rewrite <- HPT in ×.
inv H1; constructor_gso_simpl_tac; try (intros; congruence); intros.
+ eapply PMap_valid_gso_pde_unp; eauto.
eapply real_init_PTE_defined.
+ functional inversion H3.
eapply PMap_kern_gso; eauto.
+ eapply dirty_ppage´_gss; eauto.
Qed.
Lemma ptAllocPDE_low_level_inv:
∀ d d´ n vadr v n´,
big_ptAllocPDE_spec n vadr d = Some (d´, v) →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
intros. functional inversion H; subst; eauto.
- eapply palloc_low_level_inv; eauto.
- exploit palloc_low_level_inv; eauto.
intros. inv H1. constructor; eauto.
Qed.
Lemma ptAllocPDE_kernel_mode:
∀ d d´ n vadr v,
big_ptAllocPDE_spec n vadr d = Some (d´, v) →
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
- eapply palloc_kernel_mode; eauto.
- exploit palloc_kernel_mode; eauto.
Qed.
End PTPALLOCPDE.
Lemma ptInsert_high_level_inv:
∀ d d´ n vadr padr p v,
big_ptInsert_spec n vadr padr p d = Some (d´, v) →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
- eapply ptInsertPTE_high_level_inv; eassumption.
- eapply ptAllocPDE_high_level_inv; eassumption.
- eapply ptInsertPTE_high_level_inv; try eassumption.
eapply ptAllocPDE_high_level_inv; eassumption.
Qed.
Lemma ptInsert_low_level_inv:
∀ d d´ n vadr padr p n´ v,
big_ptInsert_spec n vadr padr p d = Some (d´, v) →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
intros. functional inversion H; subst; eauto.
- eapply ptInsertPTE_low_level_inv; eassumption.
- eapply ptAllocPDE_low_level_inv; eassumption.
- eapply ptInsertPTE_low_level_inv; try eassumption.
eapply ptAllocPDE_low_level_inv; eassumption.
Qed.
Lemma ptInsert_kernel_mode:
∀ d d´ n vadr padr p v,
big_ptInsert_spec n vadr padr p d = Some (d´, v) →
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
- eapply ptInsertPTE_kernel_mode; eassumption.
- eapply ptAllocPDE_kernel_mode; eassumption.
- eapply ptInsertPTE_kernel_mode; try eassumption.
eapply ptAllocPDE_kernel_mode; eassumption.
Qed.
End PTINSERT.
Section PTRESV.
Lemma ptResv_high_level_inv:
∀ d d´ n vadr p v,
big_ptResv_spec n vadr p d = Some (d´, v) →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto; clear H.
- eapply palloc_high_level_inv; eassumption.
- eapply ptInsert_high_level_inv; try eassumption.
eapply palloc_high_level_inv; eassumption.
Qed.
Lemma ptResv_low_level_inv:
∀ d d´ n vadr p n´ v,
big_ptResv_spec n vadr p d = Some (d´, v) →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
intros. functional inversion H; subst; eauto.
eapply palloc_low_level_inv; eassumption.
eapply ptInsert_low_level_inv; try eassumption.
eapply palloc_low_level_inv; eassumption.
Qed.
Lemma ptResv_kernel_mode:
∀ d d´ n vadr p v,
big_ptResv_spec n vadr p d = Some (d´, v) →
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
eapply palloc_kernel_mode; eassumption.
eapply ptInsert_kernel_mode; try eassumption.
eapply palloc_kernel_mode; eassumption.
Qed.
Global Instance ptResv_inv: PreservesInvariants big_ptResv_spec.
Proof.
preserves_invariants_simpl´.
- eapply ptResv_low_level_inv; eassumption.
- eapply ptResv_high_level_inv; eassumption.
- eapply ptResv_kernel_mode; eassumption.
Qed.
End PTRESV.
Global Instance thread_wakeup_inv: PreservesInvariants big_thread_wakeup_spec.
Proof.
preserves_invariants_simpl´.
- intros. functional inversion H2; try subst; eauto.
+ inv H0. econstructor; eauto.
+ subst me l´ l1 to. inv H0. econstructor; eauto.
+ inv H0. econstructor; eauto.
- intros. functional inversion H2; try subst; eauto.
+ inv H0. econstructor; simpl; eauto.
+ subst me l´ l1 to. inv H0. econstructor; eauto.
+ inv H0. econstructor; simpl; eauto.
- intros. functional inversion H2; try subst; eauto; simpl.
Qed.
Global Instance big_proc_create_inv: PCreateInvariants big_proc_create_spec.
Proof.
constructor; intros; inv H0;
unfold big_proc_create_spec in *;
subdestruct; inv H; simpl; auto.
-
constructor; trivial; intros; simpl in ×.
+ eapply kctxt_inject_neutral_gss_flatinj´; eauto.
eapply kctxt_inject_neutral_gss_flatinj; eauto.
+ simpl.
eapply uctxt_inject_neutral_gss; eauto.
repeat eapply uctxt_inject_neutral0_gss;
try eapply uctxt_inject_neutral0_Vint.
× eapply uctxt_inject_neutral_impl; eauto. omega.
× eapply uctxt_inject_neutral0_Vptr_flat; eauto.
-
assert (Hrdyq: 0 ≤ rdy_q_id (CPU_ID d) < num_chan).
{
unfold rdy_q_id. omega.
}
constructor; simpl; eauto 2; try congruence; intros.
+ exploit split_container_valid; eauto.
simpl. omega.
rewrite Hdestruct3.
auto.
+ generalize (valid_root0 i H); intros Htemp.
exploit Htemp; auto; intros valid_idle_cid; clear Htemp.
destruct (zeq i ((ZMap.get (CPU_ID d) (cid d)) × 8 + 1 +
Z.of_nat (length (cchildren (ZMap.get (ZMap.get (CPU_ID d) (cid d)) (AC d)))))); subst.
rewrite ZMap.gss; simpl; split; auto.
rewrite ZMap.gso; auto.
destruct (zeq i (ZMap.get (CPU_ID d) (cid d))); subst.
× rewrite ZMap.gss; simpl; auto.
× rewrite ZMap.gso; auto.
Qed.
Global Instance sched_init_inv: PreservesInvariants big_sched_init_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant.
- apply real_nps_range.
- apply AC_init_container_valid.
- eapply real_pt_PMap_valid; eauto.
- apply real_pt_PMap_kern.
- omega.
- assumption.
- apply real_idpde_init.
- unfold AC_init.
unfold init_real_container.
simpl.
case_eq (zeq i0 8); intros; subst; [rewrite ZMap.gss; auto | ].
rewrite ZMap.gso; auto.
case_eq (zeq i0 7); intros; subst; [rewrite ZMap.gss; auto | ].
rewrite ZMap.gso; auto.
case_eq (zeq i0 6); intros; subst; [rewrite ZMap.gss; auto | ].
rewrite ZMap.gso; auto.
case_eq (zeq i0 5); intros; subst; [rewrite ZMap.gss; auto | ].
rewrite ZMap.gso; auto.
case_eq (zeq i0 4); intros; subst; [rewrite ZMap.gss; auto | ].
rewrite ZMap.gso; auto.
case_eq (zeq i0 3); intros; subst; [rewrite ZMap.gss; auto | ].
rewrite ZMap.gso; auto.
case_eq (zeq i0 2); intros; subst; [rewrite ZMap.gss; auto | ].
rewrite ZMap.gso; auto.
case_eq (zeq i0 1); intros; subst; [rewrite ZMap.gss; auto | ].
rewrite ZMap.gso; auto.
case_eq (zeq i0 0); intros; subst; [rewrite ZMap.gss; auto | ].
omega.
- erewrite real_valid_cidpool; eauto.
omega.
Qed.
Local Opaque remove.
Section YIELD.
Global Instance thread_yield_inv: ThreadScheduleInvariants big_thread_yield_spec.
Proof.
constructor; intros; functional inversion H.
- inv H1.
constructor; trivial; simpl.
eapply kctxt_inject_neutral_gss_mem; eauto.
- inv H0.
constructor; auto; simpl in *; intros; try congruence; auto.
rewrite ZMap.gss. omega.
Qed.
End YIELD.
Global Instance thread_sleep_inv: ThreadTransferInvariants big_thread_sleep_spec.
Proof.
constructor; intros; functional inversion H.
- inv H1.
constructor; trivial.
eapply kctxt_inject_neutral_gss_mem; eauto.
- inv H0.
constructor; auto; simpl in *; intros; try congruence; auto.
rewrite ZMap.gss. omega.
Qed.
Global Instance uctx_set_inv: PreservesInvariants uctx_set_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
eapply uctxt_inject_neutral_gss; eauto.
eapply uctxt_inject_neutral0_gss; eauto.
eapply uctxt_inject_neutral0_Vint.
Qed.
Global Instance proc_start_user_inv: StartUserInvariants proc_start_user_spec.
Proof.
constructor; intros; functional inversion H; inv H0.
- constructor; trivial.
- subst. constructor; auto; simpl in *; intros; congruence.
Qed.
Global Instance proc_exit_user_inv: ExitUserInvariants proc_exit_user_spec.
Proof.
constructor; intros; functional inversion H; inv H0.
- constructor; trivial.
eapply uctxt_inject_neutral_gss; eauto.
repeat eapply uctxt_inject_neutral0_gss;
try eapply uctxt_inject_neutral0_Vint.
eapply uctxt_inject_neutral0_init.
- constructor; auto; simpl in *; intros.
Qed.
Global Instance proc_start_user_inv2: StartUserInvariants2 proc_start_user_spec.
Proof.
constructor; intros; functional inversion H; inv H0.
- constructor; trivial.
- subst. constructor; auto; simpl in *; intros; congruence.
Qed.
Global Instance proc_exit_user_inv2: ExitUserInvariants2 proc_exit_user_spec.
Proof.
constructor; intros; functional inversion H; inv H0.
- constructor; trivial.
eapply uctxt_inject_neutral_gss; eauto.
repeat eapply uctxt_inject_neutral0_gss;
try eapply uctxt_inject_neutral0_Vint.
eapply uctxt_inject_neutral0_init.
- constructor; auto; simpl in *; intros.
Qed.
Global Instance proc_create_postinit_inv:
PreservesInvariants proc_create_postinit_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Section INTELVMSUPPORT.
Section INTERCEPT_INTWIN.
Lemma vmx_set_intercept_intwin_high_inv:
∀ d d´ i,
vmx_set_intercept_intwin_spec i d = Some d´→
high_level_invariant d →
high_level_invariant d´.
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 d´ n i,
vmx_set_intercept_intwin_spec i d = Some d´→
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
eapply VMCS_inject_neutral_gss_Vint; eauto.
+ rewrite ZMap.gso; eauto.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
eapply VMCS_inject_neutral_gss_Vint; eauto.
+ rewrite ZMap.gso; eauto.
Qed.
Lemma vmx_set_intercept_intwin_kernel_mode:
∀ d d´ i,
vmx_set_intercept_intwin_spec i d = Some d´→
kernel_mode d →
kernel_mode d´.
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 SET_DESC1.
Lemma vmx_set_desc1_high_inv:
∀ d d´ i i0 i1,
vmx_set_desc1_spec i i0 i1 d = Some d´→
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
Qed.
Lemma vmx_set_desc1_low_inv:
∀ d d´ n i i0 i1,
vmx_set_desc1_spec i i0 i1 d = Some d´→
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
functional inversion H6.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
+ rewrite ZMap.gso; eauto.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
+ rewrite ZMap.gso; eauto.
Qed.
Lemma vmx_set_desc1_kernel_mode:
∀ d d´ i i0 i1,
vmx_set_desc1_spec i i0 i1 d = Some d´→
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; constructor; trivial.
Qed.
Global Instance vmx_set_desc1_inv: PreservesInvariants vmx_set_desc1_spec.
Proof.
preserves_invariants_simpl´.
- eapply vmx_set_desc1_low_inv; eauto.
- eapply vmx_set_desc1_high_inv; eauto.
- eapply vmx_set_desc1_kernel_mode; eauto.
Qed.
End SET_DESC1.
Section SET_DESC2.
Lemma vmx_set_desc2_high_inv:
∀ d d´ i i0 i1,
vmx_set_desc2_spec i i0 i1 d = Some d´→
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
Qed.
Lemma vmx_set_desc2_low_inv:
∀ d d´ n i i0 i1,
vmx_set_desc2_spec i i0 i1 d = Some d´→
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
functional inversion H6.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
+ rewrite ZMap.gso; eauto.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
+ rewrite ZMap.gso; eauto.
Qed.
Lemma vmx_set_desc2_kernel_mode:
∀ d d´ i i0 i1,
vmx_set_desc2_spec i i0 i1 d = Some d´→
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; constructor; trivial.
Qed.
Global Instance vmx_set_desc2_inv: PreservesInvariants vmx_set_desc2_spec.
Proof.
preserves_invariants_simpl´.
- eapply vmx_set_desc2_low_inv; eauto.
- eapply vmx_set_desc2_high_inv; eauto.
- eapply vmx_set_desc2_kernel_mode; eauto.
Qed.
End SET_DESC2.
Section INJECT_EVENT.
Lemma vmx_inject_event_high_inv:
∀ d d´ i i0 i1 i2,
vmx_inject_event_spec i i0 i1 i2 d = Some d´→
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
Qed.
Lemma vmx_inject_event_low_inv:
∀ d d´ n i i0 i1 i2,
vmx_inject_event_spec i i0 i1 i2 d = Some d´→
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
+ rewrite ZMap.gso; eauto.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
+ rewrite ZMap.gso; eauto.
Qed.
Lemma vmx_inject_event_kernel_mode:
∀ d d´ i i0 i1 i2,
vmx_inject_event_spec i i0 i1 i2 d = Some d´→
kernel_mode d →
kernel_mode d´.
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_TSC_OFFSET.
Lemma vmx_set_tsc_offset_high_inv:
∀ d d´ i,
vmx_set_tsc_offset_spec i d = Some d´→
high_level_invariant d →
high_level_invariant d´.
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 d´ n i,
vmx_set_tsc_offset_spec i d = Some d´→
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss.
repeat eapply VMCS_inject_neutral_gss_Vint_same; eauto.
+ rewrite ZMap.gso; eauto.
Qed.
Lemma vmx_set_tsc_offset_kernel_mode:
∀ d d´ i,
vmx_set_tsc_offset_spec i d = Some d´→
kernel_mode d →
kernel_mode d´.
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.
Section SET_REG.
Lemma vmx_set_reg_high_inv:
∀ d d´ i i0,
vmx_set_reg_spec i i0 d = Some d´→
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
Qed.
Lemma vmx_set_reg_low_inv:
∀ d d´ n i i0,
vmx_set_reg_spec i i0 d = Some d´→
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. inv H0; functional inversion H; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
- unfold VMXPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss; eapply VMX_inject_neutral_gss_Vint; eauto.
+ rewrite ZMap.gso; eauto.
- unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss; eapply VMCS_inject_neutral_gss_Vint; eauto.
+ rewrite ZMap.gso; eauto.
Qed.
Lemma vmx_set_reg_kernel_mode:
∀ d d´ i i0,
vmx_set_reg_spec i i0 d = Some d´→
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; constructor; trivial.
Qed.
Global Instance vmx_set_reg_inv: PreservesInvariants vmx_set_reg_spec.
Proof.
preserves_invariants_simpl´.
- eapply vmx_set_reg_low_inv; eauto.
- eapply vmx_set_reg_high_inv; eauto.
- eapply vmx_set_reg_kernel_mode; eauto.
Qed.
End SET_REG.
Section SET_MMAP.
Lemma vmx_set_mmap_high_inv:
∀ d d´ i i0 i1 v,
vmx_set_mmap_spec i i0 i1 d = Some (d´, v)→
high_level_invariant d →
high_level_invariant d´.
Proof.
intros.
inv H0; functional inversion H; functional inversion H2; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
Qed.
Lemma vmx_set_mmap_low_inv:
∀ d d´ n i i0 i1 v,
vmx_set_mmap_spec i i0 i1 d = Some (d´, v)→
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros.
inv H0; functional inversion H; functional inversion H2; subst; simpl;
constructor; simpl; intros; try congruence; eauto 2.
Qed.
Lemma vmx_set_mmap_kernel_mode:
∀ d d´ i i0 i1 v,
vmx_set_mmap_spec i i0 i1 d = Some (d´, v)→
kernel_mode d →
kernel_mode d´.
Proof.
intros; functional inversion H;
functional inversion H3; subst; constructor; trivial.
Qed.
Global Instance vmx_set_mmap_inv: PreservesInvariants vmx_set_mmap_spec.
Proof.
preserves_invariants_simpl´.
- eapply vmx_set_mmap_low_inv; eauto.
- eapply vmx_set_mmap_high_inv; eauto.
- eapply vmx_set_mmap_kernel_mode; eauto.
Qed.
End SET_MMAP.
Global Instance vmx_run_inv: VMXEnterInvariants vm_run_spec.
Proof.
constructor; intros;
unfold vm_run_spec in *;
subdestruct; inv H; simpl; auto.
-
inv H1. constructor; trivial; intros; simpl in ×.
+ unfold VMCSPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
× rewrite ZMap.gss.
eapply VMCS_inject_neutral_gss_Vint; eauto.
× rewrite ZMap.gso; eauto.
+ unfold VMXPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
× rewrite ZMap.gss.
eapply VMX_inject_neutral_gss_enter; eauto.
× rewrite ZMap.gso; eauto.
-
inv H0. constructor; simpl; eauto 2; try congruence; intros.
Qed.
Global Instance vmx_return_from_guest_inv: VMXReturnInvariants vmx_return_from_guest_spec.
Proof.
constructor; intros;
unfold vmx_return_from_guest_spec in *;
subdestruct; inv H; simpl; auto.
-
inv H0. constructor; trivial; intros; simpl in ×.
unfold VMXPool_inject_neutral; intros.
case_eq (zeq vmid (CPU_ID d)); intros; subst.
+ rewrite ZMap.gss; repeat eapply VMX_inject_neutral_gss_Vint; eauto.
+ rewrite ZMap.gso; eauto.
-
inv H0. constructor; simpl; eauto 2; try congruence; intros.
Qed.
Global Instance vmx_init_inv: VMCSSetDefaultsInvariants vmx_init_spec.
Proof.
constructor; intros; functional inversion H; simpl; auto.
-
inv H6. constructor; trivial; intros; simpl in ×.
+ eapply real_vmcs_inject_neutral; eauto.
+ eapply real_vmx_inject_neutral; eauto.
-
inv H0. constructor; simpl; try congruence; intros; eauto.
Qed.
End INTELVMSUPPORT.
End INV.
Definition exec_loadex {F V} := exec_loadex3 (F := F) (V := V).
Definition exec_storeex {F V} := exec_storeex3 (flatmem_store:= flatmem_store) (F := F) (V := V).
Global Instance flatmem_store_inv: FlatmemStoreInvariant (flatmem_store:= flatmem_store).
Proof.
split; inversion 1; intros.
- functional inversion H0. split; trivial.
- functional inversion H1.
split; simpl; try (eapply dirty_ppage´_store_unmapped´; try reflexivity; try eassumption); trivial.
Qed.
Global Instance trapinfo_set_inv: TrapinfoSetInvariant.
Proof.
split; inversion 1; intros; constructor; auto.
Qed.
Definition pbthread_passthrough : compatlayer (cdata RData) :=
vmxinfo_get ↦ gensem vmxinfo_get_spec
⊕ palloc ↦ gensem big_palloc_spec
⊕ set_pt ↦ gensem setPT_spec
⊕ pt_read ↦ gensem ptRead_spec
⊕ pt_resv ↦ gensem big_ptResv_spec
⊕ thread_wakeup ↦ gensem big_thread_wakeup_spec
⊕ thread_yield ↦ primcall_thread_schedule_compatsem big_thread_yield_spec (prim_ident:= thread_yield)
⊕ thread_sleep ↦ primcall_thread_transfer_compatsem big_thread_sleep_spec
⊕ sched_init ↦ gensem big_sched_init_spec
⊕ uctx_get ↦ gensem uctx_get_spec
⊕ uctx_set ↦ gensem uctx_set_spec
⊕ proc_create ↦ proc_create_compatsem big_proc_create_spec
⊕ container_get_nchildren ↦ gensem container_get_nchildren_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
⊕ get_CPU_ID ↦ gensem get_CPU_ID_spec
⊕ get_curid ↦ gensem get_curid_spec
⊕ acquire_lock_CHAN ↦ gensem big_acquire_lock_SC_spec
⊕ release_lock_CHAN ↦ gensem big_release_lock_SC_spec
⊕ get_sync_chan_busy ↦ gensem get_sync_chan_busy_spec
⊕ set_sync_chan_busy ↦ gensem set_sync_chan_busy_spec
⊕ ipc_send_body ↦ gensem big_ipc_send_body_spec
⊕ ipc_receive_body ↦ gensem big_ipc_receive_body_spec
⊕ rdmsr ↦ gensem rdmsr_spec
⊕ wrmsr ↦ gensem wrmsr_spec
⊕ vmx_set_intercept_intwin ↦ gensem vmx_set_intercept_intwin_spec
⊕ vmx_set_desc1 ↦ gensem vmx_set_desc1_spec
⊕ vmx_set_desc2 ↦ gensem vmx_set_desc2_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_get_exit_qualification ↦ gensem vmx_get_exit_qualification_spec
⊕ vmx_check_pending_event ↦ gensem vmx_check_pending_event_spec
⊕ vmx_check_int_shadow ↦ gensem vmx_check_int_shadow_spec
⊕ vmx_get_reg ↦ gensem vmx_get_reg_spec
⊕ vmx_set_reg ↦ gensem vmx_set_reg_spec
⊕ vmx_get_next_eip ↦ gensem vmx_get_next_eip_spec
⊕ vmx_get_io_width ↦ gensem vmx_get_io_width_spec
⊕ vmx_get_io_write ↦ gensem vmx_get_io_write_spec
⊕ vmx_get_exit_io_rep ↦ gensem vmx_get_exit_io_rep_spec
⊕ vmx_get_exit_io_str ↦ gensem vmx_get_exit_io_str_spec
⊕ vmx_get_exit_io_port ↦ gensem vmx_get_exit_io_port_spec
⊕ vmx_set_mmap ↦ gensem vmx_set_mmap_spec
⊕ vmx_run_vm ↦ primcall_vmx_enter_compatsem vm_run_spec vmx_run_vm
⊕ vmx_return_from_guest ↦ primcall_vmx_return_compatsem vmx_return_from_guest_spec
⊕ vmx_init ↦ vmcs_set_defaults_compatsem vmx_init_spec
⊕ cli ↦ gensem cli_spec
⊕ sti ↦ gensem sti_spec
⊕ serial_intr_disable ↦ gensem serial_intr_disable_spec
⊕ serial_intr_enable ↦ gensem serial_intr_enable_spec
⊕ serial_putc ↦ gensem serial_putc_spec
⊕ cons_buf_read ↦ gensem cons_buf_read_spec
⊕ host_in ↦ primcall_general_compatsem hostin_spec
⊕ host_out ↦ primcall_general_compatsem hostout_spec
⊕ proc_create_postinit ↦ gensem proc_create_postinit_spec
⊕ trap_get ↦ primcall_trap_info_get_compatsem trap_info_get_spec
⊕ trap_set ↦ primcall_trap_info_ret_compatsem trap_info_ret_spec
⊕ proc_start_user ↦ primcall_start_user_compatsem proc_start_user_spec
⊕ proc_exit_user ↦ primcall_exit_user_compatsem proc_exit_user_spec
⊕ proc_start_user2 ↦ primcall_start_user_compatsem2 proc_start_user_spec
⊕ proc_exit_user2 ↦ primcall_exit_user_compatsem2 proc_exit_user_spec
⊕ accessors ↦ {| exec_load := @exec_loadex; exec_store := @exec_storeex |}.
Definition pbthread_fresh : compatlayer (cdata RData) :=
∅ ⊕ ∅.
Definition pbthread :=
pbthread_fresh ⊕ pbthread_passthrough.
End WITHMEM.