Library mcertikos.proc.PQueueIntro


This file defines the abstract data and the primitives for the PQueueIntro layer, which will introduce abstraction of kernel context
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 INVLemmaMemory.
Require Import INVLemmaThread.
Require Import INVLemmaContainer.

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

Abstract Data and Primitives at this layer


Section WITHMEM.

  Local Open Scope Z_scope.

  Context `{oracle_prop: MultiOracleProp}.
  Context `{real_params: RealParams}.

**Definition of the invariants at this layer 0th page map is reserved for the kernel thread
  Record high_level_invariant (abd: RData) :=
    mkInvariant {
        valid_nps: pg abd = truekern_low nps abd maxpage;
        
        valid_kern: ipt abd = falsepg abd = true;
        valid_iptt: ipt abd = trueikern abd = true;
        valid_iptf: ikern abd = falseipt abd = false;
        valid_ihost: ihost abd = falsepg 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 = trueipt abd = true(PT abd) = 0;
        valid_PMap_kern: pg abd = truePMap_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 = trueIDPDE_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 = truecused (ZMap.get 0 (AC abd)) = true;

        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_cons_buf_rpos: 0 rpos (console abd) < CONSOLE_BUFFER_SIZE;
        valid_cons_buf_length: 0 Zlength (cons_buf (console abd)) < CONSOLE_BUFFER_SIZE

      }.

Definition of the abstract state ops

  Global Instance pqueueintro_data_ops : CompatDataOps RData :=
    {
      empty_data := init_adt multi_oracle_init4;
      high_level_invariant := high_level_invariant;
      low_level_invariant := low_level_invariant;
      kernel_mode adt := ikern adt = true ihost adt = true
    }.

Proofs that the initial abstract_data should satisfy the invariants

  Section Property_Abstract_Data.

    Lemma empty_data_high_level_invariant:
      high_level_invariant (init_adt multi_oracle_init4).
    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_oracle4.
      - apply valid_hlock_pool_init1.
      - apply valid_AT_oracle_pool4.
      - eapply valid_AT_log_pool_H_init.
    Qed.

Definition of the abstract state

    Global Instance pqueueinit_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}.

