(* * 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) *) theory InvariantsPre_AI imports LevityCatch_AI begin context begin interpretation Arch . requalify_types aa_type requalify_consts aa_type end (* FIXME: move *) declare ranI [intro] section "Locale Setup" locale pspace_update_eq = fixes f :: "'z::state_ext state \ 'c::state_ext state" assumes pspace: "kheap (f s) = kheap s" locale Arch_pspace_update_eq = pspace_update_eq + Arch locale arch_update_eq = fixes f :: "'z::state_ext state \ 'c::state_ext state" assumes arch: "arch_state (f s) = arch_state s" locale Arch_arch_update_eq = arch_update_eq + Arch locale arch_idle_update_eq_more = fixes f :: "'z::state_ext state \ 'c::state_ext state" assumes idle: "idle_thread (f s) = idle_thread s" assumes irq: "interrupt_irq_node (f s) = interrupt_irq_node s" locale arch_idle_update_eq = arch_update_eq + arch_idle_update_eq_more locale Arch_arch_idle_update_eq = Arch_arch_update_eq + arch_idle_update_eq_more locale p_arch_update_eq = pspace_update_eq + arch_update_eq locale Arch_p_arch_update_eq = Arch_pspace_update_eq + Arch_arch_update_eq locale p_arch_idle_update_eq = p_arch_update_eq + arch_idle_update_eq locale Arch_p_arch_idle_update_eq = Arch_p_arch_update_eq + Arch_arch_idle_update_eq locale irq_states_update_eq = fixes f :: "'z::state_ext state \ 'c::state_ext state" assumes int: "interrupt_states (f s) = interrupt_states s" locale pspace_int_update_eq = pspace_update_eq + irq_states_update_eq locale Arch_pspace_int_update_eq = Arch_pspace_update_eq + irq_states_update_eq locale p_arch_idle_update_int_eq = p_arch_idle_update_eq + pspace_int_update_eq locale Arch_p_arch_idle_update_int_eq = Arch_p_arch_idle_update_eq + Arch_pspace_int_update_eq section "Base definitions for Invariants" definition obj_at :: "(kernel_object \ bool) \ obj_ref \ 'z::state_ext state \ bool" where "obj_at P ref s \ \ko. kheap s ref = Some ko \ P ko" lemma obj_at_pspaceI: "\ obj_at P ref s; kheap s = kheap s' \ \ obj_at P ref s'" by (simp add: obj_at_def) abbreviation "ko_at k \ obj_at (op = k)" lemma obj_atE: "\ obj_at P p s; \ko. \ kheap s p = Some ko; P ko \ \ R \ \ R" by (auto simp: obj_at_def) lemma obj_at_weakenE: "\ obj_at P r s; \ko. P ko \ P' ko \ \ obj_at P' r s" by (clarsimp simp: obj_at_def) lemma ko_at_weakenE: "\ ko_at k ptr s; P k \ \ obj_at P ptr s" by (erule obj_at_weakenE, simp) text {* An alternative formulation that allows abstraction over type: *} datatype a_type = ATCB | AEndpoint | ANTFN | ACapTable nat | AGarbage | AArch aa_type definition a_type :: "kernel_object \ a_type" where "a_type ob \ case ob of CNode sz cspace \ if well_formed_cnode_n sz cspace then ACapTable sz else AGarbage | TCB tcb \ ATCB | Endpoint endpoint \ AEndpoint | Notification notification \ ANTFN | ArchObj ao \ AArch (aa_type ao)" lemmas a_type_simps = a_type_def[split_simps kernel_object.split] lemma a_type_aa_type: "(a_type (ArchObj ako) = AArch T) = (aa_type ako = T)" by (simp add: a_type_def) abbreviation "typ_at T \ obj_at (\ob. a_type ob = T)" definition pspace_aligned :: "'z::state_ext state \ bool" where "pspace_aligned s \ \x \ dom (kheap s). is_aligned x (obj_bits (the (kheap s x)))" lemma pspace_alignedD [intro?]: "\ kheap s p = Some ko; pspace_aligned s \ \ is_aligned p (obj_bits ko)" unfolding pspace_aligned_def by (drule bspec, blast, simp) text "objects don't overlap" definition pspace_distinct :: "'z::state_ext state \ bool" where "pspace_distinct \ \s. \x y ko ko'. kheap s x = Some ko \ kheap s y = Some ko' \ x \ y \ {x .. x + (2 ^ obj_bits ko - 1)} \ {y .. y + (2 ^ obj_bits ko' - 1)} = {}" definition caps_of_state :: "'z::state_ext state \ cslot_ptr \ cap option" where "caps_of_state s \ (\p. if (\cap. fst (get_cap p s) = {(cap, s)}) then Some (THE cap. fst (get_cap p s) = {(cap, s)}) else None)" definition "arch_cap_fun_lift P F c \ case c of ArchObjectCap ac \ P ac | _ \ F" lemmas arch_cap_fun_lift_simps[simp] = arch_cap_fun_lift_def[split_simps cap.split] definition "arch_obj_fun_lift P F c \ case c of ArchObj ac \ P ac | _ \ F" lemmas arch_obj_fun_lift_simps[simp] = arch_obj_fun_lift_def[split_simps kernel_object.split] lemma arch_obj_fun_lift_in_empty[dest!]: "x \ arch_obj_fun_lift f {} ko \ \ako. ko = ArchObj ako \ x \ f ako" by (cases ko; simp add: arch_obj_fun_lift_def) lemma arch_obj_fun_lift_Some[dest!]: "arch_obj_fun_lift f None ko = Some x \ \ako. ko = ArchObj ako \ f ako = Some x" by (cases ko; simp add: arch_obj_fun_lift_def) lemma arch_obj_fun_lift_True[dest!]: "arch_obj_fun_lift f False ko \ \ako. ko = ArchObj ako \ f ako" by (cases ko; simp add: arch_obj_fun_lift_def) lemma arch_cap_fun_lift_in_empty[dest!]: "x \ arch_cap_fun_lift f {} cap \ \acap. cap = ArchObjectCap acap \ x \ f acap" by (cases cap; simp add: arch_cap_fun_lift_def) lemma arch_cap_fun_lift_Some[dest!]: "arch_cap_fun_lift f None cap = Some x \ \acap. cap = ArchObjectCap acap \ f acap = Some x" by (cases cap; simp add: arch_cap_fun_lift_def) lemma arch_cap_fun_lift_True[dest!]: "arch_cap_fun_lift f False cap \ \acap. cap = ArchObjectCap acap \ f acap" by (cases cap; simp add: arch_cap_fun_lift_def) lemma arch_obj_fun_lift_non_arch[simp]: "\ako. ko \ ArchObj ako \ arch_obj_fun_lift f F ko = F" by (cases ko; fastforce) lemma arch_cap_fun_lift_non_arch[simp]: "\ako. cap \ ArchObjectCap ako \ arch_cap_fun_lift f F cap = F" by (cases cap; fastforce) end