(* * Copyright 2014, General Dynamics C4 Systems * * SPDX-License-Identifier: GPL-2.0-only *) theory DetSchedInvs_AI imports "$L4V_ARCH/ArchDeterministic_AI" begin lemma get_etcb_rev: "ekheap s p = Some etcb \ get_etcb p s = Some etcb" by (clarsimp simp: get_etcb_def) lemma get_etcb_SomeD: "get_etcb ptr s = Some v \ ekheap s ptr = Some v" apply (case_tac "ekheap s ptr", simp_all add: get_etcb_def) done definition obj_at_kh where "obj_at_kh P ref kh \ obj_at P ref ((undefined :: det_ext state)\kheap := kh\)" lemma obj_at_kh_simp[simp]: "obj_at_kh P ref (kheap st) = obj_at P ref st" apply (simp add: obj_at_def obj_at_kh_def) done definition st_tcb_at_kh where "st_tcb_at_kh test \ obj_at_kh (\ko. \tcb. ko = TCB tcb \ test (tcb_state tcb))" lemma st_tcb_at_kh_simp[simp]: "st_tcb_at_kh test t (kheap st) = st_tcb_at test t st" apply (simp add: pred_tcb_at_def st_tcb_at_kh_def) done definition is_etcb_at' where "is_etcb_at' ref ekh \ ekh ref \ None" abbreviation is_etcb_at:: "obj_ref \ det_ext state \ bool" where "is_etcb_at ref s \ is_etcb_at' ref (ekheap s)" lemmas is_etcb_at_def = is_etcb_at'_def definition etcb_at' :: "(etcb \ bool) \ obj_ref \ (obj_ref \ etcb option) \ bool" where "etcb_at' P ref ekh \ case ekh ref of Some x \ P x | _ \ True" abbreviation etcb_at :: "(etcb \ bool) \ obj_ref \ det_ext state \ bool" where "etcb_at P ref s \ etcb_at' P ref (ekheap s)" lemmas etcb_at_def = etcb_at'_def lemma etcb_at_taut[simp]: "etcb_at' \ ref ekh" apply (simp add: etcb_at'_def split: option.split) done lemma etcb_at_conj_is_etcb_at: "(is_etcb_at' t ekh \ etcb_at' P t ekh) = (case ekh t of None \ False | Some x \ P x)" by (simp add: is_etcb_at_def etcb_at_def split: option.splits) definition valid_etcbs_2 :: "(obj_ref \ etcb option) \ (obj_ref \ kernel_object option) \ bool"where "valid_etcbs_2 ekh kh \ \ptr. (st_tcb_at_kh \ ptr kh) = (is_etcb_at' ptr ekh)" abbreviation valid_etcbs :: "det_ext state \ bool" where "valid_etcbs s \ valid_etcbs_2 (ekheap s) (kheap s)" lemmas valid_etcbs_def = valid_etcbs_2_def definition valid_idle_etcb_2 :: "(obj_ref \ etcb option) \ bool" where "valid_idle_etcb_2 ekh \ etcb_at' (\etcb. tcb_domain etcb = default_domain) idle_thread_ptr ekh" abbreviation valid_idle_etcb :: "det_ext state \ bool" where "valid_idle_etcb s \ valid_idle_etcb_2 (ekheap s)" lemmas valid_idle_etcb_def = valid_idle_etcb_2_def definition not_queued_2 where "not_queued_2 qs t \ \d p. t \ set (qs d p)" abbreviation not_queued :: "obj_ref \ det_ext state \ bool" where "not_queued t s \ not_queued_2 (ready_queues s) t" definition valid_queues_2 where "valid_queues_2 queues ekh kh \ (\d p. (\t \ set (queues d p). is_etcb_at' t ekh \ etcb_at' (\t. tcb_priority t = p \ tcb_domain t = d) t ekh \ st_tcb_at_kh runnable t kh) \ distinct (queues d p))" abbreviation valid_queues :: "det_ext state \ bool" where "valid_queues s \ valid_queues_2 (ready_queues s) (ekheap s) (kheap s)" lemmas valid_queues_def = valid_queues_2_def lemma valid_queues_def2: "valid_queues_2 queues ekh kh = (\d p. (\t \ set (queues d p). is_etcb_at' t ekh \ (case ekh t of None \ False | Some x \ tcb_priority x = p \ tcb_domain x = d) \ st_tcb_at_kh runnable t kh) \ distinct (queues d p))" by (clarsimp simp: valid_queues_def conj_assoc[where P="is_etcb_at' t ekh \ (case ekh t of None \ False | Some x \ tcb_priority x = p \ tcb_domain x = d)"] etcb_at_conj_is_etcb_at[symmetric]) definition valid_blocked_2 where "valid_blocked_2 queues kh sa ct \ (\t st. not_queued_2 queues t \ st_tcb_at_kh ((=) st) t kh \ t \ ct \ sa \ switch_thread t \ (\ active st))" abbreviation valid_blocked :: "det_ext state \ bool" where "valid_blocked s \ valid_blocked_2 (ready_queues s) (kheap s) (scheduler_action s) (cur_thread s)" lemmas valid_blocked_def = valid_blocked_2_def definition valid_blocked_except_2 where "valid_blocked_except_2 thread queues kh sa ct \ (\t st. t \ thread \ not_queued_2 queues t \ st_tcb_at_kh ((=) st) t kh \ t \ ct \ sa \ switch_thread t \ (\ active st))" abbreviation valid_blocked_except :: "obj_ref \ det_ext state \ bool" where "valid_blocked_except t s \ valid_blocked_except_2 t (ready_queues s) (kheap s) (scheduler_action s) (cur_thread s)" lemmas valid_blocked_except_def = valid_blocked_except_2_def definition in_cur_domain_2 where "in_cur_domain_2 thread cdom ekh \ etcb_at' (\t. tcb_domain t = cdom) thread ekh" abbreviation in_cur_domain :: "obj_ref \ det_ext state \ bool" where "in_cur_domain thread s \ in_cur_domain_2 thread (cur_domain s) (ekheap s)" lemmas in_cur_domain_def = in_cur_domain_2_def definition ct_in_cur_domain_2 where "ct_in_cur_domain_2 thread thread' sa cdom ekh \ sa = resume_cur_thread \ thread = thread' \ in_cur_domain_2 thread cdom ekh" abbreviation ct_in_cur_domain where "ct_in_cur_domain s \ ct_in_cur_domain_2 (cur_thread s) (idle_thread s) (scheduler_action s) (cur_domain s) (ekheap s)" lemmas ct_in_cur_domain_def = ct_in_cur_domain_2_def definition is_activatable_2 where "is_activatable_2 thread sa kh \ sa = resume_cur_thread \ st_tcb_at_kh activatable thread kh" abbreviation is_activatable :: "obj_ref \ det_ext state \ bool" where "is_activatable thread s \ is_activatable_2 thread (scheduler_action s) (kheap s)" lemmas is_activatable_def = is_activatable_2_def definition weak_valid_sched_action_2 where "weak_valid_sched_action_2 sa ekh kh \ \t. sa = switch_thread t \ st_tcb_at_kh runnable t kh" abbreviation weak_valid_sched_action:: "det_ext state \ bool" where "weak_valid_sched_action s \ weak_valid_sched_action_2 (scheduler_action s) (ekheap s) (kheap s)" lemmas weak_valid_sched_action_def = weak_valid_sched_action_2_def definition switch_in_cur_domain_2 where "switch_in_cur_domain_2 sa ekh cdom \ \t. sa = switch_thread t \ in_cur_domain_2 t cdom ekh" abbreviation switch_in_cur_domain:: "det_ext state \ bool" where "switch_in_cur_domain s \ switch_in_cur_domain_2 (scheduler_action s) (ekheap s) (cur_domain s)" lemmas switch_in_cur_domain_def = switch_in_cur_domain_2_def definition valid_sched_action_2 where "valid_sched_action_2 sa ekh kh ct cdom \ is_activatable_2 ct sa kh \ weak_valid_sched_action_2 sa ekh kh \ switch_in_cur_domain_2 sa ekh cdom" abbreviation valid_sched_action :: "det_ext state \ bool" where "valid_sched_action s \ valid_sched_action_2 (scheduler_action s) (ekheap s) (kheap s) (cur_thread s) (cur_domain s)" lemmas valid_sched_action_def = valid_sched_action_2_def abbreviation ct_not_queued where "ct_not_queued s \ not_queued (cur_thread s) s" definition "ct_not_in_q_2 queues sa ct \ sa = resume_cur_thread \ not_queued_2 queues ct" abbreviation ct_not_in_q :: "det_ext state \ bool" where "ct_not_in_q s \ ct_not_in_q_2 (ready_queues s) (scheduler_action s) (cur_thread s)" lemmas ct_not_in_q_def = ct_not_in_q_2_def definition valid_sched_2 where "valid_sched_2 queues ekh sa cdom kh ct it \ valid_etcbs_2 ekh kh \ valid_queues_2 queues ekh kh \ ct_not_in_q_2 queues sa ct \ valid_sched_action_2 sa ekh kh ct cdom \ ct_in_cur_domain_2 ct it sa cdom ekh \ valid_blocked_2 queues kh sa ct \ valid_idle_etcb_2 ekh" abbreviation valid_sched :: "det_ext state \ bool" where "valid_sched s \ valid_sched_2 (ready_queues s) (ekheap s) (scheduler_action s) (cur_domain s) (kheap s) (cur_thread s) (idle_thread s)" lemmas valid_sched_def = valid_sched_2_def abbreviation einvs :: "det_ext state \ bool" where "einvs \ invs and valid_list and valid_sched" definition not_cur_thread_2 :: "obj_ref \ scheduler_action \ obj_ref \ bool" where "not_cur_thread_2 thread sa ct \ sa = resume_cur_thread \ thread \ ct" abbreviation not_cur_thread :: "obj_ref \ det_ext state \ bool" where "not_cur_thread thread s \ not_cur_thread_2 thread (scheduler_action s) (cur_thread s)" lemmas not_cur_thread_def = not_cur_thread_2_def definition simple_sched_action_2 :: "scheduler_action \ bool" where "simple_sched_action_2 action \ (case action of switch_thread t \ False | _ \ True)" abbreviation simple_sched_action :: "det_state \ bool" where "simple_sched_action s \ simple_sched_action_2 (scheduler_action s)" lemmas simple_sched_action_def = simple_sched_action_2_def definition schact_is_rct :: "det_ext state \ bool" where "schact_is_rct s \ scheduler_action s = resume_cur_thread" lemma schact_is_rct[elim!]: "schact_is_rct s \ scheduler_action s = resume_cur_thread" apply (simp add: schact_is_rct_def) done lemma schact_is_rct_simple[elim]: "schact_is_rct s \ simple_sched_action s" apply (simp add: simple_sched_action_def schact_is_rct_def) done definition scheduler_act_not_2 where "scheduler_act_not_2 sa t \ sa \ switch_thread t" abbreviation scheduler_act_not :: "obj_ref \ det_ext state \ bool" where "scheduler_act_not t s \ scheduler_act_not_2 (scheduler_action s) t" abbreviation scheduler_act_sane :: "det_ext state \ bool" where "scheduler_act_sane s \ scheduler_act_not_2 (scheduler_action s) (cur_thread s)" lemmas scheduler_act_sane_def = scheduler_act_not_2_def lemmas scheduler_act_not_def = scheduler_act_not_2_def lemmas ct_not_queued_lift = hoare_lift_Pf2[where f="cur_thread" and P="not_queued"] lemmas sch_act_sane_lift = hoare_lift_Pf2[where f="cur_thread" and P="scheduler_act_not"] lemmas not_queued_def = not_queued_2_def lemma valid_etcbs_lift: assumes a: "\P T t. \\s. P (typ_at T t s)\ f \\rv s. P (typ_at T t s)\" and b: "\P. \\s. P (ekheap s)\ f \\rv s. P (ekheap s)\" shows "\valid_etcbs\ f \\rv. valid_etcbs\" apply (rule hoare_lift_Pf[where f="\s. ekheap s", OF _ b]) apply (simp add: valid_etcbs_def) apply (simp add: tcb_at_st_tcb_at[symmetric] tcb_at_typ) apply (wp hoare_vcg_all_lift a) done lemma valid_queues_lift: assumes a: "\Q t. \\s. st_tcb_at Q t s\ f \\rv s. st_tcb_at Q t s\" and c: "\P. \\s. P (ekheap s)\ f \\rv s. P (ekheap s)\" and d: "\P. \\s. P (ready_queues s)\ f \\rv s. P (ready_queues s)\" shows "\valid_queues\ f \\rv. valid_queues\" apply (rule hoare_lift_Pf[where f="\s. ekheap s", OF _ c]) apply (rule hoare_lift_Pf[where f="\s. ready_queues s", OF _ d]) apply (simp add: valid_queues_def) apply (wp hoare_vcg_ball_lift hoare_vcg_all_lift hoare_vcg_conj_lift a) done lemma typ_at_st_tcb_at_lift: assumes typ_lift: "\P T p. \\s. P (typ_at T p s)\ f \\r s. P (typ_at T p s)\" assumes st_lift: "\P. \st_tcb_at P t\ f \\_. st_tcb_at P t\" shows "\\s. \ st_tcb_at P t s\ f \\r s. \ st_tcb_at P t s\" apply (simp add: valid_def obj_at_def st_tcb_at_def) apply clarsimp apply (case_tac "kheap s t") apply (cut_tac P="\x. \ x" and p=t and T="ATCB" in typ_lift) apply (simp add: valid_def obj_at_def) apply force apply (cut_tac P="\x. x" and p=t and T="a_type aa" in typ_lift) apply (cut_tac P="\t. \ P t" in st_lift) apply (simp add: valid_def obj_at_def st_tcb_at_def) apply (drule_tac x=s in spec) apply simp apply (drule_tac x="(a,b)" in bspec) apply simp apply simp apply (subgoal_tac "a_type aa = ATCB") apply (erule a_type_ATCBE) apply simp apply force apply simp done lemma valid_blocked_lift: assumes a: "\Q t. \\s. st_tcb_at Q t s\ f \\rv s. st_tcb_at Q t s\" assumes t: "\P T t. \\s. P (typ_at T t s)\ f \\rv s. P (typ_at T t s)\" and c: "\P. \\s. P (scheduler_action s)\ f \\rv s. P (scheduler_action s)\" and e: "\P. \\s. P (cur_thread s)\ f \\rv s. P (cur_thread s)\" and d: "\P. \\s. P (ready_queues s)\ f \\rv s. P (ready_queues s)\" shows "\valid_blocked\ f \\rv. valid_blocked\" apply (rule hoare_pre) apply (wps c e d) apply (simp add: valid_blocked_def) apply (wp hoare_vcg_ball_lift hoare_vcg_all_lift hoare_vcg_conj_lift static_imp_wp a) apply (rule hoare_convert_imp) apply (rule typ_at_st_tcb_at_lift) apply (wp a t)+ apply (simp add: valid_blocked_def) done lemma ct_not_in_q_lift: assumes a: "\P. \\s. P (scheduler_action s)\ f \\rv s. P (scheduler_action s)\" and b: "\P. \\s. P (ready_queues s)\ f \\rv s. P (ready_queues s)\" and c: "\P. \\s. P (cur_thread s)\ f \\rv s. P (cur_thread s)\" shows "\ct_not_in_q\ f \\rv. ct_not_in_q\" apply (rule hoare_lift_Pf[where f="\s. scheduler_action s", OF _ a]) apply (rule hoare_lift_Pf[where f="\s. ready_queues s", OF _ b]) apply (rule hoare_lift_Pf[where f="\s. cur_thread s", OF _ c]) apply wp done lemma ct_in_cur_domain_lift: assumes a: "\P. \\s. P (ekheap s)\ f \\rv s. P (ekheap s)\" and b: "\P. \\s. P (scheduler_action s)\ f \\rv s. P (scheduler_action s)\" and c: "\P. \\s. P (cur_domain s)\ f \\rv s. P (cur_domain s)\" and d: "\P. \\s. P (cur_thread s)\ f \\rv s. P (cur_thread s)\" and e: "\P. \\s. P (idle_thread s)\ f \\rv s. P (idle_thread s)\" shows "\ct_in_cur_domain\ f \\rv. ct_in_cur_domain\" apply (rule hoare_lift_Pf[where f="\s. ekheap s", OF _ a]) apply (rule hoare_lift_Pf[where f="\s. scheduler_action s", OF _ b]) apply (rule hoare_lift_Pf[where f="\s. cur_domain s", OF _ c]) apply (rule hoare_lift_Pf[where f="\s. cur_thread s", OF _ d]) apply (rule hoare_lift_Pf[where f="\s. idle_thread s", OF _ e]) apply wp done lemma weak_valid_sched_action_lift: assumes a: "\Q t. \\s. st_tcb_at Q t s\ f \\rv s. st_tcb_at Q t s\" assumes c: "\P. \\s. P (scheduler_action s)\ f \\rv s. P (scheduler_action s)\" shows "\weak_valid_sched_action\ f \\rv. weak_valid_sched_action\" apply (rule hoare_lift_Pf[where f="\s. scheduler_action s", OF _ c]) apply (simp add: weak_valid_sched_action_def) apply (wp hoare_vcg_all_lift static_imp_wp a) done lemma switch_in_cur_domain_lift: assumes a: "\Q t. \\s. etcb_at Q t s\ f \\rv s. etcb_at Q t s\" assumes b: "\P. \\s. P (scheduler_action s)\ f \\rv s. P (scheduler_action s)\" assumes c: "\P. \\s. P (cur_domain s)\ f \\rv s. P (cur_domain s)\" shows "\switch_in_cur_domain\ f \\rv. switch_in_cur_domain\" apply (rule hoare_lift_Pf[where f="\s. scheduler_action s", OF _ b]) apply (rule hoare_lift_Pf[where f="\s. cur_domain s", OF _ c]) apply (simp add: switch_in_cur_domain_def in_cur_domain_def) apply (wp hoare_vcg_all_lift static_imp_wp a c) done lemma valid_sched_action_lift: assumes a: "\Q t. \\s. st_tcb_at Q t s\ f \\rv s. st_tcb_at Q t s\" assumes b: "\Q t. \\s. etcb_at Q t s\ f \\rv s. etcb_at Q t s\" assumes c: "\P. \\s. P (scheduler_action s)\ f \\rv s. P (scheduler_action s)\" assumes d: "\P. \\s. P (cur_thread s)\ f \\rv s. P (cur_thread s)\" assumes e: "\Q t. \\s. Q (cur_domain s)\ f \\rv s. Q (cur_domain s)\" shows "\valid_sched_action\ f \\rv. valid_sched_action\" apply (rule hoare_lift_Pf[where f="\s. cur_thread s", OF _ d]) apply (simp add: valid_sched_action_def) apply (rule hoare_vcg_conj_lift) apply (rule hoare_lift_Pf[where f="\s. scheduler_action s", OF _ c]) apply (simp add: is_activatable_def) apply (wp weak_valid_sched_action_lift switch_in_cur_domain_lift static_imp_wp a b c d e)+ done lemma valid_sched_lift: assumes a: "\Q t. \\s. st_tcb_at Q t s\ f \\rv s. st_tcb_at Q t s\" assumes b: "\Q t. \\s. etcb_at Q t s\ f \\rv s. etcb_at Q t s\" assumes c: "\P T t. \\s. P (typ_at T t s)\ f \\rv s. P (typ_at T t s)\" assumes d: "\P. \\s. P (ekheap s)\ f \\rv s. P (ekheap s)\" assumes e: "\P. \\s. P (scheduler_action s)\ f \\rv s. P (scheduler_action s)\" assumes f: "\P. \\s. P (ready_queues s)\ f \\rv s. P (ready_queues s)\" assumes g: "\P. \\s. P (cur_domain s)\ f \\rv s. P (cur_domain s)\" assumes h: "\P. \\s. P (cur_thread s)\ f \\rv s. P (cur_thread s)\" assumes i: "\P. \\s. P (idle_thread s)\ f \\rv s. P (idle_thread s)\" shows "\valid_sched\ f \\rv. valid_sched\" apply (simp add: valid_sched_def) apply (wp valid_etcbs_lift valid_queues_lift ct_not_in_q_lift ct_in_cur_domain_lift valid_sched_action_lift valid_blocked_lift a b c d e f g h i hoare_vcg_conj_lift) done lemma valid_etcbs_tcb_etcb: "\ valid_etcbs s; kheap s ptr = Some (TCB tcb) \ \ \etcb. ekheap s ptr = Some etcb" by (force simp: valid_etcbs_def is_etcb_at_def st_tcb_at_def obj_at_def) lemma valid_etcbs_get_tcb_get_etcb: "\ valid_etcbs s; get_tcb ptr s = Some tcb \ \ \etcb. get_etcb ptr s = Some etcb" apply (clarsimp simp: valid_etcbs_def valid_etcbs_def st_tcb_at_def obj_at_def is_etcb_at_def get_etcb_def get_tcb_def split: option.splits if_split) apply (erule_tac x=ptr in allE) apply (clarsimp simp: get_etcb_def split: option.splits kernel_object.splits)+ done lemma valid_etcbs_ko_etcb: "\ valid_etcbs s; kheap s ptr = Some ko \ \ \tcb. (ko = TCB tcb = (\etcb. ekheap s ptr = Some etcb))" apply (clarsimp simp: valid_etcbs_def st_tcb_at_def obj_at_def is_etcb_at_def) apply (erule_tac x="ptr" in allE) apply auto done lemma ekheap_tcb_at: "\ekheap s x = Some y; valid_etcbs s\ \ tcb_at x s" by (fastforce simp: valid_etcbs_def is_etcb_at_def st_tcb_at_def obj_at_def is_tcb_def) lemma tcb_at_is_etcb_at: "\tcb_at t s; valid_etcbs s\ \ is_etcb_at t s" by (simp add: valid_etcbs_def tcb_at_st_tcb_at) lemma tcb_at_ekheap_dom: "\tcb_at x s; valid_etcbs s\ \ (\etcb. ekheap s x = Some etcb)" by (auto simp: is_etcb_at_def dest: tcb_at_is_etcb_at) lemma ekheap_kheap_dom: "\ekheap s x = Some etcb; valid_etcbs s\ \ \tcb. kheap s x = Some (TCB tcb)" by (fastforce simp: valid_etcbs_def st_tcb_at_def obj_at_def is_etcb_at_def) end