Library mcertikos.devdrivers.DHandlerAsm

***********************************************************************
*                                                                     *
*            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 DevicePrimSemantics.
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 INVLemmaInterrupt.

Require Import AbstractDataType.
Require Import FutureTactic.

Require Export DHandlerSw.

Abstract Data and Primitives at this layer

Section WITHMEM.

  Local Open Scope Z_scope.

  Context `{real_params: RealParams}.

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

Proofs that the primitives satisfies the invariants at this layer

  Section INV.

    Global Instance serial_intr_handler_asm_inv: SerialIntrHandlerAsmInvariants serial_intr_handler_asm_spec.
    Proof.
      constructor; intros.
      -
        inv H0. apply serial_intr_handler_asm_spec_inv in H. blast H. clear H5 H6 H7. subst.
        inv H1.
        functional inversion Hex1; subst.
        functional inversion Hex3; subst;
        functional inversion Hex6; subst;
        constructor; simpl; trivial.
        simpl in H17.
        apply app_eq_nil in H17.
        destruct H17.
        inv H17.
        rewrite removelast_app.
        simpl.
        rewrite app_nil_r.
        assumption.
        intro contra; inv contra.
        simpl in H17.
        apply app_eq_nil in H17.
        destruct H17.
        inv H17.
        rewrite removelast_app.
        simpl.
        rewrite app_nil_r.
        assumption.
        intro contra; inv contra.
      -
        inv H0. apply serial_intr_handler_asm_spec_inv in H. blast H. clear H4 H5 H6. subst.
        functional inversion Hex1; subst.
        functional inversion Hex3; subst; simpl in ×.
        {
          functional inversion Hex6; subst;
          clear Hex1 Hex3 Hex6.
          {
            constructor; simpl; eauto.
            - rewrite replace_preserves_Zlength. rewrite replace_preserves_Zlength. assumption.
            - rewrite replace_replace. intros.
              generalize (valid_ioapic_state H1 _ _ H18). intros [Hv Hex].
              split.
              + assumption.
              + assert (Hcase: (n = 4 n 4)%nat) by omega. destruct Hcase as [Hn | Hn].
                × subst. rewrite nth_error_replace_gss. true. reflexivity. repeat toZ. omega.
                × rewrite nth_error_replace_gso. assumption. apply not_eq_sym. assumption.
          }
          {
            constructor; simpl; eauto.
            - rewrite replace_preserves_Zlength. rewrite replace_preserves_Zlength. assumption.
            - rewrite replace_replace. intros.
              generalize (valid_ioapic_state H1 _ _ H18). intros [Hv Hex].
              split.
              + assumption.
              + assert (Hcase: (n = 4 n 4)%nat) by omega. destruct Hcase as [Hn | Hn].
                × subst. rewrite nth_error_replace_gss. true. reflexivity. repeat toZ. omega.
                × rewrite nth_error_replace_gso. assumption. apply not_eq_sym. assumption.
          }
        }
        {
          functional inversion Hex6; subst;
          clear Hex1 Hex3 Hex6.
          {
            constructor; simpl; eauto.
            - rewrite replace_preserves_Zlength. rewrite replace_preserves_Zlength. assumption.
            - rewrite replace_replace. intros.
              generalize (valid_ioapic_state H1 _ _ H18). intros [Hv Hex].
              split.
              + assumption.
              + assert (Hcase: (n = 4 n 4)%nat) by omega. destruct Hcase as [Hn | Hn].
                × subst. rewrite nth_error_replace_gss. true. reflexivity. repeat toZ. omega.
                × rewrite nth_error_replace_gso. assumption. apply not_eq_sym. assumption.
          }
          {
            constructor; simpl; eauto.
            - rewrite replace_preserves_Zlength. rewrite replace_preserves_Zlength. assumption.
            - rewrite replace_replace. intros.
              generalize (valid_ioapic_state H1 _ _ H18). intros [Hv Hex].
              split.
              + assumption.
              + assert (Hcase: (n = 4 n 4)%nat) by omega. destruct Hcase as [Hn | Hn].
                × subst. rewrite nth_error_replace_gss. true. reflexivity. repeat toZ. omega.
                × rewrite nth_error_replace_gso. assumption. apply not_eq_sym. assumption.
          }
        }
    Qed.

  End INV.

Layer Definition

serial device