Library mcertikos.proc.QueueIntroGenSpec

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

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

Definition of the refinement relation

Section Refinement.

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

  Notation LDATA := RData.

  Notation LDATAOps := (cdata LDATA).

  Inductive acquire_lock_TCB_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sprimcall_sem (mem := mwd LDATAOps):=
  | acquire_lock_TCB_spec_low_intro
      s m´0 m0 b b0 l id adt adt´ rs sig index
      (Hsymbol_PC: find_symbol s acquire_lock_TCB = Some b)
      (HPC: rs PC = Vptr b Int.zero)
      (Hspec: acquire_lock_spec0 (Z.of_nat local_lock_bound)
                                 lock_TCB_start (Int.unsigned index) adt = Some (adt´, id, l))
      (Hsymbol: find_symbol s id = Some b0)
      (Hstorebytes: match l with
                      | Some Mem.storebytes m´0 b0 0 (ByteList ) = Some m0
                      | _m0 = m´0
                    end)
      (Hsig: sig = mksignature (AST.Tint::nil) (Some AST.Tint) cc_default)
      (Harg: extcall_arguments rs m´0 sig (Vint index :: nil))
      (Hrange: 0 Int.unsigned index < num_chan)
      (Hasm: asm_invariant (mem := mwd LDATAOps) s rs (m´0, adt))
      (Hlow: low_level_invariant (Mem.nextblock m´0) adt):
      let rs´ := (undef_regs (CR ZF :: CR CF :: CR PF :: CR SF :: CR OF
                                 :: IR EAX :: RA :: nil)
                             (undef_regs (List.map preg_of destroyed_at_call) rs)) in
      acquire_lock_TCB_spec_low_step s rs (m´0, adt) (rs´#RA <- Vundef #PC <- (rs#RA)) (m0, adt´).

  Inductive release_lock_TCB_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sprimcall_sem (mem := mwd LDATAOps):=
  | release_lock_TCB_spec_low_intro
      s m´0 b b0 l id adt adt´ rs size sig index
      (Hsymbol_PC: find_symbol s release_lock_TCB = Some b)
      (HPC: rs PC = Vptr b Int.zero)
      (Hspec: release_lock_spec0 lock_TCB_start (Int.unsigned index) l adt = Some adt´)
      (Hsize: id2size lock_TCB_start = Some (size, id))
      (Hsymbol: find_symbol s id = Some b0)
      (Hloadbytes: Mem.loadbytes m´0 b0 0 size = Some (ByteList l))
      (Hsig: sig = mksignature (AST.Tint::nil) (Some AST.Tint) cc_default)
      (Harg: extcall_arguments rs m´0 sig (Vint index :: nil))
      (Hrange: 0 Int.unsigned index < num_chan)

      (Hasm: asm_invariant (mem := mwd LDATAOps) s rs (m´0, adt))
      (Hlow: low_level_invariant (Mem.nextblock m´0) adt):
      let rs´ := (undef_regs (CR ZF :: CR CF :: CR PF :: CR SF :: CR OF
                                 :: IR EAX :: RA :: nil)
                             (undef_regs (List.map preg_of destroyed_at_call) rs)) in
      release_lock_TCB_spec_low_step s rs (m´0, adt) (rs´#RA <- Vundef #PC <- (rs#RA)) (m´0, adt´).

  Inductive get_state_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
    get_state_spec_low_intro s (WB: _Prop) (m´0: mwd LDATAOps) b0 n v:
      find_symbol s TCBPool_LOC = Some b0
      Mem.load Mint32 m´0 b0 ((Int.unsigned n) × 16) = Some (Vint v) →
      0 (Int.unsigned n) < num_proc
      kernel_mode (snd m´0) →
      get_state_spec_low_step s WB (Vint n :: nil) m´0 (Vint v) m´0.

  Inductive tcb_get_CPU_ID_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
    tcb_get_CPU_ID_spec_low_intro s (WB: _Prop) (m´0: mwd LDATAOps) b0 n v:
      find_symbol s TCBPool_LOC = Some b0
      Mem.load Mint32 m´0 b0 ((Int.unsigned n) × 16 + 4) = Some (Vint v) →
      0 (Int.unsigned n) < num_proc
      kernel_mode (snd m´0) →
      tcb_get_CPU_ID_spec_low_step s WB (Vint n :: nil) m´0 (Vint v) m´0.

  Inductive get_prev_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
    get_prev_spec_low_intro s (WB: _Prop) (m´0: mwd LDATAOps) b0 n v:
      find_symbol s TCBPool_LOC = Some b0
      Mem.load Mint32 m´0 b0 ((Int.unsigned n) × 16 + 8) = Some (Vint v) →
      0 (Int.unsigned n) < num_proc
      kernel_mode (snd m´0) →
      get_prev_spec_low_step s WB (Vint n :: nil) m´0 (Vint v) m´0.

  Inductive get_next_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
    get_next_spec_low_intro s (WB: _Prop) (m´0: mwd LDATAOps) b0 n v:
      find_symbol s TCBPool_LOC = Some b0
      Mem.load Mint32 m´0 b0 ((Int.unsigned n) × 16 + 12) = Some (Vint v) →
      0 (Int.unsigned n) < num_proc
      kernel_mode (snd m´0) →
      get_next_spec_low_step s WB (Vint n :: nil) m´0 (Vint v) m´0.

  Inductive set_state_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
    set_state_spec_low_intro s (WB: _Prop) (m´0 m0: mwd LDATAOps) b0 n v:
      find_symbol s TCBPool_LOC = Some b0
      Mem.store Mint32 m´0 b0 (Int.unsigned n × 16) (Vint v) = Some m0
      0 (Int.unsigned n) < num_proc
      kernel_mode (snd m´0) →
      set_state_spec_low_step s WB (Vint n :: Vint v :: nil) m´0 Vundef m0.

  Inductive tcb_set_CPU_ID_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
    tcb_set_CPU_ID_spec_low_intro s (WB: _Prop) (m´0 m0: mwd LDATAOps) b0 n v:
      find_symbol s TCBPool_LOC = Some b0
      Mem.store Mint32 m´0 b0 (Int.unsigned n × 16 + 4) (Vint v) = Some m0
      0 (Int.unsigned n) < num_proc
      kernel_mode (snd m´0) →
      tcb_set_CPU_ID_spec_low_step s WB (Vint n :: Vint v :: nil) m´0 Vundef m0.

  Inductive set_prev_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
    set_prev_spec_low_intro s (WB: _Prop) (m´0 m0: mwd LDATAOps) b0 n v:
      find_symbol s TCBPool_LOC = Some b0
      Mem.store Mint32 m´0 b0 (Int.unsigned n × 16 + 8) (Vint v) = Some m0
      0 (Int.unsigned n) < num_proc
      kernel_mode (snd m´0) →
      set_prev_spec_low_step s WB (Vint n :: Vint v :: nil) m´0 Vundef m0.

  Inductive set_next_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
    set_next_spec_low_intro s (WB: _Prop) (m´0 m0: mwd LDATAOps) b0 n v:
      find_symbol s TCBPool_LOC = Some b0
      Mem.store Mint32 m´0 b0 (Int.unsigned n × 16 + 12) (Vint v) = Some m0
      0 (Int.unsigned n) < num_proc
      kernel_mode (snd m´0) →
      set_next_spec_low_step s WB (Vint n :: Vint v :: nil) m´0 Vundef m0.

  Inductive tcb_init_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
    tcb_init_spec_low_intro s (WB: _Prop) (m0 m1 m2 m3 m4: mwd LDATAOps) b0 n:
      find_symbol s TCBPool_LOC = Some b0
      Mem.store Mint32 m0 b0 (Int.unsigned n × 16) (Vint (Int.repr TSTATE_DEAD)) = Some m1
      Mem.store Mint32 m1 b0 (Int.unsigned n × 16 + 4) (Vint (Int.repr 0)) = Some m2
      Mem.store Mint32 m2 b0 (Int.unsigned n × 16 + 8) (Vint (Int.repr num_proc)) = Some m3
      Mem.store Mint32 m3 b0 (Int.unsigned n × 16 + 12) (Vint (Int.repr num_proc)) = Some m4
      0 (Int.unsigned n) < num_proc
      kernel_mode (snd m0) →
      tcb_init_spec_low_step s WB (Vint n :: nil) m0 Vundef m4.

  Inductive get_head_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
    get_head_spec_low_intro s (WB: _Prop) (m´0: mwd LDATAOps) b0 n v:
      find_symbol s TCBPool_LOC = Some b0
      Mem.load Mint32 m´0 b0 ((Int.unsigned n) × 8 + num_proc × 16) = Some (Vint v) →
      0 (Int.unsigned n) < num_chan
      kernel_mode (snd m´0) →
      get_head_spec_low_step s WB (Vint n :: nil) m´0 (Vint v) m´0.

  Inductive get_tail_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
    get_tail_spec_low_intro s (WB: _Prop) (m´0: mwd LDATAOps) b0 n v:
      find_symbol s TCBPool_LOC = Some b0
      Mem.load Mint32 m´0 b0 ((Int.unsigned n) × 8 + 4 + num_proc × 16) = Some (Vint v) →
      0 (Int.unsigned n) < num_chan
      kernel_mode (snd m´0) →
      get_tail_spec_low_step s WB (Vint n :: nil) m´0 (Vint v) m´0.

  Inductive set_head_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
    set_head_spec_low_intro s (WB: _Prop) (m´0 m0: mwd LDATAOps) b0 n v:
      find_symbol s TCBPool_LOC = Some b0
      Mem.store Mint32 m´0 b0 (Int.unsigned n × 8 + num_proc × 16) (Vint v) = Some m0
      0 (Int.unsigned n) < num_chan
      kernel_mode (snd m´0) →
      set_head_spec_low_step s WB (Vint n :: Vint v :: nil) m´0 Vundef m0.

  Inductive set_tail_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
    set_tail_spec_low_intro s (WB: _Prop) (m´0 m0: mwd LDATAOps) b0 n v:
      find_symbol s TCBPool_LOC = Some b0
      Mem.store Mint32 m´0 b0 (Int.unsigned n × 8 + 4 + num_proc × 16) (Vint v) = Some m0
      0 (Int.unsigned n) < num_chan
      kernel_mode (snd m´0) →
      set_tail_spec_low_step s WB (Vint n :: Vint v :: nil) m´0 Vundef m0.

  Inductive tdq_init_spec_low_step `{StencilOps} `{Mem.MemoryModelOps} `{UseMemWithData mem}:
    sextcall_sem (mem := mwd LDATAOps) :=
    tdq_init_spec_low_intro s (WB: _Prop) (m0 m1 m2: mwd LDATAOps) b0 n:
      find_symbol s TCBPool_LOC = Some b0
      Mem.store Mint32 m0 b0 (Int.unsigned n × 8 + num_proc × 16) (Vint (Int.repr num_proc)) = Some m1
      Mem.store Mint32 m1 b0 (Int.unsigned n × 8 + 4 + num_proc × 16) (Vint (Int.repr num_proc)) = Some m2
      0 (Int.unsigned n) < num_chan
      kernel_mode (snd m0) →
      tdq_init_spec_low_step s WB (Vint n :: nil) m0 Vundef m2.

  Section WITHMEM.

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

    Definition acquire_lock_TCB_spec_low: compatsem LDATAOps :=
      asmsem_withsig acquire_lock_TCB acquire_lock_TCB_spec_low_step (mksignature (AST.Tint::nil) None cc_default).

    Definition release_lock_TCB_spec_low: compatsem LDATAOps :=
      asmsem_withsig release_lock_TCB release_lock_TCB_spec_low_step (mksignature (AST.Tint::nil) None cc_default).

    Definition get_state_spec_low: compatsem LDATAOps :=
      csem get_state_spec_low_step (type_of_list_type (Tint32::nil)) Tint32.

    Definition set_state_spec_low: compatsem LDATAOps :=
      csem set_state_spec_low_step (type_of_list_type (Tint32::Tint32::nil)) Tvoid.

    Definition tcb_get_CPU_ID_spec_low: compatsem LDATAOps :=
      csem tcb_get_CPU_ID_spec_low_step (type_of_list_type (Tint32::nil)) Tint32.

    Definition tcb_set_CPU_ID_spec_low: compatsem LDATAOps :=
      csem tcb_set_CPU_ID_spec_low_step (type_of_list_type (Tint32::Tint32::nil)) Tvoid.

    Definition get_prev_spec_low: compatsem LDATAOps :=
      csem get_prev_spec_low_step (type_of_list_type (Tint32::nil)) Tint32.

    Definition set_prev_spec_low: compatsem LDATAOps :=
      csem set_prev_spec_low_step (type_of_list_type (Tint32::Tint32::nil)) Tvoid.

    Definition get_next_spec_low: compatsem LDATAOps :=
      csem get_next_spec_low_step (type_of_list_type (Tint32::nil)) Tint32.

    Definition set_next_spec_low: compatsem LDATAOps :=
      csem set_next_spec_low_step (type_of_list_type (Tint32::Tint32::nil)) Tvoid.

    Definition tcb_init_spec_low: compatsem LDATAOps :=
      csem tcb_init_spec_low_step (type_of_list_type (Tint32::nil)) Tvoid.

    Definition get_head_spec_low: compatsem LDATAOps :=
      csem get_head_spec_low_step (type_of_list_type (Tint32::nil)) Tint32.

    Definition set_head_spec_low: compatsem LDATAOps :=
      csem set_head_spec_low_step (type_of_list_type (Tint32::Tint32::nil)) Tvoid.

    Definition get_tail_spec_low: compatsem LDATAOps :=
      csem get_tail_spec_low_step (type_of_list_type (Tint32::nil)) Tint32.

    Definition set_tail_spec_low: compatsem LDATAOps :=
      csem set_tail_spec_low_step (type_of_list_type (Tint32::Tint32::nil)) Tvoid.

    Definition tdq_init_spec_low: compatsem LDATAOps :=
      csem tdq_init_spec_low_step (type_of_list_type (Tint32::nil)) Tvoid.

  End WITHMEM.

End Refinement.