(* * Copyright 2014, General Dynamics C4 Systems * * SPDX-License-Identifier: GPL-2.0-only *) chapter \Abstract datatype for the abstract specification\ theory ADT_AI imports ArchADT_AI begin context begin interpretation Arch . requalify_consts empty_context init_A_st ptable_lift ptable_rights addrFromPPtr ptrFromPAddr end text \ The general refinement calculus (see theory Simulation) requires the definition of a so-called ``abstract datatype'' for each refinement layer. This theory defines this datatype for the abstract specification. Along the path, the theory extends the abstract specification of the kernel with respect to user-mode transitions. \ text \ We constrain idle thread behaviour, so we distinguish three main machine control states: \ datatype mode = UserMode | KernelMode | IdleMode text \ The global state contains the current user's register context of the machine as well as the internal kernel state, the mode and the current event (if any). \ type_synonym 'k global_state = "(user_context \ 'k) \ mode \ event option" text \ As observable state, we take the global abstract state. \ type_synonym 'a observable = "('a state) global_state" text \ From the kernel's perspective, the user memory is a mapping from addresses to bytes. A virtual-memory view will later on be built on top of that. \ type_synonym user_mem = "machine_word \ word8 option" (* This is a temperary solution to model the user device memory *) type_synonym device_state = "machine_word \ word8 option" text \ A user state consists of a register context and (physical) user memory. \ type_synonym user_state = "user_context \ user_mem \ device_state" text \Virtual-memory mapping: translates virtual to physical addresses\ type_synonym vm_mapping = "machine_word \ machine_word" text \Memory rights for each virtual adress\ type_synonym mem_rights = "machine_word \ vm_rights" text \ A user transition is characterized by a function that takes the following arguments: a current thread identifier, a virtual-memory mapping (i.e., a partial function from virtual to physical addresses), a memory-rights mapping (i.e., a partial function from virtual addresses to rights). The function is then a non-deterministic state monad on user state, returning an optional kernel-entry event. Note that the current-thread identifiers are identical in both specs (i.e. @{term "Structures_A.cur_thread :: 'z state \ obj_ref"} in the abstract kernel model and @{text "KernelStateData_H.ksCurThread \ kernel_state \ machine_word"} in the executable specification). \ type_synonym user_transition = "obj_ref \ vm_mapping \ mem_rights \ (user_state, event option) nondet_monad" text \Abbreviation for user context plus additional state\ type_synonym 'k uc_state = "user_context \ 'k" text \ The following definition models machine and kernel entry/exit behaviour abstractly. Furthermore, it constrains the behaviour of user threads as well as the idle thread. The first parameter is used to check of pending interrupts, potentially modifying the machine state embedded in the kernel state. The second parameter lifts a user-thread transition into the kernel state (i.e. the new user memory should be injected into the kernel state) and the third parameter provides a kernel specification (it will later be used with the abstract specification, the executable specification, as well as the C implementation). Despite the fact that the global automaton does not distinguish different kinds of transitions to the outside world, the definition groups them into 6 kinds: 1. Kernel transition: The kernel specification is started in kernel mode and uses the event part of the state as input. The next control state will either be user mode or idle mode (idle thread running). 2. Normal user-mode executions. The computed new user states are then lifted into the kernel state using the first parameter to the definition below. 3. and 4. Event generation in user mode. In user mode, events may be generated by the user and any interrupt can be generated at any time. 5. and 6. Finally, events generated by the idle thread. These can only be interrupts. If there is no interrupt, we stay in idle mode. \ definition global_automaton :: "('k uc_state \ (bool \ 'k uc_state)) set \ ('k uc_state \ (event option \ 'k uc_state)) set \ (event \ ('k uc_state \ mode \ 'k uc_state) set) \ ('k global_state \ 'k global_state) set" where "global_automaton get_active_irq do_user_op kernel_call \ \ \ Kernel transition \ { ( (s, KernelMode, Some e), (s', m, None) ) |s s' e m. (s, m, s') \ kernel_call e } \ \ \ User to user transition, no kernel entry \ { ( (s, UserMode, None), (s', UserMode, None) ) |s s'. (s, None, s') \ do_user_op} \ \ \ User to kernel transition, potentially includes Interrupt from user mode \ { ( (s, UserMode, None), (s', KernelMode, Some e) ) |s s' e. (s, Some e, s') \ do_user_op} \ \ \ User to kernel transition, Interrupt from user mode \ { ( (s, UserMode, None), (s', KernelMode, Some Interrupt) ) |s s'. (s, True, s') \ get_active_irq} \ \ \ Idling in idle mode \ { ( (s, IdleMode, None), (s', IdleMode, None) ) |s s'. (s, False, s') \ get_active_irq} \ \ \ Interrupt while in idle mode \ { ( (s, IdleMode, None), (s', KernelMode, Some Interrupt) ) |s s'. (s, True, s') \ get_active_irq}" text \ After kernel initialisation, the machine is in UserMode, running the initial thread. \ definition Init_A :: "'a::state_ext state global_state set" where "Init_A \ {((empty_context, init_A_st), UserMode, None)}" definition "user_memory_update um \ modify (\ms. ms\underlying_memory := (\a. case um a of Some x \ x | None \ underlying_memory ms a)\)" definition "device_memory_update um \ modify (\ms. ms\device_state := (device_state ms ++ um ) \)" definition "option_to_0 x \ case x of None \ 0 | Some y \ y" text \ The below definition gives the kernel monad computation that checks for active interrupts, given the present user_context. This is necessarily a computation in the kernel monad because checking interrupts will update the interrupt state. \ definition check_active_irq :: "(bool,'z :: state_ext) s_monad" where "check_active_irq \ do irq \ do_machine_op $ getActiveIRQ False; return (irq \ None) od" definition check_active_irq_A :: "(('z :: state_ext) state uc_state \ bool \ ('z :: state_ext) state uc_state) set" where "check_active_irq_A \ {((tc, s), (irq, (tc, s'))). (irq , s') \ fst (check_active_irq s)}" text \ The definition below lifts a user transition into the kernel monad. Note that the user memory (as seen by the kernel) is converted to true physical addresses and restricted to those addresses, the current thread is permitted to access. Furthermore, user memory is updated if and only if the current thread has write permission. NOTE: An unpermitted write access would generate a page fault on the machine. The global transitions, however, model page faults non-deterministically. \ definition do_user_op :: "user_transition \ user_context \ (event option \ user_context,'z::state_ext) s_monad" where "do_user_op uop tc \ do t \ gets cur_thread; conv \ gets (ptable_lift t); rights \ gets (ptable_rights t); um \ gets (\s. (user_mem s) \ ptrFromPAddr); dm \ gets (\s. (device_mem s) \ ptrFromPAddr); ds \ gets (device_state \ machine_state); (e,tc',um',ds') \ select (fst (uop t (restrict_map conv {pa. rights pa \ {}}) rights (tc, restrict_map um {pa. \va. conv va = Some pa \ AllowRead \ rights va} ,(ds \ ptrFromPAddr) |` {pa. \va. conv va = Some pa \ AllowRead \ rights va} ) )); do_machine_op (user_memory_update ((um' |` {pa. \va. conv va = Some pa \ AllowWrite \ rights va} \ addrFromPPtr) |` (- dom ds))); do_machine_op (device_memory_update ((ds' |` {pa. \va. conv va = Some pa \ AllowWrite \ rights va} \ addrFromPPtr) |` (dom ds))); return (e, tc') od" definition monad_to_transition :: "(user_context \ ('s, event option \ user_context) nondet_monad) \ ('s uc_state \ event option \ 's uc_state) set" where "monad_to_transition m \ {((tc,s),(e,tc',s')). ((e,tc'),s') \ fst (m tc s)}" definition do_user_op_A :: "user_transition \ ('z state uc_state \ event option \ ('z::state_ext state) uc_state) set" where "do_user_op_A uop \ monad_to_transition (do_user_op uop)" text \ Kernel calls are described completely by the abstract and concrete spec. We model kernel entry by allowing an arbitrary user (register) context. The mode after a kernel call is either user or idle (see also thm in Refine.thy). \ (* FIXME: The IPC buffer pointer and TLS_BASE are stored in registers in restore_user_context, which is currently invisible to verification. The effect should be modelled in the ADT. *) definition kernel_entry :: "event \ user_context \ (user_context,'z::state_ext_sched) s_monad" where "kernel_entry e tc \ do t \ gets cur_thread; thread_set (\tcb. tcb \ tcb_arch := arch_tcb_context_set tc (tcb_arch tcb) \) t; call_kernel e; t' \ gets cur_thread; thread_get (arch_tcb_context_get o tcb_arch) t' od" definition kernel_call_A :: "event \ ((user_context \ ('a::state_ext_sched state)) \ mode \ (user_context \ 'a state)) set" where "kernel_call_A e \ {(s, m, s'). s' \ fst (split (kernel_entry e) s) \ m = (if ct_running (snd s') then UserMode else IdleMode)}" text \Putting together the final abstract datatype\ (* NOTE: the extensible abstract specification leaves the type of the extension unspecified; later on, we will instantiate this type with det_ext from the deterministic abstract specification as well as with unit. The former is used for refinement between the deterministic specification and C. The latter is used for refinement between the non-deterministic specification and C. *) definition "observable_memory ms fs \ ms \underlying_memory := option_to_0 \ fs\" definition "abs_state s \ s\machine_state:= observable_memory (machine_state s) (user_mem s)\" definition ADT_A :: "user_transition \ (('a::state_ext_sched state) global_state, 'a observable, unit) data_type" where "ADT_A uop \ \ Init = \s. Init_A, Fin = \((tc,s),m,e). ((tc, abs_state s),m,e), Step = (\u. global_automaton check_active_irq_A (do_user_op_A uop) kernel_call_A) \" text \ Lifting a state relation on kernel states to global states. \ definition "lift_state_relation sr \ { (((tc,s),m,e), ((tc,s'),m,e))|s s' m e tc. (s,s') \ sr }" lemma lift_state_relationD: "(((tc, s), m, e), ((tc', s'), m', e')) \ lift_state_relation R \ (s,s') \ R \ tc' = tc \ m' = m \ e' = e" by (simp add: lift_state_relation_def) lemma lift_state_relationI: "(s,s') \ R \ (((tc, s), m, e), ((tc, s'), m, e)) \ lift_state_relation R" by (fastforce simp: lift_state_relation_def) lemma in_lift_state_relation_eq: "(((tc, s), m, e), (tc', s'), m', e') \ lift_state_relation R \ (s, s') \ R \ tc' = tc \ m' = m \ e' = e" by (auto simp add: lift_state_relation_def) end