Proofs that the primitives satisfies the invariants at this layer

  Section INV.

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

    Global Instance set_curid_init_inv: PreservesInvariants set_curid_init_spec.
    Proof.
      preserves_invariants_simpl_auto.
    Qed.

    Global Instance release_lock_inv: ReleaseInvariants release_lock_spec2.
    Proof.
      constructor; unfold release_lock_spec; intros; subdestruct;
      inv H; inv H0; constructor; auto; simpl; intros.
      - eapply consistent_ppage_log_gso; eauto.
        eapply Shared2ID2_neq; eauto.
      - eapply LATCTable_log_gso; eauto.
        eapply Shared2ID2_neq; eauto.
      - eapply valid_hlock_pool1_gss´; eauto.
      - eapply valid_AT_log_pool_H_gso; eauto.
        eapply Shared2ID2_neq; eauto.
    Qed.

    Global Instance acquire_lock_inv: AcquireInvariants acquire_lock_spec2.
    Proof.
      constructor; unfold acquire_lock_spec; intros; subdestruct;
      inv H; inv H0; constructor; auto; simpl; intros.
      - eapply consistent_ppage_log_gso; eauto.
        eapply Shared2ID2_neq; eauto.
      - eapply LATCTable_log_gso; eauto.
        eapply Shared2ID2_neq; eauto.
      - eapply valid_hlock_pool1_gss´; eauto.
      - eapply valid_AT_log_pool_H_gso; eauto.
        eapply Shared2ID2_neq; eauto.
    Qed.


    Global Instance acquire_lock_TCB_inv: PreservesInvariants acquire_lock_TCB_spec.
    Proof.
      preserves_invariants_simpl_auto1.
      - eapply consistent_ppage_log_gso; eauto.
        eapply TCB_neq.
      - eapply LATCTable_log_gso; eauto.
        eapply TCB_neq.
      - eapply valid_hlock_pool1_gss´; eauto.
      - eapply valid_AT_log_pool_H_gso; eauto.
        eapply TCB_neq.
      - eapply consistent_ppage_log_gso; eauto.
        eapply TCB_neq.
      - eapply LATCTable_log_gso; eauto.
        eapply TCB_neq.
      - eapply valid_hlock_pool1_gss´; eauto.
      - eapply valid_AT_log_pool_H_gso; eauto.
        eapply TCB_neq.
    Qed.

    Global Instance release_lock_TCB_inv: PreservesInvariants release_lock_TCB_spec.
    Proof.
      preserves_invariants_simpl_auto1.
      - eapply consistent_ppage_log_gso; eauto.
        eapply TCB_neq.
      - eapply LATCTable_log_gso; eauto.
        eapply TCB_neq.
      - eapply valid_hlock_pool1_gss´; eauto.
      - eapply valid_AT_log_pool_H_gso; eauto.
        eapply TCB_neq.
    Qed.


    Global Instance set_state_inv: PreservesInvariants set_state_spec.
    Proof.
      preserves_invariants_simpl_auto.
    Qed.

    Global Instance set_prev_inv: PreservesInvariants set_prev_spec.
    Proof.
      preserves_invariants_simpl_auto.
    Qed.

    Global Instance set_next_inv: PreservesInvariants set_next_spec.
    Proof.
      preserves_invariants_simpl_auto.
    Qed.

    Global Instance tcb_init_inv: PreservesInvariants tcb_init_spec.
    Proof.
      preserves_invariants_simpl_auto.
    Qed.

    Global Instance tcb_set_CPU_ID_inv: PreservesInvariants tcb_set_CPU_ID_spec.
    Proof.
      preserves_invariants_simpl_auto.
    Qed.

    Global Instance set_head_inv: PreservesInvariants set_head_spec.
    Proof.
      preserves_invariants_simpl_auto.
    Qed.

    Global Instance set_tail_inv: PreservesInvariants set_tail_spec.
    Proof.
      preserves_invariants_simpl_auto.
    Qed.

    Global Instance tdq_init_inv: PreservesInvariants tdq_init_spec.
    Proof.
      preserves_invariants_simpl_auto.
    Qed.

    Section PALLOC.

      Lemma palloc_high_level_inv:
         d i n,
          palloc_spec i d = Some (, n)
          high_level_invariant d
          high_level_invariant .
      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.
        + 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.
        + 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.
      Qed.

      Lemma palloc_low_level_inv:
         d i n ,
          palloc_spec i d = Some (, n)
          low_level_invariant d
          low_level_invariant .
      Proof.
        unfold palloc_spec; intros.
        subdestruct; inv H; subst; eauto;
        inv H0; constructor; eauto.
      Qed.

      Lemma palloc_kernel_mode:
         d i n,
          palloc_spec i d = Some (, n)
          kernel_mode d
          kernel_mode .
      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 low_level_invariant high_level_invariant; eauto 2.
    Qed.

    Section PTINSERT.

      Section PTINSERT_PTE.

        Lemma ptInsertPTE_high_level_inv:
           d n vadr padr p,
            ptInsertPTE0_spec n vadr padr p d = Some
            high_level_invariant d
            high_level_invariant .
        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 n vadr padr p ,
            ptInsertPTE0_spec n vadr padr p d = Some
            low_level_invariant d
            low_level_invariant .
        Proof.
          intros. functional inversion H; subst; eauto.
          inv H0. constructor; eauto.
        Qed.

        Lemma ptInsertPTE_kernel_mode:
           d n vadr padr p,
            ptInsertPTE0_spec n vadr padr p d = Some
            kernel_mode d
            kernel_mode .
        Proof.
          intros. functional inversion H; subst; eauto.
        Qed.

      End PTINSERT_PTE.

      Section PTPALLOCPDE.

        Lemma ptAllocPDE_high_level_inv:
           d n vadr v,
            ptAllocPDE0_spec n vadr d = Some (, v)
            high_level_invariant d
            high_level_invariant .
        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 n vadr v ,
            ptAllocPDE0_spec n vadr d = Some (, v)
            low_level_invariant d
            low_level_invariant .
        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 n vadr v,
            ptAllocPDE0_spec n vadr d = Some (, v)
            kernel_mode d
            kernel_mode .
        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 n vadr padr p v,
          ptInsert0_spec n vadr padr p d = Some (, v)
          high_level_invariant d
          high_level_invariant .
      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 n vadr padr p v,
          ptInsert0_spec n vadr padr p d = Some (, v)
          low_level_invariant d
          low_level_invariant .
      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 n vadr padr p v,
          ptInsert0_spec n vadr padr p d = Some (, v)
          kernel_mode d
          kernel_mode .
      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 n vadr p v,
          ptResv_spec n vadr p d = Some (, v)
          high_level_invariant d
          high_level_invariant .
      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 n vadr p v,
          ptResv_spec n vadr p d = Some (, v)
          low_level_invariant d
          low_level_invariant .
      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 n vadr p v,
          ptResv_spec n vadr p d = Some (, v)
          kernel_mode d
          kernel_mode .
      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 n vadr p vadr´ v,
          ptResv2_spec n vadr p vadr´ d = Some (, v)
          high_level_invariant d
          high_level_invariant .
      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 n vadr p vadr´ l v,
          ptResv2_spec n vadr p vadr´ d = Some (, v)
          low_level_invariant l d
          low_level_invariant l .
      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 n vadr p vadr´ v,
          ptResv2_spec n vadr p vadr´ d = Some (, v)
          kernel_mode d
          kernel_mode .
      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.

      Global Instance ptResv2_inv: PreservesInvariants ptResv2_spec.
      Proof.
        preserves_invariants_simpl´.
        - eapply ptResv2_low_level_inv; eassumption.
        - eapply ptResv2_high_level_inv; eassumption.
        - eapply ptResv2_kernel_mode; eassumption.
      Qed.

    End PTRESV2.


    Global Instance page_copy_inv: PreservesInvariants page_copy_spec.
    Proof.
      preserves_invariants_simpl low_level_invariant high_level_invariant; 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.
    Qed.

    Global Instance page_copy_back_inv: PreservesInvariants page_copy_back_spec.
    Proof.
      preserves_invariants_simpl low_level_invariant high_level_invariant;
      try eapply dirty_ppage_gss_page_copy_back; eauto.
    Qed.

    Global Instance shared_mem_status_inv:
      PreservesInvariants ObjShareMem.shared_mem_status_spec.
    Proof.
      preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
    Qed.

    Global Instance offer_shared_mem_inv:
      PreservesInvariants ObjShareMem.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.

    Global Instance sharedmem_init_inv: PreservesInvariants sharedmem_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.
      - assumption.
      - eapply real_valid_hlock_pool1; eauto.
      - assumption.
      - eapply real_valid_AT_log_pool_H; eauto.
    Qed.

    Global Instance kctxt_switch_inv: KCtxtSwitchInvariants kctxt_switch_spec.
    Proof.
      constructor; intros; functional inversion H.
      - inv H1. constructor; trivial.
        eapply kctxt_inject_neutral_gss_mem; eauto.
      - inv H0. subst. constructor; auto; simpl in *; intros; try congruence.
    Qed.

    Global Instance kctxt_new_inv: DNewInvariants ObjThread.kctxt_new_spec.
    Proof.
      constructor; intros; inv H0;
      unfold ObjThread.kctxt_new_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.

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


    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.

