Library mcertikos.devdrivers.DHandler


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 LAsm.
Require Import LoadStoreSem1.
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 AbstractDataType.
Require Import FutureTactic.

Require Export DHandlerOp.
Require Export ObjInterruptManagement.
Require Import INVLemmaInterrupt.
Require Import DeviceStateDataType.

Abstract Data and Primitives at this layer

Section WITHMEM.

  Local Open Scope Z_scope.

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

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

Proofs that the primitives satisfies the invariants at this layer

  Section INV.

Dhandler primitives
    Global Instance get_curr_intr_num_inv: PreservesInvariants get_curr_intr_num_spec.
    Proof.
      preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
    Qed.

    Global Instance intr_default_handler_inv: PreservesInvariants intr_default_handler_spec.
    Proof.
      preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
    Qed.

passthrough primitives

    Global Instance serial_intr_disable_concrete_inv: PreservesInvariants serial_intr_disable_concrete_spec.
    Proof.
      constructor; simpl; intros; inv_generic_sem H.
      Opaque serial_intr_disable_concrete_aux.
      - functional inversion H2; subst; simpl in ×.
        assert (low_level_invariant m d0).
        {
          remember 100%nat as n.
          clear Heqn.
          Transparent serial_intr_disable_concrete_aux.
          clear H1 H3 H4 H5 H2.
          generalize dependent d.
          generalize dependent d0.
          induction n.
          {
            simpl in ×.
            intros.
            inv H7; eauto.
          }
          {
            simpl in ×.
            intros.
            subdestruct.
            eapply IHn in H7; eauto.
            functional inversion Hdestruct; subst.
            functional inversion H5; simpl in *; functional inversion H4; simpl in *; functional inversion H3; simpl in *; functional inversion H1; simpl in *; functional inversion H26; simpl in *; subst; inv H0; econstructor; simpl in *; eauto.
            inv H0; econstructor; simpl in *; eauto.
            functional inversion Hdestruct; subst.
            functional inversion H5; simpl in *; functional inversion H4; simpl in *; functional inversion H3; simpl in *; functional inversion H1; simpl in *; functional inversion H26; simpl in *; subst; inv H0; simpl; assumption.
            simpl; assumption.
            inv H7; inv H0; econstructor; simpl in *; eauto.
          }
        }
        inv H; econstructor; eauto 2 with serial_intr_disable_invariantdb.
      Opaque serial_intr_disable_concrete_aux.
      - functional inversion H2; subst; simpl in ×.
        remember 100%nat as n.
        clear Heqn.
        Transparent serial_intr_disable_concrete_aux.
        clear H4 H5 H2.
        generalize dependent d.
        generalize dependent d0.
        induction n.
        {
          simpl in ×.
          intros.
          inv H7; eauto.
          inv H0; econstructor; simpl in *; eauto.
          rewrite replace_preserves_Zlength.
          assumption.
          intros.
          generalize (valid_ioapic_state H _ _ H0); intros (vval & (b & nev)).
          assert (ncase: n = 4%nat n 4%nat) by omega.
          case_eq ncase; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ nev); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.
        }
        {
          simpl in ×.
          intros.
          subdestruct.
          eapply IHn in H7; eauto.
          functional inversion Hdestruct; subst.
          functional inversion H9; simpl in *; functional inversion H8; simpl in *; functional inversion H5; simpl in *; functional inversion H2; simpl in *; functional inversion H28; simpl in *; subst; unfold s_ioapic´ in *; inv H0; econstructor; simpl in *; eauto.

          unfold ioapic_trans_intr.
          destruct (s (ioapic d)).
          destruct (nth_error IoApicIrqs (Z.to_nat 4)).
          destruct (nth_error IoApicEnables (Z.to_nat 4)).
          destruct b.
          simpl; assumption.
          simpl; assumption.
          simpl; assumption.
          simpl; assumption.

          rewrite replace_replace.
          rewrite replace_preserves_Zlength.
          unfold ioapic_trans_intr.
          destruct (s (ioapic d)).
          destruct (nth_error IoApicIrqs (Z.to_nat 4)).
          destruct (nth_error IoApicEnables (Z.to_nat 4)).
          destruct b.
          simpl; assumption.
          simpl; assumption.
          simpl; assumption.
          simpl; assumption.

          unfold ioapic_trans_intr.
          destruct (s (ioapic d)).
          destruct (nth_error IoApicIrqs (Z.to_nat 4)).
          destruct (nth_error IoApicEnables (Z.to_nat 4)).
          destruct b.
          intros;
          simpl in ×.
          rewrite replace_replace.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H20.
          assert (n0case: n0 = 4%nat n0 4%nat) by omega.
          case_eq n0case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H20); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.

          intros;
          simpl in ×.
          rewrite replace_replace.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H20.
          assert (n0case: n0 = 4%nat n0 4%nat) by omega.
          case_eq n0case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H20); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.

          intros;
          simpl in ×.
          rewrite replace_replace.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H20.
          assert (n0case: n0 = 4%nat n0 4%nat) by omega.
          case_eq n0case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H20); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.

          intros;
          simpl in ×.
          rewrite replace_replace.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H20.
          assert (n0case: n0 = 4%nat n0 4%nat) by omega.
          case_eq n0case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H20); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.

          unfold ioapic_trans_intr.
          destruct (s (ioapic d)).
          destruct (nth_error IoApicIrqs (Z.to_nat 4)).
          destruct (nth_error IoApicEnables (Z.to_nat 4)).
          destruct b.
          simpl; assumption.
          simpl; assumption.
          simpl; assumption.
          simpl; assumption.

          rewrite replace_replace.
          rewrite replace_preserves_Zlength.
          unfold ioapic_trans_intr.
          destruct (s (ioapic d)).
          destruct (nth_error IoApicIrqs (Z.to_nat 4)).
          destruct (nth_error IoApicEnables (Z.to_nat 4)).
          destruct b.
          simpl; assumption.
          simpl; assumption.
          simpl; assumption.
          simpl; assumption.

          unfold ioapic_trans_intr.
          destruct (s (ioapic d)).
          destruct (nth_error IoApicIrqs (Z.to_nat 4)).
          destruct (nth_error IoApicEnables (Z.to_nat 4)).
          destruct b.
          intros;
          simpl in ×.
          rewrite replace_replace.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H20.
          assert (n0case: n0 = 4%nat n0 4%nat) by omega.
          case_eq n0case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H20); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.

          intros;
          simpl in ×.
          rewrite replace_replace.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H20.
          assert (n0case: n0 = 4%nat n0 4%nat) by omega.
          case_eq n0case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H20); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.

          intros;
          simpl in ×.
          rewrite replace_replace.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H20.
          assert (n0case: n0 = 4%nat n0 4%nat) by omega.
          case_eq n0case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H20); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.

          intros;
          simpl in ×.
          rewrite replace_replace.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H20.
          assert (n0case: n0 = 4%nat n0 4%nat) by omega.
          case_eq n0case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H20); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.

          inv H0; econstructor; simpl in *; eauto.

          functional inversion Hdestruct; subst.
          functional inversion H9; simpl in *; auto.
          simpl; auto.
          functional inversion Hdestruct; subst.
          functional inversion H9; simpl in *; auto.
          simpl; assumption.

          subst.
          functional inversion Hdestruct; subst.
          functional inversion H9; simpl in *; functional inversion H8; simpl in *; functional inversion H5; simpl in *; functional inversion H2; simpl in *; subst; assumption.
          simpl; assumption.
          inv H7; inv H0; econstructor; simpl in *; eauto.

          rewrite replace_preserves_Zlength.
          assumption.

          intros;
          simpl in ×.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H4.
          assert (n0case: n0 = 4%nat n0 4%nat) by omega.
          case_eq n0case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H4); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.
        }
      Opaque serial_intr_disable_concrete_aux.
      - functional inversion H2; subst; simpl in ×.
        remember 100%nat as n.
        clear Heqn.
        Transparent serial_intr_disable_concrete_aux.
        clear H1 H3 H2.
        generalize dependent d.
        generalize dependent d0.
        induction n.
        {
          simpl in ×.
          intros.
          inv H7; eauto.
        }
        {
          simpl in ×.
          intros.
          subdestruct.
          assert(ikern r = true ihost r = true init r = true).
          {
            functional inversion Hdestruct; subst.
            functional inversion H9; simpl in *; functional inversion H8; simpl in *; functional inversion H3; simpl in *; functional inversion H1; simpl in *; functional inversion H28; simpl in *; subst; auto.
            simpl; auto.
          }
          destruct H.
          destruct H1.
          eapply IHn in H7; eauto.
          inv H7; auto.
        }
    Qed.


    Global Instance serial_intr_enable_concrete_inv: PreservesInvariants serial_intr_enable_concrete_spec.
    Proof.
      constructor; simpl; intros; inv_generic_sem H.
      Opaque serial_intr_enable_concrete_aux.
      - functional inversion H2; subst; simpl in ×.
        assert (low_level_invariant m d0).
        {
          remember 100%nat as n.
          clear Heqn.
          Transparent serial_intr_enable_concrete_aux.
          clear H1 H3 H4 H5 H2.
          generalize dependent d.
          generalize dependent d0.
          induction n.
          {
            simpl in ×.
            intros.
            inv H7; eauto.
          }
          {
            simpl in ×.
            intros.
            subdestruct.
            eapply IHn in H7; eauto.
            functional inversion Hdestruct; subst.
            functional inversion H5; simpl in *; functional inversion H4; simpl in *; simpl in *; subst; functional inversion H8; functional inversion H2; functional inversion H24; simpl in *; subst; inv H0; econstructor; simpl in *; eauto.
            functional inversion Hdestruct; subst.
            functional inversion H5; simpl in *; functional inversion H4; simpl in *; simpl in *; subst; functional inversion H8; functional inversion H2; functional inversion H24; simpl in *; subst; simpl; eassumption.
            inv H7; inv H0; econstructor; simpl in *; eauto.
          }
        }
        inv H; econstructor; eauto.
      Opaque serial_intr_enable_concrete_aux.
      - functional inversion H2; subst; simpl in ×.
        remember 100%nat as n.
        clear Heqn.
        Transparent serial_intr_enable_concrete_aux.
        clear H4 H5 H2.
        generalize dependent d.
        generalize dependent d0.
        induction n.
        {
          simpl in ×.
          intros.
          inv H7; eauto.
          inv H0; econstructor; simpl in *; eauto.
          rewrite replace_preserves_Zlength.
          assumption.
          intros.
          generalize (valid_ioapic_state H _ _ H0); intros (vval & (b & nev)).
          assert (ncase: n = 4%nat n 4%nat) by omega.
          case_eq ncase; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ nev); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.
        }
        {
          simpl in ×.
          intros.
          subdestruct.
          eapply IHn in H7; eauto.
          functional inversion Hdestruct; subst.
          functional inversion H10; simpl in *; functional inversion H9; simpl in *; subst; functional inversion H8; functional inversion H4; simpl in *; subst; functional inversion H26; simpl in *; subst; unfold s_ioapic´ in *; inv H0; econstructor; simpl in *; eauto.

          unfold ioapic_trans_intr.
          destruct (s (ioapic d)).
          destruct (nth_error IoApicIrqs (Z.to_nat 4)).
          destruct (nth_error IoApicEnables (Z.to_nat 4)).
          destruct b.
          intros;
          simpl in ×.
          assumption.
          assumption.
          assumption.
          assumption.

          unfold ioapic_trans_intr.
          destruct (s (ioapic d)).
          destruct (nth_error IoApicIrqs (Z.to_nat 4)).
          destruct (nth_error IoApicEnables (Z.to_nat 4)).
          destruct b.
          intros;
          simpl in ×.
          rewrite replace_replace.
          rewrite replace_preserves_Zlength.
          assumption.
          rewrite replace_replace.
          rewrite replace_preserves_Zlength.
          assumption.
          rewrite replace_replace.
          rewrite replace_preserves_Zlength.
          assumption.
          rewrite replace_replace.
          rewrite replace_preserves_Zlength.
          assumption.

          unfold ioapic_trans_intr.
          destruct (s (ioapic d)).
          destruct (nth_error IoApicIrqs (Z.to_nat 4)).
          destruct (nth_error IoApicEnables (Z.to_nat 4)).
          destruct b.
          intros;
          simpl in ×.
          rewrite replace_replace.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H30.
          assert (n1case: n1 = 4%nat n1 4%nat) by omega.
          case_eq n1case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H30); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.

          intros;
          simpl in ×.
          rewrite replace_replace.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H30.
          assert (n1case: n1 = 4%nat n1 4%nat) by omega.
          case_eq n1case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H30); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.

          intros;
          simpl in ×.
          rewrite replace_replace.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H30.
          assert (n1case: n1 = 4%nat n1 4%nat) by omega.
          case_eq n1case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H30); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.

          intros;
          simpl in ×.
          rewrite replace_replace.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H30.
          assert (n1case: n1 = 4%nat n1 4%nat) by omega.
          case_eq n1case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H30); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.

          unfold ioapic_trans_intr.
          destruct (s (ioapic d)).
          destruct (nth_error IoApicIrqs (Z.to_nat 4)).
          destruct (nth_error IoApicEnables (Z.to_nat 4)).
          destruct b.
          simpl in *; assumption.
          simpl in *; assumption.
          simpl in *; assumption.
          simpl in *; assumption.

          rewrite replace_replace.
          rewrite replace_preserves_Zlength.
          unfold ioapic_trans_intr.
          destruct (s (ioapic d)).
          destruct (nth_error IoApicIrqs (Z.to_nat 4)).
          destruct (nth_error IoApicEnables (Z.to_nat 4)).
          destruct b.
          simpl in *; assumption.
          simpl in *; assumption.
          simpl in *; assumption.
          simpl in *; assumption.

          unfold ioapic_trans_intr.
          destruct (s (ioapic d)).
          destruct (nth_error IoApicIrqs (Z.to_nat 4)).
          destruct (nth_error IoApicEnables (Z.to_nat 4)).
          destruct b.
          intros;
          simpl in ×.
          rewrite replace_replace.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H30.
          assert (n1case: n1 = 4%nat n1 4%nat) by omega.
          case_eq n1case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H30); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.

          intros;
          simpl in ×.
          rewrite replace_replace.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H30.
          assert (n1case: n1 = 4%nat n1 4%nat) by omega.
          case_eq n1case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H30); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.

          intros;
          simpl in ×.
          rewrite replace_replace.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H30.
          assert (n1case: n1 = 4%nat n1 4%nat) by omega.
          case_eq n1case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H30); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.

          intros;
          simpl in ×.
          rewrite replace_replace.
          generalize (valid_ioapic_state H _ _ H0); intro tmp.
          destruct tmp.
          destruct H30.
          assert (n1case: n1 = 4%nat n1 4%nat) by omega.
          case_eq n1case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ H30); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.

          functional inversion Hdestruct; subst.
          functional inversion H10; simpl in *; auto.
          functional inversion Hdestruct; subst.
          functional inversion H10; simpl in *; auto.
          functional inversion Hdestruct; subst.
          functional inversion H10; simpl in *; functional inversion H9; simpl in *; functional inversion H8; simpl in *; functional inversion H4; simpl in *; subst; assumption.
          inv H7; inv H0; econstructor; simpl in *; eauto.
          rewrite replace_preserves_Zlength.
          assumption.
          intros.
          generalize (valid_ioapic_state H _ _ H0); intros (vval & (b & nev)).
          assert (n0case: n0 = 4%nat n0 4%nat) by omega.
          case_eq n0case; intros.
          subst.
          rewrite nth_error_replace_gss.
          split.
          reflexivity.
          esplit; reflexivity.
          generalize (nth_error_range _ _ _ _ nev); intro.
          xomega.
          rewrite nth_error_replace_gso.
          eapply valid_ioapic_state; eauto.
          change (Pos.to_nat 4%positive) with 4%nat.
          omega.
        }
      Opaque serial_intr_enable_concrete_aux.
      - functional inversion H2; subst; simpl in ×.
        remember 100%nat as n.
        clear Heqn.
        Transparent serial_intr_enable_concrete_aux.
        clear H1 H3 H2.
        generalize dependent d.
        generalize dependent d0.
        induction n.
        {
          simpl in ×.
          intros.
          inv H7; eauto.
        }
        {
          simpl in ×.
          intros.
          subdestruct.
          assert(ikern r = true ihost r = true init r = true).
          {
            functional inversion Hdestruct; subst.
            functional inversion H10; simpl in *; functional inversion H9; simpl in *; subst; functional inversion H8; functional inversion H2; simpl in *; subst; functional inversion H26; simpl in *; subst; eauto.
          }
          destruct H.
          destruct H1.
          eapply IHn in H7; eauto.
          inv H7; auto.
        }
    Qed.

  End INV.

Layer Definition

  Definition dhandler_fresh : compatlayer (cdata RData) :=
    serial_intr_disable gensem serial_intr_disable_concrete_spec
                         serial_intr_enable gensem serial_intr_enable_concrete_spec.

  Definition dhandler_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 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
           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_spec0
           acquire_lock primcall_acquire_lock_compatsem acquire_lock_spec0
           ticket_lock_init gensem ticket_lock_init0_spec
           cli gensem cli_spec
           sti gensem sti_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 _ _ _ _ _ Hmwd);
                           exec_store := (@exec_storeex _ _ _ _ _ Hmwd) |}.

  Definition dhandler : compatlayer (cdata RData) :=
    dhandler_fresh dhandler_passthrough.

End WITHMEM.