lh-l4v/proof/drefine/Refine_D.thy

70 lines
2.5 KiB
Plaintext
Raw Normal View History

2014-07-14 19:32:44 +00:00
(*
* Copyright 2014, NICTA
*
* 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(NICTA_GPL)
*)
(*
* Toplevel capDL refinement theorem.
*)
theory Refine_D
imports Syscall_DR
begin
2016-05-01 00:09:19 +00:00
context begin interpretation Arch . (*FIXME: arch_split*)
2014-07-14 19:32:44 +00:00
text {*
2017-07-12 05:13:51 +00:00
Toplevel @{text dcorres} theorem.
2014-07-14 19:32:44 +00:00
*}
lemma valid_etcbs_sched: "valid_sched s \<longrightarrow> valid_etcbs s" by (fastforce simp: valid_sched_def)
2014-07-14 19:32:44 +00:00
lemma handle_event_invs_and_valid_sched:
"\<lbrace>invs and valid_sched and (\<lambda>s. e \<noteq> Interrupt \<longrightarrow> ct_active s)
and (\<lambda>s. scheduler_action s = resume_cur_thread)\<rbrace> Syscall_A.handle_event e
2014-07-14 19:32:44 +00:00
\<lbrace>\<lambda>rv. invs and valid_sched\<rbrace>"
by ((wp he_invs handle_event_valid_sched), clarsimp)
2014-07-14 19:32:44 +00:00
lemma dcorres_call_kernel:
"dcorres dc \<top>
(invs and valid_sched and valid_pdpt_objs
and (\<lambda>s. e \<noteq> Interrupt \<longrightarrow> ct_running s)
2014-07-14 19:32:44 +00:00
and (\<lambda>s. scheduler_action s = resume_cur_thread))
(Syscall_D.call_kernel e) (Syscall_A.call_kernel e)"
apply (simp_all add: Syscall_D.call_kernel_def Syscall_A.call_kernel_def)
apply (rule corres_guard_imp)
apply (rule corres_split)
prefer 2
apply (rule corres_split_handle [OF _ handle_event_corres])
prefer 4
apply (subst bind_return[symmetric])
apply (rule corres_split)
apply (rule activate_thread_corres[unfolded fun_app_def])
apply simp
apply (rule schedule_dcorres)
apply (wp schedule_valid_sched | strengthen valid_etcbs_sched)+
apply (simp add: handle_pending_interrupts_def)
apply (rule corres_split [OF _ get_active_irq_corres])
apply (clarsimp simp: when_def split: option.splits)
apply (rule handle_interrupt_corres)
apply ((wp | simp)+)[3]
apply (rule hoare_post_imp_dc2E, rule handle_event_invs_and_valid_sched)
apply (clarsimp simp: invs_def valid_state_def)
apply (simp add: conj_comms if_apply_def2 non_kernel_IRQs_def
| wp | strengthen valid_etcbs_sched valid_idle_invs_strg)+
2014-07-23 04:30:30 +00:00
apply (rule valid_validE2)
2014-07-14 19:32:44 +00:00
apply (rule hoare_vcg_conj_lift)
apply (rule he_invs)
apply (rule handle_event_valid_sched)
apply (fastforce intro: active_from_running simp: valid_sched_def)+
2018-03-14 00:48:48 +00:00
done
2014-07-14 19:32:44 +00:00
end
2016-05-01 00:09:19 +00:00
end