Specification of primitives that will be implemented at this layer

  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.

Layer Definition

  Definition pqueueintro_fresh : compatlayer (cdata RData) :=
    (get_state gensem get_state_spec
               get_prev gensem get_prev_spec
               get_next gensem get_next_spec
               tcb_get_CPU_ID gensem tcb_get_CPU_ID_spec
               set_state gensem set_state_spec
               set_prev gensem set_prev_spec
               set_next gensem set_next_spec
               tcb_set_CPU_ID gensem tcb_set_CPU_ID_spec
               tcb_init gensem tcb_init_spec

               get_head gensem get_head_spec
               get_tail gensem get_tail_spec
               set_head gensem set_head_spec
               set_tail gensem set_tail_spec
               tdq_init gensem tdq_init_spec)

               acquire_lock_TCB gensem acquire_lock_TCB_spec
               release_lock_TCB gensem release_lock_TCB_spec
              .

  Definition pqueueintro_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
           palloc gensem palloc_spec
          
           set_pt gensem setPT_spec
           pt_read gensem ptRead_spec
           pt_resv gensem ptResv_spec
           kctxt_new dnew_compatsem ObjThread.kctxt_new_spec
          
           shared_mem_init gensem sharedmem_init_spec
           shared_mem_status gensem shared_mem_status_spec
           offer_shared_mem gensem offer_shared_mem_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
           set_curid_init gensem set_curid_init_spec

           sleeper_inc gensem sleeper_inc_spec
           sleeper_dec gensem sleeper_dec_spec
           sleeper_zzz gensem sleeper_zzz_spec
           release_lock primcall_release_lock_compatsem release_lock release_lock_spec2
           acquire_lock primcall_acquire_lock_compatsem acquire_lock_spec2
          

           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
           kctxt_switch primcall_kctxt_switch_compatsem kctxt_switch_spec
           accessors {| exec_load := @exec_loadex; exec_store := @exec_storeex |}.

  Definition pqueueintro : compatlayer (cdata RData) := pqueueintro_fresh pqueueintro_passthrough.

