(* * Copyright 2014, General Dynamics C4 Systems * * This software may be distributed and modified according to the terms of * the GNU General Public License version 2. Note that NO WARRANTY is provided. * See "LICENSE_GPLv2.txt" for details. * * @TAG(GD_GPL) *) (* Properties of machine operations. *) theory Machine_R imports Bits_R begin definition "irq_state_independent_H (P :: kernel_state \ bool)\ \(f :: nat \ nat) (s :: kernel_state). P s \ P (s\ksMachineState := ksMachineState s \irq_state := f (irq_state (ksMachineState s))\\)" lemma irq_state_independent_HI[intro!, simp]: "\\s f. P (s\ksMachineState := ksMachineState s \irq_state := f (irq_state (ksMachineState s))\\) = P s\ \ irq_state_independent_H P" by (simp add: irq_state_independent_H_def) lemma irq_state_independent_H_conjI[intro!]: "\irq_state_independent_H P; irq_state_independent_H Q\ \ irq_state_independent_H (P and Q)" "\irq_state_independent_H P; irq_state_independent_H Q\ \ irq_state_independent_H (\s. P s \ Q s)" by (simp add: irq_state_independent_H_def)+ lemma irq_state_independent_H_disjI[intro]: "\irq_state_independent_H P; irq_state_independent_H Q\ \ irq_state_independent_H (P or Q)" "\irq_state_independent_H P; irq_state_independent_H Q\ \ irq_state_independent_H (\s. P s \ Q s)" by (simp add: irq_state_independent_H_def)+ lemma dmo_getirq_inv[wp]: "irq_state_independent_H P \ \P\ doMachineOp getActiveIRQ \\rv. P\" apply (simp add: getActiveIRQ_def doMachineOp_def split_def exec_gets select_f_select[simplified liftM_def] select_modify_comm gets_machine_state_modify) apply wp apply (clarsimp simp: irq_state_independent_H_def in_monad return_def split: if_splits) done lemma getActiveIRQ_masked: "\\s. valid_irq_masks' table (irq_masks s)\ getActiveIRQ \\rv s. \irq. rv = Some irq \ table irq \ IRQInactive\" apply (simp add: getActiveIRQ_def) apply wp apply (clarsimp simp: valid_irq_masks'_def) done lemma dmo_maskInterrupt: "\\s. P (ksMachineState_update (irq_masks_update (\t. t (irq := m))) s)\ doMachineOp (maskInterrupt m irq) \\_. P\" apply (simp add: doMachineOp_def split_def) apply wp apply (clarsimp simp: maskInterrupt_def in_monad) apply (erule rsubst [where P=P]) apply simp done lemma dmo_maskInterrupt_True: "\invs'\ doMachineOp (maskInterrupt True irq) \\_. invs'\" apply (wp dmo_maskInterrupt) apply (clarsimp simp: invs'_def valid_state'_def) apply (simp add: valid_irq_masks'_def valid_machine_state'_def ct_not_inQ_def ct_idle_or_in_cur_domain'_def tcb_in_cur_domain'_def) done lemma setIRQState_irq_states': "\valid_irq_states'\ setIRQState state irq \\rv. valid_irq_states'\" apply (simp add: setIRQState_def setInterruptState_def getInterruptState_def) apply (wp dmo_maskInterrupt) apply (simp add: valid_irq_masks'_def) done lemma getActiveIRQ_le_maxIRQ: "\irqs_masked' and valid_irq_states'\ doMachineOp getActiveIRQ \\rv s. \x. rv = Some x \ x \ maxIRQ\" apply (simp add: doMachineOp_def split_def) apply wp apply clarsimp apply (drule use_valid, rule getActiveIRQ_le_maxIRQ') prefer 2 apply simp apply (simp add: irqs_masked'_def valid_irq_states'_def) done (* FIXME: follows already from getActiveIRQ_le_maxIRQ *) lemma getActiveIRQ_neq_Some0xFF: "\\\ doMachineOp getActiveIRQ \\rv s. rv \ Some 0x3FF\" apply (simp add: doMachineOp_def split_def) apply wp apply clarsimp apply (drule use_valid, rule getActiveIRQ_neq_Some0xFF') apply auto done end