Library mcertikos.proc.PThreadSched
This file defines the abstract data and the primitives for the PThreadSched 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 LoadStoreSem2.
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 INVLemmaContainer.
Require Import INVLemmaMemory.
Require Import INVLemmaThread.
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 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 LoadStoreSem2.
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 INVLemmaContainer.
Require Import INVLemmaMemory.
Require Import INVLemmaThread.
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.
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);
valid_pperm_ppage: consistent_ppage_log (multi_log abd) (pperm abd) (nps 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_pperm_pmap: weak_consistent_pmap (ptpool abd) (pperm abd) (LAT abd) (nps abd);
valid_pmap_domain: consistent_pmap_domain (ptpool abd) (pperm abd) (LAT abd) (nps abd);
valid_lat_domain: consistent_lat_domain (ptpool abd) (LAT abd) (nps abd);
valid_LATable_nil: LATCTable_log (multi_log abd) (LAT abd);
valid_root: pg abd = true → cused (ZMap.get 0 (AC abd)) = true;
valid_TCB: pg abd = true → AbTCBCorrect_range (abtcb abd);
valid_TDQ: pg abd = true → AbQCorrect_range (abq abd);
valid_count: pg abd = true → QCount (abtcb abd) (abq abd);
valid_inQ: pg abd = true → InQ (abtcb abd) (abq abd);
valid_multi_oracle_pool_inv: valid_multi_oracle_pool_H1 (multi_oracle abd);
valid_hlock_pool_inv: valid_hlock_pool1 (multi_log abd);
valid_AT_oracle_pool_inv: valid_AT_oracle_pool_H (multi_oracle abd);
valid_AT_log_pool_inv: valid_AT_log_pool_H (multi_log abd);
valid_ABTCB_oracle_pool_inv: valid_ABTCB_oracle_pool (multi_oracle abd);
valid_ABTCB_log_pool_inv: valid_ABTCB_log_pool (multi_log 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;
CPU_ID_range: 0 ≤ (CPU_ID abd) < TOTAL_CPU;
valid_curid: 0 ≤ ZMap.get (CPU_ID abd) (cid abd) < num_proc
}.
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);
valid_pperm_ppage: consistent_ppage_log (multi_log abd) (pperm abd) (nps 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_pperm_pmap: weak_consistent_pmap (ptpool abd) (pperm abd) (LAT abd) (nps abd);
valid_pmap_domain: consistent_pmap_domain (ptpool abd) (pperm abd) (LAT abd) (nps abd);
valid_lat_domain: consistent_lat_domain (ptpool abd) (LAT abd) (nps abd);
valid_LATable_nil: LATCTable_log (multi_log abd) (LAT abd);
valid_root: pg abd = true → cused (ZMap.get 0 (AC abd)) = true;
valid_TCB: pg abd = true → AbTCBCorrect_range (abtcb abd);
valid_TDQ: pg abd = true → AbQCorrect_range (abq abd);
valid_count: pg abd = true → QCount (abtcb abd) (abq abd);
valid_inQ: pg abd = true → InQ (abtcb abd) (abq abd);
valid_multi_oracle_pool_inv: valid_multi_oracle_pool_H1 (multi_oracle abd);
valid_hlock_pool_inv: valid_hlock_pool1 (multi_log abd);
valid_AT_oracle_pool_inv: valid_AT_oracle_pool_H (multi_oracle abd);
valid_AT_log_pool_inv: valid_AT_log_pool_H (multi_log abd);
valid_ABTCB_oracle_pool_inv: valid_ABTCB_oracle_pool (multi_oracle abd);
valid_ABTCB_log_pool_inv: valid_ABTCB_log_pool (multi_log 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;
CPU_ID_range: 0 ≤ (CPU_ID abd) < TOTAL_CPU;
valid_curid: 0 ≤ ZMap.get (CPU_ID abd) (cid abd) < num_proc
}.
Global Instance pthreadsched_data_ops : CompatDataOps RData :=
{
empty_data := init_adt multi_oracle_init6;
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_init6;
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_init6).
Proof.
constructor; simpl; intros; auto; try inv H.
- apply empty_container_valid.
- eapply consistent_ppage_log_init.
- eapply dirty_ppage_init.
- eapply weak_consistent_pmap_init.
- eapply consistent_pmap_domain_init.
- eapply consistent_lat_domain_init.
- eapply LATCTable_log_init.
- eapply valid_ticket_oracle6.
- apply valid_hlock_pool_init1.
- apply valid_AT_oracle_pool6.
- eapply valid_AT_log_pool_H_init.
- eapply valid_ABTCB_oracle_pool6; eauto.
- eapply valid_ABTCB_log_pool_init; eauto.
- apply current_CPU_ID_range.
- rewrite ZMap.gi; intuition.
Qed.
Lemma empty_data_high_level_invariant:
high_level_invariant (init_adt multi_oracle_init6).
Proof.
constructor; simpl; intros; auto; try inv H.
- apply empty_container_valid.
- eapply consistent_ppage_log_init.
- eapply dirty_ppage_init.
- eapply weak_consistent_pmap_init.
- eapply consistent_pmap_domain_init.
- eapply consistent_lat_domain_init.
- eapply LATCTable_log_init.
- eapply valid_ticket_oracle6.
- apply valid_hlock_pool_init1.
- apply valid_AT_oracle_pool6.
- eapply valid_AT_log_pool_H_init.
- eapply valid_ABTCB_oracle_pool6; eauto.
- eapply valid_ABTCB_log_pool_init; eauto.
- apply current_CPU_ID_range.
- rewrite ZMap.gi; intuition.
Qed.
Global Instance pthreadsched_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,
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,
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,
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,
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´,
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,
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 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,
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.
- eapply consistent_ppage_log_gso; eauto.
eapply Shared2ID2_neq; eauto.
reflexivity.
- eapply LATCTable_log_gso; eauto.
eapply Shared2ID2_neq; eauto.
reflexivity.
- eapply valid_hlock_pool1_gss´; eauto.
- eapply valid_AT_log_pool_H_gso; eauto.
eapply Shared2ID2_neq; eauto.
reflexivity.
- eapply valid_ABTCB_log_pool_gso´; eauto.
reflexivity.
Qed.
Lemma page_copy_low_level_inv:
∀ d d´ chid cound addr n,
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,
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,
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´,
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,
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 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.
Global Instance set_curid_inv: PreservesInvariants set_curid_spec.
Proof.
preserves_invariants_simpl_auto.
rewrite ZMap.gss. trivial.
Qed.
Global Instance sleeper_inc_inv:
PreservesInvariants sleeper_inc_spec.
Proof.
preserves_invariants_simpl_auto.
Qed.
Global Instance sleeper_dec_inv:
PreservesInvariants sleeper_dec_spec.
Proof.
preserves_invariants_simpl_auto.
Qed.
Section SC_lock.
Lemma acquire_lock_SC_high_level_inv:
∀ d d´ i,
acquire_lock_SC_spec i d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
unfold acquire_lock_SC_spec; intros.
subdestruct; inv H; eauto.
- inv H0. constructor; simpl; eauto; intros.
+ eapply consistent_ppage_log_gso; eauto.
eapply SC_neq´; omega.
+ eapply LATCTable_log_gso; eauto.
eapply SC_neq´; omega.
+ eapply valid_hlock_pool1_gss´; eauto.
+ eapply valid_AT_log_pool_H_gso; eauto.
eapply SC_neq´; omega.
+ eapply valid_ABTCB_log_pool_gso; eauto.
eapply SC_neq0; eauto.
- inv H0. constructor; simpl; eauto; intros.
+ eapply consistent_ppage_log_gso; eauto.
eapply SC_neq´; omega.
+ eapply LATCTable_log_gso; eauto.
eapply SC_neq´; omega.
+ eapply valid_hlock_pool1_gss´; eauto.
+ eapply valid_AT_log_pool_H_gso; eauto.
eapply SC_neq´; omega.
+ eapply valid_ABTCB_log_pool_gso; eauto.
eapply SC_neq0; eauto.
Qed.
Lemma acquire_lock_SC_low_level_inv:
∀ d d´ i n,
acquire_lock_SC_spec i d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
unfold 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,
acquire_lock_SC_spec i d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
unfold acquire_lock_SC_spec; intros.
subdestruct; inv H; eauto.
Qed.
Global Instance acquire_lock_SC_inv: PreservesInvariants 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,
release_lock_SC_spec i d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. unfold release_lock_SC_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto; intros.
+ eapply consistent_ppage_log_gso; eauto.
eapply SC_neq´. omega.
+ eapply LATCTable_log_gso; eauto.
eapply SC_neq´; omega.
+ eapply valid_hlock_pool1_gss´; eauto.
+ eapply valid_AT_log_pool_H_gso; eauto.
eapply SC_neq´; omega.
+ eapply valid_ABTCB_log_pool_gso; eauto.
eapply SC_neq0; eauto.
Qed.
Lemma release_lock_SC_low_level_inv:
∀ d d´ i n,
release_lock_SC_spec i d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. unfold release_lock_SC_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto.
Qed.
Lemma release_lock_SC_kernel_mode:
∀ d d´ i,
release_lock_SC_spec i d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
intros. unfold release_lock_SC_spec in *; subdestruct; inv H.
- inv H0; eauto.
Qed.
Global Instance release_lock_SC_inv: PreservesInvariants 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 TCB_lock.
Lemma acquire_lock_ABTCB_high_level_inv:
∀ d d´ i,
acquire_lock_ABTCB_spec i d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
unfold acquire_lock_ABTCB_spec; intros.
subdestruct; inv H; eauto.
- inv H0. constructor; simpl; eauto; intros.
+ eapply consistent_ppage_log_gso; eauto.
eapply TCB_neq´; omega.
+ eapply LATCTable_log_gso; eauto.
eapply TCB_neq´; omega.
+ eapply valid_ABTCB_oracle_pool_inv0; eauto.
+ exploit valid_ABTCB_oracle_pool_inv0; eauto.
intros (Hd & _). eapply Hd; eauto.
+ exploit valid_ABTCB_oracle_pool_inv0; eauto.
intros (Hd & _). eapply Hd; eauto.
+ exploit valid_ABTCB_oracle_pool_inv0; eauto.
intros (Hd & _). eapply Hd; eauto.
+ eapply valid_hlock_pool1_gss´; eauto.
+ eapply valid_AT_log_pool_H_gso; eauto.
eapply TCB_neq´; omega.
+ eapply valid_ABTCB_log_pool_gss; eauto.
eapply valid_ABTCB_log_pull; eauto.
eapply valid_ABTCB_log_wait; eauto.
eapply valid_ABTCB_oracle_pool_inv0; eauto.
- inv H0. constructor; simpl; eauto; intros.
+ eapply consistent_ppage_log_gso; eauto.
eapply TCB_neq´; omega.
+ eapply LATCTable_log_gso; eauto.
eapply TCB_neq´; omega.
+ eapply valid_hlock_pool1_gss´; eauto.
+ eapply valid_AT_log_pool_H_gso; eauto.
eapply TCB_neq´; omega.
+ eapply valid_ABTCB_log_pool_gss; eauto.
eapply valid_ABTCB_log_pull; eauto.
eapply valid_ABTCB_log_wait; eauto.
eapply valid_ABTCB_oracle_pool_inv0; eauto.
Qed.
Lemma acquire_lock_ABTCB_low_level_inv:
∀ d d´ i n,
acquire_lock_ABTCB_spec i d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
unfold acquire_lock_ABTCB_spec; intros.
subdestruct; inv H; eauto.
- inv H0. constructor; simpl; eauto.
- inv H0. constructor; simpl; eauto.
Qed.
Lemma acquire_lock_ABTCB_kernel_mode:
∀ d d´ i,
acquire_lock_ABTCB_spec i d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
unfold acquire_lock_ABTCB_spec; intros.
subdestruct; inv H; eauto.
Qed.
Lemma release_lock_ABTCB_high_level_inv:
∀ d d´ i,
release_lock_ABTCB_spec i d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. unfold release_lock_ABTCB_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto; intros.
+ eapply consistent_ppage_log_gso; eauto.
eapply TCB_neq´; omega.
+ eapply LATCTable_log_gso; eauto.
eapply TCB_neq´; omega.
+ eapply valid_hlock_pool1_gss´; eauto.
+ eapply valid_AT_log_pool_H_gso; eauto.
eapply TCB_neq´; omega.
+ eapply valid_ABTCB_log_pool_gss; eauto.
eapply valid_ABTCB_log_rel.
eapply valid_ABTCB_log_shared; eauto.
Qed.
Lemma release_lock_ABTCB_low_level_inv:
∀ d d´ i n,
release_lock_ABTCB_spec i d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. unfold release_lock_ABTCB_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto.
Qed.
Lemma release_lock_ABTCB_kernel_mode:
∀ d d´ i,
release_lock_ABTCB_spec i d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
intros. unfold release_lock_ABTCB_spec in *; subdestruct; inv H.
- inv H0; eauto.
Qed.
End TCB_lock.
Section PALLOC.
Lemma palloc_high_level_inv:
∀ d d´ i n,
palloc_spec i d = Some (d´, n) →
high_level_invariant d →
high_level_invariant d´.
Proof.
unfold palloc_spec; intros.
subdestruct; inv H; subst; eauto;
inv H0; constructor; simpl; eauto; intros.
+ rewrite <- Hdestruct3.
eapply alloc_container_valid´; eauto.
+ eapply consistent_ppage_log_norm_alloc; eauto. omega.
+ subst; simpl;
intros; congruence.
+ eapply dirty_ppage_gso_alloc; eauto.
+ eapply (weak_consistent_pmap_gso_at_palloc n); eauto; try apply a0.
+ eapply consistent_pmap_domain_gso_at; eauto.
intros. intro HF.
exploit (LATCTable_log_gss (ZMap.get 0 (multi_oracle d) (CPU_ID d) l) _ _ _
valid_LATable_nil0 Hdestruct6); eauto.
× rewrite ZMap.gss. trivial.
× eapply a0.
+ eapply consistent_lat_domain_gss_nil; eauto.
+ eapply LATCTable_log_alloc´; eauto.
+ intros. destruct (zeq i 0); subst.
× rewrite ZMap.gss; trivial.
× rewrite ZMap.gso; auto.
+ eapply valid_hlock_pool1_gso; eauto.
+ eapply valid_AT_log_pool_H_n; eauto.
eapply a0.
+ eapply valid_ABTCB_log_pool_gso_AT; eauto.
+ rewrite app_comm_cons.
eapply consistent_ppage_log_gss; eauto.
+ eapply LATCTable_log_alloc; eauto.
+ eapply valid_hlock_pool1_gso; eauto.
+ eapply valid_AT_log_pool_H_0; eauto.
+ eapply valid_ABTCB_log_pool_gso_AT; eauto.
+ rewrite app_comm_cons.
eapply consistent_ppage_log_gss; eauto.
+ eapply LATCTable_log_alloc; eauto.
+ eapply valid_hlock_pool1_gso; eauto.
+ eapply valid_AT_log_pool_H_0´; eauto.
+ eapply valid_ABTCB_log_pool_gso_AT; eauto.
Qed.
Lemma palloc_low_level_inv:
∀ d d´ i n n´,
palloc_spec i d = Some (d´, n) →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
unfold palloc_spec; intros.
subdestruct; inv H; subst; eauto;
inv H0; constructor; eauto.
Qed.
Lemma palloc_kernel_mode:
∀ d d´ i n,
palloc_spec i d = Some (d´, n) →
kernel_mode d →
kernel_mode d´.
Proof.
unfold palloc_spec; intros.
subdestruct; inv H; simpl; eauto.
Qed.
Global Instance palloc_inv: PreservesInvariants 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 trapin_inv: PrimInvariants trapin_spec.
Proof.
PrimInvariants_simpl_auto.
Qed.
Global Instance trapout_inv: PrimInvariants trapout_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 ptin_inv: PrimInvariants ptin_spec.
Proof.
PrimInvariants_simpl_auto.
Qed.
Global Instance ptout_inv: PrimInvariants ptout_spec.
Proof.
PrimInvariants_simpl_auto.
Qed.
Global Instance fstore_inv: PreservesInvariants fstore_spec.
Proof.
split; intros; inv_generic_sem H; inv H0; functional inversion H2.
- functional inversion H. split; trivial.
- functional inversion H.
split; subst; simpl;
try (eapply dirty_ppage_store_unmaped; try reflexivity; try eassumption); trivial.
- functional inversion H0.
split; simpl; try assumption.
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.
- functional inversion H2. functional inversion H0.
eapply weak_consistent_pmap_ptp_same; try eassumption.
eapply weak_consistent_pmap_gso_pperm_alloc´; eassumption.
- functional inversion H2.
eapply consistent_pmap_domain_append; eauto.
destruct (ZMap.get pti pdt); try contradiction;
red; intros (v0 & p0 & He); contra_inv.
- eapply consistent_lat_domain_gss_append; eauto.
subst pti; destruct (ZMap.get (PTX vadr) pdt); try contradiction;
red; intros (v0 & p0 & He); contra_inv.
- eapply LATCTable_log_not_nil_gso_true; eauto.
functional inversion H2. omega.
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 ptAllocPDE_high_level_inv:
∀ d d´ n vadr v,
ptAllocPDE0_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 palloc_inv_prop; eauto. intros (HPT & Halloc & Hpg).
clear H11.
rewrite <- HPT in ×.
inv H1; constructor_gso_simpl_tac; try (intros; congruence); intros.
+ apply consistent_ppage_log_alloc_hide; eauto.
eapply Halloc; eauto.
+ 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.
+ eapply weak_consistent_pmap_ptp_gss0; eauto; apply Halloc; eauto.
+ eapply consistent_pmap_domain_gso_at_00; eauto; try apply Halloc; eauto.
eapply consistent_pmap_domain_ptp_unp; eauto.
apply real_init_PTE_unp.
+ apply consistent_lat_domain_gso_p; eauto.
Qed.
Lemma ptAllocPDE_low_level_inv:
∀ d d´ n vadr v n´,
ptAllocPDE0_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,
ptAllocPDE0_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,
ptInsert0_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,
ptInsert0_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,
ptInsert0_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,
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,
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,
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 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.
Section PTRESV2.
Lemma ptResv2_high_level_inv:
∀ d d´ n vadr p n´ vadr´ p´ v,
ptResv2_spec n vadr p 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.
- eapply ptInsert_high_level_inv; try eassumption.
eapply ptInsert_high_level_inv; try eassumption.
eapply palloc_high_level_inv; eassumption.
Qed.
Lemma ptResv2_low_level_inv:
∀ d d´ n vadr p n´ vadr´ p´ l v,
ptResv2_spec n vadr p n´ vadr´ p´ d = Some (d´, v) →
low_level_invariant l d →
low_level_invariant l 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.
- eapply ptInsert_low_level_inv; try eassumption.
eapply ptInsert_low_level_inv; try eassumption.
eapply palloc_low_level_inv; eassumption.
Qed.
Lemma ptResv2_kernel_mode:
∀ d d´ n vadr p n´ vadr´ p´ v,
ptResv2_spec n vadr p 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.
- eapply ptInsert_kernel_mode; try eassumption.
eapply ptInsert_kernel_mode; try eassumption.
eapply palloc_kernel_mode; eassumption.
Qed.
End PTRESV2.
Section OFFER_SHARE.
Global Instance offer_shared_mem_inv:
PreservesInvariants offer_shared_mem_spec.
Proof.
preserves_invariants_simpl´;
functional inversion H2; subst; eauto 2; try (inv H0; constructor; trivial; fail).
- exploit ptResv2_low_level_inv; eauto.
intros HP; inv HP. constructor; trivial.
- exploit ptResv2_low_level_inv; eauto.
intros HP; inv HP. constructor; trivial.
- exploit ptResv2_high_level_inv; eauto.
intros HP; inv HP. constructor; trivial.
- exploit ptResv2_high_level_inv; eauto.
intros HP; inv HP. constructor; trivial.
- exploit ptResv2_kernel_mode; eauto.
- exploit ptResv2_kernel_mode; eauto.
Qed.
End OFFER_SHARE.
Global Instance shared_mem_status_inv:
PreservesInvariants shared_mem_status_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Section QUEUE.
Lemma enqueue_high_level_inv:
∀ d d´ n i,
enqueue0_spec n i d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
unfold enqueue0_spec; intros.
subdestruct; inv H; eauto.
inv H0.
constructor; simpl; eauto; intros;
functional inversion Hdestruct3.
- eapply AbTCBCorrect_range_gss; eauto. omega.
- eapply AbQCorrect_range_gss_enqueue; eauto.
- eapply QCount_gss_enqueue; eauto.
- eapply InQ_gss_enqueue; eauto.
Qed.
Lemma enqueue_low_level_inv:
∀ d d´ i r n,
enqueue0_spec r i d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
unfold enqueue0_spec; intros.
subdestruct; inv H; eauto.
inv H0. constructor; simpl; eauto.
Qed.
Lemma enqueue_kernel_mode:
∀ d d´ n i,
enqueue0_spec n i d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
unfold enqueue0_spec; intros.
subdestruct; inv H; eauto.
Qed.
Global Instance enqueue_inv: PreservesInvariants enqueue0_spec.
Proof.
preserves_invariants_simpl´.
- eapply enqueue_low_level_inv; eassumption.
- eapply enqueue_high_level_inv; eassumption.
- eapply enqueue_kernel_mode; eassumption.
Qed.
Lemma dequeue_high_level_inv:
∀ d d´ n r,
dequeue0_spec n d = Some (d´, r) →
high_level_invariant d →
high_level_invariant d´.
Proof.
unfold dequeue0_spec; intros.
subdestruct; inv H; eauto.
inv H0.
constructor; simpl; eauto; intros.
- eapply AbTCBCorrect_range_gss; eauto.
- eapply AbQCorrect_range_gss_remove; eauto.
- eapply QCount_gss_remove; eauto.
- eapply InQ_gss_remove; eauto.
Qed.
Lemma dequeue_low_level_inv:
∀ d d´ i r n,
dequeue0_spec i d = Some (d´, r) →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
unfold dequeue0_spec; intros.
subdestruct; inv H; eauto.
inv H0. constructor; simpl; eauto.
Qed.
Lemma dequeue_kernel_mode:
∀ d d´ n i,
dequeue0_spec n d = Some (d´, i) →
kernel_mode d →
kernel_mode d´.
Proof.
unfold dequeue0_spec; intros.
subdestruct; inv H; eauto.
Qed.
Global Instance dequeue_inv: PreservesInvariants dequeue0_spec.
Proof.
preserves_invariants_simpl´.
- eapply dequeue_low_level_inv; eassumption.
- eapply dequeue_high_level_inv; eassumption.
- eapply dequeue_kernel_mode; eassumption.
Qed.
End QUEUE.
Section QUEUE_ATOMIC.
Lemma enqueue_atomic_high_level_inv:
∀ d d´ i n,
enqueue_atomic_spec n i d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
eapply release_lock_ABTCB_high_level_inv; eauto.
eapply enqueue_high_level_inv; eauto.
eapply acquire_lock_ABTCB_high_level_inv; eauto.
Qed.
Lemma enqueue_atomic_low_level_inv:
∀ d d´ i n n´,
enqueue_atomic_spec n i d = Some d´ →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
intros. functional inversion H; subst; eauto.
eapply release_lock_ABTCB_low_level_inv; eauto.
eapply enqueue_low_level_inv; eauto.
eapply acquire_lock_ABTCB_low_level_inv; eauto.
Qed.
Lemma enqueue_atomic_kernel_mode:
∀ d d´ n i,
enqueue_atomic_spec n i d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
unfold release_lock_ABTCB_spec in ×.
subdestruct; inv H1; simpl; eauto.
Qed.
Global Instance enqueue_atomic_inv: PreservesInvariants enqueue_atomic_spec.
Proof.
preserves_invariants_simpl´.
- eapply enqueue_atomic_low_level_inv; eassumption.
- eapply enqueue_atomic_high_level_inv; eassumption.
- eapply enqueue_atomic_kernel_mode; eassumption.
Qed.
Lemma dequeue_atomic_high_level_inv:
∀ d d´ i n,
dequeue_atomic_spec n d = Some (d´, i) →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
eapply release_lock_ABTCB_high_level_inv; eauto.
eapply dequeue_high_level_inv; eauto.
eapply acquire_lock_ABTCB_high_level_inv; eauto.
Qed.
Lemma dequeue_atomic_low_level_inv:
∀ d d´ i n n´,
dequeue_atomic_spec n d = Some (d´, i) →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
intros. functional inversion H; subst; eauto.
eapply release_lock_ABTCB_low_level_inv; eauto.
eapply dequeue_low_level_inv; eauto.
eapply acquire_lock_ABTCB_low_level_inv; eauto.
Qed.
Lemma dequeue_atomic_kernel_mode:
∀ d d´ n i,
dequeue_atomic_spec n d = Some (d´, i) →
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
unfold release_lock_ABTCB_spec in ×.
subdestruct; inv H5; simpl; eauto.
Qed.
Global Instance dequeue_atomic_inv: PreservesInvariants dequeue_atomic_spec.
Proof.
preserves_invariants_simpl´.
- eapply dequeue_atomic_low_level_inv; eassumption.
- eapply dequeue_atomic_high_level_inv; eassumption.
- eapply dequeue_atomic_kernel_mode; eassumption.
Qed.
End QUEUE_ATOMIC.
Global Instance set_state_inv: PreservesInvariants set_state1_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto 2.
- eapply AbTCBCorrect_range_gss; eauto.
eapply AbTCBCorrect_range_valid_b; eauto.
- eapply QCount_gso_state; eauto.
- eapply InQ_gso_state; eauto.
Qed.
Global Instance sched_init_inv: PreservesInvariants sched_init_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant.
- apply real_nps_range.
- apply AC_init_container_valid.
- rewrite init_pperm0; [|try assumption].
apply real_pperm_log_valid.
- eapply real_pt_PMap_valid; eauto.
- apply real_pt_PMap_kern.
- omega.
- assumption.
- apply real_idpde_init.
- apply real_pt_weak_consistent_pmap.
- apply real_pt_consistent_pmap_domain.
- apply Lreal_at_consistent_lat_domain.
- eapply LATCTable_log_real; eauto.
- assumption.
- apply real_abtcb_range´; auto.
- apply real_abq_range; auto.
- eapply real_abtcb_abq_QCount´; eauto.
- eapply real_abq_tcb_inQ; eauto.
- assumption.
- eapply real_valid_hlock_pool1; eauto.
- assumption.
- eapply real_valid_AT_log_pool_H; eauto.
- assumption.
- eapply real_valid_ABTCB_log_pool; eauto.
- erewrite real_valid_cidpool with (cidpool := cid d); eauto.
omega.
Qed.
Lemma dequeue_atomic_tcb:
∀ n d d´ la s c b i´
(Hde: dequeue_atomic_spec n d = Some (d´, la))
(Htcb: ZMap.get la (abtcb d´) = AbTCBValid s c b)
(Hrange: 0 ≤ la < num_proc)
(Hi´: 0 ≤ i´ < num_chan),
b = -1.
Proof.
intros. functional inversion Hde; subst.
functional inversion H3; subst. subst l´.
functional inversion H2.
- omega.
- subst cpu. inv H; simpl in ×.
rewrite ZMap.gss in Htcb. inv Htcb. trivial.
Qed.
Lemma dequeue_atomic_tcb´:
∀ n d d´ la i´
(Hde: dequeue_atomic_spec n d = Some (d´, la))
(Hrange: 0 ≤ la < num_proc)
(Hi´: 0 ≤ i´ < num_chan),
∃ s c, ZMap.get la (abtcb d´) = AbTCBValid s c (-1).
Proof.
intros. functional inversion Hde; subst.
functional inversion H3; subst. subst l´.
functional inversion H2; simpl.
- omega.
- subst cpu. rewrite ZMap.gss. eauto.
Qed.
Global Instance thread_wakeup_inv: PreservesInvariants thread_wakeup_spec.
Proof.
preserves_invariants_simpl´.
- intros. functional inversion H2; subst; eauto.
+ exploit dequeue_atomic_low_level_inv; eauto.
intros Hlow. inversion Hlow.
econstructor; eauto.
+ eapply enqueue_atomic_low_level_inv; eauto.
eapply dequeue_atomic_low_level_inv; eauto.
+ eapply dequeue_atomic_low_level_inv; eauto.
- intros.
functional inversion H2; subst; eauto.
+ assert (Hrdyq: 0 ≤ rdyq < num_chan).
{
unfold rdyq in *; subst.
unfold rdy_q_id.
inv H0.
omega.
}
exploit dequeue_atomic_high_level_inv; eauto.
intros Hlow. inversion Hlow.
econstructor; eauto; simpl; intros.
× eapply AbTCBCorrect_range_gss; eauto.
omega.
× eapply AbQCorrect_range_gss_enqueue; eauto.
× eapply QCount_gss_wakeup´; eauto.
{ eapply dequeue_atomic_tcb; eauto. }
× eapply InQ_gss_wakeup´; eauto.
{ eapply dequeue_atomic_tcb; eauto. }
+ eapply enqueue_atomic_high_level_inv; eauto.
eapply dequeue_atomic_high_level_inv; eauto.
+ eapply dequeue_atomic_high_level_inv; eauto.
- intros. functional inversion H2; subst; eauto; simpl.
+ eapply dequeue_atomic_kernel_mode; eauto.
+ eapply enqueue_atomic_kernel_mode; eauto.
eapply dequeue_atomic_kernel_mode; eauto.
+ eapply dequeue_atomic_kernel_mode; eauto.
Qed.
Global Instance thread_sched_inv: ThreadScheduleInvariants thread_sched_spec.
Proof.
constructor; intros; functional inversion H.
- inv H1. constructor; trivial.
eapply kctxt_inject_neutral_gss_mem; eauto.
- inv H0. subst.
assert (HOS: 0≤ rdyq < num_chan).
{
unfold rdyq, cpu in ×.
unfold rdy_q_id.
omega.
}
constructor; auto; simpl in *; intros; try congruence.
+ eapply AbTCBCorrect_range_gss; eauto.
+ eapply AbQCorrect_range_gss_remove; eauto.
+ eapply QCount_gss_remove; eauto.
+ eapply InQ_gss_remove; eauto.
+ rewrite ZMap.gss. exploit valid_TDQ0; eauto. rewrite H7.
unfold AbQCorrect. intros (l0 & Heq & Hrange). inv Heq.
eapply Hrange. left; trivial.
Qed.
Global Instance thread_spawn_inv: DNewInvariants thread_spawn_spec.
Proof.
constructor; intros; inv H0;
unfold thread_spawn_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.
-
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.
+ destruct (zeq 0 (id × max_children + 1 + Z.of_nat (length (cchildren (ZMap.get id (AC d)))))); subst.
rewrite e.
rewrite ZMap.gss; simpl; split; auto.
rewrite ZMap.gso; auto.
destruct (zeq 0 id); subst.
rewrite ZMap.gss; simpl; auto.
rewrite ZMap.gso; auto.
+ eapply AbTCBCorrect_range_gss; eauto.
unfold rdy_q_id.
omega.
+ eapply AbQCorrect_range_gss_enqueue; eauto. omega.
+ eapply QCount_gss_enqueue; eauto.
omega.
+ eapply InQ_gss_enqueue; eauto.
omega.
Qed.
Global Instance thread_poll_pending_inv: PreservesInvariants thread_poll_pending_spec.
Proof.
preserves_invariants_simpl´.
- intros. functional inversion H2; subst; eauto.
+ exploit dequeue_atomic_low_level_inv; eauto.
intros Hlow. inv Hlow. econstructor; eauto.
+ exploit dequeue_atomic_low_level_inv; eauto.
- intros. functional inversion H2; subst; eauto.
+ assert (Hrdyq: 0 ≤ rdyq < num_chan).
{
unfold rdyq in *; subst.
unfold rdy_q_id. subst cpu.
inv H0.
omega.
}
exploit dequeue_atomic_tcb´; eauto.
intros (s & c & Htcb).
exploit dequeue_atomic_high_level_inv; eauto.
intro Hhigh. inv Hhigh.
econstructor; eauto; simpl; intros.
× eapply AbTCBCorrect_range_gss; eauto.
omega.
× eapply AbQCorrect_range_gss_enqueue; eauto.
× eapply QCount_gss_wakeup´; eauto.
× eapply InQ_gss_wakeup´; eauto.
+ exploit dequeue_atomic_high_level_inv; eauto.
- intros. functional inversion H2; subst; eauto; simpl.
+ eapply dequeue_atomic_kernel_mode; eauto.
+ eapply dequeue_atomic_kernel_mode; eauto.
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_loadex2 (F := F) (V := V).
Definition exec_storeex {F V} := exec_storeex2 (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_unmaped´; try reflexivity; try eassumption); trivial.
Qed.
Global Instance trapinfo_set_inv: TrapinfoSetInvariant.
Proof.
split; inversion 1; intros; constructor; auto.
Qed.
Section RECEIVE.
Lemma page_copy_back_high_level_inv:
∀ d d´ chid cound addr,
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,
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,
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,
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´,
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,
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 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,
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.
- eapply consistent_ppage_log_gso; eauto.
eapply Shared2ID2_neq; eauto.
reflexivity.
- eapply LATCTable_log_gso; eauto.
eapply Shared2ID2_neq; eauto.
reflexivity.
- eapply valid_hlock_pool1_gss´; eauto.
- eapply valid_AT_log_pool_H_gso; eauto.
eapply Shared2ID2_neq; eauto.
reflexivity.
- eapply valid_ABTCB_log_pool_gso´; eauto.
reflexivity.
Qed.
Lemma page_copy_low_level_inv:
∀ d d´ chid cound addr n,
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,
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,
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´,
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,
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 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.
Global Instance set_curid_inv: PreservesInvariants set_curid_spec.
Proof.
preserves_invariants_simpl_auto.
rewrite ZMap.gss. trivial.
Qed.
Global Instance sleeper_inc_inv:
PreservesInvariants sleeper_inc_spec.
Proof.
preserves_invariants_simpl_auto.
Qed.
Global Instance sleeper_dec_inv:
PreservesInvariants sleeper_dec_spec.
Proof.
preserves_invariants_simpl_auto.
Qed.
Section SC_lock.
Lemma acquire_lock_SC_high_level_inv:
∀ d d´ i,
acquire_lock_SC_spec i d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
unfold acquire_lock_SC_spec; intros.
subdestruct; inv H; eauto.
- inv H0. constructor; simpl; eauto; intros.
+ eapply consistent_ppage_log_gso; eauto.
eapply SC_neq´; omega.
+ eapply LATCTable_log_gso; eauto.
eapply SC_neq´; omega.
+ eapply valid_hlock_pool1_gss´; eauto.
+ eapply valid_AT_log_pool_H_gso; eauto.
eapply SC_neq´; omega.
+ eapply valid_ABTCB_log_pool_gso; eauto.
eapply SC_neq0; eauto.
- inv H0. constructor; simpl; eauto; intros.
+ eapply consistent_ppage_log_gso; eauto.
eapply SC_neq´; omega.
+ eapply LATCTable_log_gso; eauto.
eapply SC_neq´; omega.
+ eapply valid_hlock_pool1_gss´; eauto.
+ eapply valid_AT_log_pool_H_gso; eauto.
eapply SC_neq´; omega.
+ eapply valid_ABTCB_log_pool_gso; eauto.
eapply SC_neq0; eauto.
Qed.
Lemma acquire_lock_SC_low_level_inv:
∀ d d´ i n,
acquire_lock_SC_spec i d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
unfold 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,
acquire_lock_SC_spec i d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
unfold acquire_lock_SC_spec; intros.
subdestruct; inv H; eauto.
Qed.
Global Instance acquire_lock_SC_inv: PreservesInvariants 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,
release_lock_SC_spec i d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. unfold release_lock_SC_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto; intros.
+ eapply consistent_ppage_log_gso; eauto.
eapply SC_neq´. omega.
+ eapply LATCTable_log_gso; eauto.
eapply SC_neq´; omega.
+ eapply valid_hlock_pool1_gss´; eauto.
+ eapply valid_AT_log_pool_H_gso; eauto.
eapply SC_neq´; omega.
+ eapply valid_ABTCB_log_pool_gso; eauto.
eapply SC_neq0; eauto.
Qed.
Lemma release_lock_SC_low_level_inv:
∀ d d´ i n,
release_lock_SC_spec i d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. unfold release_lock_SC_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto.
Qed.
Lemma release_lock_SC_kernel_mode:
∀ d d´ i,
release_lock_SC_spec i d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
intros. unfold release_lock_SC_spec in *; subdestruct; inv H.
- inv H0; eauto.
Qed.
Global Instance release_lock_SC_inv: PreservesInvariants 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 TCB_lock.
Lemma acquire_lock_ABTCB_high_level_inv:
∀ d d´ i,
acquire_lock_ABTCB_spec i d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
unfold acquire_lock_ABTCB_spec; intros.
subdestruct; inv H; eauto.
- inv H0. constructor; simpl; eauto; intros.
+ eapply consistent_ppage_log_gso; eauto.
eapply TCB_neq´; omega.
+ eapply LATCTable_log_gso; eauto.
eapply TCB_neq´; omega.
+ eapply valid_ABTCB_oracle_pool_inv0; eauto.
+ exploit valid_ABTCB_oracle_pool_inv0; eauto.
intros (Hd & _). eapply Hd; eauto.
+ exploit valid_ABTCB_oracle_pool_inv0; eauto.
intros (Hd & _). eapply Hd; eauto.
+ exploit valid_ABTCB_oracle_pool_inv0; eauto.
intros (Hd & _). eapply Hd; eauto.
+ eapply valid_hlock_pool1_gss´; eauto.
+ eapply valid_AT_log_pool_H_gso; eauto.
eapply TCB_neq´; omega.
+ eapply valid_ABTCB_log_pool_gss; eauto.
eapply valid_ABTCB_log_pull; eauto.
eapply valid_ABTCB_log_wait; eauto.
eapply valid_ABTCB_oracle_pool_inv0; eauto.
- inv H0. constructor; simpl; eauto; intros.
+ eapply consistent_ppage_log_gso; eauto.
eapply TCB_neq´; omega.
+ eapply LATCTable_log_gso; eauto.
eapply TCB_neq´; omega.
+ eapply valid_hlock_pool1_gss´; eauto.
+ eapply valid_AT_log_pool_H_gso; eauto.
eapply TCB_neq´; omega.
+ eapply valid_ABTCB_log_pool_gss; eauto.
eapply valid_ABTCB_log_pull; eauto.
eapply valid_ABTCB_log_wait; eauto.
eapply valid_ABTCB_oracle_pool_inv0; eauto.
Qed.
Lemma acquire_lock_ABTCB_low_level_inv:
∀ d d´ i n,
acquire_lock_ABTCB_spec i d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
unfold acquire_lock_ABTCB_spec; intros.
subdestruct; inv H; eauto.
- inv H0. constructor; simpl; eauto.
- inv H0. constructor; simpl; eauto.
Qed.
Lemma acquire_lock_ABTCB_kernel_mode:
∀ d d´ i,
acquire_lock_ABTCB_spec i d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
unfold acquire_lock_ABTCB_spec; intros.
subdestruct; inv H; eauto.
Qed.
Lemma release_lock_ABTCB_high_level_inv:
∀ d d´ i,
release_lock_ABTCB_spec i d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. unfold release_lock_ABTCB_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto; intros.
+ eapply consistent_ppage_log_gso; eauto.
eapply TCB_neq´; omega.
+ eapply LATCTable_log_gso; eauto.
eapply TCB_neq´; omega.
+ eapply valid_hlock_pool1_gss´; eauto.
+ eapply valid_AT_log_pool_H_gso; eauto.
eapply TCB_neq´; omega.
+ eapply valid_ABTCB_log_pool_gss; eauto.
eapply valid_ABTCB_log_rel.
eapply valid_ABTCB_log_shared; eauto.
Qed.
Lemma release_lock_ABTCB_low_level_inv:
∀ d d´ i n,
release_lock_ABTCB_spec i d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
intros. unfold release_lock_ABTCB_spec in *; subdestruct; inv H.
- inv H0. constructor; simpl; eauto.
Qed.
Lemma release_lock_ABTCB_kernel_mode:
∀ d d´ i,
release_lock_ABTCB_spec i d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
intros. unfold release_lock_ABTCB_spec in *; subdestruct; inv H.
- inv H0; eauto.
Qed.
End TCB_lock.
Section PALLOC.
Lemma palloc_high_level_inv:
∀ d d´ i n,
palloc_spec i d = Some (d´, n) →
high_level_invariant d →
high_level_invariant d´.
Proof.
unfold palloc_spec; intros.
subdestruct; inv H; subst; eauto;
inv H0; constructor; simpl; eauto; intros.
+ rewrite <- Hdestruct3.
eapply alloc_container_valid´; eauto.
+ eapply consistent_ppage_log_norm_alloc; eauto. omega.
+ subst; simpl;
intros; congruence.
+ eapply dirty_ppage_gso_alloc; eauto.
+ eapply (weak_consistent_pmap_gso_at_palloc n); eauto; try apply a0.
+ eapply consistent_pmap_domain_gso_at; eauto.
intros. intro HF.
exploit (LATCTable_log_gss (ZMap.get 0 (multi_oracle d) (CPU_ID d) l) _ _ _
valid_LATable_nil0 Hdestruct6); eauto.
× rewrite ZMap.gss. trivial.
× eapply a0.
+ eapply consistent_lat_domain_gss_nil; eauto.
+ eapply LATCTable_log_alloc´; eauto.
+ intros. destruct (zeq i 0); subst.
× rewrite ZMap.gss; trivial.
× rewrite ZMap.gso; auto.
+ eapply valid_hlock_pool1_gso; eauto.
+ eapply valid_AT_log_pool_H_n; eauto.
eapply a0.
+ eapply valid_ABTCB_log_pool_gso_AT; eauto.
+ rewrite app_comm_cons.
eapply consistent_ppage_log_gss; eauto.
+ eapply LATCTable_log_alloc; eauto.
+ eapply valid_hlock_pool1_gso; eauto.
+ eapply valid_AT_log_pool_H_0; eauto.
+ eapply valid_ABTCB_log_pool_gso_AT; eauto.
+ rewrite app_comm_cons.
eapply consistent_ppage_log_gss; eauto.
+ eapply LATCTable_log_alloc; eauto.
+ eapply valid_hlock_pool1_gso; eauto.
+ eapply valid_AT_log_pool_H_0´; eauto.
+ eapply valid_ABTCB_log_pool_gso_AT; eauto.
Qed.
Lemma palloc_low_level_inv:
∀ d d´ i n n´,
palloc_spec i d = Some (d´, n) →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
unfold palloc_spec; intros.
subdestruct; inv H; subst; eauto;
inv H0; constructor; eauto.
Qed.
Lemma palloc_kernel_mode:
∀ d d´ i n,
palloc_spec i d = Some (d´, n) →
kernel_mode d →
kernel_mode d´.
Proof.
unfold palloc_spec; intros.
subdestruct; inv H; simpl; eauto.
Qed.
Global Instance palloc_inv: PreservesInvariants 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 trapin_inv: PrimInvariants trapin_spec.
Proof.
PrimInvariants_simpl_auto.
Qed.
Global Instance trapout_inv: PrimInvariants trapout_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 ptin_inv: PrimInvariants ptin_spec.
Proof.
PrimInvariants_simpl_auto.
Qed.
Global Instance ptout_inv: PrimInvariants ptout_spec.
Proof.
PrimInvariants_simpl_auto.
Qed.
Global Instance fstore_inv: PreservesInvariants fstore_spec.
Proof.
split; intros; inv_generic_sem H; inv H0; functional inversion H2.
- functional inversion H. split; trivial.
- functional inversion H.
split; subst; simpl;
try (eapply dirty_ppage_store_unmaped; try reflexivity; try eassumption); trivial.
- functional inversion H0.
split; simpl; try assumption.
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.
- functional inversion H2. functional inversion H0.
eapply weak_consistent_pmap_ptp_same; try eassumption.
eapply weak_consistent_pmap_gso_pperm_alloc´; eassumption.
- functional inversion H2.
eapply consistent_pmap_domain_append; eauto.
destruct (ZMap.get pti pdt); try contradiction;
red; intros (v0 & p0 & He); contra_inv.
- eapply consistent_lat_domain_gss_append; eauto.
subst pti; destruct (ZMap.get (PTX vadr) pdt); try contradiction;
red; intros (v0 & p0 & He); contra_inv.
- eapply LATCTable_log_not_nil_gso_true; eauto.
functional inversion H2. omega.
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 ptAllocPDE_high_level_inv:
∀ d d´ n vadr v,
ptAllocPDE0_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 palloc_inv_prop; eauto. intros (HPT & Halloc & Hpg).
clear H11.
rewrite <- HPT in ×.
inv H1; constructor_gso_simpl_tac; try (intros; congruence); intros.
+ apply consistent_ppage_log_alloc_hide; eauto.
eapply Halloc; eauto.
+ 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.
+ eapply weak_consistent_pmap_ptp_gss0; eauto; apply Halloc; eauto.
+ eapply consistent_pmap_domain_gso_at_00; eauto; try apply Halloc; eauto.
eapply consistent_pmap_domain_ptp_unp; eauto.
apply real_init_PTE_unp.
+ apply consistent_lat_domain_gso_p; eauto.
Qed.
Lemma ptAllocPDE_low_level_inv:
∀ d d´ n vadr v n´,
ptAllocPDE0_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,
ptAllocPDE0_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,
ptInsert0_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,
ptInsert0_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,
ptInsert0_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,
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,
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,
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 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.
Section PTRESV2.
Lemma ptResv2_high_level_inv:
∀ d d´ n vadr p n´ vadr´ p´ v,
ptResv2_spec n vadr p 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.
- eapply ptInsert_high_level_inv; try eassumption.
eapply ptInsert_high_level_inv; try eassumption.
eapply palloc_high_level_inv; eassumption.
Qed.
Lemma ptResv2_low_level_inv:
∀ d d´ n vadr p n´ vadr´ p´ l v,
ptResv2_spec n vadr p n´ vadr´ p´ d = Some (d´, v) →
low_level_invariant l d →
low_level_invariant l 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.
- eapply ptInsert_low_level_inv; try eassumption.
eapply ptInsert_low_level_inv; try eassumption.
eapply palloc_low_level_inv; eassumption.
Qed.
Lemma ptResv2_kernel_mode:
∀ d d´ n vadr p n´ vadr´ p´ v,
ptResv2_spec n vadr p 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.
- eapply ptInsert_kernel_mode; try eassumption.
eapply ptInsert_kernel_mode; try eassumption.
eapply palloc_kernel_mode; eassumption.
Qed.
End PTRESV2.
Section OFFER_SHARE.
Global Instance offer_shared_mem_inv:
PreservesInvariants offer_shared_mem_spec.
Proof.
preserves_invariants_simpl´;
functional inversion H2; subst; eauto 2; try (inv H0; constructor; trivial; fail).
- exploit ptResv2_low_level_inv; eauto.
intros HP; inv HP. constructor; trivial.
- exploit ptResv2_low_level_inv; eauto.
intros HP; inv HP. constructor; trivial.
- exploit ptResv2_high_level_inv; eauto.
intros HP; inv HP. constructor; trivial.
- exploit ptResv2_high_level_inv; eauto.
intros HP; inv HP. constructor; trivial.
- exploit ptResv2_kernel_mode; eauto.
- exploit ptResv2_kernel_mode; eauto.
Qed.
End OFFER_SHARE.
Global Instance shared_mem_status_inv:
PreservesInvariants shared_mem_status_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Section QUEUE.
Lemma enqueue_high_level_inv:
∀ d d´ n i,
enqueue0_spec n i d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
unfold enqueue0_spec; intros.
subdestruct; inv H; eauto.
inv H0.
constructor; simpl; eauto; intros;
functional inversion Hdestruct3.
- eapply AbTCBCorrect_range_gss; eauto. omega.
- eapply AbQCorrect_range_gss_enqueue; eauto.
- eapply QCount_gss_enqueue; eauto.
- eapply InQ_gss_enqueue; eauto.
Qed.
Lemma enqueue_low_level_inv:
∀ d d´ i r n,
enqueue0_spec r i d = Some d´ →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
unfold enqueue0_spec; intros.
subdestruct; inv H; eauto.
inv H0. constructor; simpl; eauto.
Qed.
Lemma enqueue_kernel_mode:
∀ d d´ n i,
enqueue0_spec n i d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
unfold enqueue0_spec; intros.
subdestruct; inv H; eauto.
Qed.
Global Instance enqueue_inv: PreservesInvariants enqueue0_spec.
Proof.
preserves_invariants_simpl´.
- eapply enqueue_low_level_inv; eassumption.
- eapply enqueue_high_level_inv; eassumption.
- eapply enqueue_kernel_mode; eassumption.
Qed.
Lemma dequeue_high_level_inv:
∀ d d´ n r,
dequeue0_spec n d = Some (d´, r) →
high_level_invariant d →
high_level_invariant d´.
Proof.
unfold dequeue0_spec; intros.
subdestruct; inv H; eauto.
inv H0.
constructor; simpl; eauto; intros.
- eapply AbTCBCorrect_range_gss; eauto.
- eapply AbQCorrect_range_gss_remove; eauto.
- eapply QCount_gss_remove; eauto.
- eapply InQ_gss_remove; eauto.
Qed.
Lemma dequeue_low_level_inv:
∀ d d´ i r n,
dequeue0_spec i d = Some (d´, r) →
low_level_invariant n d →
low_level_invariant n d´.
Proof.
unfold dequeue0_spec; intros.
subdestruct; inv H; eauto.
inv H0. constructor; simpl; eauto.
Qed.
Lemma dequeue_kernel_mode:
∀ d d´ n i,
dequeue0_spec n d = Some (d´, i) →
kernel_mode d →
kernel_mode d´.
Proof.
unfold dequeue0_spec; intros.
subdestruct; inv H; eauto.
Qed.
Global Instance dequeue_inv: PreservesInvariants dequeue0_spec.
Proof.
preserves_invariants_simpl´.
- eapply dequeue_low_level_inv; eassumption.
- eapply dequeue_high_level_inv; eassumption.
- eapply dequeue_kernel_mode; eassumption.
Qed.
End QUEUE.
Section QUEUE_ATOMIC.
Lemma enqueue_atomic_high_level_inv:
∀ d d´ i n,
enqueue_atomic_spec n i d = Some d´ →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
eapply release_lock_ABTCB_high_level_inv; eauto.
eapply enqueue_high_level_inv; eauto.
eapply acquire_lock_ABTCB_high_level_inv; eauto.
Qed.
Lemma enqueue_atomic_low_level_inv:
∀ d d´ i n n´,
enqueue_atomic_spec n i d = Some d´ →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
intros. functional inversion H; subst; eauto.
eapply release_lock_ABTCB_low_level_inv; eauto.
eapply enqueue_low_level_inv; eauto.
eapply acquire_lock_ABTCB_low_level_inv; eauto.
Qed.
Lemma enqueue_atomic_kernel_mode:
∀ d d´ n i,
enqueue_atomic_spec n i d = Some d´ →
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
unfold release_lock_ABTCB_spec in ×.
subdestruct; inv H1; simpl; eauto.
Qed.
Global Instance enqueue_atomic_inv: PreservesInvariants enqueue_atomic_spec.
Proof.
preserves_invariants_simpl´.
- eapply enqueue_atomic_low_level_inv; eassumption.
- eapply enqueue_atomic_high_level_inv; eassumption.
- eapply enqueue_atomic_kernel_mode; eassumption.
Qed.
Lemma dequeue_atomic_high_level_inv:
∀ d d´ i n,
dequeue_atomic_spec n d = Some (d´, i) →
high_level_invariant d →
high_level_invariant d´.
Proof.
intros. functional inversion H; subst; eauto.
eapply release_lock_ABTCB_high_level_inv; eauto.
eapply dequeue_high_level_inv; eauto.
eapply acquire_lock_ABTCB_high_level_inv; eauto.
Qed.
Lemma dequeue_atomic_low_level_inv:
∀ d d´ i n n´,
dequeue_atomic_spec n d = Some (d´, i) →
low_level_invariant n´ d →
low_level_invariant n´ d´.
Proof.
intros. functional inversion H; subst; eauto.
eapply release_lock_ABTCB_low_level_inv; eauto.
eapply dequeue_low_level_inv; eauto.
eapply acquire_lock_ABTCB_low_level_inv; eauto.
Qed.
Lemma dequeue_atomic_kernel_mode:
∀ d d´ n i,
dequeue_atomic_spec n d = Some (d´, i) →
kernel_mode d →
kernel_mode d´.
Proof.
intros. functional inversion H; subst; eauto.
unfold release_lock_ABTCB_spec in ×.
subdestruct; inv H5; simpl; eauto.
Qed.
Global Instance dequeue_atomic_inv: PreservesInvariants dequeue_atomic_spec.
Proof.
preserves_invariants_simpl´.
- eapply dequeue_atomic_low_level_inv; eassumption.
- eapply dequeue_atomic_high_level_inv; eassumption.
- eapply dequeue_atomic_kernel_mode; eassumption.
Qed.
End QUEUE_ATOMIC.
Global Instance set_state_inv: PreservesInvariants set_state1_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto 2.
- eapply AbTCBCorrect_range_gss; eauto.
eapply AbTCBCorrect_range_valid_b; eauto.
- eapply QCount_gso_state; eauto.
- eapply InQ_gso_state; eauto.
Qed.
Global Instance sched_init_inv: PreservesInvariants sched_init_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant.
- apply real_nps_range.
- apply AC_init_container_valid.
- rewrite init_pperm0; [|try assumption].
apply real_pperm_log_valid.
- eapply real_pt_PMap_valid; eauto.
- apply real_pt_PMap_kern.
- omega.
- assumption.
- apply real_idpde_init.
- apply real_pt_weak_consistent_pmap.
- apply real_pt_consistent_pmap_domain.
- apply Lreal_at_consistent_lat_domain.
- eapply LATCTable_log_real; eauto.
- assumption.
- apply real_abtcb_range´; auto.
- apply real_abq_range; auto.
- eapply real_abtcb_abq_QCount´; eauto.
- eapply real_abq_tcb_inQ; eauto.
- assumption.
- eapply real_valid_hlock_pool1; eauto.
- assumption.
- eapply real_valid_AT_log_pool_H; eauto.
- assumption.
- eapply real_valid_ABTCB_log_pool; eauto.
- erewrite real_valid_cidpool with (cidpool := cid d); eauto.
omega.
Qed.
Lemma dequeue_atomic_tcb:
∀ n d d´ la s c b i´
(Hde: dequeue_atomic_spec n d = Some (d´, la))
(Htcb: ZMap.get la (abtcb d´) = AbTCBValid s c b)
(Hrange: 0 ≤ la < num_proc)
(Hi´: 0 ≤ i´ < num_chan),
b = -1.
Proof.
intros. functional inversion Hde; subst.
functional inversion H3; subst. subst l´.
functional inversion H2.
- omega.
- subst cpu. inv H; simpl in ×.
rewrite ZMap.gss in Htcb. inv Htcb. trivial.
Qed.
Lemma dequeue_atomic_tcb´:
∀ n d d´ la i´
(Hde: dequeue_atomic_spec n d = Some (d´, la))
(Hrange: 0 ≤ la < num_proc)
(Hi´: 0 ≤ i´ < num_chan),
∃ s c, ZMap.get la (abtcb d´) = AbTCBValid s c (-1).
Proof.
intros. functional inversion Hde; subst.
functional inversion H3; subst. subst l´.
functional inversion H2; simpl.
- omega.
- subst cpu. rewrite ZMap.gss. eauto.
Qed.
Global Instance thread_wakeup_inv: PreservesInvariants thread_wakeup_spec.
Proof.
preserves_invariants_simpl´.
- intros. functional inversion H2; subst; eauto.
+ exploit dequeue_atomic_low_level_inv; eauto.
intros Hlow. inversion Hlow.
econstructor; eauto.
+ eapply enqueue_atomic_low_level_inv; eauto.
eapply dequeue_atomic_low_level_inv; eauto.
+ eapply dequeue_atomic_low_level_inv; eauto.
- intros.
functional inversion H2; subst; eauto.
+ assert (Hrdyq: 0 ≤ rdyq < num_chan).
{
unfold rdyq in *; subst.
unfold rdy_q_id.
inv H0.
omega.
}
exploit dequeue_atomic_high_level_inv; eauto.
intros Hlow. inversion Hlow.
econstructor; eauto; simpl; intros.
× eapply AbTCBCorrect_range_gss; eauto.
omega.
× eapply AbQCorrect_range_gss_enqueue; eauto.
× eapply QCount_gss_wakeup´; eauto.
{ eapply dequeue_atomic_tcb; eauto. }
× eapply InQ_gss_wakeup´; eauto.
{ eapply dequeue_atomic_tcb; eauto. }
+ eapply enqueue_atomic_high_level_inv; eauto.
eapply dequeue_atomic_high_level_inv; eauto.
+ eapply dequeue_atomic_high_level_inv; eauto.
- intros. functional inversion H2; subst; eauto; simpl.
+ eapply dequeue_atomic_kernel_mode; eauto.
+ eapply enqueue_atomic_kernel_mode; eauto.
eapply dequeue_atomic_kernel_mode; eauto.
+ eapply dequeue_atomic_kernel_mode; eauto.
Qed.
Global Instance thread_sched_inv: ThreadScheduleInvariants thread_sched_spec.
Proof.
constructor; intros; functional inversion H.
- inv H1. constructor; trivial.
eapply kctxt_inject_neutral_gss_mem; eauto.
- inv H0. subst.
assert (HOS: 0≤ rdyq < num_chan).
{
unfold rdyq, cpu in ×.
unfold rdy_q_id.
omega.
}
constructor; auto; simpl in *; intros; try congruence.
+ eapply AbTCBCorrect_range_gss; eauto.
+ eapply AbQCorrect_range_gss_remove; eauto.
+ eapply QCount_gss_remove; eauto.
+ eapply InQ_gss_remove; eauto.
+ rewrite ZMap.gss. exploit valid_TDQ0; eauto. rewrite H7.
unfold AbQCorrect. intros (l0 & Heq & Hrange). inv Heq.
eapply Hrange. left; trivial.
Qed.
Global Instance thread_spawn_inv: DNewInvariants thread_spawn_spec.
Proof.
constructor; intros; inv H0;
unfold thread_spawn_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.
-
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.
+ destruct (zeq 0 (id × max_children + 1 + Z.of_nat (length (cchildren (ZMap.get id (AC d)))))); subst.
rewrite e.
rewrite ZMap.gss; simpl; split; auto.
rewrite ZMap.gso; auto.
destruct (zeq 0 id); subst.
rewrite ZMap.gss; simpl; auto.
rewrite ZMap.gso; auto.
+ eapply AbTCBCorrect_range_gss; eauto.
unfold rdy_q_id.
omega.
+ eapply AbQCorrect_range_gss_enqueue; eauto. omega.
+ eapply QCount_gss_enqueue; eauto.
omega.
+ eapply InQ_gss_enqueue; eauto.
omega.
Qed.
Global Instance thread_poll_pending_inv: PreservesInvariants thread_poll_pending_spec.
Proof.
preserves_invariants_simpl´.
- intros. functional inversion H2; subst; eauto.
+ exploit dequeue_atomic_low_level_inv; eauto.
intros Hlow. inv Hlow. econstructor; eauto.
+ exploit dequeue_atomic_low_level_inv; eauto.
- intros. functional inversion H2; subst; eauto.
+ assert (Hrdyq: 0 ≤ rdyq < num_chan).
{
unfold rdyq in *; subst.
unfold rdy_q_id. subst cpu.
inv H0.
omega.
}
exploit dequeue_atomic_tcb´; eauto.
intros (s & c & Htcb).
exploit dequeue_atomic_high_level_inv; eauto.
intro Hhigh. inv Hhigh.
econstructor; eauto; simpl; intros.
× eapply AbTCBCorrect_range_gss; eauto.
omega.
× eapply AbQCorrect_range_gss_enqueue; eauto.
× eapply QCount_gss_wakeup´; eauto.
× eapply InQ_gss_wakeup´; eauto.
+ exploit dequeue_atomic_high_level_inv; eauto.
- intros. functional inversion H2; subst; eauto; simpl.
+ eapply dequeue_atomic_kernel_mode; eauto.
+ eapply dequeue_atomic_kernel_mode; eauto.
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_loadex2 (F := F) (V := V).
Definition exec_storeex {F V} := exec_storeex2 (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_unmaped´; try reflexivity; try eassumption); trivial.
Qed.
Global Instance trapinfo_set_inv: TrapinfoSetInvariant.
Proof.
split; inversion 1; intros; constructor; auto.
Qed.
Definition pthreadsched_fresh_c : compatlayer (cdata RData) :=
thread_spawn ↦ dnew_compatsem thread_spawn_spec
⊕ thread_wakeup ↦ gensem thread_wakeup_spec
⊕ sched_init ↦ gensem sched_init_spec
⊕ thread_poll_pending ↦ gensem thread_poll_pending_spec.
Definition pthreadsched_fresh_asm : compatlayer (cdata RData) :=
thread_sched ↦ primcall_thread_schedule_compatsem thread_sched_spec (prim_ident:= thread_sched).
Definition pthreadsched_passthrough : compatlayer (cdata RData) :=
fload ↦ gensem fload_spec
⊕ fstore ↦ gensem fstore_spec
⊕ vmxinfo_get ↦ gensem vmxinfo_get_spec
⊕ palloc ↦ gensem palloc_spec
⊕ set_pt ↦ gensem setPT_spec
⊕ pt_read ↦ gensem ptRead_spec
⊕ pt_resv ↦ gensem ptResv_spec
⊕ shared_mem_status ↦ gensem shared_mem_status_spec
⊕ offer_shared_mem ↦ gensem offer_shared_mem_spec
⊕ get_state ↦ gensem get_state0_spec
⊕ set_state ↦ gensem set_state1_spec
⊕ tcb_get_CPU_ID ↦ gensem get_abtcb_CPU_ID_spec
⊕ enqueue ↦ gensem enqueue0_spec
⊕ enqueue_atomic ↦ gensem enqueue_atomic_spec
⊕ pt_in ↦ primcall_general_compatsem´ ptin_spec (prim_ident:= pt_in)
⊕ pt_out ↦ primcall_general_compatsem´ ptout_spec (prim_ident:= pt_out)
⊕ 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
⊕ set_curid ↦ gensem set_curid_spec
⊕ sleeper_inc ↦ gensem sleeper_inc_spec
⊕ sleeper_dec ↦ gensem sleeper_dec_spec
⊕ sleeper_zzz ↦ gensem sleeper_zzz_spec
⊕ acquire_lock_CHAN ↦ gensem acquire_lock_SC_spec
⊕ release_lock_CHAN ↦ gensem 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 ipc_send_body_spec
⊕ ipc_receive_body ↦ gensem ipc_receive_body_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
⊕ trap_in ↦ primcall_general_compatsem trapin_spec
⊕ trap_out ↦ primcall_general_compatsem trapout_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 pthreadsched : compatlayer (cdata RData) :=
(pthreadsched_fresh_c ⊕ pthreadsched_fresh_asm) ⊕ pthreadsched_passthrough.
End WITHMEM.
thread_spawn ↦ dnew_compatsem thread_spawn_spec
⊕ thread_wakeup ↦ gensem thread_wakeup_spec
⊕ sched_init ↦ gensem sched_init_spec
⊕ thread_poll_pending ↦ gensem thread_poll_pending_spec.
Definition pthreadsched_fresh_asm : compatlayer (cdata RData) :=
thread_sched ↦ primcall_thread_schedule_compatsem thread_sched_spec (prim_ident:= thread_sched).
Definition pthreadsched_passthrough : compatlayer (cdata RData) :=
fload ↦ gensem fload_spec
⊕ fstore ↦ gensem fstore_spec
⊕ vmxinfo_get ↦ gensem vmxinfo_get_spec
⊕ palloc ↦ gensem palloc_spec
⊕ set_pt ↦ gensem setPT_spec
⊕ pt_read ↦ gensem ptRead_spec
⊕ pt_resv ↦ gensem ptResv_spec
⊕ shared_mem_status ↦ gensem shared_mem_status_spec
⊕ offer_shared_mem ↦ gensem offer_shared_mem_spec
⊕ get_state ↦ gensem get_state0_spec
⊕ set_state ↦ gensem set_state1_spec
⊕ tcb_get_CPU_ID ↦ gensem get_abtcb_CPU_ID_spec
⊕ enqueue ↦ gensem enqueue0_spec
⊕ enqueue_atomic ↦ gensem enqueue_atomic_spec
⊕ pt_in ↦ primcall_general_compatsem´ ptin_spec (prim_ident:= pt_in)
⊕ pt_out ↦ primcall_general_compatsem´ ptout_spec (prim_ident:= pt_out)
⊕ 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
⊕ set_curid ↦ gensem set_curid_spec
⊕ sleeper_inc ↦ gensem sleeper_inc_spec
⊕ sleeper_dec ↦ gensem sleeper_dec_spec
⊕ sleeper_zzz ↦ gensem sleeper_zzz_spec
⊕ acquire_lock_CHAN ↦ gensem acquire_lock_SC_spec
⊕ release_lock_CHAN ↦ gensem 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 ipc_send_body_spec
⊕ ipc_receive_body ↦ gensem ipc_receive_body_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
⊕ trap_in ↦ primcall_general_compatsem trapin_spec
⊕ trap_out ↦ primcall_general_compatsem trapout_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 pthreadsched : compatlayer (cdata RData) :=
(pthreadsched_fresh_c ⊕ pthreadsched_fresh_asm) ⊕ pthreadsched_passthrough.
End WITHMEM.