Library mcertikos.devdrivers.DSerial
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 XOmega.
Require Import liblayers.logic.PTreeModules.
Require Import liblayers.logic.LayerLogicImpl.
Require Import liblayers.compat.CompatLayers.
Require Import liblayers.compat.CompatGenSem.
Require Import INVLemmaDevice.
Require Import INVLemmaDriver.
Require Export ObjCPU.
Require Export ObjMM.
Require Export ObjFlatMem.
Require Export ObjSerialDevice.
Require Export ObjSerialDriver.
Require Export ObjInterruptController.
Require Export ObjInterruptDriver.
Require Export ObjConsole.
Require Export ObjX86.
Require Export ObjQLock.
Require Export ObjMultiprocessor.
Require Import AbstractDataType.
Require Import DeviceStateDataType.
Require Import AbstractDataType.
Require Import FutureTactic.
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 XOmega.
Require Import liblayers.logic.PTreeModules.
Require Import liblayers.logic.LayerLogicImpl.
Require Import liblayers.compat.CompatLayers.
Require Import liblayers.compat.CompatGenSem.
Require Import INVLemmaDevice.
Require Import INVLemmaDriver.
Require Export ObjCPU.
Require Export ObjMM.
Require Export ObjFlatMem.
Require Export ObjSerialDevice.
Require Export ObjSerialDriver.
Require Export ObjInterruptController.
Require Export ObjInterruptDriver.
Require Export ObjConsole.
Require Export ObjX86.
Require Export ObjQLock.
Require Export ObjMultiprocessor.
Require Import AbstractDataType.
Require Import DeviceStateDataType.
Require Import AbstractDataType.
Require Import FutureTactic.
Section WITHMEM.
Local Open Scope Z_scope.
Context `{real_params: RealParams}.
Context `{oracle_prop: MultiOracleProp}.
Local Open Scope Z_scope.
Context `{real_params: RealParams}.
Context `{oracle_prop: MultiOracleProp}.
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;
valid_ioapic_enables: Zlength (IoApicEnables (s (ioapic abd))) = IoApicMaxIntr (s (ioapic abd)) + 1;
valid_ioapic_irqs: Zlength (IoApicIrqs (s (ioapic abd))) = IoApicMaxIntr (s (ioapic abd)) + 1;
valid_ioapic_state_range: ∀ n, 0 ≤ nth n (IoApicIrqs (s (ioapic abd))) 0 ≤ 255
}.
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;
valid_ioapic_enables: Zlength (IoApicEnables (s (ioapic abd))) = IoApicMaxIntr (s (ioapic abd)) + 1;
valid_ioapic_irqs: Zlength (IoApicIrqs (s (ioapic abd))) = IoApicMaxIntr (s (ioapic abd)) + 1;
valid_ioapic_state_range: ∀ n, 0 ≤ nth n (IoApicIrqs (s (ioapic abd))) 0 ≤ 255
}.
Global Instance dserial_data_ops : CompatDataOps RData :=
{
empty_data := (init_adt multi_oracle_init1);
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 multi_oracle_init1);
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 multi_oracle_init1)).
Proof.
constructor; simpl; intros; auto; try inv H.
omega.
generalize max_unsigned_val; intro muval.
repeat (destruct n; [omega|]); omega.
Qed.
Lemma empty_data_high_level_invariant:
high_level_invariant ((init_adt multi_oracle_init1)).
Proof.
constructor; simpl; intros; auto; try inv H.
omega.
generalize max_unsigned_val; intro muval.
repeat (destruct n; [omega|]); omega.
Qed.
Global Instance dserial_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.MemoryModelX}.
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.MemoryModelX}.
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 page_copy_inv: PreservesInvariants page_copy´_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto.
Qed.
Global Instance page_copy_back_inv: PreservesInvariants page_copy_back´_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 ticket_lock_init0_inv: PreservesInvariants ticket_lock_init0_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
{
assert (l1: IoApicMaxIntr (s (ioapic d)) < Zlength (IoApicIrqs (s (ioapic d)))) by omega.
assert (l2: IoApicMaxIntr (s (ioapic d)) < Zlength (IoApicEnables (s (ioapic d)))) by omega.
generalize l1 l2.
generalize valid_ioapic_max_intr0.
replace (IoApicMaxIntr (s (ioapic d))) with (Z.of_nat (Z.to_nat (IoApicMaxIntr (s (ioapic d)) + 1 - 1))) by (rewrite Z2Nat.id; omega).
revert valid_ioapic_state_range0.
repeat match goal with [H: _ |- _] ⇒ clear H end.
generalize max_unsigned_val; intros muval.
induction (Z.to_nat (IoApicMaxIntr (s (ioapic d)) + 1 - 1)).
simpl.
intros.
assert(ncase: n = O ∨ n ≠ O) by omega.
case_eq ncase; intros.
rewrite e.
assert(lcase: (0 < length (IoApicIrqs (s (ioapic d))))%nat ∨ (0 ≥ length (IoApicIrqs (s (ioapic d))))%nat) by omega.
case_eq lcase; intros.
rewrite nth_replace_same.
omega.
assumption.
rewrite nth_overflow.
omega.
rewrite <- replace_preserves_length.
omega.
rewrite nth_replace_different.
eapply valid_ioapic_state_range0.
omega.
Opaque Z.of_nat.
intros.
replace (Z.of_nat (S n0) + 1 - 1) with (Z.of_nat (S n0)) by omega.
rewrite Nat2Z.id.
Opaque nth replace Z.of_nat Z.add.
simpl.
replace (Z.of_nat n0 + 1 - 1) with (Z.of_nat n0) in IHn0 by omega.
rewrite Nat2Z.id in IHn0.
assert(ncase: n = S n0 ∨ n ≠ S n0) by omega.
case_eq ncase; intros.
rewrite e.
assert(lcase: (S n0 < length (IoApicIrqs (ioapic_init_aux (s (ioapic d)) n0)))%nat ∨ (S n0 ≥ length (IoApicIrqs (ioapic_init_aux (s (ioapic d)) n0)))%nat) by omega.
case_eq lcase; intros.
rewrite nth_replace_same.
omega.
assumption.
rewrite nth_overflow.
omega.
rewrite <- replace_preserves_length.
omega.
rewrite nth_replace_different.
eapply IHn0; eauto; xomega.
xomega.
}
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.
rewrite ioapic_trans_intr_reserves_irqs. auto.
Qed.
Global Instance ic_intr_inv: PreservesInvariants ic_intr_spec.
Proof.
preserves_invariants_nested low_level_invariant high_level_invariant; eauto.
rewrite ioapic_trans_intr_reserves_irqs. auto.
Qed.
Global Instance ioapic_read_inv: PreservesInvariants ioapic_read_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
Qed.
Global Instance ioapic_write_inv: PreservesInvariants ioapic_write_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
{
unfold ioapic_trans_cpu.
destruct (s (ioapic d)).
destruct (zeq (Int.unsigned i) 0).
simpl in ×.
eapply valid_ioapic_state_range0.
destruct (zle_le 16 (Int.unsigned i) (IoApicMaxIntr × 2 + 17)).
destruct (addr_to_idx (Int.unsigned i)).
simpl in ×.
assert(xcase: Z.to_nat x = n ∨ Z.to_nat x ≠ n) by omega.
case_eq xcase; intros.
rewrite e.
assert(ncase: (n < length IoApicIrqs)%nat ∨ (n ≥ length IoApicIrqs)%nat) by omega.
case_eq ncase; intros.
rewrite nth_replace_same.
unfold low32_to_irq.
xomega.
assumption.
rewrite nth_overflow.
omega.
rewrite <- replace_preserves_length.
omega.
rewrite nth_replace_different.
eapply valid_ioapic_state_range0.
omega.
simpl in *; eapply valid_ioapic_state_range0.
simpl in *; eapply valid_ioapic_state_range0.
}
{
generalize max_unsigned_val; intro muval.
unfold ioapic_trans_cpu.
destruct (s (ioapic d)).
destruct (zeq (Int.unsigned i) 0).
simpl in ×.
eapply valid_ioapic_state_range0.
destruct (zle_le 16 (Int.unsigned i) (IoApicMaxIntr × 2 + 17)).
destruct (addr_to_idx (Int.unsigned i)).
simpl in ×.
assert(xcase: Z.to_nat x = n ∨ Z.to_nat x ≠ n) by omega.
case_eq xcase; intros.
rewrite e.
assert(ncase: (n < length IoApicIrqs)%nat ∨ (n ≥ length IoApicIrqs)%nat) by omega.
case_eq ncase; intros.
rewrite nth_replace_same.
unfold low32_to_irq.
xomega.
assumption.
rewrite nth_overflow.
omega.
rewrite <- replace_preserves_length.
omega.
rewrite nth_replace_different.
eapply valid_ioapic_state_range0.
omega.
simpl in *; eapply valid_ioapic_state_range0.
simpl in *; eapply valid_ioapic_state_range0.
}
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 1.
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 local_irq_save_inv:
PreservesInvariants local_irq_save_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance local_irq_restore_inv:
PreservesInvariants local_irq_restore_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 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.
Global Instance set_curid_inv: PreservesInvariants set_curid_spec.
Proof.
preserves_invariants_simpl_auto; eauto 2.
Qed.
Global Instance set_curid_init_inv: PreservesInvariants set_curid_init_spec.
Proof.
preserves_invariants_simpl_auto; eauto 2.
Qed.
Global Instance acquire_lock_inv: AcquireInvariants acquire_lock_spec0.
Proof.
constructor; unfold acquire_lock_spec; intros; subdestruct;
inv H; inv H0; constructor; auto; simpl; intros.
Qed.
Global Instance release_lock_inv: ReleaseInvariants release_lock_spec0.
Proof.
constructor; unfold release_lock_spec; intros; subdestruct;
inv H; inv H0; constructor; auto; simpl; intros.
Qed.
Global Instance proc_create_postinit_inv:
PreservesInvariants proc_create_postinit_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
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 page_copy_inv: PreservesInvariants page_copy´_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto.
Qed.
Global Instance page_copy_back_inv: PreservesInvariants page_copy_back´_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 ticket_lock_init0_inv: PreservesInvariants ticket_lock_init0_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
{
assert (l1: IoApicMaxIntr (s (ioapic d)) < Zlength (IoApicIrqs (s (ioapic d)))) by omega.
assert (l2: IoApicMaxIntr (s (ioapic d)) < Zlength (IoApicEnables (s (ioapic d)))) by omega.
generalize l1 l2.
generalize valid_ioapic_max_intr0.
replace (IoApicMaxIntr (s (ioapic d))) with (Z.of_nat (Z.to_nat (IoApicMaxIntr (s (ioapic d)) + 1 - 1))) by (rewrite Z2Nat.id; omega).
revert valid_ioapic_state_range0.
repeat match goal with [H: _ |- _] ⇒ clear H end.
generalize max_unsigned_val; intros muval.
induction (Z.to_nat (IoApicMaxIntr (s (ioapic d)) + 1 - 1)).
simpl.
intros.
assert(ncase: n = O ∨ n ≠ O) by omega.
case_eq ncase; intros.
rewrite e.
assert(lcase: (0 < length (IoApicIrqs (s (ioapic d))))%nat ∨ (0 ≥ length (IoApicIrqs (s (ioapic d))))%nat) by omega.
case_eq lcase; intros.
rewrite nth_replace_same.
omega.
assumption.
rewrite nth_overflow.
omega.
rewrite <- replace_preserves_length.
omega.
rewrite nth_replace_different.
eapply valid_ioapic_state_range0.
omega.
Opaque Z.of_nat.
intros.
replace (Z.of_nat (S n0) + 1 - 1) with (Z.of_nat (S n0)) by omega.
rewrite Nat2Z.id.
Opaque nth replace Z.of_nat Z.add.
simpl.
replace (Z.of_nat n0 + 1 - 1) with (Z.of_nat n0) in IHn0 by omega.
rewrite Nat2Z.id in IHn0.
assert(ncase: n = S n0 ∨ n ≠ S n0) by omega.
case_eq ncase; intros.
rewrite e.
assert(lcase: (S n0 < length (IoApicIrqs (ioapic_init_aux (s (ioapic d)) n0)))%nat ∨ (S n0 ≥ length (IoApicIrqs (ioapic_init_aux (s (ioapic d)) n0)))%nat) by omega.
case_eq lcase; intros.
rewrite nth_replace_same.
omega.
assumption.
rewrite nth_overflow.
omega.
rewrite <- replace_preserves_length.
omega.
rewrite nth_replace_different.
eapply IHn0; eauto; xomega.
xomega.
}
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.
rewrite ioapic_trans_intr_reserves_irqs. auto.
Qed.
Global Instance ic_intr_inv: PreservesInvariants ic_intr_spec.
Proof.
preserves_invariants_nested low_level_invariant high_level_invariant; eauto.
rewrite ioapic_trans_intr_reserves_irqs. auto.
Qed.
Global Instance ioapic_read_inv: PreservesInvariants ioapic_read_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
Qed.
Global Instance ioapic_write_inv: PreservesInvariants ioapic_write_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; auto.
{
unfold ioapic_trans_cpu.
destruct (s (ioapic d)).
destruct (zeq (Int.unsigned i) 0).
simpl in ×.
eapply valid_ioapic_state_range0.
destruct (zle_le 16 (Int.unsigned i) (IoApicMaxIntr × 2 + 17)).
destruct (addr_to_idx (Int.unsigned i)).
simpl in ×.
assert(xcase: Z.to_nat x = n ∨ Z.to_nat x ≠ n) by omega.
case_eq xcase; intros.
rewrite e.
assert(ncase: (n < length IoApicIrqs)%nat ∨ (n ≥ length IoApicIrqs)%nat) by omega.
case_eq ncase; intros.
rewrite nth_replace_same.
unfold low32_to_irq.
xomega.
assumption.
rewrite nth_overflow.
omega.
rewrite <- replace_preserves_length.
omega.
rewrite nth_replace_different.
eapply valid_ioapic_state_range0.
omega.
simpl in *; eapply valid_ioapic_state_range0.
simpl in *; eapply valid_ioapic_state_range0.
}
{
generalize max_unsigned_val; intro muval.
unfold ioapic_trans_cpu.
destruct (s (ioapic d)).
destruct (zeq (Int.unsigned i) 0).
simpl in ×.
eapply valid_ioapic_state_range0.
destruct (zle_le 16 (Int.unsigned i) (IoApicMaxIntr × 2 + 17)).
destruct (addr_to_idx (Int.unsigned i)).
simpl in ×.
assert(xcase: Z.to_nat x = n ∨ Z.to_nat x ≠ n) by omega.
case_eq xcase; intros.
rewrite e.
assert(ncase: (n < length IoApicIrqs)%nat ∨ (n ≥ length IoApicIrqs)%nat) by omega.
case_eq ncase; intros.
rewrite nth_replace_same.
unfold low32_to_irq.
xomega.
assumption.
rewrite nth_overflow.
omega.
rewrite <- replace_preserves_length.
omega.
rewrite nth_replace_different.
eapply valid_ioapic_state_range0.
omega.
simpl in *; eapply valid_ioapic_state_range0.
simpl in *; eapply valid_ioapic_state_range0.
}
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 1.
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 local_irq_save_inv:
PreservesInvariants local_irq_save_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance local_irq_restore_inv:
PreservesInvariants local_irq_restore_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 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.
Global Instance set_curid_inv: PreservesInvariants set_curid_spec.
Proof.
preserves_invariants_simpl_auto; eauto 2.
Qed.
Global Instance set_curid_init_inv: PreservesInvariants set_curid_init_spec.
Proof.
preserves_invariants_simpl_auto; eauto 2.
Qed.
Global Instance acquire_lock_inv: AcquireInvariants acquire_lock_spec0.
Proof.
constructor; unfold acquire_lock_spec; intros; subdestruct;
inv H; inv H0; constructor; auto; simpl; intros.
Qed.
Global Instance release_lock_inv: ReleaseInvariants release_lock_spec0.
Proof.
constructor; unfold release_lock_spec; intros; subdestruct;
inv H; inv H0; constructor; auto; simpl; intros.
Qed.
Global Instance proc_create_postinit_inv:
PreservesInvariants proc_create_postinit_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
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 dserial_fresh : compatlayer (cdata RData) :=
serial_init ↦ gensem serial_init_spec
⊕ serial_putc ↦ gensem serial_putc_spec
⊕ serial_getc ↦ gensem serial_getc_spec.
Definition dserial_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
⊕ 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
⊕ local_irq_save ↦ gensem local_irq_save_spec
⊕ local_irq_restore ↦ gensem local_irq_restore_spec
⊕ serial_hw_intr ↦ gensem serial_hw_intr_spec
serial_init ↦ gensem serial_init_spec
⊕ serial_putc ↦ gensem serial_putc_spec
⊕ serial_getc ↦ gensem serial_getc_spec.
Definition dserial_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
⊕ 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
⊕ local_irq_save ↦ gensem local_irq_save_spec
⊕ local_irq_restore ↦ gensem local_irq_restore_spec
⊕ serial_hw_intr ↦ gensem serial_hw_intr_spec
serial device
ioapic device
ioapic device
lapic device
lapic device
⊕ cons_buf_init ↦ gensem cons_buf_init_spec
⊕ cons_buf_write ↦ gensem cons_buf_write_spec
⊕ cons_buf_read ↦ gensem cons_buf_read_spec
⊕ cons_buf_wpos ↦ gensem cons_buf_wpos_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; exec_store := @exec_storeex |}.
Definition dserial : compatlayer (cdata RData) :=
dserial_fresh ⊕ dserial_passthrough.
End WITHMEM.
⊕ cons_buf_write ↦ gensem cons_buf_write_spec
⊕ cons_buf_read ↦ gensem cons_buf_read_spec
⊕ cons_buf_wpos ↦ gensem cons_buf_wpos_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; exec_store := @exec_storeex |}.
Definition dserial : compatlayer (cdata RData) :=
dserial_fresh ⊕ dserial_passthrough.
End WITHMEM.