Library mcertikos.proc.EPTOpGenSpec


This file provide the contextual refinement proof between MBoot layer and MALInit layer
Require Import Coqlib.
Require Import Errors.
Require Import AST.
Require Import Integers.
Require Import Floats.
Require Import Op.
Require Import Asm.
Require Import Events.
Require Import Globalenvs.
Require Import Smallstep.
Require Import Values.
Require Import Memory.
Require Import Maps.
Require Import AuxLemma.
Require Import FlatMemory.
Require Import AuxStateDataType.
Require Import Constant.
Require Import GlobIdent.
Require Import RealParams.
Require Import AsmImplLemma.
Require Import GenSem.

Require Import liblayers.logic.PTreeModules.
Require Import liblayers.logic.LayerLogicImpl.
Require Import liblayers.compcertx.Stencil.
Require Import liblayers.compcertx.MakeProgram.
Require Import liblayers.compat.CompatLayers.
Require Import liblayers.compat.CompatGenSem.
Require Import compcert.cfrontend.Ctypes.
Require Import Conventions.
Require Import VEPTIntro.
Require Import AbstractDataType.
Require Import mCertiKOSType.

Local Open Scope string_scope.
Local Open Scope error_monad_scope.
Local Open Scope Z_scope.

Definition of the low level specification

Section SPECIFICATION.

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

  Notation LDATA := RData.

  Notation LDATAOps := (cdata LDATA).

  Inductive ept_init_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps):=
  | ept_init_spec_low_intro s (WB: _Prop) m´0 labd labd´:
      ept_init_spec labd = Some labd´
      kernel_mode labd
      high_level_invariant labd
      ept_init_spec_low_step s WB nil (m´0, labd) Vundef (m´0, labd´).

  Inductive ept_get_page_entry_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps):=
  | ept_get_page_entry_spec_low_intro s (WB: _Prop) m´0 labd gpa hpa:
      ept_get_page_entry_spec (Int.unsigned gpa) labd = Some (Int.unsigned hpa) →
      kernel_mode labd
      high_level_invariant labd
      ept_get_page_entry_spec_low_step s WB (Vint gpa :: nil) (m´0, labd) (Vint hpa) (m´0, labd).

  Inductive ept_set_page_entry_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps):=
  | ept_set_page_entry_spec_low_intro s (WB: _Prop) m´0 labd labd´ gpa hpa:
      ept_set_page_entry_spec (Int.unsigned gpa) (Int.unsigned hpa) labd = Some labd´
      kernel_mode labd
      high_level_invariant labd
      ept_set_page_entry_spec_low_step s WB (Vint gpa :: Vint hpa :: nil) (m´0, labd) Vundef (m´0, labd´).

  Inductive ept_add_mapping_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps):=
  | ept_add_mapping_spec_low_intro s (WB: _Prop) m´0 labd labd´ gpa hpa memtype res:
      ept_add_mapping_spec (Int.unsigned gpa) (Int.unsigned hpa)
                           (Int.unsigned memtype) labd = Some (labd´, Int.unsigned res)
      kernel_mode labd
      high_level_invariant labd
      ept_add_mapping_spec_low_step s WB (Vint gpa :: Vint hpa :: Vint memtype :: nil) (m´0, labd) (Vint res) (m´0, labd´).

  Section WITHMEM.

    Context `{Hstencil: Stencil}.
    Context `{Hmem: Mem.MemoryModelX}.
    Context `{Hmwd: UseMemWithData mem}.

    Definition ept_init_spec_low: compatsem LDATAOps :=
      csem ept_init_spec_low_step (type_of_list_type nil) Tvoid.

    Definition ept_get_page_entry_spec_low: compatsem LDATAOps :=
      csem ept_get_page_entry_spec_low_step (type_of_list_type (Tint32::nil)) Tint32.

    Definition ept_set_page_entry_spec_low: compatsem LDATAOps :=
      csem ept_set_page_entry_spec_low_step (type_of_list_type (Tint32::Tint32::nil)) Tvoid.

    Definition ept_add_mapping_spec_low: compatsem LDATAOps :=
      csem ept_add_mapping_spec_low_step (type_of_list_type (Tint32::Tint32::Tint32::nil)) Tint32.

  End WITHMEM.

End SPECIFICATION.