(* * 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) *) (* The refinement relation between abstract and concrete states *) theory StateRelation imports Invariants_H begin definition cte_map :: "cslot_ptr \ word32" where "cte_map \ \(oref, cref). oref + (of_bl cref * 16)" definition lookup_failure_map :: "ExceptionTypes_A.lookup_failure \ Fault_H.lookup_failure" where "lookup_failure_map \ \lf. case lf of ExceptionTypes_A.InvalidRoot \ Fault_H.InvalidRoot | ExceptionTypes_A.MissingCapability n \ Fault_H.MissingCapability n | ExceptionTypes_A.DepthMismatch n m \ Fault_H.DepthMismatch n m | ExceptionTypes_A.GuardMismatch n g \ Fault_H.GuardMismatch n (of_bl g) (length g)" primrec fault_map :: "ExceptionTypes_A.fault \ Fault_H.fault" where "fault_map (ExceptionTypes_A.CapFault ref bool failure) = Fault_H.CapFault ref bool (lookup_failure_map failure)" | "fault_map (ExceptionTypes_A.VMFault ptr bool) = Fault_H.VMFault ptr bool" | "fault_map (ExceptionTypes_A.UnknownSyscallException n) = Fault_H.UnknownSyscallException n" | "fault_map (ExceptionTypes_A.UserException x y) = Fault_H.UserException x y" text {* A pspace and a tree are related if every object in the pspace corresponds to an object in the tree. Some abstract objects like CapTables correspond to multiple concrete ones, thus we have to make cuts. *} type_synonym obj_relation_cut = "Structures_A.kernel_object \ Structures_H.kernel_object \ bool" type_synonym obj_relation_cuts = "(word32 \ obj_relation_cut) set" definition vmrights_map :: "rights set \ vmrights" where "vmrights_map S \ if AllowRead \ S then (if AllowWrite \ S then VMReadWrite else VMReadOnly) else VMKernelOnly" definition zbits_map :: "nat option \ zombie_type" where "zbits_map N \ case N of Some n \ ZombieCNode n | None \ ZombieTCB" primrec acap_relation :: "arch_cap \ arch_capability \ bool" where "acap_relation (arch_cap.ASIDPoolCap x y) c = (c = arch_capability.ASIDPoolCap x y)" | "acap_relation (arch_cap.ASIDControlCap) c = (c = arch_capability.ASIDControlCap)" | "acap_relation (arch_cap.PageCap dev word rghts sz data) c = (c = arch_capability.PageCap dev word (vmrights_map rghts) sz data)" | "acap_relation (arch_cap.PageTableCap word data) c = (c = arch_capability.PageTableCap word data)" | "acap_relation (arch_cap.PageDirectoryCap word data) c = (c = arch_capability.PageDirectoryCap word data)" primrec cap_relation :: "cap \ capability \ bool" where "cap_relation Structures_A.NullCap c = (c = Structures_H.NullCap)" | "cap_relation Structures_A.DomainCap c = (c = Structures_H.DomainCap)" | "cap_relation (Structures_A.UntypedCap dev ref n f) c = (c = Structures_H.UntypedCap dev ref n f)" | "cap_relation (Structures_A.EndpointCap ref b r) c = (c = Structures_H.EndpointCap ref b (AllowSend \ r) (AllowRecv \ r) (AllowGrant \ r))" | "cap_relation (Structures_A.NotificationCap ref b r) c = (c = Structures_H.NotificationCap ref b (AllowSend \ r) (AllowRecv \ r))" | "cap_relation (Structures_A.CNodeCap ref n L) c = (c = Structures_H.CNodeCap ref n (of_bl L) (length L))" | "cap_relation (Structures_A.ThreadCap ref) c = (c = Structures_H.ThreadCap ref)" | "cap_relation (Structures_A.ReplyCap ref master) c = (c = Structures_H.ReplyCap ref master)" | "cap_relation (Structures_A.IRQControlCap) c = (c = Structures_H.IRQControlCap)" | "cap_relation (Structures_A.IRQHandlerCap irq) c = (c = Structures_H.IRQHandlerCap irq)" | "cap_relation (Structures_A.ArchObjectCap a) c = (\a'. acap_relation a a' \ c = Structures_H.ArchObjectCap a')" | "cap_relation (Structures_A.Zombie p b n) c = (c = Structures_H.Zombie p (zbits_map b) n)" definition cte_relation :: "cap_ref \ obj_relation_cut" where "cte_relation y \ \ko ko'. \sz cs cte cap. ko = CNode sz cs \ ko' = KOCTE cte \ cs y = Some cap \ cap_relation cap (cteCap cte)" definition asid_pool_relation :: "(10 word \ word32) \ ArchStructures_H.asidpool \ bool" where "asid_pool_relation \ \p p'. p = inv ASIDPool p' o ucast" definition ntfn_relation :: "Structures_A.notification \ Structures_H.notification \ bool" where "ntfn_relation \ \ntfn ntfn'. (case ntfn_obj ntfn of Structures_A.IdleNtfn \ ntfnObj ntfn' = Structures_H.IdleNtfn | Structures_A.WaitingNtfn q \ ntfnObj ntfn' = Structures_H.WaitingNtfn q | Structures_A.ActiveNtfn b \ ntfnObj ntfn' = Structures_H.ActiveNtfn b) \ ntfn_bound_tcb ntfn = ntfnBoundTCB ntfn'" definition ep_relation :: "Structures_A.endpoint \ Structures_H.endpoint \ bool" where "ep_relation \ \ep ep'. case ep of Structures_A.IdleEP \ ep' = Structures_H.IdleEP | Structures_A.RecvEP q \ ep' = Structures_H.RecvEP q | Structures_A.SendEP q \ ep' = Structures_H.SendEP q" definition fault_rel_optionation :: "ExceptionTypes_A.fault option \ Fault_H.fault option \ bool" where "fault_rel_optionation \ \f f'. f' = option_map fault_map f" primrec thread_state_relation :: "Structures_A.thread_state \ Structures_H.thread_state \ bool" where "thread_state_relation (Structures_A.Running) ts' = (ts' = Structures_H.Running)" | "thread_state_relation (Structures_A.Restart) ts' = (ts' = Structures_H.Restart)" | "thread_state_relation (Structures_A.Inactive) ts' = (ts' = Structures_H.Inactive)" | "thread_state_relation (Structures_A.IdleThreadState) ts' = (ts' = Structures_H.IdleThreadState)" | "thread_state_relation (Structures_A.BlockedOnReply) ts' = (ts' = Structures_H.BlockedOnReply)" | "thread_state_relation (Structures_A.BlockedOnReceive oref) ts' = (ts' = Structures_H.BlockedOnReceive oref)" | "thread_state_relation (Structures_A.BlockedOnSend oref sp) ts' = (ts' = Structures_H.BlockedOnSend oref (sender_badge sp) (sender_can_grant sp) (sender_is_call sp))" | "thread_state_relation (Structures_A.BlockedOnNotification oref) ts' = (ts' = Structures_H.BlockedOnNotification oref)" definition tcb_relation :: "Structures_A.tcb \ Structures_H.tcb \ bool" where "tcb_relation \ \tcb tcb'. tcb_fault_handler tcb = to_bl (tcbFaultHandler tcb') \ tcb_ipc_buffer tcb = tcbIPCBuffer tcb' \ tcb_context tcb = tcbContext tcb' \ thread_state_relation (tcb_state tcb) (tcbState tcb') \ fault_rel_optionation (tcb_fault tcb) (tcbFault tcb') \ cap_relation (tcb_ctable tcb) (cteCap (tcbCTable tcb')) \ cap_relation (tcb_vtable tcb) (cteCap (tcbVTable tcb')) \ cap_relation (tcb_reply tcb) (cteCap (tcbReply tcb')) \ cap_relation (tcb_caller tcb) (cteCap (tcbCaller tcb')) \ cap_relation (tcb_ipcframe tcb) (cteCap (tcbIPCBufferFrame tcb')) \ tcb_bound_notification tcb = tcbBoundNotification tcb'" definition other_obj_relation :: "Structures_A.kernel_object \ Structures_H.kernel_object \ bool" where "other_obj_relation obj obj' \ (case (obj, obj') of (TCB tcb, KOTCB tcb') \ tcb_relation tcb tcb' | (Endpoint ep, KOEndpoint ep') \ ep_relation ep ep' | (Notification ntfn, KONotification ntfn') \ ntfn_relation ntfn ntfn' | (ArchObj (Arch_Structs_A.ASIDPool pool), KOArch (KOASIDPool pool')) \ asid_pool_relation pool pool' | _ \ False)" primrec pde_relation' :: "Arch_Structs_A.pde \ Hardware_H.pde \ bool" where "pde_relation' Arch_Structs_A.InvalidPDE x = (x = Hardware_H.InvalidPDE)" | "pde_relation' (Arch_Structs_A.PageTablePDE ptr atts domain) x = (x = Hardware_H.PageTablePDE ptr (ParityEnabled \ atts) domain)" | "pde_relation' (Arch_Structs_A.SectionPDE ptr atts domain rghts) x = (x = Hardware_H.SectionPDE ptr (ParityEnabled \ atts) domain (PageCacheable \ atts) (Global \ atts) (XNever \ atts) (vmrights_map rghts))" | "pde_relation' (Arch_Structs_A.SuperSectionPDE ptr atts rghts) x = (x = Hardware_H.SuperSectionPDE ptr (ParityEnabled \ atts) (PageCacheable \ atts) (Global \ atts) (XNever \ atts) (vmrights_map rghts))" primrec pte_relation' :: "Arch_Structs_A.pte \ Hardware_H.pte \ bool" where "pte_relation' Arch_Structs_A.InvalidPTE x = (x = Hardware_H.InvalidPTE)" | "pte_relation' (Arch_Structs_A.LargePagePTE ptr atts rghts) x = (x = Hardware_H.LargePagePTE ptr (PageCacheable \ atts) (Global \ atts) (XNever \ atts) (vmrights_map rghts))" | "pte_relation' (Arch_Structs_A.SmallPagePTE ptr atts rghts) x = (x = Hardware_H.SmallPagePTE ptr (PageCacheable \ atts) (Global \ atts) (XNever \ atts) (vmrights_map rghts))" definition pde_align' :: "Hardware_H.pde \ nat" where "pde_align' pde \ case pde of Hardware_H.pde.SuperSectionPDE _ _ _ _ _ _ \ 4 | _ \ 0" lemmas pde_align_simps[simp] = pde_align'_def[split_simps Arch_Structs_A.pde.split] definition pte_align' :: "Hardware_H.pte \ nat" where "pte_align' pte \ case pte of Hardware_H.pte.LargePagePTE _ _ _ _ _ \ 4 | _ \ 0" lemmas pte_align_simps[simp] = pte_align'_def[split_simps Arch_Structs_A.pte.split] definition "pde_relation_aligned y pde pde' \ if is_aligned y (pde_align' pde') then pde_relation' pde pde' else pde = Arch_Structs_A.InvalidPDE" definition "pte_relation_aligned y pte pte' \ if is_aligned y (pte_align' pte') then pte_relation' pte pte' else pte = Arch_Structs_A.InvalidPTE" definition "pte_relation y \ \ko ko'. \pt pte. ko = ArchObj (PageTable pt) \ ko' = KOArch (KOPTE pte) \ pte_relation_aligned y (pt y) pte" definition "pde_relation y \ \ko ko'. \pd pde. ko = ArchObj (PageDirectory pd) \ ko' = KOArch (KOPDE pde) \ pde_relation_aligned y (pd y) pde" primrec aobj_relation_cuts :: "Arch_Structs_A.arch_kernel_obj \ word32 \ obj_relation_cuts" where "aobj_relation_cuts (DataPage sz) x = {(x + n * 2 ^ pageBits, \_ obj. obj = KOUserData \ obj = KOUserDataDevice) | n . n < 2 ^ (pageBitsForSize sz - pageBits) }" | "aobj_relation_cuts (Arch_Structs_A.ASIDPool pool) x = {(x, other_obj_relation)}" | "aobj_relation_cuts (PageTable pt) x = (\y. (x + (ucast y << 2), pte_relation y)) ` UNIV" | "aobj_relation_cuts (PageDirectory pd) x = (\y. (x + (ucast y << 2), pde_relation y)) ` UNIV" primrec obj_relation_cuts :: "Structures_A.kernel_object \ word32 \ obj_relation_cuts" where "obj_relation_cuts (CNode sz cs) x = (if well_formed_cnode_n sz cs then {(cte_map (x, y), cte_relation y) | y. y \ dom cs} else {(x, \\)})" | "obj_relation_cuts (TCB tcb) x = {(x, other_obj_relation)}" | "obj_relation_cuts (Endpoint ep) x = {(x, other_obj_relation)}" | "obj_relation_cuts (Notification ntfn) x = {(x, other_obj_relation)}" | "obj_relation_cuts (ArchObj ao) x = aobj_relation_cuts ao x" lemma obj_relation_cuts_def2: "obj_relation_cuts ko x = (case ko of CNode sz cs \ if well_formed_cnode_n sz cs then {(cte_map (x, y), cte_relation y) | y. y \ dom cs} else {(x, \\)} | ArchObj (PageTable pt) \ (\y. (x + (ucast y << 2), pte_relation y)) ` (UNIV :: word8 set) | ArchObj (PageDirectory pd) \ (\y. (x + (ucast y << 2), pde_relation y)) ` (UNIV :: 12 word set) | ArchObj (DataPage sz) \ {(x + n * 2 ^ pageBits, \_ obj. obj = KOUserData \ obj = KOUserDataDevice) | n . n < 2 ^ (pageBitsForSize sz - pageBits) } | _ \ {(x, other_obj_relation)})" by (simp split: Structures_A.kernel_object.split Arch_Structs_A.arch_kernel_obj.split) lemma obj_relation_cuts_def3: "obj_relation_cuts ko x = (case (a_type ko) of ACapTable n \ {(cte_map (x, y), cte_relation y) | y. length y = n} | AArch APageTable \ (\y. (x + (ucast y << 2), pte_relation y)) ` (UNIV :: word8 set) | AArch APageDirectory \ (\y. (x + (ucast y << 2), pde_relation y)) ` (UNIV :: 12 word set) | AArch (AIntData sz) \ {(x + n * 2 ^ pageBits, \_ obj. obj = KOUserData \ obj = KOUserDataDevice) | n . n < 2 ^ (pageBitsForSize sz - pageBits) } | AGarbage \ {(x, \\)} | _ \ {(x, other_obj_relation)})" apply (simp add: obj_relation_cuts_def2 a_type_def split: Structures_A.kernel_object.split Arch_Structs_A.arch_kernel_obj.split) apply (clarsimp simp: well_formed_cnode_n_def length_set_helper) done definition "is_other_obj_relation_type tp \ case tp of ACapTable n \ False | AArch APageTable \ False | AArch APageDirectory \ False | AArch (AIntData _) \ False | AGarbage \ False | _ \ True" lemma is_other_obj_relation_type_CapTable: "\ is_other_obj_relation_type (ACapTable n)" by (simp add: is_other_obj_relation_type_def) lemma is_other_obj_relation_type_IntData: "\ is_other_obj_relation_type (AArch (AIntData sz))" unfolding is_other_obj_relation_type_def by simp lemma is_other_obj_relation_type: "is_other_obj_relation_type (a_type ko) \ obj_relation_cuts ko x = {(x, other_obj_relation)}" by (simp add: obj_relation_cuts_def3 is_other_obj_relation_type_def split: a_type.splits aa_type.splits) definition pspace_dom :: "Structures_A.kheap \ word32 set" where "pspace_dom ps \ \x\dom ps. fst ` (obj_relation_cuts (the (ps x)) x)" definition pspace_relation :: "Structures_A.kheap \ (word32 \ Structures_H.kernel_object) \ bool" where "pspace_relation ab con \ (pspace_dom ab = dom con) \ (\x \ dom ab. \(y, P) \ obj_relation_cuts (the (ab x)) x. P (the (ab x)) (the (con y)))" definition etcb_relation :: "etcb \ Structures_H.tcb \ bool" where "etcb_relation \ \etcb tcb'. tcb_priority etcb = tcbPriority tcb' \ tcb_time_slice etcb = tcbTimeSlice tcb' \ tcb_domain etcb = tcbDomain tcb'" definition ekheap_relation :: "(obj_ref \ etcb option) \ (word32 \ Structures_H.kernel_object) \ bool" where "ekheap_relation ab con \ \x \ dom ab. \tcb'. con x = Some (KOTCB tcb') \ etcb_relation (the (ab x)) tcb'" primrec sched_act_relation :: "Deterministic_A.scheduler_action \ Structures_H.scheduler_action \ bool" where "sched_act_relation resume_cur_thread a' = (a' = ResumeCurrentThread)" | "sched_act_relation choose_new_thread a' = (a' = ChooseNewThread)" | "sched_act_relation (switch_thread x) a' = (a' = SwitchToThread x)" definition ready_queues_relation :: "(Deterministic_A.domain \ Deterministic_A.priority \ Deterministic_A.ready_queue) \ (domain \ priority \ KernelStateData_H.ready_queue) \ bool" where "ready_queues_relation qs qs' \ \d p. (qs d p = qs' (d, p))" definition ghost_relation :: "Structures_A.kheap \ (word32 \ vmpage_size) \ (word32 \ nat) \ bool" where "ghost_relation h ups cns \ (\a sz. h a = Some (ArchObj (DataPage sz)) \ ups a = Some sz) \ (\a n. (\cs. h a = Some (CNode n cs) \ well_formed_cnode_n n cs) \ cns a = Some n)" definition cdt_relation :: "(cslot_ptr \ bool) \ cdt \ cte_heap \ bool" where "cdt_relation \ \cte_at m m'. \c. cte_at c \ cte_map ` descendants_of c m = descendants_of' (cte_map c) m'" definition cdt_list_relation :: "cdt_list \ cdt \ cte_heap \ bool" where "cdt_list_relation \ \t m m'. \c cap node. m' (cte_map c) = Some (CTE cap node) \ (case next_slot c t m of None \ True | Some next \ mdbNext node = cte_map next)" definition revokable_relation :: "(cslot_ptr \ bool) \ (cslot_ptr \ cap option) \ cte_heap \ bool" where "revokable_relation revo cs m' \ \c cap node. cs c \ None \ m' (cte_map c) = Some (CTE cap node) \ revo c = mdbRevocable node" definition irq_state_relation :: "irq_state \ irqstate \ bool" where "irq_state_relation irq irq' \ case (irq, irq') of (irq_state.IRQInactive, irqstate.IRQInactive) \ True | (irq_state.IRQSignal, irqstate.IRQSignal) \ True | (irq_state.IRQTimer, irqstate.IRQTimer) \ True | _ \ False" definition interrupt_state_relation :: "(irq \ obj_ref) \ (irq \ irq_state) \ interrupt_state \ bool" where "interrupt_state_relation node_map irqs is \ (\node irqs'. is = InterruptState node irqs' \ (\irq. node_map irq = node + (ucast irq << cte_level_bits)) \ (\irq. irq_state_relation (irqs irq) (irqs' irq)))" definition arch_state_relation :: "(arch_state \ ArchStateData_H.kernel_state) set" where "arch_state_relation \ {(s, s') . arm_globals_frame s = armKSGlobalsFrame s' \ arm_asid_table s = armKSASIDTable s' o ucast \ arm_global_pd s = armKSGlobalPD s' \ arm_hwasid_table s = armKSHWASIDTable s' \ arm_global_pts s = armKSGlobalPTs s' \ arm_next_asid s = armKSNextASID s' \ arm_asid_map s = armKSASIDMap s' \ arm_kernel_vspace s = armKSKernelVSpace s'}" definition (* NOTE: this map discards the Ident right, needed on endpoints only *) rights_mask_map :: "rights set \ Types_H.cap_rights" where "rights_mask_map \ \rs. CapRights (AllowWrite \ rs) (AllowRead \ rs) (AllowGrant \ rs)" lemma obj_relation_cutsE: "\ (y, P) \ obj_relation_cuts ko x; P ko ko'; \sz cs z cap cte. \ ko = CNode sz cs; well_formed_cnode_n sz cs; y = cte_map (x, z); ko' = KOCTE cte; cs z = Some cap; cap_relation cap (cteCap cte) \ \ R; \pt (z :: word8) pte'. \ ko = ArchObj (PageTable pt); y = x + (ucast z << 2); ko' = KOArch (KOPTE pte'); pte_relation_aligned z (pt z) pte' \ \ R; \pd (z :: 12 word) pde'. \ ko = ArchObj (PageDirectory pd); y = x + (ucast z << 2); ko' = KOArch (KOPDE pde'); pde_relation_aligned z (pd z) pde' \ \ R; \sz n. \ ko = ArchObj (DataPage sz); ko' = KOUserData \ ko' = KOUserDataDevice; y = x + n * 2 ^ pageBits; n < 2 ^ (pageBitsForSize sz - pageBits) \ \ R; \ y = x; other_obj_relation ko ko'; is_other_obj_relation_type (a_type ko) \ \ R \ \ R" apply (simp add: obj_relation_cuts_def2 is_other_obj_relation_type_def a_type_def split: Structures_A.kernel_object.split_asm split_if_asm Arch_Structs_A.arch_kernel_obj.split_asm) apply (force simp: cte_relation_def pte_relation_def pde_relation_def)+ done lemma eq_trans_helper: "\ x = y; P y = Q \ \ P x = Q" by simp lemma cap_relation_case': "cap_relation cap cap' = (case cap of cap.ArchObjectCap arch_cap.ASIDControlCap \ cap_relation cap cap' | _ \ cap_relation cap cap')" by (simp split: cap.split arch_cap.split) schematic_goal cap_relation_case: "cap_relation cap cap' = ?P" apply (subst cap_relation_case') apply (clarsimp cong: cap.case_cong arch_cap.case_cong) apply (rule refl) done lemmas cap_relation_split = eq_trans_helper [where P=P, OF cap_relation_case cap.split[where P=P]] for P lemmas cap_relation_split_asm = eq_trans_helper [where P=P, OF cap_relation_case cap.split_asm[where P=P]] for P text {* Relations on other data types that aren't stored but used as intermediate values in the specs. *} primrec message_info_map :: "Structures_A.message_info \ Types_H.message_info" where "message_info_map (Structures_A.MI a b c d) = (Types_H.MI a b c d)" lemma mi_map_label[simp]: "msgLabel (message_info_map mi) = mi_label mi" by (cases mi, simp) primrec syscall_error_map :: "ExceptionTypes_A.syscall_error \ Fault_H.syscall_error" where "syscall_error_map (ExceptionTypes_A.InvalidArgument n) = Fault_H.InvalidArgument n" | "syscall_error_map (ExceptionTypes_A.InvalidCapability n) = (Fault_H.InvalidCapability n)" | "syscall_error_map ExceptionTypes_A.IllegalOperation = Fault_H.IllegalOperation" | "syscall_error_map (ExceptionTypes_A.RangeError n m) = Fault_H.RangeError n m" | "syscall_error_map ExceptionTypes_A.AlignmentError = Fault_H.AlignmentError" | "syscall_error_map (ExceptionTypes_A.FailedLookup b lf) = Fault_H.FailedLookup b (lookup_failure_map lf)" | "syscall_error_map ExceptionTypes_A.TruncatedMessage = Fault_H.TruncatedMessage" | "syscall_error_map ExceptionTypes_A.DeleteFirst = Fault_H.DeleteFirst" | "syscall_error_map ExceptionTypes_A.RevokeFirst = Fault_H.RevokeFirst" | "syscall_error_map (ExceptionTypes_A.NotEnoughMemory n) = Fault_H.syscall_error.NotEnoughMemory n" definition APIType_map :: "Structures_A.apiobject_type \ ArchTypes_H.object_type" where "APIType_map ty \ case ty of Structures_A.Untyped \ APIObjectType ArchTypes_H.Untyped | Structures_A.TCBObject \ APIObjectType ArchTypes_H.TCBObject | Structures_A.EndpointObject \ APIObjectType ArchTypes_H.EndpointObject | Structures_A.NotificationObject \ APIObjectType ArchTypes_H.NotificationObject | Structures_A.CapTableObject \ APIObjectType ArchTypes_H.CapTableObject | ArchObject ao \ (case ao of SmallPageObj \ SmallPageObject | LargePageObj \ LargePageObject | SectionObj \ SectionObject | SuperSectionObj \ SuperSectionObject | PageTableObj \ PageTableObject | PageDirectoryObj \ PageDirectoryObject)" lemma get_tcb_at: "tcb_at t s \ (\tcb. get_tcb t s = Some tcb)" by (simp add: tcb_at_def) definition state_relation :: "(det_state \ kernel_state) set" where "state_relation \ {(s, s'). pspace_relation (kheap s) (ksPSpace s') \ ekheap_relation (ekheap s) (ksPSpace s') \ sched_act_relation (scheduler_action s) (ksSchedulerAction s') \ ready_queues_relation (ready_queues s) (ksReadyQueues s') \ ghost_relation (kheap s) (gsUserPages s') (gsCNodes s') \ cdt_relation (swp cte_at s) (cdt s) (ctes_of s') \ cdt_list_relation (cdt_list s) (cdt s) (ctes_of s') \ revokable_relation (is_original_cap s) (null_filter (caps_of_state s)) (ctes_of s') \ (arch_state s, ksArchState s') \ arch_state_relation \ interrupt_state_relation (interrupt_irq_node s) (interrupt_states s) (ksInterruptState s') \ (cur_thread s = ksCurThread s') \ (idle_thread s = ksIdleThread s') \ (machine_state s = ksMachineState s') \ (work_units_completed s = ksWorkUnitsCompleted s') \ (domain_index s = ksDomScheduleIdx s') \ (domain_list s = ksDomSchedule s') \ (cur_domain s = ksCurDomain s') \ (domain_time s = ksDomainTime s')}" text {* Rules for using states in the relation. *} lemma curthread_relation: "(a, b) \ state_relation \ ksCurThread b = cur_thread a" by (simp add: state_relation_def) lemma workunitscompleted_relation: "(a, b) \ state_relation \ ksWorkUnitsCompleted b = work_units_completed a" by (simp add: state_relation_def) lemma state_relation_pspace_relation[elim!]: "(s,s') \ state_relation \ pspace_relation (kheap s) (ksPSpace s')" by (simp add: state_relation_def) lemma state_relation_ekheap_relation[elim!]: "(s,s') \ state_relation \ ekheap_relation (ekheap s) (ksPSpace s')" by (simp add: state_relation_def) (* intro/dest/elim for state_relation *) lemma state_relationI [intro?]: "\s s'. \ pspace_relation (kheap s) (ksPSpace s'); ekheap_relation (ekheap s) (ksPSpace s'); sched_act_relation (scheduler_action s) (ksSchedulerAction s'); ready_queues_relation (ready_queues s) (ksReadyQueues s'); ghost_relation (kheap s) (gsUserPages s') (gsCNodes s'); cdt_relation (swp cte_at s) (cdt s) (ctes_of s'); cdt_list_relation (cdt_list s) (cdt s) (ctes_of s'); revokable_relation (is_original_cap s) (null_filter (caps_of_state s)) (ctes_of s'); (arch_state s, ksArchState s') \ arch_state_relation; interrupt_state_relation (interrupt_irq_node s) (interrupt_states s) (ksInterruptState s'); cur_thread s = ksCurThread s'; idle_thread s = ksIdleThread s'; machine_state s = ksMachineState s'; work_units_completed s = ksWorkUnitsCompleted s'; domain_index s = ksDomScheduleIdx s'; domain_list s = ksDomSchedule s'; cur_domain s = ksCurDomain s'; domain_time s = ksDomainTime s' \ \ (s, s') \ state_relation" unfolding state_relation_def by blast lemma state_relationD: assumes sr: "(s, s') \ state_relation" shows "pspace_relation (kheap s) (ksPSpace s') \ ekheap_relation (ekheap s) (ksPSpace s') \ sched_act_relation (scheduler_action s) (ksSchedulerAction s') \ ready_queues_relation (ready_queues s) (ksReadyQueues s') \ ghost_relation (kheap s) (gsUserPages s') (gsCNodes s') \ cdt_relation (swp cte_at s) (cdt s) (ctes_of s') \ cdt_list_relation (cdt_list s) (cdt s) (ctes_of s') \ revokable_relation (is_original_cap s) (null_filter (caps_of_state s)) (ctes_of s') \ (arch_state s, ksArchState s') \ arch_state_relation \ interrupt_state_relation (interrupt_irq_node s) (interrupt_states s) (ksInterruptState s') \ cur_thread s = ksCurThread s' \ idle_thread s = ksIdleThread s' \ machine_state s = ksMachineState s' \ work_units_completed s = ksWorkUnitsCompleted s' \ domain_index s = ksDomScheduleIdx s' \ domain_list s = ksDomSchedule s' \ cur_domain s = ksCurDomain s' \ domain_time s = ksDomainTime s'" using sr unfolding state_relation_def by simp lemma state_relationE [elim?]: assumes sr: "(s, s') \ state_relation" and rl: "\pspace_relation (kheap s) (ksPSpace s'); ekheap_relation (ekheap s) (ksPSpace s'); sched_act_relation (scheduler_action s) (ksSchedulerAction s'); ready_queues_relation (ready_queues s) (ksReadyQueues s'); ghost_relation (kheap s) (gsUserPages s') (gsCNodes s'); cdt_relation (swp cte_at s) (cdt s) (ctes_of s') \ revokable_relation (is_original_cap s) (null_filter (caps_of_state s)) (ctes_of s'); cdt_list_relation (cdt_list s) (cdt s) (ctes_of s'); (arch_state s, ksArchState s') \ arch_state_relation; interrupt_state_relation (interrupt_irq_node s) (interrupt_states s) (ksInterruptState s'); cur_thread s = ksCurThread s'; idle_thread s = ksIdleThread s'; machine_state s = ksMachineState s'; work_units_completed s = ksWorkUnitsCompleted s'; domain_index s = ksDomScheduleIdx s'; domain_list s = ksDomSchedule s'; cur_domain s = ksCurDomain s'; domain_time s = ksDomainTime s' \ \ R" shows "R" using sr by (blast intro!: rl dest: state_relationD) text {* This isn't defined for arch objects *} lemma objBits_obj_bits: assumes rel: "other_obj_relation obj obj'" shows "obj_bits obj = objBitsKO obj'" using rel by (simp add: other_obj_relation_def objBitsKO_simps pageBits_def archObjSize_def split: Structures_A.kernel_object.split_asm Arch_Structs_A.arch_kernel_obj.split_asm Structures_H.kernel_object.split_asm ArchStructures_H.arch_kernel_object.split_asm) lemma replicate_length_cong: "x = y \ replicate x n = replicate y n" by simp lemmas isCap_defs = isZombie_def isArchObjectCap_def isThreadCap_def isCNodeCap_def isNotificationCap_def isEndpointCap_def isUntypedCap_def isNullCap_def isIRQHandlerCap_def isIRQControlCap_def isReplyCap_def isPageCap_def isPageTableCap_def isPageDirectoryCap_def isASIDControlCap_def isASIDPoolCap_def isArchPageCap_def isDomainCap_def lemma isCNodeCap_cap_map [simp]: "cap_relation c c' \ isCNodeCap c' = is_cnode_cap c" apply (cases c, simp_all add: isCap_defs split: sum.splits) apply clarsimp+ done lemma isNullCap_cap_map [simp]: "cap_relation c c' \ isNullCap c' = (c = cap.NullCap)" apply (cases c, simp_all add: isCap_defs split: sum.splits) apply clarsimp+ done lemma revokable_relation_eqI: assumes r: "revokable_relation (is_original_cap s) (null_filter (caps_of_state s)) m" assumes c: "\P p. cte_wp_at P p s'' \ cte_wp_at P p s" assumes m: "\p. is_original_cap s' p = is_original_cap s p" shows "revokable_relation (is_original_cap s') (null_filter (caps_of_state s'')) m" using r apply (clarsimp simp add: m revokable_relation_def null_filter_def) apply (drule caps_of_state_cteD) apply (drule c) apply (simp add: cte_wp_at_caps_of_state) done lemma sts_rel_idle : "thread_state_relation st IdleThreadState = (st = Structures_A.IdleThreadState)" by (cases st, auto) lemma pspace_relation_absD: "\ ab x = Some y; pspace_relation ab con \ \ \(x', P) \ obj_relation_cuts y x. \z. con x' = Some z \ P y z" apply (clarsimp simp add: pspace_relation_def) apply (drule bspec, erule domI) apply simp apply (drule(1) bspec) apply (subgoal_tac "a \ pspace_dom ab") apply clarsimp apply (simp(no_asm) add: pspace_dom_def) apply (rule rev_bexI, erule domI) apply (simp add: image_def rev_bexI) done lemma ekheap_relation_absD: "\ ab x = Some y; ekheap_relation ab con \ \ \tcb'. con x = Some (KOTCB tcb') \ etcb_relation y tcb'" by (force simp add: ekheap_relation_def) lemma in_related_pspace_dom: "\ s' x = Some y; pspace_relation s s' \ \ x \ pspace_dom s" by (clarsimp simp add: pspace_relation_def) lemma pspace_dom_revE: "\ x \ pspace_dom ps; \ko y P. \ ps y = Some ko; (x, P) \ obj_relation_cuts ko y \ \ R \ \ R" by (clarsimp simp add: pspace_dom_def) lemma pspace_dom_relatedE: "\ s' x = Some ko'; pspace_relation s s'; \y ko P. \ s y = Some ko; (x, P) \ obj_relation_cuts ko y; P ko ko' \ \ R \ \ R" apply (rule pspace_dom_revE [OF in_related_pspace_dom], assumption+) apply (frule(1) pspace_relation_absD) apply fastforce done lemma ghost_relation_typ_at: "ghost_relation (kheap s) ups cns \ (\a sz. typ_at (AArch (AIntData sz)) a s = (ups a = Some sz)) \ (\a n. typ_at (ACapTable n) a s = (cns a = Some n))" by (rule eq_reflection) (clarsimp simp: ghost_relation_def typ_at_eq_kheap_obj) end