End WITHMEM.

Section WITHPARAM.

  Context `{oracle_prop: MultiOracleProp}.
  Context `{real_params: RealParams}.

  Local Open Scope Z_scope.

  Section Impl.

primitve: enqueue
    Function enqueue_spec (n i: Z) (adt: RData): option RData :=
      match (ikern adt, pg adt, ihost adt, ipt adt) with
        | (true, true, true, true)
          if Queue_arg n i then
            match (ZMap.get n (tdq adt), ZMap.get i (tcb adt)) with
              | (TDQValid h t, TCBValid st c _ _)
                if zeq t num_proc then
                  Some adt {tcb: ZMap.set i (TCBValid st c num_proc num_proc) (tcb adt)}
                       {tdq: ZMap.set n (TDQValid i i) (tdq adt)}
                else
                  if zle_lt 0 t num_proc then
                    match (ZMap.get t (tcb adt)) with
                      | TCBValid st´ prev´ _
                        let tcb´:= ZMap.set t (TCBValid st´ prev´ i) (tcb adt) in
                        Some adt {tcb: ZMap.set i (TCBValid st c t num_proc) tcb´}
                             {tdq: ZMap.set n (TDQValid h i) (tdq adt)}
                      | _None
                    end
                  else None
              | _None
            end
          else None
        | _None
      end.

primitve: dequeue
    Function dequeue_spec (n: Z) (adt: RData): option (RData× Z) :=
      match (ikern adt, pg adt, ihost adt, ipt adt) with
        | (true, true, true, true)
          if zle_lt 0 n num_chan then
              match (ZMap.get n (tdq adt)) with
                | TDQValid h t
                  if zeq h num_proc then
                    Some (adt, num_proc)
                  else
                    if zle_lt 0 h num_proc then
                      match (ZMap.get h (tcb adt)) with
                        | TCBValid st c _ next
                          if zeq next num_proc then
                            Some (adt {tdq: ZMap.set n (TDQValid num_proc num_proc) (tdq adt)}, h)
                            else
                              if zle_lt 0 next num_proc then
                                match (ZMap.get next (tcb adt)) with
                                  | TCBValid st´ _ next´
                                    Some (adt {tcb: ZMap.set next (TCBValid st´ num_proc next´) (tcb adt)}
                                              {tdq: ZMap.set n (TDQValid next t) (tdq adt)} , h)
                                  | _None
                                end
                              else None
                        | _None
                      end
                    else None
                | _None
              end
          else None
        | _None
      end.

primitive: initialize the allocation table, set up the paging mechanism, and initialize the page table pool
    Function tdqueue_init_spec (mbi_adr:Z) (abd: RData): option RData :=
      match sharedmem_init_spec mbi_adr abd with
        | Some adt
          Some adt {tcb: real_tcb (tcb adt)}
               {tdq: real_tdq (tdq adt)}
        | _None
      end.

  End Impl.

End WITHPARAM.