Library mcertikos.mm.ContainerIntroGenSpec


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

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 AbstractDataType.
Require Import DAbsHandler.

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

Definition of the refinement relation

Section CONTAINERGEN_DEFINE.

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

  Notation LDATAOps := (cdata RData).

  Definition QUOTA := 0.
  Definition USAGE := 4.
  Definition PARENT := 8.
  Definition NCHILDREN := 12.
  Definition USED := 16.
  Definition CSIZE := 20.

  Inductive container_init_node_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
  | container_init_node_low_intro s (WB: _Prop) (m m1 m2 m3 m4 : mem) labd b0 i quota_val parent_val:
      find_symbol s AC_LOC = Some b0
      Mem.store Mint32 m b0 (Int.unsigned i × CSIZE + QUOTA) (Vint quota_val) = Some m1
      Mem.store Mint32 m1 b0 (Int.unsigned i × CSIZE + USAGE) (Vint Int.zero) = Some m2
      Mem.store Mint32 m2 b0 (Int.unsigned i × CSIZE + PARENT) (Vint parent_val) = Some m3
      Mem.store Mint32 m3 b0 (Int.unsigned i × CSIZE + NCHILDREN) (Vint Int.zero) = Some m4
      Mem.store Mint32 m4 b0 (Int.unsigned i × CSIZE + USED) (Vint Int.one) = Some
      0 Int.unsigned i < num_proc
      kernel_mode labd
      container_init_node_low_step s WB (Vint i::Vint quota_val::Vint parent_val::nil) (m,labd) Vundef (,labd).

  Inductive container_set_usage_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
  | container_set_usage_low_intro s (WB: _Prop) (m : mem) labd b0 i usage_val :
      find_symbol s AC_LOC = Some b0
      Mem.store Mint32 m b0 (Int.unsigned i × CSIZE + USAGE) (Vint usage_val) = Some
      0 Int.unsigned i < num_proc
      kernel_mode labd
      container_set_usage_low_step s WB (Vint i::Vint usage_val::nil) (m,labd) Vundef (,labd).

  Inductive container_set_nchildren_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
  | container_set_nchildren_low_intro s (WB: _Prop) (m : mem) labd b0 i next_child cur_num:
      find_symbol s AC_LOC = Some b0
      Mem.load Mint32 m b0 (Int.unsigned i × CSIZE + NCHILDREN) = Some (Vint cur_num) →
      Mem.store Mint32 m b0 (Int.unsigned i × CSIZE + NCHILDREN) (Vint (Int.add cur_num Int.one)) = Some
      0 Int.unsigned i < num_proc
      kernel_mode labd
      container_set_nchildren_low_step s WB (Vint i::Vint next_child::nil) (m,labd) Vundef (,labd).

  Inductive container_get_parent_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
  | container_get_parent_low_intro s (WB: _Prop) (m´0: mwd LDATAOps) b0 i v:
      find_symbol s AC_LOC = Some b0
      Mem.load Mint32 m´0 b0 (Int.unsigned i × CSIZE + PARENT) = Some (Vint v) →
      0 Int.unsigned i < num_proc
      kernel_mode (snd m´0) →
      container_get_parent_low_step s WB (Vint i :: nil) m´0 (Vint v) m´0.

  Inductive container_get_quota_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
  | container_get_quota_low_intro s (WB: _Prop) (m´0: mwd LDATAOps) b0 i v:
      find_symbol s AC_LOC = Some b0
      Mem.load Mint32 m´0 b0 (Int.unsigned i × CSIZE + QUOTA) = Some (Vint v) →
      0 Int.unsigned i < num_proc
      kernel_mode (snd m´0) →
      container_get_quota_low_step s WB (Vint i :: nil) m´0 (Vint v) m´0.

  Inductive container_get_usage_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
  | container_get_usage_low_intro s (WB: _Prop) (m´0: mwd LDATAOps) b0 i v:
      find_symbol s AC_LOC = Some b0
      Mem.load Mint32 m´0 b0 (Int.unsigned i × CSIZE + USAGE) = Some (Vint v) →
      0 Int.unsigned i < num_proc
      kernel_mode (snd m´0) →
      container_get_usage_low_step s WB (Vint i :: nil) m´0 (Vint v) m´0.

  Inductive container_get_nchildren_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
  | container_get_nchildren_low_intro s (WB: _Prop) (m´0: mwd LDATAOps) b0 i v:
      find_symbol s AC_LOC = Some b0
      Mem.load Mint32 m´0 b0 (Int.unsigned i × CSIZE + NCHILDREN) = Some (Vint v) →
      0 Int.unsigned i < num_proc
      kernel_mode (snd m´0) →
      container_get_nchildren_low_step s WB (Vint i :: nil) m´0 (Vint v) m´0.

  Inductive container_can_consume_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
  | container_can_consume_low_intro s (WB: _Prop) (m´0: mwd LDATAOps) b0 i n q u:
      find_symbol s AC_LOC = Some b0
      Mem.load Mint32 m´0 b0 (Int.unsigned i × CSIZE + QUOTA) = Some (Vint q) →
      Mem.load Mint32 m´0 b0 (Int.unsigned i × CSIZE + USAGE) = Some (Vint u) →
      0 Int.unsigned i < num_proc
      kernel_mode (snd m´0) →
      container_can_consume_low_step s WB (Vint i :: Vint n :: nil) m´0
        (Vint (match (Int.unsigned n <=? Int.unsigned q,
                      Int.unsigned u <=? Int.unsigned q - Int.unsigned n) with
               | (true, true)Int.one
               | _Int.zero end)) m´0.

  Inductive container_alloc_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
  | container_alloc_low_step_intro0 s (WB: _Prop) (m: mwd LDATAOps) b0 i u:
      find_symbol s AC_LOC = Some b0
      0 Int.unsigned i < num_proc
      Mem.load Mint32 m b0 (Int.unsigned i × CSIZE + USAGE) = Some (Vint u) →
      Mem.load Mint32 m b0 (Int.unsigned i × CSIZE + QUOTA) = Some (Vint u) →
      kernel_mode (snd m) →
      container_alloc_low_step s WB (Vint i :: nil) m Vzero m
  | container_alloc_low_step_intro1 s (WB: _Prop) (m : mwd LDATAOps) b0 i u q:
      find_symbol s AC_LOC = Some b0
      0 Int.unsigned i < num_procInt.unsigned u < Int.unsigned q
      Mem.load Mint32 m b0 (Int.unsigned i × CSIZE + USAGE) = Some (Vint u) →
      Mem.load Mint32 m b0 (Int.unsigned i × CSIZE + QUOTA) = Some (Vint q) →
      Mem.store Mint32 m b0 (Int.unsigned i × CSIZE + USAGE) (Vint (Int.add u Int.one)) = Some
      kernel_mode (snd m) →
      container_alloc_low_step s WB (Vint i :: nil) m Vone .

  Section WITHMEM.

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

    Definition container_init_node_spec_low: compatsem LDATAOps :=
      csem container_init_node_low_step (type_of_list_type (Tint32::Tint32::Tint32::nil)) Tvoid.

    Definition container_set_usage_spec_low: compatsem LDATAOps :=
      csem container_set_usage_low_step (type_of_list_type (Tint32::Tint32::nil)) Tvoid.

    Definition container_set_nchildren_spec_low: compatsem LDATAOps :=
      csem container_set_nchildren_low_step (type_of_list_type (Tint32::Tint32::nil)) Tvoid.

    Definition container_get_parent_spec_low: compatsem LDATAOps :=
      csem container_get_parent_low_step (type_of_list_type (Tint32::nil)) Tint32.

    Definition container_get_quota_spec_low: compatsem LDATAOps :=
      csem container_get_quota_low_step (type_of_list_type (Tint32::nil)) Tint32.

    Definition container_get_usage_spec_low: compatsem LDATAOps :=
      csem container_get_usage_low_step (type_of_list_type (Tint32::nil)) Tint32.

    Definition container_get_nchildren_spec_low: compatsem LDATAOps :=
      csem container_get_nchildren_low_step (type_of_list_type (Tint32::nil)) Tint32.

    Definition container_can_consume_spec_low: compatsem LDATAOps :=
      csem container_can_consume_low_step (type_of_list_type (Tint32::Tint32::nil)) Tint32.

    Definition container_alloc_spec_low: compatsem LDATAOps :=
      csem container_alloc_low_step (type_of_list_type (Tint32::nil)) Tint32.

  End WITHMEM.

End CONTAINERGEN_DEFINE.