Library mcertikos.mm.MPTIntro


This file defines the abstract data and the primitives for the MPTIntro layer, which will abstract the page table pool from the memory
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 liblayers.logic.PTreeModules.
Require Import liblayers.logic.LayerLogicImpl.
Require Import liblayers.compat.CompatLayers.
Require Import liblayers.compat.CompatGenSem.

Require Import INVLemmaContainer.
Require Import INVLemmaMemory.
Require Import CalRealIDPDE.
Require Import INVLemmaQLock.
Require Export INVLemmaInterrupt.
Require Import INVLemmaDriver.

Require Import AbstractDataType.

Require Import DeviceStateDataType.

Require Export ObjCPU.
Require Export ObjMM.
Require Export ObjFlatMem.
Require Export ObjContainer.
Require Export ObjMultiprocessor.
Require Export ObjQLock.
Require Export ObjPMM.

Require Export ObjInterruptManagement.
Require Export ObjInterruptController.
Require Export ObjConsole.
Require Export ObjSerialDriver.
Require Export ObjVMM.

Require Import FutureTactic.

Section WITHMEM.

  Local Open Scope Z_scope.

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

Raw Abstract Data



Invariants at this layer

  Record high_level_invariant (abd: RData) :=
    mkInvariant {
        CPU_ID_range: 0 (CPU_ID abd) < TOTAL_CPU;
        valid_curid: 0 ZMap.get (CPU_ID abd) (cid abd) < num_proc;

        valid_nps: init abd= truekern_low nps abd maxpage;
        
        valid_kern: ipt abd = falsepg abd = true init abd = true;
        valid_iptt: ipt abd = trueikern abd = true;
        valid_iptf: ikern abd = falseipt abd = false;
        valid_ihost: ihost abd = falsepg abd = true init 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: init abd = false(pperm abd) = ZMap.init PGUndef;
        valid_PMap: pg abd = truePMap_valid (ZMap.get (PT abd) (ptpool abd));
        valid_PMap_kern: pg abd = trueipt abd = true
                         PMap_kern (ZMap.get (PT abd) (ptpool abd));
        valid_PT: pg abd = true → 0 PT abd < num_proc;
        valid_dirty: dirty_ppage (pperm abd) (HP abd);

        valid_uninitialized: init abd = falsePMap_uninitialized (ptpool 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_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 mptintro_data_ops : CompatDataOps RData :=
    {
      empty_data := init_adt multi_oracle_init3;
      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_init3).
    Proof.
      constructor; simpl; intros; auto; try inv H.
      - apply current_CPU_ID_range.
      - rewrite ZMap.gi; intuition.
      -
        apply empty_container_valid.
      - eapply consistent_ppage_log_init.
      - eapply dirty_ppage_init.
      - eapply PMap_uninitialized_undef.
      - eapply valid_ticket_oracle3.
      - apply valid_hlock_pool_init1.
      - apply valid_AT_oracle_pool3.
      - eapply valid_AT_log_pool_H_init.
    Qed.

Definition of the abstract state

    Global Instance mptintro_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.
      - eauto 2 with serial_intr_disable_invariantdb.
   Qed.

    Global Instance serial_intr_enable_inv:
      PreservesInvariants serial_intr_enable_spec.
    Proof.
      constructor; simpl; intros; inv_generic_sem H.
      - inversion H0; econstructor; eauto 2 with serial_intr_enable_invariantdb.
        generalize (serial_intr_enable_preserves_tf _ _ H2); intro tmprw; rewrite <- tmprw; assumption.
      - inversion H0; econstructor; eauto 2 with serial_intr_enable_invariantdb.
      - eauto 2 with serial_intr_enable_invariantdb.
    Qed.

    Global Instance set_curid_inv: PreservesInvariants set_curid_spec.
    Proof.
      preserves_invariants_simpl_auto.
      - rewrite ZMap.gss. auto.
    Qed.

    Global Instance set_curid_init_inv: PreservesInvariants set_curid_init_spec.
    Proof.
      preserves_invariants_simpl_auto; eauto 2.
      case_eq (zeq (CPU_ID d) (Int.unsigned i)); intros; subst.
      - rewrite e; rewrite ZMap.gss; omega.
      - rewrite ZMap.gso; auto.
    Qed.

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

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


    Section PALLOC.


      Lemma lpalloc_high_level_inv:
         d i n,
          lpalloc_spec i d = Some (, n)
          high_level_invariant d
          high_level_invariant .
      Proof.
        unfold lpalloc_spec; intros.
        subdestruct; inv H; subst; eauto;
        inv H0; constructor; simpl; eauto.
        + rewrite <- Hdestruct2.
          eapply alloc_container_valid´; eauto.
        + eapply consistent_ppage_log_norm_alloc; eauto. omega.
        + subst; simpl;
          intros; congruence.
        + eapply dirty_ppage_gso_alloc; eauto.
        + 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 valid_hlock_pool1_gso; eauto.
        + eapply valid_AT_log_pool_H_0; eauto.
        + rewrite app_comm_cons.
          eapply consistent_ppage_log_gss; eauto.
        + eapply valid_hlock_pool1_gso; eauto.
        + eapply valid_AT_log_pool_H_0´; eauto.
      Qed.


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


      Lemma lpalloc_kernel_mode:
         d i n,
          lpalloc_spec i d = Some (, n)
          kernel_mode d
          kernel_mode .
      Proof.
        unfold lpalloc_spec; intros.
        subdestruct; inv H; simpl; eauto.
      Qed.

      Global Instance lpalloc_inv: PreservesInvariants lpalloc_spec.
      Proof.
        preserves_invariants_simpl´.
        - eapply lpalloc_low_level_inv; eassumption.
        - eapply lpalloc_high_level_inv; eassumption.
        - eapply lpalloc_kernel_mode; eassumption.
      Qed.

    End PALLOC.


    Global Instance container_split_inv: PreservesInvariants container_split_spec.
    Proof.
      preserves_invariants_simpl_auto.
      rewrite <- H1 in H0.
      exploit split_container_valid; eauto.
    Qed.

    Global Instance 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 mem_init_inv: PreservesInvariants mem_init0_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.
      - assumption.
      - assumption.
      - eapply real_valid_hlock_pool1; eauto.
      - assumption.
      - eapply real_valid_AT_log_pool_H; eauto.
    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 setPG_inv: PreservesInvariants setPG1_spec.
    Proof.
      preserves_invariants_simpl_auto.
    Qed.

    Global Instance set_at_c_inv: PreservesInvariants set_at_c0_spec.
    Proof.
      preserves_invariants_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.
      - inv H0. functional inversion H2. functional inversion H.
        split; subst; simpl;
        try (eapply dirty_ppage_store_unmaped; try reflexivity; try eassumption); trivial.
      - inv H0. functional inversion H2. functional inversion H0.
        split; simpl; try assumption.
    Qed.

    Global Instance setPT_inv: PreservesInvariants setPT´_spec.
    Proof.
      preserves_invariants_simpl_auto.
      - rewrite H5 in ×. contradiction.
    Qed.

    Global Instance setPDE_inv: PreservesInvariants setPDE_spec.
    Proof.
      preserves_invariants_simpl_auto.
    Qed.

    Global Instance rmvPDE_inv: PreservesInvariants rmvPDE_spec.
    Proof.
      preserves_invariants_simpl_auto1;
      try (rewrite ZMap.gso; eauto; fail).
      - eapply consistent_ppage_log_hide_alloc; eassumption.
      - eapply dirty_ppage_gso_alloc; eauto.
      - eapply consistent_ppage_log_hide_alloc; eassumption.
      - eapply dirty_ppage_gso_alloc; eauto.
    Qed.

    Global Instance setPDEU_inv: PreservesInvariants setPDEU_spec.
    Proof.
      preserves_invariants_simpl_auto;
      try (rewrite ZMap.gso; eauto; fail).
      - apply consistent_ppage_log_alloc_hide; assumption.
      - apply dirty_ppage_gss. assumption.
    Qed.

    Global Instance setPTE_inv:
      PreservesInvariants setPTE_spec.
    Proof.
      preserves_invariants_simpl_auto;
      try (rewrite ZMap.gso; eauto; fail).
    Qed.

    Global Instance rmvPTE_inv: PreservesInvariants rmvPTE_spec.
    Proof.
      preserves_invariants_simpl_auto;
      try (rewrite ZMap.gso; eauto; fail).
    Qed.

    Global Instance setIDPTE_inv:
      PreservesInvariants setIDPTE_spec.
    Proof.
      preserves_invariants_simpl_auto.
    Qed.


    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 Shared2ID1_neq; eauto.
        reflexivity.
      - eapply valid_hlock_pool1_gss´; eauto.
      - eapply valid_AT_log_pool_H_gso; eauto.
        eapply Shared2ID1_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 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

Layer Definition newly introduced

  Definition mptintro_fresh : compatlayer (cdata RData) :=
    (set_pt gensem setPT´_spec
             get_PDE gensem getPDE_spec
             set_PDE gensem setPDE_spec
             rmv_PDE gensem rmvPDE_spec
             set_PDEU gensem setPDEU_spec
             get_PTE gensem getPTE_spec
             set_PTE gensem setPTE_spec
             rmv_PTE gensem rmvPTE_spec
             set_IDPTE gensem setIDPTE_spec)
       (pt_in primcall_general_compatsem´ (prim_ident:= pt_in) ptin´_spec
                pt_out primcall_general_compatsem´ (prim_ident:= pt_out) ptout_spec).

Layer Definition passthrough

  Definition mptintro_passthrough : compatlayer (cdata RData) :=
    fload gensem fload_spec
           fstore gensem fstore_spec
          
           page_copy gensem page_copy_spec
           page_copy_back gensem page_copy_back_spec

           vmxinfo_get gensem vmxinfo_get_spec
           set_pg gensem setPG1_spec
           at_get_c gensem get_at_c_spec
           at_set_c gensem set_at_c0_spec
           palloc gensem lpalloc_spec
          
           mem_init gensem mem_init0_spec
           container_get_parent gensem container_get_parent_spec
           container_get_nchildren gensem container_get_nchildren_spec
           container_get_quota gensem container_get_quota_spec
           container_get_usage gensem container_get_usage_spec
           container_can_consume gensem container_can_consume_spec
           container_split gensem container_split_spec
           get_CPU_ID gensem get_CPU_ID_spec
           get_curid gensem get_curid_spec
           set_curid gensem set_curid_spec
           set_curid_init gensem set_curid_init_spec

           release_lock primcall_release_lock_compatsem release_lock release_lock_spec1
           acquire_lock primcall_acquire_lock_compatsem acquire_lock_spec1
          

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

           trap_in primcall_general_compatsem trapin_spec
           trap_out primcall_general_compatsem 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 |}.

Layer Definition

  Definition mptintro : compatlayer (cdata RData) := mptintro_fresh mptintro_passthrough.

End WITHMEM.