Library mcertikos.devdrivers.DHandlerOp
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 INVLemmaInterrupt.
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 ObjInterruptManagement.
Require Export ObjX86.
Require Export ObjQLock.
Require Export ObjMultiprocessor.
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 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 ObjInterruptDriver.
Require Export ObjCPU.
Require Export ObjMM.
Require Export ObjFlatMem.
Require Export ObjSerialDevice.
Require Export ObjSerialDriver.
Require Export ObjInterruptController.
Require Export ObjConsole.
Require Export ObjInterruptManagement.
Require Export ObjX86.
Require Export ObjQLock.
Require Export ObjMultiprocessor.
Require Import DeviceStateDataType.
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_cons_buf_rpos: 0 ≤ rpos (console abd) < CONSOLE_BUFFER_SIZE;
valid_cons_buf_length: 0 ≤ Zlength (cons_buf (console abd)) < CONSOLE_BUFFER_SIZE;
valid_ioapic_max_intr: 23 ≤ IoApicMaxIntr (s (ioapic abd)) < 32;
valid_ioapic_enables: 23 ≤ Zlength (IoApicEnables (s (ioapic abd))) < 32;
valid_ioapic_state:
init abd = true →
∀ n v, nth_error (IoApicIrqs (s (ioapic abd))) n = value v →
( v = Z.of_nat n + IRQ0
∧ ∃ b, nth_error (IoApicEnables (s (ioapic abd))) n = value b);
valid_intr_state: in_intr abd = false →
(LapicEoi (s (lapic abd)) = NoIntr ∧
LapicEsrClrPen (s (lapic abd)) = false ∧
CurrentIrq (s (ioapic abd)) = None ∧
curr_intr_num abd = 256)
}.
mkInvariant {
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_ioapic_max_intr: 23 ≤ IoApicMaxIntr (s (ioapic abd)) < 32;
valid_ioapic_enables: 23 ≤ Zlength (IoApicEnables (s (ioapic abd))) < 32;
valid_ioapic_state:
init abd = true →
∀ n v, nth_error (IoApicIrqs (s (ioapic abd))) n = value v →
( v = Z.of_nat n + IRQ0
∧ ∃ b, nth_error (IoApicEnables (s (ioapic abd))) n = value b);
valid_intr_state: in_intr abd = false →
(LapicEoi (s (lapic abd)) = NoIntr ∧
LapicEsrClrPen (s (lapic abd)) = false ∧
CurrentIrq (s (ioapic abd)) = None ∧
curr_intr_num abd = 256)
}.
Global Instance dhandlerop_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.
vm_compute.
split.
intro contra; inv contra.
reflexivity.
Qed.
Lemma empty_data_high_level_invariant:
high_level_invariant (init_adt multi_oracle_init1).
Proof.
constructor; simpl; intros; auto; try inv H.
vm_compute.
split.
intro contra; inv contra.
reflexivity.
Qed.
Global Instance dhandlerop_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}.
DhandlerOp 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.
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 ticket_lock_init0_inv: PreservesInvariants ticket_lock_init0_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto.
{
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.
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_max_intr0.
repeat match goal with [H: _ |- _] ⇒ clear H end.
induction (Z.to_nat (IoApicMaxIntr (s (ioapic d)) + 1 - 1)).
simpl.
intros.
omega.
Opaque Z.of_nat.
intros.
replace (Z.of_nat (S n) + 1 - 1) with (Z.of_nat (S n)) by omega.
rewrite Nat2Z.id.
simpl.
replace (Z.of_nat n + 1 - 1) with (Z.of_nat n) in IHn by omega.
rewrite Nat2Z.id in IHn.
eapply IHn; eauto; xomega.
}
{
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_enables0.
replace (IoApicMaxIntr (s (ioapic d))) with (Z.of_nat (Z.to_nat (IoApicMaxIntr (s (ioapic d)) + 1 - 1))) by (rewrite Z2Nat.id; omega).
repeat match goal with [H: _ |- _] ⇒ clear H end.
induction (Z.to_nat (IoApicMaxIntr (s (ioapic d)) + 1 - 1)).
Transparent Z.of_nat.
simpl.
intros.
rewrite replace_preserves_Zlength.
assumption.
Opaque Z.of_nat.
intros.
replace (Z.of_nat (S n) + 1 - 1) with (Z.of_nat (S n)) by omega.
rewrite Nat2Z.id.
simpl.
replace (Z.of_nat n + 1 - 1) with (Z.of_nat n) in IHn by omega.
rewrite Nat2Z.id in IHn.
rewrite replace_preserves_Zlength.
eapply IHn; eauto; xomega.
}
{
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.
assert(0 ≤ Z.of_nat n ≤ Z.of_nat (Z.to_nat (IoApicMaxIntr (s (ioapic d)) + 1 - 1))).
{
split.
xomega.
eapply nth_error_range in H8.
rewrite Z2Nat.id by xomega.
assert (Z.of_nat (length
(IoApicIrqs
(ioapic_init_aux (s (ioapic d))
(Z.to_nat (IoApicMaxIntr (s (ioapic d)) + 1 - 1))))) = IoApicMaxIntr (s (ioapic d)) + 1).
{
rewrite <- ioapic_init_aux_preserves_ioapicirq_length.
rewrite _x.
rewrite length_to_zlength.
reflexivity.
}
rewrite H9 in H8.
omega.
}
generalize H8, H9.
replace (IoApicMaxIntr (s (ioapic d))) with (Z.of_nat (Z.to_nat (IoApicMaxIntr (s (ioapic d)) + 1 - 1))) by (rewrite Z2Nat.id; omega).
generalize _x _x0.
repeat match goal with [H: _ |- _] ⇒ clear H end.
induction (Z.to_nat (IoApicMaxIntr (s (ioapic d)) + 1 - 1)).
{
simpl.
intros.
assert(Z.of_nat n = 0) by xomega.
assert(n = O).
{
rewrite <- Nat2Z.id at 1.
rewrite H.
reflexivity.
}
rewrite H0 in H8.
simpl in ×.
subdestruct.
inv H8.
replace (Z.to_nat (Z.of_nat 0 + 1 - 1)) with O in × by xomega.
simpl in ×.
destruct (IoApicIrqs (s (ioapic d))).
simpl in ×.
inv Hdestruct.
simpl in ×.
inv Hdestruct.
split.
reflexivity.
simpl.
destruct (IoApicEnables (s (ioapic d))); simpl in ×.
rewrite _x0 in _x.
change (Zlength nil) with 0 in _x.
generalize (Zlength_ge_1 z l); intro.
omega.
esplit; reflexivity.
}
{
Opaque Z.of_nat Z.add.
simpl.
intros.
simpl.
assert(ncase: n = S n0 ∨ n ≠ S n0) by omega.
case_eq ncase; intros.
{
subst.
replace (Z.of_nat (S n0) + 1 - 1) with (Z.of_nat (S n0)) in × by omega.
rewrite Z2Nat.id in × by omega.
rewrite Nat2Z.id in × by omega.
Opaque nth_error.
simpl in ×.
rewrite replace_gss.
rewrite replace_gss in H8.
split.
inv H8.
omega.
esplit; reflexivity.
rewrite <- ioapic_init_aux_preserves_ioapicirq_length.
rewrite length_to_zlength.
omega.
rewrite <- ioapic_init_aux_preserves_ioapicenables_length.
rewrite length_to_zlength.
omega.
}
{
Transparent nth_error.
replace (Z.of_nat (S n0) + 1 - 1) with (Z.of_nat (S n0)) in × by omega.
replace (Z.of_nat n0 + 1 - 1) with (Z.of_nat n0) in × by omega.
rewrite Z2Nat.id in × by omega.
rewrite Nat2Z.id in × by omega.
simpl.
assert(nth_error
(replace false (S n0)
(IoApicEnables (ioapic_init_aux (s (ioapic d)) n0))) n = nth_error (IoApicEnables (ioapic_init_aux (s (ioapic d)) n0)) n).
{
eapply ntheq; eauto.
}
rewrite H0.
eapply IHn0; eauto.
unfold disable_irq in H8.
Opaque Z.add.
simpl in H8.
erewrite ntheq in H8; eauto.
xomega.
xomega.
xomega.
}
}
}
{
repeat (split; try reflexivity).
generalize (valid_intr_state0 H); intro tmp.
destruct tmp.
destruct H9.
destruct H10; assumption.
}
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 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.
Global Instance cli_inv: PreservesInvariants cli_spec.
Proof.
preserves_invariants_direct low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance sti_inv: PreservesInvariants sti_spec.
Proof.
preserves_invariants_direct low_level_invariant high_level_invariant; eauto 2.
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 serial_putc_inv:
PreservesInvariants serial_putc_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
Qed.
Global Instance serial_intr_disable_concrete_aux_inv: PreservesInvariants serial_intr_disable_concrete_aux_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; 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.
}
{
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.
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_state0 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.
omega.
rewrite nth_error_replace_gso.
eapply valid_ioapic_state0; eauto.
change (Pos.to_nat 4%positive) with 4%nat.
omega.
intros;
simpl in ×.
rewrite replace_replace.
generalize (valid_ioapic_state0 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.
omega.
rewrite nth_error_replace_gso.
eapply valid_ioapic_state0; eauto.
change (Pos.to_nat 4%positive) with 4%nat.
omega.
intros;
simpl in ×.
rewrite replace_replace.
generalize (valid_ioapic_state0 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.
omega.
rewrite nth_error_replace_gso.
eapply valid_ioapic_state0; eauto.
change (Pos.to_nat 4%positive) with 4%nat.
omega.
intros;
simpl in ×.
rewrite replace_replace.
generalize (valid_ioapic_state0 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.
omega.
rewrite nth_error_replace_gso.
eapply valid_ioapic_state0; 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_state0 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.
omega.
rewrite nth_error_replace_gso.
eapply valid_ioapic_state0; eauto.
change (Pos.to_nat 4%positive) with 4%nat.
omega.
intros;
simpl in ×.
rewrite replace_replace.
generalize (valid_ioapic_state0 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.
omega.
rewrite nth_error_replace_gso.
eapply valid_ioapic_state0; eauto.
change (Pos.to_nat 4%positive) with 4%nat.
omega.
intros;
simpl in ×.
rewrite replace_replace.
generalize (valid_ioapic_state0 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.
omega.
rewrite nth_error_replace_gso.
eapply valid_ioapic_state0; eauto.
change (Pos.to_nat 4%positive) with 4%nat.
omega.
intros;
simpl in ×.
rewrite replace_replace.
generalize (valid_ioapic_state0 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.
omega.
rewrite nth_error_replace_gso.
eapply valid_ioapic_state0; 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.
}
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_aux_inv: PreservesInvariants serial_intr_enable_concrete_aux_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.
}
{
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_state0 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.
omega.
rewrite nth_error_replace_gso.
eapply valid_ioapic_state0; eauto.
change (Pos.to_nat 4%positive) with 4%nat.
omega.
intros;
simpl in ×.
rewrite replace_replace.
generalize (valid_ioapic_state0 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.
omega.
rewrite nth_error_replace_gso.
eapply valid_ioapic_state0; eauto.
change (Pos.to_nat 4%positive) with 4%nat.
omega.
intros;
simpl in ×.
rewrite replace_replace.
generalize (valid_ioapic_state0 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.
omega.
rewrite nth_error_replace_gso.
eapply valid_ioapic_state0; eauto.
change (Pos.to_nat 4%positive) with 4%nat.
omega.
intros;
simpl in ×.
rewrite replace_replace.
generalize (valid_ioapic_state0 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.
omega.
rewrite nth_error_replace_gso.
eapply valid_ioapic_state0; 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_state0 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.
omega.
rewrite nth_error_replace_gso.
eapply valid_ioapic_state0; eauto.
change (Pos.to_nat 4%positive) with 4%nat.
omega.
intros;
simpl in ×.
rewrite replace_replace.
generalize (valid_ioapic_state0 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.
omega.
rewrite nth_error_replace_gso.
eapply valid_ioapic_state0; eauto.
change (Pos.to_nat 4%positive) with 4%nat.
omega.
intros;
simpl in ×.
rewrite replace_replace.
generalize (valid_ioapic_state0 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.
omega.
rewrite nth_error_replace_gso.
eapply valid_ioapic_state0; eauto.
change (Pos.to_nat 4%positive) with 4%nat.
omega.
intros;
simpl in ×.
rewrite replace_replace.
generalize (valid_ioapic_state0 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.
omega.
rewrite nth_error_replace_gso.
eapply valid_ioapic_state0; 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.
}
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.
Global Instance ioapic_mask_inv:
PreservesInvariants ioapic_mask_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
rewrite replace_preserves_Zlength.
assumption.
eapply valid_ioapic_state0 in H5; eauto.
destruct H5.
destruct H6.
assert (ncase: (Z.to_nat (Int.unsigned i)) = n ∨ (Z.to_nat (Int.unsigned i)) ≠ n) by omega.
case_eq ncase; intros.
rewrite e.
split.
assumption.
rewrite replace_gss.
esplit; reflexivity.
apply (nth_error_range _ _ _ _ H6).
split.
assumption.
rewrite nth_error_replace_gso by assumption.
esplit; eassumption.
Qed.
Global Instance ioapic_unmask_inv:
PreservesInvariants ioapic_unmask_spec.
Proof.
preserves_invariants_simpl low_level_invariant high_level_invariant; eauto 2.
rewrite replace_preserves_Zlength.
assumption.
eapply valid_ioapic_state0 in H5; eauto.
destruct H5.
destruct H6.
assert (ncase: (Z.to_nat (Int.unsigned i)) = n ∨ (Z.to_nat (Int.unsigned i)) ≠ n) by omega.
case_eq ncase; intros.
rewrite e.
split.
assumption.
rewrite replace_gss.
esplit; reflexivity.
apply (nth_error_range _ _ _ _ H6).
split.
assumption.
rewrite nth_error_replace_gso by assumption.
esplit; eassumption.
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.
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 dhandlerop_fresh : compatlayer (cdata RData) :=
serial_intr_disable_handler ↦ gensem serial_intr_disable_concrete_aux_spec
⊕ serial_intr_enable_handler ↦ gensem serial_intr_enable_concrete_aux_spec.
Definition dhandlerop_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
⊕ ioapic_mask ↦ gensem ioapic_mask_spec
⊕ ioapic_unmask ↦ gensem ioapic_unmask_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; exec_store := @exec_storeex |}.
Definition dhandlerop : compatlayer (cdata RData) :=
dhandlerop_fresh ⊕ dhandlerop_passthrough.
End WITHMEM.
serial_intr_disable_handler ↦ gensem serial_intr_disable_concrete_aux_spec
⊕ serial_intr_enable_handler ↦ gensem serial_intr_enable_concrete_aux_spec.
Definition dhandlerop_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
⊕ ioapic_mask ↦ gensem ioapic_mask_spec
⊕ ioapic_unmask ↦ gensem ioapic_unmask_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; exec_store := @exec_storeex |}.
Definition dhandlerop : compatlayer (cdata RData) :=
dhandlerop_fresh ⊕ dhandlerop_passthrough.
End WITHMEM.