Library mcertikos.devdrivers.DIoApic
*********************************************************************** * * * The CertiKOS Certified Kit Operating System * * * * The FLINT Group, Yale University * * * * Copyright The FLINT Group, Yale University. All rights reserved. * * This file is distributed under the terms of the Yale University * * Non-Commercial License Agreement. * * * ***********************************************************************
This file defines the abstract data and the primitives for the
ConsoleBufOp layer, which will introduce the primtives of console buffer
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 XOmega.
Require Import liblayers.logic.PTreeModules.
Require Import liblayers.logic.LayerLogicImpl.
Require Import liblayers.compat.CompatLayers.
Require Import liblayers.compat.CompatGenSem.
Require Import INVLemmaDriver.
Require Import INVLemmaDevice.
Require Import AbstractDataType.
Require Import FutureTactic.
Require Export ObjInterruptDriver.
Require Export ObjCPU.
Require Export ObjMM.
Require Export ObjFlatMem.
Require Export ObjSerialDevice.
Require Export ObjSerialDriver.
Require Export ObjInterruptController.
Require Export ObjConsole.
Require Export ObjX86.
Require Import AbstractDataType.
Require Import DeviceStateDataType.
Require Import AbstractDataType.
Require Import FutureTactic.
Require Import DeviceStateDataType.
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 XOmega.
Require Import liblayers.logic.PTreeModules.
Require Import liblayers.logic.LayerLogicImpl.
Require Import liblayers.compat.CompatLayers.
Require Import liblayers.compat.CompatGenSem.
Require Import INVLemmaDriver.
Require Import INVLemmaDevice.
Require Import AbstractDataType.
Require Import FutureTactic.
Require Export ObjInterruptDriver.
Require Export ObjCPU.
Require Export ObjMM.
Require Export ObjFlatMem.
Require Export ObjSerialDevice.
Require Export ObjSerialDriver.
Require Export ObjInterruptController.
Require Export ObjConsole.
Require Export ObjX86.
Require Import AbstractDataType.
Require Import DeviceStateDataType.
Require Import AbstractDataType.
Require Import FutureTactic.
Require Import DeviceStateDataType.
Record high_level_invariant (abd: RData) :=
mkInvariant {
valid_com1_port: abd.(com1).(s).(Base) = 1016;
valid_cons_buf_rpos: 0 ≤ rpos (console abd) < CONSOLE_BUFFER_SIZE;
valid_cons_buf_length: 0 ≤ Zlength (cons_buf (console abd)) < CONSOLE_BUFFER_SIZE;
valid_serial_exists: abd.(drv_serial).(serial_exists) = 0 ∨ abd.(drv_serial).(serial_exists) = 1;
valid_ioapic_max_intr: 0 ≤ IoApicMaxIntr (s (ioapic abd)) < 32
}.
mkInvariant {
valid_com1_port: abd.(com1).(s).(Base) = 1016;
valid_cons_buf_rpos: 0 ≤ rpos (console abd) < CONSOLE_BUFFER_SIZE;
valid_cons_buf_length: 0 ≤ Zlength (cons_buf (console abd)) < CONSOLE_BUFFER_SIZE;
valid_serial_exists: abd.(drv_serial).(serial_exists) = 0 ∨ abd.(drv_serial).(serial_exists) = 1;
valid_ioapic_max_intr: 0 ≤ IoApicMaxIntr (s (ioapic abd)) < 32
}.
Global Instance dioapic_data_ops : CompatDataOps RData :=
{
empty_data := init_adt;
high_level_invariant := high_level_invariant;
low_level_invariant := low_level_invariant;
kernel_mode adt := ikern adt = true ∧ ihost adt = true
}.
{
empty_data := init_adt;
high_level_invariant := high_level_invariant;
low_level_invariant := low_level_invariant;
kernel_mode adt := ikern adt = true ∧ ihost adt = true
}.
Section Property_Abstract_Data.
Lemma empty_data_high_level_invariant:
high_level_invariant init_adt.
Proof.
constructor; simpl; intros; auto; try inv H; omega.
Qed.
Lemma empty_data_high_level_invariant:
high_level_invariant init_adt.
Proof.
constructor; simpl; intros; auto; try inv H; omega.
Qed.
Global Instance dioapic_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.MemoryModel}.
Context `{Hmwd: UseMemWithData mem}.
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.MemoryModel}.
Context `{Hmwd: UseMemWithData mem}.
Section INV.
Global Instance serial_putc_inv:
PreservesInvariants serial_putc_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
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 flatmem_copy_inv: PreservesInvariants flatmem_copy´_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto.
Qed.
Global Instance fstore_inv: PreservesInvariants fstore´_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
Qed.
Global Instance setPG_inv: PreservesInvariants setPG_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
Qed.
Global Instance setCR3_inv: SetCR3Invariants setCR3_spec.
Proof.
constructor; intros.
- functional inversion H; inv H0; constructor; trivial.
- functional inversion H; inv H0; constructor; auto.
- functional inversion H; simpl in *; assumption.
Qed.
Global Instance cli_inv: PreservesInvariants cli_spec.
Proof.
econstructor; intros; inv H; functional inversion H1; subst.
- inv H0; constructor; trivial.
- inv H0; constructor; auto.
- simpl in *; assumption.
Qed.
Global Instance sti_inv: PreservesInvariants sti_spec.
Proof.
econstructor; intros; inv H; functional inversion H1; subst.
- inv H0; constructor; trivial.
- inv H0; constructor; auto.
- simpl in *; assumption.
Qed.
Global Instance bootloader0_inv: PreservesInvariants bootloader0_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
Qed.
Require Import FutureTactic.
Global Instance serial_hw_intr_inv: PreservesInvariants serial_hw_intr_spec.
Proof.
preserves_invariants_nested low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance ic_intr_inv: PreservesInvariants ic_intr_spec.
Proof.
preserves_invariants_nested low_level_invariant high_level_invariant; eauto.
Qed.
Global Instance lapic_read_inv: PreservesInvariants lapic_read_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
Qed.
Global Instance lapic_write_inv: PreservesInvariants lapic_write_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
Qed.
Global Instance trapin_inv: PrimInvariants trapin_spec.
Proof.
PrimInvariants_simpl H H0.
Qed.
Global Instance trapout_inv: PrimInvariants trapout´_spec.
Proof.
PrimInvariants_simpl H H0.
Qed.
Global Instance hostin_inv: PrimInvariants hostin_spec.
Proof.
PrimInvariants_simpl H H0.
Qed.
Global Instance hostout_inv: PrimInvariants hostout´_spec.
Proof.
PrimInvariants_simpl H H0.
Qed.
Lemma bichoose_1: 1 = 0 ∨ 1 = 1. right; reflexivity. Qed.
Lemma bichoose_0: 0 = 0 ∨ 0 = 1. left; reflexivity. Qed.
Hint Resolve bichoose_1.
Hint Resolve bichoose_0.
Global Instance serial_init_inv:
PreservesInvariants serial_init_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto.
Qed.
Global Instance serial_getc_inv:
PreservesInvariants serial_getc_spec.
Proof.
preserves_invariants_nested low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance cons_buf_init_inv:
PreservesInvariants cons_buf_init_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance cons_buf_write_inv:
PreservesInvariants cons_buf_write_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance cons_buf_wpos_inv:
PreservesInvariants cons_buf_wpos_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance iret_inv: PreservesInvariants iret_spec.
Proof.
preserves_invariants_direct low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance ioapic_init_inv:
PreservesInvariants ioapic_init_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant;
eauto 2; try omega.
Qed.
Global Instance ioapic_enable_inv:
PreservesInvariants ioapic_enable_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant;
eauto 2; try omega.
Qed.
Global Instance ioapic_mask_inv:
PreservesInvariants ioapic_mask_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance ioapic_unmask_inv:
PreservesInvariants ioapic_unmask_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance save_context_inv: SaveContextInvariants save_context_spec.
Proof.
constructor; intros.
-
inv H0. unfold save_context_spec; simpl. inv H1.
constructor; simpl; trivial.
induction (tf d).
simpl.
constructor.
constructor.
simpl.
apply inv_reg_wt.
simpl.
inv inv_inject_neutral.
eapply inv_reg_inject_neutral; auto.
constructor.
constructor.
inv tf_INJECT.
assumption.
eapply IHt; auto.
inv tf_INJECT.
assumption.
-
inv H0. unfold save_context_spec; simpl.
constructor; simpl; trivial.
Qed.
Global Instance restore_context_inv: RestoreContextInvariants restore_context_spec.
Proof.
constructor; intros.
-
inv H0. functional inversion H; inv H1;
constructor; simpl; trivial.
rewrite H4 in ×.
clear H4 H.
generalize dependent _x.
induction _x0.
intros.
simpl in ×.
constructor.
intros.
simpl in ×.
constructor.
inv tf_INJECT.
assumption.
eapply IH_x0; auto.
inv tf_INJECT; assumption.
-
inv H0. functional inversion H; subst;
constructor; simpl; trivial.
Qed.
End INV.
Require Import LoadStoreSem1.
Definition exec_loadex {F V} := exec_loadex1 (F := F) (V := V).
Definition exec_storeex {F V} := exec_storeex1 (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; constructor; auto.
- functional inversion H1; constructor; auto.
Qed.
Global Instance trapinfo_set_inv: TrapinfoSetInvariant.
Proof.
split; inversion 1; intros; constructor; auto.
Qed.
Global Instance serial_putc_inv:
PreservesInvariants serial_putc_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
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 flatmem_copy_inv: PreservesInvariants flatmem_copy´_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto.
Qed.
Global Instance fstore_inv: PreservesInvariants fstore´_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
Qed.
Global Instance setPG_inv: PreservesInvariants setPG_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
Qed.
Global Instance setCR3_inv: SetCR3Invariants setCR3_spec.
Proof.
constructor; intros.
- functional inversion H; inv H0; constructor; trivial.
- functional inversion H; inv H0; constructor; auto.
- functional inversion H; simpl in *; assumption.
Qed.
Global Instance cli_inv: PreservesInvariants cli_spec.
Proof.
econstructor; intros; inv H; functional inversion H1; subst.
- inv H0; constructor; trivial.
- inv H0; constructor; auto.
- simpl in *; assumption.
Qed.
Global Instance sti_inv: PreservesInvariants sti_spec.
Proof.
econstructor; intros; inv H; functional inversion H1; subst.
- inv H0; constructor; trivial.
- inv H0; constructor; auto.
- simpl in *; assumption.
Qed.
Global Instance bootloader0_inv: PreservesInvariants bootloader0_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
Qed.
Require Import FutureTactic.
Global Instance serial_hw_intr_inv: PreservesInvariants serial_hw_intr_spec.
Proof.
preserves_invariants_nested low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance ic_intr_inv: PreservesInvariants ic_intr_spec.
Proof.
preserves_invariants_nested low_level_invariant high_level_invariant; eauto.
Qed.
Global Instance lapic_read_inv: PreservesInvariants lapic_read_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
Qed.
Global Instance lapic_write_inv: PreservesInvariants lapic_write_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
Qed.
Global Instance trapin_inv: PrimInvariants trapin_spec.
Proof.
PrimInvariants_simpl H H0.
Qed.
Global Instance trapout_inv: PrimInvariants trapout´_spec.
Proof.
PrimInvariants_simpl H H0.
Qed.
Global Instance hostin_inv: PrimInvariants hostin_spec.
Proof.
PrimInvariants_simpl H H0.
Qed.
Global Instance hostout_inv: PrimInvariants hostout´_spec.
Proof.
PrimInvariants_simpl H H0.
Qed.
Lemma bichoose_1: 1 = 0 ∨ 1 = 1. right; reflexivity. Qed.
Lemma bichoose_0: 0 = 0 ∨ 0 = 1. left; reflexivity. Qed.
Hint Resolve bichoose_1.
Hint Resolve bichoose_0.
Global Instance serial_init_inv:
PreservesInvariants serial_init_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto.
Qed.
Global Instance serial_getc_inv:
PreservesInvariants serial_getc_spec.
Proof.
preserves_invariants_nested low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance cons_buf_init_inv:
PreservesInvariants cons_buf_init_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance cons_buf_write_inv:
PreservesInvariants cons_buf_write_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance cons_buf_wpos_inv:
PreservesInvariants cons_buf_wpos_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance iret_inv: PreservesInvariants iret_spec.
Proof.
preserves_invariants_direct low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance ioapic_init_inv:
PreservesInvariants ioapic_init_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant;
eauto 2; try omega.
Qed.
Global Instance ioapic_enable_inv:
PreservesInvariants ioapic_enable_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant;
eauto 2; try omega.
Qed.
Global Instance ioapic_mask_inv:
PreservesInvariants ioapic_mask_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance ioapic_unmask_inv:
PreservesInvariants ioapic_unmask_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance save_context_inv: SaveContextInvariants save_context_spec.
Proof.
constructor; intros.
-
inv H0. unfold save_context_spec; simpl. inv H1.
constructor; simpl; trivial.
induction (tf d).
simpl.
constructor.
constructor.
simpl.
apply inv_reg_wt.
simpl.
inv inv_inject_neutral.
eapply inv_reg_inject_neutral; auto.
constructor.
constructor.
inv tf_INJECT.
assumption.
eapply IHt; auto.
inv tf_INJECT.
assumption.
-
inv H0. unfold save_context_spec; simpl.
constructor; simpl; trivial.
Qed.
Global Instance restore_context_inv: RestoreContextInvariants restore_context_spec.
Proof.
constructor; intros.
-
inv H0. functional inversion H; inv H1;
constructor; simpl; trivial.
rewrite H4 in ×.
clear H4 H.
generalize dependent _x.
induction _x0.
intros.
simpl in ×.
constructor.
intros.
simpl in ×.
constructor.
inv tf_INJECT.
assumption.
eapply IH_x0; auto.
inv tf_INJECT; assumption.
-
inv H0. functional inversion H; subst;
constructor; simpl; trivial.
Qed.
End INV.
Require Import LoadStoreSem1.
Definition exec_loadex {F V} := exec_loadex1 (F := F) (V := V).
Definition exec_storeex {F V} := exec_storeex1 (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; constructor; auto.
- functional inversion H1; constructor; auto.
Qed.
Global Instance trapinfo_set_inv: TrapinfoSetInvariant.
Proof.
split; inversion 1; intros; constructor; auto.
Qed.
Definition dioapic_fresh : compatlayer (cdata RData) :=
ioapic_init ↦ gensem ioapic_init_spec
⊕ ioapic_enable ↦ gensem ioapic_enable_spec
⊕ ioapic_mask ↦ gensem ioapic_mask_spec
⊕ ioapic_unmask ↦ gensem ioapic_unmask_spec.
Definition dioapic_passthrough : compatlayer (cdata RData) :=
fload ↦ gensem fload´_spec
⊕ fstore ↦ gensem fstore´_spec
⊕ flatmem_copy ↦ gensem flatmem_copy´_spec
⊕ vmxinfo_get ↦ gensem vmxinfo_get_spec
⊕ set_pg ↦ gensem setPG_spec
⊕ set_cr3 ↦ setCR3_compatsem setCR3_spec
⊕ get_size ↦ gensem MMSize
⊕ is_usable ↦ gensem is_mm_usable_spec
⊕ get_mms ↦ gensem get_mm_s_spec
⊕ get_mml ↦ gensem get_mm_l_spec
⊕ boot_loader ↦ gensem bootloader0_spec
⊕ serial_irq_check ↦ gensem serial_irq_check_spec
⊕ iret ↦ gensem iret_spec
⊕ cli ↦ gensem cli_spec
⊕ sti ↦ gensem sti_spec
⊕ serial_irq_current ↦ gensem serial_irq_current_spec
⊕ ic_intr ↦ gensem ic_intr_spec
⊕ save_context ↦ primcall_save_context_compatsem save_context_spec
⊕ restore_context ↦ primcall_restore_context_compatsem restore_context_spec
⊕ cons_buf_init ↦ gensem cons_buf_init_spec
⊕ serial_init ↦ gensem serial_init_spec
⊕ serial_getc ↦ gensem serial_getc_spec
⊕ serial_putc ↦ gensem serial_putc_spec
⊕ serial_hw_intr ↦ gensem serial_hw_intr_spec
ioapic_init ↦ gensem ioapic_init_spec
⊕ ioapic_enable ↦ gensem ioapic_enable_spec
⊕ ioapic_mask ↦ gensem ioapic_mask_spec
⊕ ioapic_unmask ↦ gensem ioapic_unmask_spec.
Definition dioapic_passthrough : compatlayer (cdata RData) :=
fload ↦ gensem fload´_spec
⊕ fstore ↦ gensem fstore´_spec
⊕ flatmem_copy ↦ gensem flatmem_copy´_spec
⊕ vmxinfo_get ↦ gensem vmxinfo_get_spec
⊕ set_pg ↦ gensem setPG_spec
⊕ set_cr3 ↦ setCR3_compatsem setCR3_spec
⊕ get_size ↦ gensem MMSize
⊕ is_usable ↦ gensem is_mm_usable_spec
⊕ get_mms ↦ gensem get_mm_s_spec
⊕ get_mml ↦ gensem get_mm_l_spec
⊕ boot_loader ↦ gensem bootloader0_spec
⊕ serial_irq_check ↦ gensem serial_irq_check_spec
⊕ iret ↦ gensem iret_spec
⊕ cli ↦ gensem cli_spec
⊕ sti ↦ gensem sti_spec
⊕ serial_irq_current ↦ gensem serial_irq_current_spec
⊕ ic_intr ↦ gensem ic_intr_spec
⊕ save_context ↦ primcall_save_context_compatsem save_context_spec
⊕ restore_context ↦ primcall_restore_context_compatsem restore_context_spec
⊕ cons_buf_init ↦ gensem cons_buf_init_spec
⊕ serial_init ↦ gensem serial_init_spec
⊕ serial_getc ↦ gensem serial_getc_spec
⊕ serial_putc ↦ gensem serial_putc_spec
⊕ serial_hw_intr ↦ gensem serial_hw_intr_spec
serial device
lapic device
lapic device
console buffer
⊕ 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
⊕ 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 |}.
Definition dioapic : compatlayer (cdata RData) :=
dioapic_fresh ⊕ dioapic_passthrough.
End WITHMEM.
⊕ trap_out ↦ primcall_general_compatsem trapout´_spec
⊕ host_in ↦ primcall_general_compatsem hostin_spec
⊕ host_out ↦ primcall_general_compatsem hostout´_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 |}.
Definition dioapic : compatlayer (cdata RData) :=
dioapic_fresh ⊕ dioapic_passthrough.
End WITHMEM.