Library mcertikos.mm.MALOp
This file defines the abstract data and the primitives for the MALOp layer, which will hide the MMTable and initialize the allocation table
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 LoadStoreSem1.
Require Import liblayers.logic.PTreeModules.
Require Import liblayers.logic.LayerLogicImpl.
Require Import liblayers.compat.CompatLayers.
Require Import liblayers.compat.CompatGenSem.
Require Import FutureTactic.
Require Import INVLemmaContainer.
Require Import INVLemmaMemory.
Require Import INVLemmaQLock.
Require Import INVLemmaInterrupt.
Require Import INVLemmaDriver.
Require Import AbstractDataType.
Require Import DeviceStateDataType.
Require Export ObjCPU.
Require Export ObjMM.
Require Export ObjFlatMem.
Require Export ObjContainer.
Require Export ObjMultiprocessor.
Require Export ObjQLock.
Require Export ObjPMM.
Require Export ObjInterruptManagement.
Require Export ObjInterruptController.
Require Export ObjConsole.
Require Export ObjSerialDriver.
Section WITHMEM.
Local Open Scope Z_scope.
Context `{real_params: RealParams}.
Context `{multi_oracle_prop: MultiOracleProp}.
Record high_level_invariant (abd: RData) :=
mkInvariant {
CPU_ID_range: 0 ≤ (CPU_ID abd) < TOTAL_CPU;
valid_curid: 0 ≤ ZMap.get (CPU_ID abd) (cid abd) < num_proc;
valid_nps: init abd= true → kern_low ≤ nps abd ≤ maxpage;
valid_kern: ikern abd = false → pg abd = true ∧ init abd = true;
valid_AT_kern: init abd = true → AT_kern (AT abd) (nps abd);
valid_AT_usr: init abd = true → AT_usr (AT abd) (nps abd);
valid_CR3: pg abd= true → CR3_valid (CR3 abd);
valid_ihost: ihost abd = false → pg abd = true ∧ init abd = true ∧ ikern abd = true;
valid_container: Container_valid (AC abd);
valid_multi_oracle_pool_inv: valid_multi_oracle_pool_H (multi_oracle abd);
valid_hlock_pool_inv: valid_hlock_pool (multi_log abd);
valid_AT_oracle_pool_inv: valid_AT_oracle_pool (multi_oracle abd);
valid_AT_log_pool_inv: init abd = true → valid_AT_log_pool (multi_log abd) (nps abd);
valid_cons_buf_rpos: 0 ≤ rpos (console abd) < CONSOLE_BUFFER_SIZE;
valid_cons_buf_length: 0 ≤ Zlength (cons_buf (console abd)) < CONSOLE_BUFFER_SIZE
}.
mkInvariant {
CPU_ID_range: 0 ≤ (CPU_ID abd) < TOTAL_CPU;
valid_curid: 0 ≤ ZMap.get (CPU_ID abd) (cid abd) < num_proc;
valid_nps: init abd= true → kern_low ≤ nps abd ≤ maxpage;
valid_kern: ikern abd = false → pg abd = true ∧ init abd = true;
valid_AT_kern: init abd = true → AT_kern (AT abd) (nps abd);
valid_AT_usr: init abd = true → AT_usr (AT abd) (nps abd);
valid_CR3: pg abd= true → CR3_valid (CR3 abd);
valid_ihost: ihost abd = false → pg abd = true ∧ init abd = true ∧ ikern abd = true;
valid_container: Container_valid (AC abd);
valid_multi_oracle_pool_inv: valid_multi_oracle_pool_H (multi_oracle abd);
valid_hlock_pool_inv: valid_hlock_pool (multi_log abd);
valid_AT_oracle_pool_inv: valid_AT_oracle_pool (multi_oracle abd);
valid_AT_log_pool_inv: init abd = true → valid_AT_log_pool (multi_log abd) (nps abd);
valid_cons_buf_rpos: 0 ≤ rpos (console abd) < CONSOLE_BUFFER_SIZE;
valid_cons_buf_length: 0 ≤ Zlength (cons_buf (console abd)) < CONSOLE_BUFFER_SIZE
}.
Global Instance malop_data_ops : CompatDataOps RData :=
{
empty_data := init_adt multi_oracle_init2;
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_init2;
high_level_invariant := high_level_invariant;
low_level_invariant := low_level_invariant;
kernel_mode adt := ikern adt = true ∧ ihost adt = true
}.
Lemma empty_data_high_level_invariant:
high_level_invariant (init_adt multi_oracle_init2).
Proof.
constructor; simpl; intros; auto; try inv H.
- apply current_CPU_ID_range.
- rewrite ZMap.gi; intuition.
-
apply empty_container_valid.
- apply valid_ticket_oracle2.
- apply valid_hlock_pool_init.
- eapply valid_AT_oracle_pool2.
Qed.
high_level_invariant (init_adt multi_oracle_init2).
Proof.
constructor; simpl; intros; auto; try inv H.
- apply current_CPU_ID_range.
- rewrite ZMap.gi; intuition.
-
apply empty_container_valid.
- apply valid_ticket_oracle2.
- apply valid_hlock_pool_init.
- eapply valid_AT_oracle_pool2.
Qed.
Global Instance malop_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.
Global Instance release_lock_inv: ReleaseInvariants release_lock_spec1.
Proof.
constructor; unfold release_lock_spec; intros; subdestruct;
inv H; inv H0; constructor; auto; simpl; intros.
- eapply valid_hlock_pool_gss´; eauto.
- eapply valid_AT_log_pool_gso; eauto.
eapply Shared2ID1_neq; eauto.
Qed.
Global Instance acquire_lock_inv: AcquireInvariants acquire_lock_spec1.
Proof.
constructor; unfold acquire_lock_spec; intros; subdestruct;
inv H; inv H0; constructor; auto; simpl; intros.
- eapply valid_hlock_pool_gss´; eauto.
- eapply valid_AT_log_pool_gso; eauto.
eapply Shared2ID1_neq; eauto.
Qed.
Lemma acquire_lock_AT_high_level_inv:
∀ d d´,
acquire_lock_AT_spec d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. unfold acquire_lock_AT_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto; intros.
+ eapply valid_AT_log_imply; eauto.
+ eapply valid_AT_log_imply; eauto.
+ eapply valid_hlock_pool_gss´; eauto.
+ eapply valid_AT_log_pool_gss´; eauto.
eapply valid_AT_log_pull; eauto.
eapply valid_AT_log_TICKET_imply; eauto.
- inv H0. constructor; simpl; eauto; intros.
+ eapply valid_hlock_pool_gss´; eauto.
+ eapply valid_AT_log_pool_gss´; eauto.
eapply valid_AT_log_pull; eauto.
eapply valid_AT_log_TICKET_imply; eauto.
Qed.
Lemma acquire_lock_AT_low_level_inv:
∀ d d´ n,
acquire_lock_AT_spec d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. unfold acquire_lock_AT_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto.
- inv H0. constructor; simpl; eauto.
Qed.
Lemma acquire_lock_AT_kernel_mode:
∀ d d´,
acquire_lock_AT_spec d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
intros. unfold acquire_lock_AT_spec in *; subdestruct; inv H.
- inv H0; eauto.
- inv H0; eauto.
Qed.
Global Instance acquire_lock_AT_inv: PreservesInvariants acquire_lock_AT_spec.
Proof.
preserves_invariants_simpl´.
- eapply acquire_lock_AT_low_level_inv; eassumption.
- eapply acquire_lock_AT_high_level_inv; eassumption.
- eapply acquire_lock_AT_kernel_mode; eassumption.
Qed.
Lemma release_lock_AT_high_level_inv:
∀ d d´,
release_lock_AT_spec d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. unfold release_lock_AT_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto; intros.
+ eapply valid_hlock_pool_gss´; eauto.
+ eapply valid_AT_log_pool_TICKET; eauto.
Qed.
Lemma release_lock_AT_low_level_inv:
∀ d d´ n,
release_lock_AT_spec d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. unfold release_lock_AT_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto.
Qed.
Lemma release_lock_AT_kernel_mode:
∀ d d´,
release_lock_AT_spec d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
intros. unfold release_lock_AT_spec in *; subdestruct; inv H.
- inv H0; eauto.
Qed.
Global Instance release_lock_AT_inv: PreservesInvariants release_lock_AT_spec.
Proof.
preserves_invariants_simpl´.
- eapply release_lock_AT_low_level_inv; eassumption.
- eapply release_lock_AT_high_level_inv; eassumption.
- eapply release_lock_AT_kernel_mode; eassumption.
Qed.
Global Instance set_curid_inv: PreservesInvariants set_curid_spec.
Proof.
preserves_invariants_simpl_auto.
- rewrite ZMap.gss. auto.
Qed.
Global Instance set_curid_init_inv: PreservesInvariants set_curid_init_spec.
Proof.
preserves_invariants_simpl_auto; eauto 2.
case_eq (zeq (CPU_ID d) (Int.unsigned i)); intros; subst.
- rewrite e; rewrite ZMap.gss; omega.
- rewrite ZMap.gso; auto.
Qed.
Global Instance mem_init_inv: PreservesInvariants mem_init_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant.
- apply real_nps_range.
- apply real_at_kern_valid.
- apply real_at_usr_valid.
- apply AC_init_container_valid.
- assumption.
- eapply real_valid_hlock_pool; eauto.
- assumption.
- eapply real_valid_AT_log_pool; eauto.
Qed.
Global Instance setPG_inv: PreservesInvariants setPG0_spec.
Proof.
preserves_invariants_simpl_auto.
Qed.
Global Instance set_at_u_inv: PreservesInvariants set_at_u_spec.
Proof.
preserves_invariants_simpl_auto.
- intros; eapply AT_kern_norm; eauto.
- intros; eapply AT_usr_norm; eauto.
Qed.
Global Instance set_at_c_inv: PreservesInvariants set_at_c_spec.
Proof.
preserves_invariants_simpl_auto.
Qed.
Global Instance setCR3_inv: SetCR3Invariants setCR30_spec.
Proof.
constructor; intros; functional inversion H.
- inv H0; constructor; trivial.
- inv H0; constructor; auto.
- assumption.
Qed.
Global Instance container_split_inv: PreservesInvariants container_split_spec.
Proof.
preserves_invariants_simpl_auto.
rewrite <- H1 in H0.
exploit split_container_valid; eauto.
Qed.
Global Instance container_alloc_inv: PreservesInvariants container_alloc_spec.
Proof.
preserves_invariants_simpl_auto.
exploit alloc_container_valid; eauto.
Qed.
Global Instance trapin_inv: PrimInvariants trapin_spec.
Proof.
PrimInvariants_simpl_auto.
Qed.
Global Instance trapout_inv: PrimInvariants trapout0_spec.
Proof.
PrimInvariants_simpl_auto.
Qed.
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 fstore_inv: PreservesInvariants fstore´_spec.
Proof.
preserves_invariants_simpl_auto.
Qed.
Global Instance page_copy_inv: PreservesInvariants page_copy´_spec.
Proof.
preserves_invariants_simpl_auto.
- eapply valid_hlock_pool_gss´; eauto.
- eapply valid_AT_log_pool_gso; eauto.
eapply Shared2ID1_neq; eauto.
unfold Shared2ID1; simpl; reflexivity.
Qed.
Global Instance page_copy_back_inv: PreservesInvariants page_copy_back´_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 low_level_invariant high_level_invariant; eauto 2.
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.
- 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.
- eauto 2 with serial_intr_enable_invariantdb.
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.
End INV.
Global Instance release_lock_inv: ReleaseInvariants release_lock_spec1.
Proof.
constructor; unfold release_lock_spec; intros; subdestruct;
inv H; inv H0; constructor; auto; simpl; intros.
- eapply valid_hlock_pool_gss´; eauto.
- eapply valid_AT_log_pool_gso; eauto.
eapply Shared2ID1_neq; eauto.
Qed.
Global Instance acquire_lock_inv: AcquireInvariants acquire_lock_spec1.
Proof.
constructor; unfold acquire_lock_spec; intros; subdestruct;
inv H; inv H0; constructor; auto; simpl; intros.
- eapply valid_hlock_pool_gss´; eauto.
- eapply valid_AT_log_pool_gso; eauto.
eapply Shared2ID1_neq; eauto.
Qed.
Lemma acquire_lock_AT_high_level_inv:
∀ d d´,
acquire_lock_AT_spec d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. unfold acquire_lock_AT_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto; intros.
+ eapply valid_AT_log_imply; eauto.
+ eapply valid_AT_log_imply; eauto.
+ eapply valid_hlock_pool_gss´; eauto.
+ eapply valid_AT_log_pool_gss´; eauto.
eapply valid_AT_log_pull; eauto.
eapply valid_AT_log_TICKET_imply; eauto.
- inv H0. constructor; simpl; eauto; intros.
+ eapply valid_hlock_pool_gss´; eauto.
+ eapply valid_AT_log_pool_gss´; eauto.
eapply valid_AT_log_pull; eauto.
eapply valid_AT_log_TICKET_imply; eauto.
Qed.
Lemma acquire_lock_AT_low_level_inv:
∀ d d´ n,
acquire_lock_AT_spec d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. unfold acquire_lock_AT_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto.
- inv H0. constructor; simpl; eauto.
Qed.
Lemma acquire_lock_AT_kernel_mode:
∀ d d´,
acquire_lock_AT_spec d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
intros. unfold acquire_lock_AT_spec in *; subdestruct; inv H.
- inv H0; eauto.
- inv H0; eauto.
Qed.
Global Instance acquire_lock_AT_inv: PreservesInvariants acquire_lock_AT_spec.
Proof.
preserves_invariants_simpl´.
- eapply acquire_lock_AT_low_level_inv; eassumption.
- eapply acquire_lock_AT_high_level_inv; eassumption.
- eapply acquire_lock_AT_kernel_mode; eassumption.
Qed.
Lemma release_lock_AT_high_level_inv:
∀ d d´,
release_lock_AT_spec d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. unfold release_lock_AT_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto; intros.
+ eapply valid_hlock_pool_gss´; eauto.
+ eapply valid_AT_log_pool_TICKET; eauto.
Qed.
Lemma release_lock_AT_low_level_inv:
∀ d d´ n,
release_lock_AT_spec d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. unfold release_lock_AT_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto.
Qed.
Lemma release_lock_AT_kernel_mode:
∀ d d´,
release_lock_AT_spec d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
intros. unfold release_lock_AT_spec in *; subdestruct; inv H.
- inv H0; eauto.
Qed.
Global Instance release_lock_AT_inv: PreservesInvariants release_lock_AT_spec.
Proof.
preserves_invariants_simpl´.
- eapply release_lock_AT_low_level_inv; eassumption.
- eapply release_lock_AT_high_level_inv; eassumption.
- eapply release_lock_AT_kernel_mode; eassumption.
Qed.
Global Instance set_curid_inv: PreservesInvariants set_curid_spec.
Proof.
preserves_invariants_simpl_auto.
- rewrite ZMap.gss. auto.
Qed.
Global Instance set_curid_init_inv: PreservesInvariants set_curid_init_spec.
Proof.
preserves_invariants_simpl_auto; eauto 2.
case_eq (zeq (CPU_ID d) (Int.unsigned i)); intros; subst.
- rewrite e; rewrite ZMap.gss; omega.
- rewrite ZMap.gso; auto.
Qed.
Global Instance mem_init_inv: PreservesInvariants mem_init_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant.
- apply real_nps_range.
- apply real_at_kern_valid.
- apply real_at_usr_valid.
- apply AC_init_container_valid.
- assumption.
- eapply real_valid_hlock_pool; eauto.
- assumption.
- eapply real_valid_AT_log_pool; eauto.
Qed.
Global Instance setPG_inv: PreservesInvariants setPG0_spec.
Proof.
preserves_invariants_simpl_auto.
Qed.
Global Instance set_at_u_inv: PreservesInvariants set_at_u_spec.
Proof.
preserves_invariants_simpl_auto.
- intros; eapply AT_kern_norm; eauto.
- intros; eapply AT_usr_norm; eauto.
Qed.
Global Instance set_at_c_inv: PreservesInvariants set_at_c_spec.
Proof.
preserves_invariants_simpl_auto.
Qed.
Global Instance setCR3_inv: SetCR3Invariants setCR30_spec.
Proof.
constructor; intros; functional inversion H.
- inv H0; constructor; trivial.
- inv H0; constructor; auto.
- assumption.
Qed.
Global Instance container_split_inv: PreservesInvariants container_split_spec.
Proof.
preserves_invariants_simpl_auto.
rewrite <- H1 in H0.
exploit split_container_valid; eauto.
Qed.
Global Instance container_alloc_inv: PreservesInvariants container_alloc_spec.
Proof.
preserves_invariants_simpl_auto.
exploit alloc_container_valid; eauto.
Qed.
Global Instance trapin_inv: PrimInvariants trapin_spec.
Proof.
PrimInvariants_simpl_auto.
Qed.
Global Instance trapout_inv: PrimInvariants trapout0_spec.
Proof.
PrimInvariants_simpl_auto.
Qed.
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 fstore_inv: PreservesInvariants fstore´_spec.
Proof.
preserves_invariants_simpl_auto.
Qed.
Global Instance page_copy_inv: PreservesInvariants page_copy´_spec.
Proof.
preserves_invariants_simpl_auto.
- eapply valid_hlock_pool_gss´; eauto.
- eapply valid_AT_log_pool_gso; eauto.
eapply Shared2ID1_neq; eauto.
unfold Shared2ID1; simpl; reflexivity.
Qed.
Global Instance page_copy_back_inv: PreservesInvariants page_copy_back´_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 low_level_invariant high_level_invariant; eauto 2.
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.
- 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.
- eauto 2 with serial_intr_enable_invariantdb.
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.
End INV.
Definition exec_loadex {F V} := exec_loadex1 (F := F) (V := V).
Definition exec_storeex {F V} := exec_storeex1 (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; constructor; auto.
- functional inversion H1; constructor; auto.
Qed.
Global Instance trapinfo_set_inv: TrapinfoSetInvariant.
Proof.
split; inversion 1; intros; constructor; auto.
Qed.
Definition exec_storeex {F V} := exec_storeex1 (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; constructor; auto.
- functional inversion H1; constructor; auto.
Qed.
Global Instance trapinfo_set_inv: TrapinfoSetInvariant.
Proof.
split; inversion 1; intros; constructor; auto.
Qed.
Definition malop_passthrough : compatlayer (cdata RData) :=
fload ↦ gensem fload´_spec
⊕ fstore ↦ gensem fstore´_spec
⊕ page_copy ↦ gensem page_copy´_spec
⊕ page_copy_back ↦ gensem page_copy_back´_spec
⊕ vmxinfo_get ↦ gensem vmxinfo_get_spec
⊕ set_pg ↦ gensem setPG0_spec
⊕ set_cr3 ↦ setCR3_compatsem setCR30_spec
⊕ get_nps ↦ gensem get_nps_spec
⊕ is_norm ↦ gensem is_at_norm_spec
⊕ at_get ↦ gensem get_at_u_spec
⊕ at_get_c ↦ gensem get_at_c_spec
⊕ at_set ↦ gensem set_at_u_spec
⊕ at_set_c ↦ gensem set_at_c_spec
⊕ container_get_parent ↦ gensem container_get_parent_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
⊕ container_split ↦ gensem container_split_spec
⊕ container_alloc ↦ gensem container_alloc_spec
⊕ acquire_lock_AT ↦ gensem acquire_lock_AT_spec
⊕ release_lock_AT ↦ gensem release_lock_AT_spec
⊕ get_CPU_ID ↦ gensem get_CPU_ID_spec
⊕ get_curid ↦ gensem get_curid_spec
⊕ set_curid ↦ gensem set_curid_spec
⊕ set_curid_init ↦ gensem set_curid_init_spec
⊕ release_lock ↦ primcall_release_lock_compatsem release_lock release_lock_spec1
⊕ acquire_lock ↦ primcall_acquire_lock_compatsem acquire_lock_spec1
⊕ 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
⊕ trap_in ↦ primcall_general_compatsem trapin_spec
⊕ trap_out ↦ primcall_general_compatsem trapout0_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
⊕ accessors ↦ {| exec_load := @exec_loadex; exec_store := @exec_storeex |}.
fload ↦ gensem fload´_spec
⊕ fstore ↦ gensem fstore´_spec
⊕ page_copy ↦ gensem page_copy´_spec
⊕ page_copy_back ↦ gensem page_copy_back´_spec
⊕ vmxinfo_get ↦ gensem vmxinfo_get_spec
⊕ set_pg ↦ gensem setPG0_spec
⊕ set_cr3 ↦ setCR3_compatsem setCR30_spec
⊕ get_nps ↦ gensem get_nps_spec
⊕ is_norm ↦ gensem is_at_norm_spec
⊕ at_get ↦ gensem get_at_u_spec
⊕ at_get_c ↦ gensem get_at_c_spec
⊕ at_set ↦ gensem set_at_u_spec
⊕ at_set_c ↦ gensem set_at_c_spec
⊕ container_get_parent ↦ gensem container_get_parent_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
⊕ container_split ↦ gensem container_split_spec
⊕ container_alloc ↦ gensem container_alloc_spec
⊕ acquire_lock_AT ↦ gensem acquire_lock_AT_spec
⊕ release_lock_AT ↦ gensem release_lock_AT_spec
⊕ get_CPU_ID ↦ gensem get_CPU_ID_spec
⊕ get_curid ↦ gensem get_curid_spec
⊕ set_curid ↦ gensem set_curid_spec
⊕ set_curid_init ↦ gensem set_curid_init_spec
⊕ release_lock ↦ primcall_release_lock_compatsem release_lock release_lock_spec1
⊕ acquire_lock ↦ primcall_acquire_lock_compatsem acquire_lock_spec1
⊕ 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
⊕ trap_in ↦ primcall_general_compatsem trapin_spec
⊕ trap_out ↦ primcall_general_compatsem trapout0_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
⊕ accessors ↦ {| exec_load := @exec_loadex; exec_store := @exec_storeex |}.
Definition malop : compatlayer (cdata RData) := malop_fresh ⊕ malop_passthrough.
End WITHMEM.
Section WITHPARAM.
Context `{real_params: RealParams}.
Local Open Scope Z_scope.
Section Impl.
Function palloc_aux_spec (n: Z) (adt: RData): option (RData × Z) :=
match (ikern adt, init adt, ihost adt) with
| (true, true, true) ⇒
let c := ZMap.get n (AC adt) in
match cused c with
| true ⇒
if (cusage c) <? (cquota c) then
let cur := mkContainer (cquota c) (cusage c + 1) (cparent c)
(cchildren c) (cused c) in
match first_free (AT adt) (nps adt) with
| inleft (exist i _) ⇒
match ZMap.get i (ATC adt) with
| ATCValid _ ⇒ Some (adt {AT: ZMap.set i (ATValid true ATNorm) (AT adt)}
{ATC: ZMap.set i (ATCValid 0) (ATC adt)}
{AC: ZMap.set n cur (AC adt)},i)
| _ ⇒ None
end
| _ ⇒ Some (adt, 0)
end
else Some (adt, 0)
| _ ⇒ None
end
| _ ⇒ None
end.
Function palloc_spec (n: Z) (adt: RData): option (RData × Z) :=
match acquire_lock_AT_spec adt with
| Some adt´ ⇒
match palloc_aux_spec n adt´ with
| Some (adt0, i) ⇒
match release_lock_AT_spec adt0 with
| Some adt´´ ⇒ Some (adt´´, i)
| _ ⇒ None
end
| _ ⇒ None
end
| _ ⇒ None
end.
Function pfree_spec (i: Z) (adt: RData): option RData :=
match (ikern adt, ihost adt, init adt, ZMap.get lock_AT_start (lock adt)) with
| (true, true, true, LockOwn true) ⇒
if zle_lt 0 i maxpage then
match (ZMap.get i (AT adt), ZMap.get i (ATC adt)) with
| (ATValid true ATNorm, ATCValid 0) ⇒
Some adt {AT: ZMap.set i (ATValid false ATNorm) (AT adt)}
| _ ⇒ None
end
else None
| _ ⇒ None
end.
End Impl.
End WITHPARAM.
End WITHMEM.
Section WITHPARAM.
Context `{real_params: RealParams}.
Local Open Scope Z_scope.
Section Impl.
Function palloc_aux_spec (n: Z) (adt: RData): option (RData × Z) :=
match (ikern adt, init adt, ihost adt) with
| (true, true, true) ⇒
let c := ZMap.get n (AC adt) in
match cused c with
| true ⇒
if (cusage c) <? (cquota c) then
let cur := mkContainer (cquota c) (cusage c + 1) (cparent c)
(cchildren c) (cused c) in
match first_free (AT adt) (nps adt) with
| inleft (exist i _) ⇒
match ZMap.get i (ATC adt) with
| ATCValid _ ⇒ Some (adt {AT: ZMap.set i (ATValid true ATNorm) (AT adt)}
{ATC: ZMap.set i (ATCValid 0) (ATC adt)}
{AC: ZMap.set n cur (AC adt)},i)
| _ ⇒ None
end
| _ ⇒ Some (adt, 0)
end
else Some (adt, 0)
| _ ⇒ None
end
| _ ⇒ None
end.
Function palloc_spec (n: Z) (adt: RData): option (RData × Z) :=
match acquire_lock_AT_spec adt with
| Some adt´ ⇒
match palloc_aux_spec n adt´ with
| Some (adt0, i) ⇒
match release_lock_AT_spec adt0 with
| Some adt´´ ⇒ Some (adt´´, i)
| _ ⇒ None
end
| _ ⇒ None
end
| _ ⇒ None
end.
Function pfree_spec (i: Z) (adt: RData): option RData :=
match (ikern adt, ihost adt, init adt, ZMap.get lock_AT_start (lock adt)) with
| (true, true, true, LockOwn true) ⇒
if zle_lt 0 i maxpage then
match (ZMap.get i (AT adt), ZMap.get i (ATC adt)) with
| (ATValid true ATNorm, ATCValid 0) ⇒
Some adt {AT: ZMap.set i (ATValid false ATNorm) (AT adt)}
| _ ⇒ None
end
else None
| _ ⇒ None
end.
End Impl.
End WITHPARAM.