diff --git a/proof/crefine/RISCV64/SchedContext_C.thy b/proof/crefine/RISCV64/SchedContext_C.thy index df06eaca2e..b53b8efd9d 100644 --- a/proof/crefine/RISCV64/SchedContext_C.thy +++ b/proof/crefine/RISCV64/SchedContext_C.thy @@ -11,90 +11,6 @@ begin context kernel_m begin -lemma getRefillSize_exs_valid[wp]: - "\(=) s\ getRefillSize scPtr \\r. (=) s\" - by (wpsimp simp: getRefillSize_def) - -crunch getRefillSize - for (empty_fail) empty_fail[wp] - -lemma refill_add_tail_ccorres: - "ccorres dc xfdc - (active_sc_at' scPtr and invs') - (\\sc = Ptr scPtr\ \ {s'. crefill_relation new (refill_' s')}) [] - (refillAddTail scPtr new) (Call refill_add_tail_'proc)" - supply sched_context_C_size[simp del] refill_C_size[simp del] len_bit0[simp del] - - apply (simp add: refillAddTail_def) - apply (rule ccorres_symb_exec_l'[rotated, OF _ getRefillSize_sp]; wpsimp) - apply (rule ccorres_symb_exec_l'[rotated, OF _ get_sc_sp']; wpsimp?) - apply (rule ccorres_symb_exec_l'[rotated, OF _ assert_sp]; wpsimp) - - apply (cinit' lift: sc_' refill_' simp: updateRefillIndex_def) - apply (rule ccorres_move_c_guard_sc) - apply (ctac add: refill_next_ccorres) - apply (rule ccorres_move_c_guard_sc) - apply clarsimp - apply (rule ccorres_split_nothrow[where r'=dc and xf'=xfdc]) - apply (rule_tac P=\ in updateSchedContext_ccorres_lemma2) - apply vcg - apply fastforce - apply clarsimp - apply (rule_tac sc'="scRefillTail_C_update (\_. new_tail) sc'" - in rf_sr_sc_update_no_refill_buffer_update2; - fastforce?) - apply (clarsimp simp: typ_heap_simps') - apply (clarsimp simp: csched_context_relation_def) - apply (fastforce intro: refill_buffer_relation_sc_no_refills_update) - apply ceqv - apply (rule_tac P="\sc. scRefillTail sc = unat new_tail" - and P'="active_sc_at' scPtr and invs'" - in updateSchedContext_ccorres_lemma3) - apply vcg - apply fastforce - apply normalise_obj_at' - apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) - apply (intro conjI) - apply (fastforce dest: sc_at_h_t_valid[rotated]) - apply (rule disjCI2) - apply (rule_tac n="length (scRefills sc)" in array_assertion_shrink_right) - apply (fastforce intro: sc_at_array_assertion' - simp: valid_sched_context'_def MIN_REFILLS_def) - apply (clarsimp simp: typ_heap_simps valid_sched_context'_def obj_at_simps - active_sc_at'_def csched_context_relation_def) - apply (frule rf_sr_refill_buffer_relation) - apply (frule_tac n="scRefillTail sc" in h_t_valid_refill; fastforce?) - apply (clarsimp simp: valid_sched_context'_def active_sc_at'_def obj_at'_def) - apply fastforce - apply (clarsimp simp: sc_ptr_to_crefill_ptr_def typ_heap_simps csched_context_relation_def) - apply (clarsimp simp: ret__ptr_to_struct_refill_C_'_update_rf_sr_helper - cong: StateSpace.state.fold_congs) - apply (rule_tac old_sc=sc and n="scRefillTail sc" and refill=new and refill'=refill - in rf_sr_refill_update2; - fastforce?) - apply (clarsimp simp: valid_sched_context'_def active_sc_at'_def obj_at'_def) - apply (fastforce simp: sched_context.expand) - apply (clarsimp simp: typ_heap_simps sc_ptr_to_crefill_ptr_def csched_context_relation_def) - apply (clarsimp simp: csched_context_relation_def) - apply ((rule hoare_vcg_conj_lift - | wpsimp wp: updateSchedContext_active_sc_at' updateSchedContext_refills_invs' - | wpsimp wp: updateSchedContext_wp)+)[1] - apply vcg - apply (wpsimp wp: getRefillNext_wp) - apply vcg - apply normalise_obj_at' - apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) - apply (elim conjE) - apply (frule (1) length_scRefills_bounded) - apply (intro conjI) - apply fastforce - apply (clarsimp simp: valid_sched_context'_def word_bits_len_of refillSizeBytes_def) - apply (fastforce simp: obj_at'_def objBits_simps opt_map_def ps_clear_def) - apply (fastforce simp: valid_sched_context'_def refillNext_def refillSize_def split: if_splits) - apply (clarsimp simp: valid_sched_context_size'_def) - apply (fastforce dest: obj_at_cslift_sc simp: csched_context_relation_def typ_heap_simps) - done - lemma maybe_add_empty_tail_ccorres: "ccorres dc xfdc (active_sc_at' scPtr and invs') \\sc = Ptr scPtr\ [] diff --git a/proof/crefine/RISCV64/Schedule_C.thy b/proof/crefine/RISCV64/Schedule_C.thy index a312834f69..5b60a6cf79 100644 --- a/proof/crefine/RISCV64/Schedule_C.thy +++ b/proof/crefine/RISCV64/Schedule_C.thy @@ -514,7 +514,7 @@ lemma refill_size_length_scRefills_helper: simp: refillSizeBytes_def word_bits_def) done -crunch getSchedContext +crunch getSchedContext, isRoundRobin for (empty_fail) empty_fail[wp] lemma refill_size_ccorres: @@ -593,35 +593,957 @@ lemma length_scRefills_bounded: apply (clarsimp simp: word_bits_def untypedBits_defs) done -lemma sc_released_ccorres: +lemma switchSchedContext_ccorres: + "ccorres dc xfdc \ UNIV [] switchSchedContext (Call switchSchedContext_'proc)" +sorry (* FIXME RT: switchSchedContext_ccorres *) + +lemma head_refill_overrun_ccorres: "ccorres (\rv rv'. rv = to_bool rv') ret__unsigned_long_' - (active_sc_at' scPtr and valid_objs') \\sc = Ptr scPtr\ [] - (scReleased scPtr) (Call sc_released_'proc)" - apply (cinit simp: readScReleased_def scActive_def[symmetric] gets_the_if_distrib) - apply (ctac add: sc_active_ccorres) - apply (rule ccorres_cond[where R=\]) - apply (clarsimp simp: to_bool_def) - apply (simp flip: refillReady_def) - apply (rule ccorres_add_return2) - apply (ctac add: refill_ready_ccorres) - apply (fastforce intro: ccorres_return_C) - apply wpsimp + (active_sc_at' scPtr and valid_objs' and no_0_obj') + (\\sc = Ptr scPtr\ \ \usage = \usage\) [] + (gets_the (headRefillOverrun scPtr usage)) + (Call head_refill_overrun_'proc)" + supply sched_context_C_size[simp del] refill_C_size[simp del] + Collect_const [simp del] + apply (clarsimp simp: headRefillOverrun_def getRefillHead_def[symmetric]) + apply (cinit' lift: sc_' usage_' simp: max_minus_one_word64) + apply (rule_tac xf'="\s. h_val (hrs_mem (t_hrs_' (globals s))) (ret__ptr_to_struct_refill_C_' s)" + in ccorres_split_nothrow) + apply (rule ccorres_call) + apply (rule refill_head_ccorres) + apply clarsimp + apply clarsimp + apply clarsimp + apply ceqv + apply (rename_tac head head') + apply (rule ccorres_Guard_Seq) + apply csymbr + apply csymbr + apply (rule_tac P="rAmount head \ usage" in ccorres_cases) + apply (rule ccorres_cond_seq) + apply (rule ccorres_cond_true) + apply (rule ccorres_rhs_assoc) + apply (rule_tac val="usToTicks maxPeriodUs" + and xf'=ret__unsigned_longlong_' + in ccorres_symb_exec_r_known_rv[where R=\ and R'=UNIV]) + apply (rule conseqPre, vcg) + apply (clarsimp simp: maxPeriodUs_def timer_defs unat_max_word) + apply ceqv + apply csymbr + apply (fastforce intro: ccorres_return_C) apply vcg + apply (rule ccorres_cond_seq) + apply (rule ccorres_cond_false) + apply ccorres_rewrite apply (fastforce intro: ccorres_return_C) apply wpsimp apply vcg + apply (clarsimp simp: active_sc_at'_rewrite) + apply (normalise_obj_at', rename_tac sc) + apply (frule (1) sc_ko_at_valid_objs_valid_sc') apply clarsimp + apply (frule (1) length_scRefills_bounded) + apply (frule (1) obj_at_cslift_sc) + apply (frule rf_sr_refill_buffer_relation) + apply (frule_tac n="scRefillHead sc" in h_t_valid_refill; fastforce?) + apply (fastforce simp: valid_sched_context'_def is_active_sc'_def obj_at'_def + opt_map_def opt_pred_def) + apply fastforce + apply (frule_tac n=0 in h_t_valid_refill; fastforce?) + apply (fastforce simp: valid_sched_context'_def is_active_sc'_def obj_at'_def + opt_map_def opt_pred_def) + apply fastforce + apply (clarsimp simp: typ_heap_simps csched_context_relation_def crefill_relation_def + sc_ptr_to_crefill_ptr_def maxReleaseTime_def timer_defs unat_max_word + maxBound_word split: if_splits) done -lemma switchSchedContext_ccorres: - "ccorres dc xfdc \ UNIV [] switchSchedContext (Call switchSchedContext_'proc)" -sorry (* FIXME RT: switchSchedContext_ccorres *) +lemma refillPopHead_scPeriod[wp]: + "refillPopHead scPtr' \obj_at' (\sc. P (scPeriod sc)) scPtr\" + unfolding refillPopHead_def + apply (wpsimp wp: updateSchedContext_wp getRefillNext_wp) + by (clarsimp simp: obj_at'_def opt_map_def objBits_simps ps_clear_upd split: if_splits) + +crunch getRefillSize + for (empty_fail) empty_fail[wp] + +lemma getRefillSize_exs_valid[wp]: + "\(=) s\ getRefillSize scPtr \\r. (=) s\" + by (wpsimp simp: getRefillSize_def) + +lemma refill_add_tail_ccorres: + "ccorres dc xfdc + (active_sc_at' scPtr and invs') + (\\sc = Ptr scPtr\ \ {s'. crefill_relation new (refill_' s')}) [] + (refillAddTail scPtr new) (Call refill_add_tail_'proc)" + supply sched_context_C_size[simp del] refill_C_size[simp del] len_bit0[simp del] + + unfolding refillAddTail_def haskell_assert_def + apply (rule ccorres_symb_exec_l'[rotated, OF _ getRefillSize_sp]; (solves wpsimp)?) + apply (rule ccorres_symb_exec_l'[rotated, OF _ get_sc_sp']; (solves wpsimp)?) + apply (rule ccorres_symb_exec_l'[rotated, OF _ assert_sp]; (solves wpsimp)?) + + apply (cinit' lift: sc_' refill_' simp: updateRefillIndex_def) + apply (rule ccorres_move_c_guard_sc) + apply (ctac add: refill_next_ccorres) + apply (rule ccorres_move_c_guard_sc) + apply clarsimp + apply (rule ccorres_split_nothrow[where r'=dc and xf'=xfdc]) + apply (rule_tac P=\ in updateSchedContext_ccorres_lemma2) + apply vcg + apply fastforce + apply clarsimp + apply (rule_tac sc'="scRefillTail_C_update (\_. new_tail) sc'" + in rf_sr_sc_update_no_refill_buffer_update2; + fastforce?) + apply (clarsimp simp: typ_heap_simps') + apply (clarsimp simp: csched_context_relation_def) + apply (fastforce intro: refill_buffer_relation_sc_no_refills_update) + apply ceqv + apply (rule_tac P="\sc. scRefillTail sc = unat new_tail" + and P'="active_sc_at' scPtr and invs'" + in updateSchedContext_ccorres_lemma3) + apply vcg + apply fastforce + apply normalise_obj_at' + apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) + apply (intro conjI) + apply (fastforce dest: sc_at_h_t_valid[rotated]) + apply (rule disjCI2) + apply (rule_tac n="length (scRefills sc)" in array_assertion_shrink_right) + apply (fastforce intro: sc_at_array_assertion' + simp: valid_sched_context'_def MIN_REFILLS_def) + apply (clarsimp simp: typ_heap_simps valid_sched_context'_def obj_at_simps + active_sc_at'_def csched_context_relation_def) + apply (frule rf_sr_refill_buffer_relation) + apply (frule_tac n="scRefillTail sc" in h_t_valid_refill; fastforce?) + apply (clarsimp simp: valid_sched_context'_def active_sc_at'_def obj_at'_def) + apply fastforce + apply (clarsimp simp: sc_ptr_to_crefill_ptr_def typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: ret__ptr_to_struct_refill_C_'_update_rf_sr_helper + cong: StateSpace.state.fold_congs) + apply (rule_tac old_sc=sc and n="scRefillTail sc" and refill=new and refill'=refill + in rf_sr_refill_update2; + fastforce?) + apply (clarsimp simp: valid_sched_context'_def active_sc_at'_def obj_at'_def) + apply (fastforce simp: sched_context.expand) + apply (clarsimp simp: typ_heap_simps sc_ptr_to_crefill_ptr_def csched_context_relation_def) + apply (clarsimp simp: csched_context_relation_def) + apply ((rule hoare_vcg_conj_lift + | wpsimp wp: updateSchedContext_active_sc_at' updateSchedContext_refills_invs' + | wpsimp wp: updateSchedContext_wp)+)[1] + apply vcg + apply (wpsimp wp: getRefillNext_wp) + apply vcg + apply normalise_obj_at' + apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) + apply (elim conjE) + apply (frule (1) length_scRefills_bounded) + apply (intro conjI) + apply fastforce + apply (clarsimp simp: valid_sched_context'_def word_bits_len_of refillSizeBytes_def) + apply (fastforce simp: obj_at'_def objBits_simps opt_map_def ps_clear_def) + apply (fastforce simp: valid_sched_context'_def refillNext_def refillSize_def split: if_splits) + apply (clarsimp simp: valid_sched_context_size'_def) + apply (fastforce dest: obj_at_cslift_sc simp: csched_context_relation_def typ_heap_simps) + done + +lemma schedule_used_ccorres: + "ccorres dc xfdc + (active_sc_at' scPtr and invs') + (\\sc = Ptr scPtr\ \ \crefill_relation new \new\) hs + (scheduleUsed scPtr new) (Call schedule_used_'proc)" + (is "ccorres _ _ ?abs _ _ _ _") + supply sched_context_C_size[simp del] refill_C_size[simp del] + unfolding scheduleUsed_def haskell_assert_def + apply (rule ccorres_symb_exec_l'[OF _ _ scActive_sp]; (solves wpsimp)?) + apply (rule ccorres_symb_exec_l'[OF _ _ assert_sp]; (solves wpsimp)?) + apply (cinit' lift: sc_' new_') + apply (rule_tac xf'="\s. h_val (hrs_mem (t_hrs_' (globals s))) (ret__ptr_to_struct_refill_C_' s)" + in ccorres_split_nothrow_call) + apply (fastforce intro: refill_tail_ccorres, fastforce+) + apply ceqv + apply (rule ccorres_Guard_Seq) + apply csymbr + apply (rule ccorres_cond[where R=\]) + apply (clarsimp simp: crefill_relation_def) + apply (clarsimp simp: setRefillTl_def updateRefillTl_def) + apply (rule updateSchedContext_ccorres_lemma3 + [where P="\sc. scRefillTail sc < length (scRefills sc)" and P'="?abs"]) + apply vcg + apply fastforce + apply clarsimp + apply (rename_tac sc sc') + apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) + apply (frule rf_sr_refill_buffer_relation) + apply (frule_tac n="scRefillTail sc" in h_t_valid_refill, fastforce+) + apply (intro conjI) + apply (clarsimp simp: typ_heap_simps) + apply (rule disjCI2) + apply (rule_tac n="length (scRefills sc)" in array_assertion_shrink_right) + apply (fastforce intro: sc_at_array_assertion') + apply (clarsimp simp: typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: sc_ptr_to_crefill_ptr_def typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: crefill_relation_def csched_context_relation_def) + apply (frule rf_sr_refill_buffer_relation) + apply (frule h_t_valid_clift_Some_iff[THEN iffD1]) + apply (clarsimp cong: StateSpace.state.fold_congs) + apply (rule_tac old_sc=sc and n="scRefillTail sc" + and fa="rAmount_update (\_. rAmount rv + rAmount new)" + in rf_sr_refill_update2; fastforce?) + apply (fastforce simp: sched_context.expand) + apply (clarsimp simp: typ_heap_simps) + apply (fastforce simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def crefill_relation_def) + apply (fastforce simp: csched_context_relation_def) + apply (fastforce dest: crefill_relationD + simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def crefill_relation_def) + apply (ctac add: refill_full_ccorres) + apply (rule ccorres_cond[where R=\]) + apply(clarsimp simp: to_bool_def) + apply (ctac add: refill_add_tail_ccorres) + apply (rule ccorres_rhs_assoc) + apply (rule ccorres_split_nothrow) + apply (clarsimp simp: setRefillTl_def updateRefillTl_def) + apply (rule updateSchedContext_ccorres_lemma3 + [where P="\sc. scRefillTail sc < length (scRefills sc)" and P'="?abs"]) + apply vcg + apply fastforce + apply clarsimp + apply (rename_tac sc sc') + apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) + apply (frule rf_sr_refill_buffer_relation) + apply (frule_tac n="scRefillTail sc" in h_t_valid_refill, fastforce+) + apply (intro conjI) + apply (clarsimp simp: typ_heap_simps) + apply (rule disjCI2) + apply (rule_tac n="length (scRefills sc)" in array_assertion_shrink_right) + apply (fastforce intro: sc_at_array_assertion') + apply (clarsimp simp: typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: sc_ptr_to_crefill_ptr_def typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: crefill_relation_def csched_context_relation_def) + apply (frule rf_sr_refill_buffer_relation) + apply (frule h_t_valid_clift_Some_iff[THEN iffD1]) + apply (clarsimp cong: StateSpace.state.fold_congs) + apply (rule_tac old_sc=sc and n="scRefillTail sc" + and fa="rTime_update (\_. rTime new - rAmount rv)" + in rf_sr_refill_update2; fastforce?) + apply (fastforce simp: sched_context.expand) + apply (clarsimp simp: typ_heap_simps) + apply (fastforce simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def crefill_relation_def) + apply (fastforce simp: csched_context_relation_def) + apply (fastforce dest: crefill_relationD + simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def crefill_relation_def) + apply ceqv + apply (clarsimp simp: setRefillTl_def updateRefillTl_def) + apply (rule updateSchedContext_ccorres_lemma3 + [where P="\sc. scRefillTail sc < length (scRefills sc)" and P'="?abs"]) + apply vcg + apply fastforce + apply clarsimp + apply (rename_tac sc sc') + apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) + apply (frule rf_sr_refill_buffer_relation) + apply (frule_tac n="scRefillTail sc" in h_t_valid_refill, fastforce+) + apply (intro conjI) + apply (clarsimp simp: typ_heap_simps) + apply (rule disjCI2) + apply (rule_tac n="length (scRefills sc)" in array_assertion_shrink_right) + apply (fastforce intro: sc_at_array_assertion') + apply (clarsimp simp: typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: sc_ptr_to_crefill_ptr_def typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: crefill_relation_def csched_context_relation_def) + apply (frule rf_sr_refill_buffer_relation) + apply (frule h_t_valid_clift_Some_iff[THEN iffD1]) + apply (clarsimp cong: StateSpace.state.fold_congs) + apply (rule_tac old_sc=sc and n="scRefillTail sc" + and fa="rAmount_update (\_. rAmount rv + rAmount new)" + in rf_sr_refill_update2; fastforce?) + apply (fastforce simp: sched_context.expand) + apply (clarsimp simp: typ_heap_simps) + apply (fastforce simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def crefill_relation_def) + apply (fastforce simp: csched_context_relation_def) + apply (fastforce dest: crefill_relationD + simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def crefill_relation_def) + apply (rule_tac Q'="\_. ?abs and active_sc_at' scPtr" in hoare_post_imp) + apply (clarsimp simp: active_sc_at'_def) + apply normalise_obj_at' + apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) + apply (clarsimp simp: valid_sched_context'_def) + apply (wpsimp wp: updateRefillHd_valid_objs') + apply vcg + apply (wpsimp wp: getRefillFull_wp) + apply vcg + apply (wpsimp wp: getRefillTail_wp) + apply vcg + apply (normalise_obj_at', rename_tac sc) + apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) + apply clarsimp + apply (frule (1) length_scRefills_bounded) + apply (frule (1) obj_at_cslift_sc) + apply (frule rf_sr_refill_buffer_relation) + apply (frule_tac n="scRefillTail sc" in h_t_valid_refill; fastforce?) + apply (clarsimp simp: valid_sched_context'_def) + apply fastforce + apply (fastforce simp: crefill_relation_def typ_heap_simps csched_context_relation_def + sc_ptr_to_crefill_ptr_def valid_sched_context'_def active_sc_at'_rewrite + intro: valid_objs'_valid_refills') + done + +lemma charge_entire_head_refill_ccorres: + "ccorres (=) ret__unsigned_longlong_' + (active_sc_at' scPtr and invs') + (\\sc = Ptr scPtr\ \ \\usage = usage\) [] + (chargeEntireHeadRefill scPtr usage) (Call charge_entire_head_refill_'proc)" + (is "ccorres _ _ ?abs _ _ _ _") + supply sched_context_C_size[simp del] refill_C_size[simp del] + Collect_const [simp del] + apply (cinit lift: sc_' usage_') + apply (rule_tac xf'="\s. h_val (hrs_mem (t_hrs_' (globals s))) (ret__ptr_to_struct_refill_C_' s)" + in ccorres_split_nothrow) + apply (rule ccorres_call) + apply (rule refill_head_ccorres) + apply clarsimp + apply clarsimp + apply clarsimp + apply ceqv + apply (rename_tac head head') + apply (rule ccorres_Guard_Seq) + apply csymbr + apply (rule ccorres_pre_getObject_sc, rename_tac abs_sc) + apply (ctac add: refill_single_ccorres) + apply (clarsimp simp: if_to_top_of_bind) + apply (rule ccorres_cond_seq) + apply (rule ccorres_cond[where R=\]) + apply (fastforce simp: to_bool_def) + apply (rule ccorres_split_nothrow) + apply (clarsimp simp: setRefillHd_def updateRefillHd_def) + apply (rule_tac P="\sc. scRefillHead sc < length (scRefills sc) + \ scPeriod sc = scPeriod abs_sc" + in updateSchedContext_ccorres_lemma3[where P'="?abs"]) + apply vcg + apply fastforce + apply clarsimp + apply (rename_tac sc sc') + apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) + apply (frule rf_sr_refill_buffer_relation) + apply (frule_tac n="scRefillHead sc" in h_t_valid_refill, fastforce+) + apply (intro conjI) + apply (clarsimp simp: typ_heap_simps) + apply (rule disjCI2) + apply (rule_tac n="length (scRefills sc)" in array_assertion_shrink_right) + apply (fastforce intro: sc_at_array_assertion') + apply (clarsimp simp: typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def csched_context_relation_def) + apply (clarsimp simp: typ_heap_simps csched_context_relation_def) + apply normalise_obj_at' + apply (clarsimp simp: crefill_relation_def csched_context_relation_def) + apply (frule rf_sr_refill_buffer_relation) + apply (frule h_t_valid_clift_Some_iff[THEN iffD1]) + apply (clarsimp cong: StateSpace.state.fold_congs) + apply (rule_tac old_sc=sc and n="scRefillHead sc" + and fa="rTime_update (\_. rTime head + scPeriod sc)" + in rf_sr_refill_update2; fastforce?) + apply (fastforce simp: sched_context.expand) + apply (clarsimp simp: typ_heap_simps) + apply (fastforce simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def crefill_relation_def) + apply (fastforce simp: csched_context_relation_def) + apply normalise_obj_at' + apply (fastforce dest: crefill_relationD + simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def crefill_relation_def) + apply ceqv + apply (fastforce intro: ccorres_return_C) + apply wpsimp + apply vcg + apply (clarsimp simp: bind_assoc) + apply (rule ccorres_rhs_assoc)+ + apply (ctac add: refill_pop_head_ccorres) + apply (rule ccorres_move_c_guard_sc) + apply (rule_tac val="refill_C (rTime rv + scPeriod abs_sc) (rAmount rv)" + and xf'=old_head_' + and R="obj_at' (\sc. scPeriod sc = scPeriod abs_sc) scPtr" + in ccorres_symb_exec_r_known_rv[where R'=UNIV]) + apply (rule conseqPre, vcg) + apply normalise_obj_at' + apply (frule (1) obj_at_cslift_sc) + apply (clarsimp simp: crefill_relation_def typ_heap_simps csched_context_relation_def) + apply (metis rTime_C_update.simps refill_C_idupdates(1)) + apply ceqv + apply (ctac add: schedule_used_ccorres) + apply (fastforce intro: ccorres_return_C) + apply wpsimp + apply (vcg exspec=schedule_used_modifies) + apply vcg + apply wpsimp + apply (vcg exspec=refill_pop_head_modifies) + apply (wpsimp simp: refillSingle_def) + apply (vcg exspec=refill_single_modifies) + apply wpsimp + apply vcg + apply (clarsimp simp: active_sc_at'_rewrite) + apply (normalise_obj_at', rename_tac sc) + apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) + apply clarsimp + apply (frule (1) length_scRefills_bounded) + apply (frule (1) obj_at_cslift_sc) + apply (frule rf_sr_refill_buffer_relation) + apply (frule_tac n="scRefillHead sc" in h_t_valid_refill; fastforce?) + apply (fastforce simp: valid_sched_context'_def is_active_sc'_def obj_at'_def + opt_map_def opt_pred_def) + apply fastforce + apply (frule_tac n=0 in h_t_valid_refill; fastforce?) + apply (fastforce simp: valid_sched_context'_def is_active_sc'_def obj_at'_def + opt_map_def opt_pred_def) + apply fastforce + apply (auto simp: typ_heap_simps csched_context_relation_def crefill_relation_def + sc_ptr_to_crefill_ptr_def is_active_sc'_def obj_at'_def opt_pred_def opt_map_def + valid_sched_context'_def) + done + +lemma no_ofail_headRefillOverrun[wp]: + "no_ofail (sc_at' scPtr) (headRefillOverrun scPtr usage)" + unfolding headRefillOverrun_def + by wpsimp + +lemma no_ofail_readRefillFull[wp]: + "no_ofail (sc_at' scPtr) (readRefillFull scPtr)" + unfolding readRefillFull_def + by (wpsimp wp: no_ofail_readSchedContext) + +lemmas no_fail_getRefillFull[wp] = + no_ofail_gets_the[OF no_ofail_readRefillFull, simplified getRefillFull_def[symmetric]] + +lemma no_ofail_readRefillTail[wp]: + "no_ofail (sc_at' scPtr) (readRefillTail scPtr)" + unfolding readRefillTail_def + by (wpsimp wp: no_ofail_readSchedContext) + +lemmas no_fail_getRefillTail[wp] = + no_ofail_gets_the[OF no_ofail_readRefillTail, simplified getRefillTail_def[symmetric]] + +lemma readRefillSingle_wp[wp]: + "\\s. \sc. scs_of' s scPtr = Some sc \ Q (scRefillHead sc = scRefillTail sc) s\ + readRefillSingle scPtr + \Q\" + unfolding readRefillTail_def readSchedContext_def + apply (wpsimp wp: set_sc'.readObject_wp) + apply (clarsimp simp: opt_map_def obj_at'_def) + done + +lemmas refillSingle_wp[wp] = + ovalid_gets_the[OF readRefillSingle_wp, simplified refillSingle_def[symmetric]] + +lemmas no_fail_refillSingle[wp] = + no_ofail_gets_the[OF no_ofail_readRefillSingle, simplified refillSingle_def[symmetric]] + +lemma no_fail_updateRefillTl[wp]: + "no_fail (sc_at' scPtr) (updateRefillTl scPtr f)" + unfolding updateRefillTl_def + apply wpsimp + by (clarsimp simp: opt_pred_def opt_map_def obj_at'_def objBits_simps) + +lemma no_fail_updateRefillIndex[wp]: + "no_fail (sc_at' scPtr) (updateRefillIndex scPtr f idx)" + unfolding updateRefillIndex_def + apply wpsimp + by (clarsimp simp: opt_pred_def opt_map_def obj_at'_def objBits_simps) + +lemma no_fail_chargeEntireHeadRefill: + "no_fail (active_sc_at' scPtr and valid_objs') (chargeEntireHeadRefill scPtr usage)" + unfolding chargeEntireHeadRefill_def scheduleUsed_def refillAddTail_def + apply (wpsimp wp: getRefillNext_wp getRefillSize_wp getRefillFull_wp) + apply (rule_tac Q'="\_. active_sc_at' scPtr and valid_objs'" in hoare_post_imp) + apply (clarsimp simp: active_sc_at'_rewrite) + apply normalise_obj_at' + apply (frule (1) sc_ko_at_valid_objs_valid_sc') + apply (clarsimp simp: opt_pred_def opt_map_def obj_at'_def objBits_simps + valid_sched_context'_def is_active_sc'_def + split: if_splits) + apply wpsimp+ + apply (clarsimp simp: active_sc_at'_rewrite split: if_splits) + done + +crunch chargeEntireHeadRefill + for invs'[wp]: invs' + (wp: crunch_wps) + +lemma handle_overrun_ccorres: + "ccorres (=) ret__unsigned_longlong_' + (active_sc_at' scPtr and invs') (\\sc = Ptr scPtr\ \ \\usage = usage\) hs + (handleOverrun scPtr usage) (Call handle_overrun_'proc)" + supply sched_context_C_size[simp del] refill_C_size[simp del] + Collect_const [simp del] + apply (cinit lift: sc_' usage_' simp: runReaderT_def whileAnno_def) + apply (rule ccorres_symb_exec_r) + apply (rule ccorres_add_return2) + apply (rule ccorres_rhs_assoc2) + apply (rule ccorres_split_nothrow_novcg) + apply (rule_tac rrel="(=)" and xf="current_usage_'" + and G="\_. active_sc_at' scPtr and invs'" + in ccorres_While'[where G'=UNIV]) + apply (rule ccorres_guard_imp) + apply (ctac add: charge_entire_head_refill_ccorres) + apply fastforce + apply fastforce + apply (rule ccorres_guard_imp) + apply (ctac add: head_refill_overrun_ccorres) + apply fastforce + apply fastforce + apply wpsimp + apply (clarsimp simp: active_sc_at'_rewrite) + apply wpsimp + apply (rule conseqPre) + apply clarsimp + apply (rule_tac s=s and xf=current_usage_' in ccorres_to_vcg_Normal) + apply (fastforce intro: ccorres_call[OF charge_entire_head_refill_ccorres]) + apply (wpsimp wp: no_fail_chargeEntireHeadRefill) + apply fastforce + apply (rule conseqPre, vcg) + apply (clarsimp simp: active_sc_at'_rewrite) + apply normalise_obj_at' + apply (rename_tac sc) + apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) + apply (frule rf_sr_refill_buffer_relation) + apply (frule_tac n="scRefillHead sc" in h_t_valid_refill, fastforce+) + apply (clarsimp simp: valid_sched_context'_def is_active_sc'_def obj_at'_def + opt_pred_def opt_map_def + split: option.splits) + apply fastforce + apply (frule (1) obj_at_cslift_sc) + apply (intro conjI) + apply (clarsimp simp: typ_heap_simps) + apply (rule disjCI2) + apply (rule_tac n="length (scRefills sc)" in array_assertion_shrink_right) + apply (fastforce intro: sc_at_array_assertion' + simp: valid_sched_context'_def is_active_sc'_def obj_at'_def + opt_pred_def opt_map_def + split: option.splits) + apply (clarsimp simp: valid_sched_context'_def is_active_sc'_def obj_at'_def + opt_pred_def opt_map_def typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def csched_context_relation_def) + apply (clarsimp simp: timer_defs unat_max_word) + apply ceqv + apply (fastforce intro: ccorres_return_C') + apply wpsimp + apply (clarsimp simp: guard_is_UNIV_def) + apply vcg + apply (rule conseqPre, vcg) + apply clarsimp + apply fastforce + done + +lemma head_refill_insufficient_ccorres: + "ccorres (\rv rv'. rv = to_bool rv') ret__unsigned_long_' + (active_sc_at' scPtr and valid_objs' and no_0_obj') + \\sc = Ptr scPtr\ [] + (gets_the (refillHdInsufficient scPtr )) + (Call head_refill_insufficient_'proc)" + supply sched_context_C_size[simp del] refill_C_size[simp del] + Collect_const [simp del] + unfolding refillHdInsufficient_def + apply (cinit' lift: sc_' simp: getRefillHead_def[symmetric]) + apply (rule_tac xf'="\s. h_val (hrs_mem (t_hrs_' (globals s))) (ret__ptr_to_struct_refill_C_' s)" + in ccorres_split_nothrow) + apply (fastforce intro: ccorres_call[OF refill_head_ccorres]) + apply ceqv + apply csymbr + apply (rule ccorres_Guard) + apply (fastforce intro: ccorres_return_C) + apply wpsimp + apply vcg + apply (clarsimp simp: active_sc_at'_rewrite) + apply (normalise_obj_at', rename_tac sc) + apply (frule (1) sc_ko_at_valid_objs_valid_sc') + apply clarsimp + apply (frule (1) length_scRefills_bounded) + apply (frule (1) obj_at_cslift_sc) + apply (frule rf_sr_refill_buffer_relation) + apply (frule_tac n="scRefillHead sc" in h_t_valid_refill; fastforce?) + apply (fastforce simp: valid_sched_context'_def is_active_sc'_def obj_at'_def + opt_map_def opt_pred_def) + apply fastforce + apply (frule_tac n="0" in h_t_valid_refill; fastforce?) + apply (fastforce simp: valid_sched_context'_def is_active_sc'_def obj_at'_def + opt_map_def opt_pred_def) + apply fastforce + apply (clarsimp simp: to_bool_def typ_heap_simps crefill_relation_def minBudget_def + h_val_field_from_bytes' sc_ptr_to_crefill_ptr_def + csched_context_relation_def + split: if_splits) + done + +lemma merge_nonoverlapping_head_refill_ccorres: + "ccorres dc xfdc + (active_sc_at' scPtr and valid_objs' and no_0_obj' + and pspace_aligned' and pspace_distinct' and pspace_bounded') + \\sc = Ptr scPtr\ hs + (mergeNonoverlappingHeadRefill scPtr) (Call merge_nonoverlapping_head_refill_'proc)" + (is "ccorres _ _ ?abs _ _ _ _") + supply sched_context_C_size[simp del] refill_C_size[simp del] + apply (cinit lift: sc_') + apply (ctac add: refill_pop_head_ccorres) + apply (rule ccorres_rhs_assoc2) + apply (rule ccorres_rhs_assoc2) + apply (rule ccorres_split_nothrow) + apply (clarsimp simp: setRefillHd_def updateRefillHd_def) + apply (rule updateSchedContext_ccorres_lemma3 + [where P="\sc. scRefillHead sc < length (scRefills sc)" and P'="?abs"]) + apply vcg + apply fastforce + apply clarsimp + apply (rename_tac sc sc') + apply (frule (1) sc_ko_at_valid_objs_valid_sc') + apply (frule rf_sr_refill_buffer_relation) + apply (frule_tac n="scRefillHead sc" in h_t_valid_refill, fastforce+) + apply (intro conjI) + apply (clarsimp simp: typ_heap_simps) + apply (rule disjCI2) + apply (rule_tac n="length (scRefills sc)" in array_assertion_shrink_right) + apply (fastforce intro: sc_at_array_assertion') + apply (clarsimp simp: typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: typ_heap_simps) + apply (rule disjCI2) + apply (rule_tac n="length (scRefills sc)" in array_assertion_shrink_right) + apply (fastforce intro: sc_at_array_assertion') + apply (clarsimp simp: typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: sc_ptr_to_crefill_ptr_def typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: crefill_relation_def csched_context_relation_def) + apply (frule rf_sr_refill_buffer_relation) + apply (frule h_t_valid_clift_Some_iff[THEN iffD1]) + apply (clarsimp cong: StateSpace.state.fold_congs) + apply (rule_tac old_sc=sc and n="scRefillHead sc" + and fa="\head. rAmount_update (\_. rAmount head + rAmount rv) head" + in rf_sr_refill_update2; fastforce?) + apply (fastforce simp: sched_context.expand) + apply (clarsimp simp: typ_heap_simps) + apply (fastforce simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def crefill_relation_def) + apply (fastforce simp: csched_context_relation_def) + apply (fastforce dest: crefill_relationD + simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def crefill_relation_def) + apply ceqv + apply (clarsimp simp: setRefillHd_def updateRefillHd_def) + apply (rule updateSchedContext_ccorres_lemma3 + [where P="\sc. scRefillHead sc < length (scRefills sc)" and P'="?abs"]) + apply vcg + apply fastforce + apply clarsimp + apply (rename_tac sc sc') + apply (frule (1) sc_ko_at_valid_objs_valid_sc') + apply (frule rf_sr_refill_buffer_relation) + apply (frule_tac n="scRefillHead sc" in h_t_valid_refill, fastforce+) + apply (intro conjI) + apply (clarsimp simp: typ_heap_simps) + apply (rule disjCI2) + apply (rule_tac n="length (scRefills sc)" in array_assertion_shrink_right) + apply (fastforce intro: sc_at_array_assertion') + apply (clarsimp simp: typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: typ_heap_simps) + apply (rule disjCI2) + apply (rule_tac n="length (scRefills sc)" in array_assertion_shrink_right) + apply (fastforce intro: sc_at_array_assertion') + apply (clarsimp simp: typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: sc_ptr_to_crefill_ptr_def typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: crefill_relation_def csched_context_relation_def) + apply (frule rf_sr_refill_buffer_relation) + apply (frule h_t_valid_clift_Some_iff[THEN iffD1]) + apply (clarsimp cong: StateSpace.state.fold_congs) + apply (rule_tac old_sc=sc and n="scRefillHead sc" + and fa="\head. rTime_update (\_. rTime head - rAmount rv) head" + in rf_sr_refill_update2; fastforce?) + apply (fastforce simp: sched_context.expand) + apply (clarsimp simp: typ_heap_simps) + apply (fastforce simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def crefill_relation_def) + apply (fastforce simp: csched_context_relation_def) + apply (fastforce dest: crefill_relationD + simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def crefill_relation_def) + apply (rule_tac Q'="\_. ?abs and active_sc_at' scPtr" in hoare_post_imp) + apply (clarsimp simp: active_sc_at'_def) + apply normalise_obj_at' + apply (frule (1) sc_ko_at_valid_objs_valid_sc') + apply (clarsimp simp: valid_sched_context'_def) + apply (wpsimp wp: updateRefillHd_valid_objs') + apply vcg + apply (rule_tac Q'="\_. ?abs and active_sc_at' scPtr" in hoare_post_imp) + apply (clarsimp simp: active_sc_at'_def) + apply normalise_obj_at' + apply (frule (1) sc_ko_at_valid_objs_valid_sc') + apply (clarsimp simp: valid_sched_context'_def) + apply (wpsimp wp: updateRefillHd_valid_objs') + apply vcg + apply clarsimp + done + +lemma no_ofail_refillHdInsufficient[wp]: + "no_ofail (sc_at' scPtr) (refillHdInsufficient scPtr)" + unfolding refillHdInsufficient_def + by wpsimp + +crunch mergeNonoverlappingHeadRefill + for invs'[wp]: invs' + (wp: crunch_wps) + +lemma handleOverrun_invs'[wp]: + "\invs' and active_sc_at' scPtr\ handleOverrun scPtr usage \\_. invs'\" + unfolding handleOverrun_def + by (wpsimp wp: valid_whileLoop[where I="\_. invs' and active_sc_at' scPtr"]; fastforce?) + +crunch mergeNonoverlappingHeadRefill + for (no_fail) no_fail[wp] lemma refill_budget_check_ccorres: "ccorres dc xfdc - \ \\usage = usage\ [] + invs' \\usage = usage\ hs (refillBudgetCheck usage) (Call refill_budget_check_'proc)" -sorry (* FIXME RT: refill_budget_check_ccorres *) + supply sched_context_C_size[simp del] refill_C_size[simp del] + Collect_const [simp del] + unfolding refillBudgetCheck_def haskell_assert_def K_bind_def + apply (rule ccorres_symb_exec_l'[rotated, OF _ getCurSc_sp]; (solves wpsimp)?) + apply (rule ccorres_symb_exec_l'[rotated, OF _ scActive_sp]; (solves wpsimp)?) + apply (rule ccorres_symb_exec_l'[rotated, OF _ assert_sp]; (solves wpsimp)?) + apply (rule ccorres_symb_exec_l'[rotated, OF _ isRoundRobin_sp]; (solves wpsimp)?) + apply (rule ccorres_symb_exec_l'[rotated, OF _ assert_sp]; (solves wpsimp)?) + apply (cinit' simp: max_minus_one_word64 runReaderT_def whileAnno_def) + apply (rule_tac xf'=sc_' + and val="Ptr scPtr" + and R="\s. scPtr = ksCurSc s" + in ccorres_symb_exec_r_known_rv[where R'=UNIV]) + apply (rule conseqPre, vcg) + apply (clarsimp simp: cstate_relation_def rf_sr_def Let_def) + apply ceqv + apply (ctac (no_vcg) add: handle_overrun_ccorres) + apply (rule ccorres_rhs_assoc2) + apply (rule_tac r'=crefill_relation and xf'="head_'" in ccorres_split_nothrow_novcg) + apply (rule_tac P="active_sc_at' scPtr and invs'" and P'=UNIV in ccorres_from_vcg) + apply (rule allI, rule conseqPre, vcg) + apply (clarsimp simp: active_sc_at'_rewrite) + apply normalise_obj_at' + apply (rename_tac sc) + apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) + apply (frule rf_sr_refill_buffer_relation) + apply (frule_tac n="scRefillHead sc" in h_t_valid_refill, fastforce+) + apply (clarsimp simp: valid_sched_context'_def is_active_sc'_def obj_at'_def + opt_pred_def opt_map_def + split: option.splits) + apply fastforce + apply (frule (1) obj_at_cslift_sc) + apply (intro conjI) + apply (clarsimp simp: typ_heap_simps) + apply (rule disjCI2) + apply (rule_tac n="length (scRefills sc)" in array_assertion_shrink_right) + apply (rule sc_at_array_assertion') + apply fastforce + apply fastforce + apply (clarsimp simp: valid_sched_context'_def is_active_sc'_def obj_at'_def + opt_pred_def opt_map_def) + apply (clarsimp simp: valid_sched_context'_def is_active_sc'_def obj_at'_def + opt_pred_def opt_map_def typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: sc_ptr_to_crefill_ptr_def typ_heap_simps csched_context_relation_def) + apply (rename_tac s s' sc) + apply (clarsimp simp: getRefillHead_def readRefillHead_def readRefillIndex_def + refillIndex_def readSchedContext_def getObject_def[symmetric] + bind_def return_def) + apply (rule_tac x="(sc, s)" in bexI[rotated]) + apply (fastforce intro: getObject_eq) + apply (drule refill_buffer_relation_crefill_relation) + apply (clarsimp simp: Let_def) + apply (drule_tac x=scPtr in spec) + apply (clarsimp simp: dyn_array_list_rel_pointwise obj_at'_def) + apply (drule_tac x="scRefillHead sc" in spec) + apply (clarsimp simp: valid_sched_context'_def is_active_sc'_def opt_pred_def opt_map_def + typ_heap_simps csched_context_relation_def refillHd_def + sc_ptr_to_crefill_ptr_def) + apply ceqv + apply csymbr + apply clarsimp + apply (rename_tac head head') + apply (rule ccorres_rhs_assoc2) + apply (rule ccorres_split_nothrow_novcg[where r'=dc and xf'=xfdc]) + apply (rule_tac P="0 < usage" in ccorres_cases) + apply clarsimp + apply ccorres_rewrite + apply (rule ccorres_rhs_assoc) + apply (rule_tac val="usToTicks maxPeriodUs" + and xf'=ret__unsigned_longlong_' + in ccorres_symb_exec_r_known_rv[where R=\ and R'=UNIV]) + apply (rule conseqPre, vcg) + apply (clarsimp simp: maxPeriodUs_def timer_defs unat_max_word) + apply ceqv + apply (rule_tac val="from_bool (rTime head < maxReleaseTime)" + and xf'=ret__int_' + in ccorres_symb_exec_r_known_rv[where R=\ and R'=UNIV]) + apply (rule conseqPre, vcg) + apply (clarsimp simp: crefill_relation_def maxReleaseTime_def maxBound_word + split: if_splits) + apply ceqv + apply (clarsimp simp: when_def) + apply (rule ccorres_cond[where R=\]) + apply fastforce + apply (rule ccorres_pre_getObject_sc, rename_tac sc) + apply (rule ccorres_rhs_assoc)+ + apply (rule ccorres_move_c_guard_sc) + apply (rule_tac val="refill_C (rTime head + scPeriod sc) usage" + and xf'=used___struct_refill_C_' + and R="obj_at' (\sc'. scPeriod sc' = scPeriod sc) scPtr" + in ccorres_symb_exec_r_known_rv[where R'=UNIV]) + apply (rule conseqPre, vcg) + apply normalise_obj_at' + apply (frule (1) obj_at_cslift_sc) + apply (clarsimp simp: crefill_relation_def typ_heap_simps csched_context_relation_def) + apply ceqv + apply (rule ccorres_rhs_assoc2) + apply (rule ccorres_split_nothrow_novcg[where r'=dc and xf'=xfdc]) + apply (clarsimp simp: setRefillHd_def updateRefillHd_def) + apply (rule_tac P'="sc_at' scPtr and invs'" + in updateSchedContext_ccorres_lemma3 + [where P="\sc. scRefillHead sc < length (scRefills sc)"]) + apply vcg + apply fastforce + apply clarsimp + apply (rename_tac sc sc') + apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) + apply (frule rf_sr_refill_buffer_relation) + apply (frule_tac n="scRefillHead sc" in h_t_valid_refill, fastforce+) + apply (intro conjI) + apply (clarsimp simp: typ_heap_simps) + apply (rule disjCI2) + apply (rule_tac n="length (scRefills sc)" in array_assertion_shrink_right) + apply (fastforce intro: sc_at_array_assertion') + apply (clarsimp simp: typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def + csched_context_relation_def) + apply (clarsimp simp: typ_heap_simps csched_context_relation_def) + apply normalise_obj_at' + apply (clarsimp simp: crefill_relation_def csched_context_relation_def) + apply (frule rf_sr_refill_buffer_relation) + apply (frule h_t_valid_clift_Some_iff[THEN iffD1]) + apply (clarsimp cong: StateSpace.state.fold_congs) + apply (rule_tac old_sc=sc and n="scRefillHead sc" + and fa="rAmount_update (\_. rAmount head - usage)" + in rf_sr_refill_update2; fastforce?) + apply (fastforce simp: sched_context.expand) + apply (clarsimp simp: typ_heap_simps) + apply (fastforce simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def) + apply (fastforce simp: csched_context_relation_def) + apply normalise_obj_at' + apply (fastforce dest: crefill_relationD + simp: typ_heap_simps' crefill_relation_def) + apply ceqv + apply (rule ccorres_rhs_assoc2) + apply (rule ccorres_split_nothrow_novcg[where r'=dc and xf'=xfdc]) + apply (clarsimp simp: setRefillHd_def updateRefillHd_def) + apply (rule_tac P'="sc_at' scPtr and invs'" + in updateSchedContext_ccorres_lemma3 + [where P="\sc. scRefillHead sc < length (scRefills sc)"]) + apply vcg + apply fastforce + apply clarsimp + apply (rename_tac sc sc') + apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) + apply (frule rf_sr_refill_buffer_relation) + apply (frule_tac n="scRefillHead sc" in h_t_valid_refill, fastforce+) + apply (intro conjI) + apply (clarsimp simp: typ_heap_simps) + apply (rule disjCI2) + apply (rule_tac n="length (scRefills sc)" in array_assertion_shrink_right) + apply (fastforce intro: sc_at_array_assertion') + apply (clarsimp simp: typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def + csched_context_relation_def) + apply normalise_obj_at' + apply (clarsimp simp: crefill_relation_def csched_context_relation_def) + apply (frule rf_sr_refill_buffer_relation) + apply (frule h_t_valid_clift_Some_iff[THEN iffD1]) + apply (clarsimp cong: StateSpace.state.fold_congs) + apply (rule_tac old_sc=sc and n="scRefillHead sc" + and fa="rTime_update (\_. rTime head + usage)" + in rf_sr_refill_update2; fastforce?) + apply (fastforce simp: sched_context.expand) + apply (clarsimp simp: typ_heap_simps) + apply (fastforce simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def) + apply (fastforce simp: csched_context_relation_def) + apply normalise_obj_at' + apply (fastforce dest: crefill_relationD + simp: typ_heap_simps' crefill_relation_def) + apply ceqv + apply (ctac (no_vcg) add: schedule_used_ccorres) + apply (wpsimp wp: updateRefillHd_invs') + apply (clarsimp simp: guard_is_UNIV_def crefill_relation_def) + apply (clarsimp simp: guard_is_UNIV_def crefill_relation_def) + apply (rule_tac Q'="\_. invs' and active_sc_at' scPtr" in hoare_post_imp) + apply (clarsimp simp: active_sc_at'_def) + apply normalise_obj_at' + apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) + apply (clarsimp simp: valid_sched_context'_def) + apply (wpsimp wp: updateRefillHd_invs') + apply (clarsimp simp: guard_is_UNIV_def crefill_relation_def) + apply vcg + apply (rule ccorres_return_Skip) + apply vcg + apply vcg + apply clarsimp + apply ccorres_rewrite + apply (rule ccorres_cond_false) + apply (rule ccorres_return_Skip) + apply ceqv + apply (rule ccorres_handlers_weaken) + apply (clarsimp simp: headInsufficientLoop_def runReaderT_def) + apply (rule_tac rrel=dc and xf=xfdc + and G="\_. active_sc_at' scPtr and invs'" + in ccorres_While'[where G'=UNIV]) + apply (rule ccorres_guard_imp) + apply (ctac add: merge_nonoverlapping_head_refill_ccorres) + apply fastforce + apply fastforce + apply (rule ccorres_guard_imp) + apply (ctac (no_vcg) add: head_refill_insufficient_ccorres) + apply fastforce + apply fastforce + apply wpsimp + apply (clarsimp simp: active_sc_at'_rewrite) + apply wpsimp + apply (rule conseqPre) + apply clarsimp + apply (rule_tac s=s and xf=xfdc in ccorres_to_vcg_Normal) + apply (fastforce intro: ccorres_call[OF merge_nonoverlapping_head_refill_ccorres]) + apply (wpsimp wp: no_fail_mergeNonoverlappingHeadRefill) + apply (fastforce simp: active_sc_at'_rewrite) + apply (rule conseqPre, vcg) + apply (clarsimp simp: active_sc_at'_rewrite) + apply normalise_obj_at' + apply (rename_tac sc) + apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) + apply (frule rf_sr_refill_buffer_relation) + apply (frule_tac n="scRefillHead sc" in h_t_valid_refill, fastforce+) + apply (clarsimp simp: valid_sched_context'_def is_active_sc'_def obj_at'_def + opt_pred_def opt_map_def + split: option.splits) + apply fastforce + apply (frule (1) obj_at_cslift_sc) + apply (intro conjI) + apply (clarsimp simp: typ_heap_simps) + apply (rule disjCI2) + apply (rule_tac n="length (scRefills sc)" in array_assertion_shrink_right) + apply (fastforce intro: sc_at_array_assertion' + simp: valid_sched_context'_def is_active_sc'_def obj_at'_def + opt_pred_def opt_map_def + split: option.splits) + apply (clarsimp simp: valid_sched_context'_def is_active_sc'_def obj_at'_def + opt_pred_def opt_map_def typ_heap_simps csched_context_relation_def) + apply (clarsimp simp: typ_heap_simps' sc_ptr_to_crefill_ptr_def csched_context_relation_def) + apply (wpsimp wp: updateRefillHd_invs') + apply (clarsimp simp: guard_is_UNIV_def) + apply (rule_tac Q'="\_. active_sc_at' scPtr and invs'" in hoare_post_imp) + apply (clarsimp simp: active_sc_at'_rewrite split: if_splits) + apply normalise_obj_at' + apply (frule (1) sc_ko_at_valid_objs_valid_sc'[OF _ invs_valid_objs']) + apply (clarsimp simp: valid_sched_context'_def is_active_sc'_def obj_at'_def + opt_pred_def opt_map_def) + apply wpsimp + apply (clarsimp simp: guard_is_UNIV_def timer_defs unat_max_word) + apply (wpsimp wp: hoare_drop_imps) + apply (vcg exspec=handle_overrun_modifies) + apply (clarsimp simp: active_sc_at'_rewrite is_active_sc'_def obj_at'_def + opt_pred_def opt_map_def) + done lemma checkDomainTime_ccorres: "ccorres dc xfdc \ UNIV [] checkDomainTime (Call checkDomainTime_'proc)" diff --git a/proof/invariant-abstract/DetSchedSchedule_AI.thy b/proof/invariant-abstract/DetSchedSchedule_AI.thy index 79f069d9ae..dc368fff2c 100644 --- a/proof/invariant-abstract/DetSchedSchedule_AI.thy +++ b/proof/invariant-abstract/DetSchedSchedule_AI.thy @@ -3819,7 +3819,7 @@ lemma refill_unblock_check_is_refill_sufficient[wp]: apply (cases "sc_ptr' \ sc_ptr") apply (wpsimp wp: set_refills_wp get_refills_wp whileLoop_valid_inv simp: merge_overlapping_refills_def refill_pop_head_def head_insufficient_loop_def - refill_head_overlapping_loop_def non_overlapping_merge_refills_def + refill_head_overlapping_loop_def merge_nonoverlapping_head_refill_def vs_all_heap_simps update_refill_hd_rewrite update_sched_context_set_refills_rewrite) apply (rename_tac refills) @@ -4327,7 +4327,7 @@ lemma refill_unblock_check_valid_refills[wp]: apply (case_tac "sc_ptr \ p") apply (wpsimp wp: set_refills_wp get_refills_wp whileLoop_valid_inv simp: merge_overlapping_refills_def refill_pop_head_def head_insufficient_loop_def - refill_head_overlapping_loop_def non_overlapping_merge_refills_def + refill_head_overlapping_loop_def merge_nonoverlapping_head_refill_def vs_all_heap_simps update_refill_hd_rewrite update_sched_context_set_refills_rewrite) apply (rule_tac P'1="\s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) @@ -4531,40 +4531,67 @@ lemma refill_budget_check_round_robin_cur_sc_active[wp]: "refill_budget_check_round_robin consumed \cur_sc_active\" by (rule cur_sc_active_lift; wpsimp) -method handle_overrun_loop_body_simple - = ((clarsimp simp: handle_overrun_loop_body_def)? - , wpsimp wp: whileLoop_valid_inv set_refills_wp get_refills_wp - simp: refill_pop_head_def vs_all_heap_simps obj_at_def refill_single_def refill_size_def - update_sched_context_set_refills_rewrite update_refill_hd_rewrite - schedule_used_defs) +method charge_entire_head_refill_simple + = (clarsimp simp: charge_entire_head_refill_def)?, + wpsimp wp: whileLoop_valid_inv set_refills_wp get_refills_wp + simp: refill_pop_head_def vs_all_heap_simps obj_at_def refill_single_def refill_size_def + update_sched_context_set_refills_rewrite update_refill_hd_rewrite + schedule_used_defs sc_at_ppred_def -method handle_overrun_loop_simple - = ((clarsimp simp: handle_overrun_loop_def)? - , wpsimp wp: whileLoop_valid_inv set_refills_wp get_refills_wp - simp: handle_overrun_loop_body_def merge_overlapping_refills_def round_robin_def refill_pop_head_def - sc_valid_refills_def vs_all_heap_simps obj_at_def refill_budget_check_defs - update_sched_context_set_refills_rewrite schedule_used_defs) +method handle_overrun_simple + = (clarsimp simp: handle_overrun_def)?, + wpsimp wp: whileLoop_valid_inv set_refills_wp get_refills_wp + simp: refill_budget_check_defs round_robin_def update_sched_context_set_refills_rewrite + sc_valid_refills_def vs_all_heap_simps obj_at_def method head_insufficient_loop_simple - = ((clarsimp simp: head_insufficient_loop_def)? - , wpsimp wp: whileLoop_valid_inv set_refills_wp get_refills_wp - simp: non_overlapping_merge_refills_def refill_pop_head_def schedule_used_defs - update_sched_context_set_refills_rewrite update_refill_hd_rewrite - , clarsimp simp: vs_all_heap_simps obj_at_def sc_valid_refills_def round_robin_def - unat_MAX_RELEASE_TIME - split: if_splits - , fastforce?) + = (clarsimp simp: head_insufficient_loop_def)?, + wpsimp wp: whileLoop_valid_inv set_refills_wp get_refills_wp + simp: merge_nonoverlapping_head_refill_def refill_pop_head_def schedule_used_defs + update_sched_context_set_refills_rewrite update_refill_hd_rewrite; + ((clarsimp simp: vs_all_heap_simps obj_at_def sc_valid_refills_def round_robin_def + unat_MAX_RELEASE_TIME obj_at_kh_kheap_simps + split: if_splits)?, + fastforce?) + +lemma update_refill_hd_wp: + "\\s. \sc n. + ko_at (SchedContext sc n) sc_ptr s + \ P (s\kheap := (kheap s)(sc_ptr \ SchedContext + (sc_refills_update (\rf. f (hd rf) # tl rf) sc) n)\)\ + update_refill_hd sc_ptr f + \\r. P\" + unfolding update_refill_hd_def + by (wpsimp wp: update_sched_context_wp) + +lemma set_refills_is_active_sc[wp]: + "set_refills sc_ptr new \\s. P (is_active_sc sc_ptr' s)\" + apply (wpsimp wp: set_refills_wp) + apply (erule rsubst[where P=P]) + apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps) + done + +lemma update_refill_hd_is_active_sc[wp]: + "update_refill_hd sc_ptr f \\s. P (is_active_sc sc_ptr' s)\" + apply (wpsimp wp: update_refill_hd_wp) + apply (erule rsubst[where P=P]) + apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps) + done + +crunch charge_entire_head_refill, merge_nonoverlapping_head_refill + for is_active_sc[wp]: "\s. P (is_active_sc sc_ptr s)" + (wp: crunch_wps simp: update_sched_context_set_refills_rewrite ignore: update_sched_context) lemma refill_budget_check_is_active_sc[wp]: "refill_budget_check usage \\s. P (is_active_sc sc_ptr s)\" apply (clarsimp simp: refill_budget_check_def update_refill_hd_rewrite schedule_used_defs) apply (rule bind_wp_fwd_skip, solves wpsimp)+ apply (rule bind_wp_fwd_skip) - apply handle_overrun_loop_simple + apply handle_overrun_simple apply (rule bind_wp_fwd_skip, solves wpsimp)+ apply (rule bind_wp_fwd_skip) - apply (wpsimp wp: set_refills_wp get_refills_wp simp: update_sched_context_set_refills_rewrite) - apply (clarsimp simp: obj_at_def cur_sc_tcb_def sc_at_pred_n_def vs_all_heap_simps) + apply (wpsimp wp: get_refills_wp hoare_drop_imps hoare_vcg_all_lift + simp: update_sched_context_set_refills_rewrite set_refill_hd_def) apply head_insufficient_loop_simple done @@ -8404,98 +8431,57 @@ lemma refill_update_is_refill_sufficient[wp]: done lemma no_ofail_head_insufficient: - "no_ofail (\s. \sc n. kheap s sc_ptr = Some (SchedContext sc n)) (head_insufficient sc_ptr)" - unfolding head_insufficient_def no_ofail_def - apply (clarsimp simp: obind_def obj_at_def read_sched_context_def) - done - -lemma bound_head_insufficient: - "kheap s sc_ptr = Some (SchedContext sc n) \ bound (head_insufficient sc_ptr s)" - apply (clarsimp simp: head_insufficient_def obind_def read_sched_context_def) - done + "no_ofail (\s. sc_refills_sc_at (\refills. refills \ []) sc_ptr s) + (refill_head_insufficient sc_ptr)" + unfolding refill_head_insufficient_def + by wpsimp lemma head_insufficient_true_imp_insufficient: - "\the (head_insufficient sc_ptr s); pred_map \ (scs_of s) sc_ptr\ + "refill_head_insufficient sc_ptr s = Some True \ pred_map (\cfg. r_amount (scrc_refill_hd cfg) < MIN_BUDGET) (sc_refill_cfgs_of s) sc_ptr" - apply (clarsimp simp: vs_all_heap_simps) - apply (frule bound_head_insufficient) - apply (fastforce simp: head_insufficient_def obind_def vs_all_heap_simps read_sched_context_def) - done + by (clarsimp simp: refill_head_insufficient_def vs_all_heap_simps read_sched_context_def + read_refill_head_def + split: option.splits kernel_object.splits) lemma head_insufficient_false_imp_sufficient: - "\\ (the (head_insufficient sc_ptr s)); pred_map \ (scs_of s) sc_ptr\ + "refill_head_insufficient sc_ptr s = Some False \ pred_map (\cfg. MIN_BUDGET \ r_amount (scrc_refill_hd cfg)) (sc_refill_cfgs_of s) sc_ptr" - apply (clarsimp simp: vs_all_heap_simps) - apply (frule bound_head_insufficient) - apply (fastforce simp: head_insufficient_def obind_def vs_all_heap_simps read_sched_context_def) - done - -method non_overlapping_merge_refills_simple - = (clarsimp simp: non_overlapping_merge_refills_def refill_pop_head_def merge_refill_def - update_refill_hd_rewrite update_sched_context_set_refills_rewrite - , wpsimp wp: set_refills_wp get_refills_wp - , (clarsimp simp: vs_all_heap_simps obj_at_def)?) - -lemma non_overlapping_merge_refills_nonempty_refills: - "non_overlapping_merge_refills sc_ptr' - \\s. pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr\" - apply non_overlapping_merge_refills_simple - done - -lemma head_insufficient_loop_nonempty_refills: - "head_insufficient_loop sc_ptr' - \\s. pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr\" - (is "valid ?pre _ _") - apply (clarsimp simp: head_insufficient_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) - apply (wpsimp wp: non_overlapping_merge_refills_nonempty_refills) - done + by (fastforce simp: refill_head_insufficient_def vs_all_heap_simps read_sched_context_def + read_refill_head_def + split: option.splits kernel_object.splits) -lemma head_insufficient_length_at_least_two: +lemma head_insufficient_length: "\pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr; - the ((head_insufficient sc_ptr) s)\ - \ pred_map (\cfg. 1 < length (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr" + refill_head_insufficient sc_ptr s = Some True\ + \ sc_refills_sc_at (\refills. Suc 0 < length refills) sc_ptr s" apply (prop_tac "\ pred_map (\cfg. MIN_BUDGET \ r_amount (scrc_refill_hd cfg)) (sc_refill_cfgs_of s) sc_ptr") apply (fastforce dest: head_insufficient_true_imp_insufficient simp: vs_all_heap_simps refills_unat_sum_def) apply (clarsimp simp: vs_all_heap_simps refills_unat_sum_def) apply (rename_tac sc n) - apply (case_tac "sc_refills sc"; clarsimp simp: word_le_nat_alt) + apply (case_tac "sc_refills sc"; clarsimp simp: word_le_nat_alt sc_at_ppred_def obj_at_def) done -lemma non_overlapping_merge_refills_refills_unat_sum: - "\\s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) \ unat max_time) - (sc_refill_cfgs_of s) sc_ptr - \ (sc_ptr' = sc_ptr \ pred_map (\cfg. 1 < length (scrc_refills cfg)) - (sc_refill_cfgs_of s) sc_ptr)\ - non_overlapping_merge_refills sc_ptr' - \\_ s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) \ unat max_time) - (sc_refill_cfgs_of s) sc_ptr\" - apply non_overlapping_merge_refills_simple - apply (clarsimp simp: refills_unat_sum_def) - apply (rename_tac sc n) - apply (case_tac "sc_refills sc"; simp?) - apply (rename_tac list) - apply (subst unat_add_lem') - apply (prop_tac "unat (r_amount (hd list)) \ sum_list (map unat (map r_amount list))") - apply (fastforce intro: member_le_sum_list) - apply (clarsimp simp: unat_minus_one_word) - apply (case_tac list; simp) - done +method merge_nonoverlapping_head_refill_simple + = clarsimp simp: merge_nonoverlapping_head_refill_def refill_pop_head_def merge_refill_def + update_refill_hd_rewrite update_sched_context_set_refills_rewrite, + wpsimp wp: set_refills_wp get_refills_wp, + (clarsimp simp: vs_all_heap_simps obj_at_def)? -lemma non_overlapping_merge_refills_refills_unat_sum_lower_bound: +lemma merge_nonoverlapping_head_refill_refills_unat_sum_lower_bound: "\\s. pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr \ (sc_ptr' = sc_ptr \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) \ unat max_time) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. 1 < length (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr)\ - non_overlapping_merge_refills sc_ptr' + \ sc_refills_sc_at (\rfls. 1 < length rfls) sc_ptr s)\ + merge_nonoverlapping_head_refill sc_ptr' \\_ s. pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr\" - apply non_overlapping_merge_refills_simple + apply merge_nonoverlapping_head_refill_simple + apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps pred_map_def) apply (clarsimp simp: refills_unat_sum_def) apply (subst unat_add_lem') apply (case_tac "sc_refills sc"; clarsimp) @@ -8506,9 +8492,57 @@ lemma non_overlapping_merge_refills_refills_unat_sum_lower_bound: apply (case_tac list; clarsimp) done +lemma merge_nonoverlapping_head_refill_refills_unat_sum: + "\\s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) \ unat max_time) + (sc_refill_cfgs_of s) sc_ptr + \ (sc_ptr' = sc_ptr \ sc_refills_sc_at (\rfls. 1 < length rfls) sc_ptr s)\ + merge_nonoverlapping_head_refill sc_ptr' + \\_ s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) \ unat max_time) + (sc_refill_cfgs_of s) sc_ptr\" + apply merge_nonoverlapping_head_refill_simple + apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps pred_map_def) + apply (simp add: refills_unat_sum_cons) + apply (rename_tac sc n) + apply (case_tac "sc_refills sc"; simp) + apply (rename_tac list) + apply (case_tac list; simp add: unat_minus_one_word unat_add_lem' refills_unat_sum_def) + done + +lemma merge_nonoverlapping_head_refill_nonempty_refills: + "\\s. sc_refills_sc_at (\refills. 1 < length refills) sc_ptr s\ + merge_nonoverlapping_head_refill sc_ptr + \\_ s. sc_refills_sc_at (\refills. refills \ []) sc_ptr s\" + apply merge_nonoverlapping_head_refill_simple + apply (clarsimp simp: sc_at_pred_n_def obj_at_def neq_Nil_lengthI) + done + +lemma head_insufficient_loop_nonempty_refills: + "\\s. sc_refills_sc_at (\refills. refills \ []) sc_ptr s + \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) \ unat max_time) + (sc_refill_cfgs_of s) sc_ptr + \ pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) + (sc_refill_cfgs_of s) sc_ptr\ + head_insufficient_loop sc_ptr' + \\_ s. sc_refills_sc_at (\refills. refills \ []) sc_ptr s\" + (is "valid ?pre _ _") + apply (clarsimp simp: head_insufficient_loop_def) + apply (cases "sc_ptr' = sc_ptr") + apply clarsimp + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) + apply (intro hoare_vcg_conj_lift_pre_fix) + apply (wpsimp wp: merge_nonoverlapping_head_refill_nonempty_refills) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum_lower_bound) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply head_insufficient_loop_simple + done + lemma head_insufficient_loop_refills_sum: "\\s. pred_map (\cfg. refills_sum (scrc_refills cfg) = scrc_budget cfg) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ (sc_ptr' = sc_ptr \ pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr @@ -8521,37 +8555,41 @@ lemma head_insufficient_loop_refills_sum: apply (cases "sc_ptr' \ sc_ptr") apply head_insufficient_loop_simple apply (clarsimp simp: head_insufficient_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) apply (intro hoare_vcg_conj_lift_pre_fix; (solves wpsimp)?) - defer - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum_lower_bound) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (clarsimp simp: non_overlapping_merge_refills_def refill_pop_head_def merge_refill_def + defer + apply (wpsimp wp: merge_nonoverlapping_head_refill_nonempty_refills) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum_lower_bound) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (clarsimp simp: merge_nonoverlapping_head_refill_def refill_pop_head_def merge_refill_def update_refill_hd_rewrite update_sched_context_set_refills_rewrite) apply (wpsimp wp: set_refills_wp get_refills_wp) - apply (frule head_insufficient_length_at_least_two, clarsimp) + apply (frule no_ofailD[OF no_ofail_head_insufficient]) + apply (frule head_insufficient_length) + apply clarsimp apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps pred_map_simps) - apply (rename_tac sc n) + apply (rename_tac sc n bool) apply (case_tac "sc_refills sc"; clarsimp) apply (rename_tac list) apply (case_tac list; clarsimp) apply (clarsimp simp: add_ac) done -lemma non_overlapping_merge_refills_refills_unat_sum_equals_budget: +lemma merge_nonoverlapping_head_refill_refills_unat_sum_equals_budget: "\\s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr \ (sc_ptr' = sc_ptr \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) \ unat max_time) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. 1 < length (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr)\ - non_overlapping_merge_refills sc_ptr' + \ sc_refills_sc_at (\refills. 1 < length refills) sc_ptr s)\ + merge_nonoverlapping_head_refill sc_ptr' \\_ s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr\" - apply non_overlapping_merge_refills_simple + apply merge_nonoverlapping_head_refill_simple + apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps pred_map_simps) apply (clarsimp simp: refills_unat_sum_def) apply (rename_tac sc n) apply (case_tac "sc_refills sc"; simp?) @@ -8567,6 +8605,7 @@ lemma non_overlapping_merge_refills_refills_unat_sum_equals_budget: lemma head_insufficient_loop_refills_unat_sum_equals_budget: "\\s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ (sc_ptr' = sc_ptr \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) \ unat max_time) (sc_refill_cfgs_of s) sc_ptr @@ -8579,19 +8618,22 @@ lemma head_insufficient_loop_refills_unat_sum_equals_budget: apply (cases "sc_ptr' \ sc_ptr") apply head_insufficient_loop_simple apply (clarsimp simp: head_insufficient_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) apply (intro hoare_vcg_conj_lift_pre_fix; (solves wpsimp)?) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum_equals_budget) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (clarsimp simp: non_overlapping_merge_refills_def refill_pop_head_def merge_refill_def + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum_equals_budget) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_nonempty_refills) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (clarsimp simp: merge_nonoverlapping_head_refill_def refill_pop_head_def merge_refill_def update_sched_context_set_refills_rewrite update_refill_hd_rewrite) apply (wpsimp wp: set_refills_wp get_refills_wp) - apply (frule head_insufficient_length_at_least_two, clarsimp) + apply (frule no_ofailD[OF no_ofail_head_insufficient]) + apply (frule head_insufficient_length) + apply fastforce apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps pred_map_simps) - apply (rename_tac sc n) + apply (rename_tac sc n bool) apply (case_tac "sc_refills sc"; clarsimp) apply (rename_tac list) apply (case_tac list; clarsimp) @@ -8602,7 +8644,7 @@ lemma head_insufficient_loop_refills_unat_sum_equals_budget: apply presburger done -lemma non_overlapping_merge_refills_ordered_disjoint_helper: +lemma merge_nonoverlapping_head_refill_ordered_disjoint_helper: "\ordered_disjoint (a # r1 # rs); sum_list (map unat (map r_amount (a # r1 # rs))) \ unat max_time\ \ ordered_disjoint (\r_time = r_time r1 - r_amount a, r_amount = r_amount a + r_amount r1\ # rs)" apply (rule_tac left="[\r_time = r_time r1 - r_amount a, r_amount = r_amount a + r_amount r1\]" @@ -8624,25 +8666,26 @@ lemma non_overlapping_merge_refills_ordered_disjoint_helper: apply simp done -lemma non_overlapping_merge_refills_ordered_disjoint: +lemma merge_nonoverlapping_head_refill_ordered_disjoint: "\\s. pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr \ (sc_ptr' = sc_ptr \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) \ unat max_time) (sc_refill_cfgs_of s) sc_ptr)\ - non_overlapping_merge_refills sc_ptr' + merge_nonoverlapping_head_refill sc_ptr' \\_ s. pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr\" - apply non_overlapping_merge_refills_simple - apply (clarsimp simp: refills_unat_sum_def) + apply merge_nonoverlapping_head_refill_simple + apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps pred_map_simps refills_unat_sum_def) apply (rename_tac sc n) apply (case_tac "sc_refills sc"; (solves simp)?) apply (rename_tac list) apply (case_tac list; (solves simp)?) - apply (fastforce dest: non_overlapping_merge_refills_ordered_disjoint_helper + apply (fastforce dest: merge_nonoverlapping_head_refill_ordered_disjoint_helper simp: add_ac cong: refill.ext_split) done lemma head_insufficient_loop_ordered_disjoint: "\\s. pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ (sc_ptr' = sc_ptr \ pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr @@ -8654,15 +8697,16 @@ lemma head_insufficient_loop_ordered_disjoint: apply (cases "sc_ptr' \ sc_ptr") apply head_insufficient_loop_simple apply (clarsimp simp: head_insufficient_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) apply (intro hoare_vcg_conj_lift_pre_fix; (solves wpsimp)?) defer - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum_lower_bound) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (wpsimp wp: non_overlapping_merge_refills_ordered_disjoint) + apply (wpsimp wp: merge_nonoverlapping_head_refill_nonempty_refills) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum_lower_bound) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_ordered_disjoint) done lemma head_time_plus_budget_bounded: @@ -8684,7 +8728,7 @@ lemma head_time_plus_budget_bounded: apply assumption done -lemma non_overlapping_merge_refills_no_overflow_helper: +lemma merge_nonoverlapping_head_refill_no_overflow_helper: "\no_overflow (a # r1 # rs); ordered_disjoint (a # r1 # rs)\ \ no_overflow (\r_time = r_time r1 - r_amount a, r_amount = r_amount a + r_amount r1\ # rs)" apply (frule (1) head_time_plus_budget_bounded) @@ -8702,25 +8746,28 @@ lemma non_overlapping_merge_refills_no_overflow_helper: apply simp done -lemma non_overlapping_merge_refills_no_overflow: +lemma merge_nonoverlapping_head_refill_no_overflow: "\\s. pred_map (\cfg. no_overflow (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ (sc_ptr' = sc_ptr \ pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. 1 < length (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr)\ - non_overlapping_merge_refills sc_ptr' + \ sc_refills_sc_at (\refills. 1 < length refills) sc_ptr s)\ + merge_nonoverlapping_head_refill sc_ptr' \\_ s. pred_map (\cfg. no_overflow (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr\" - apply non_overlapping_merge_refills_simple + apply merge_nonoverlapping_head_refill_simple + apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps pred_map_simps) apply (rename_tac sc n) apply (case_tac "sc_refills sc"; (solves simp)?) apply (rename_tac list) apply (case_tac list; (solves simp)?) - apply (fastforce dest: non_overlapping_merge_refills_no_overflow_helper + apply (fastforce dest: merge_nonoverlapping_head_refill_no_overflow_helper simp: add_ac cong: refill.ext_split) done lemma head_insufficient_loop_no_overflow: "\\s. pred_map (\cfg. no_overflow (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ (sc_ptr' = sc_ptr \ pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) @@ -8733,21 +8780,23 @@ lemma head_insufficient_loop_no_overflow: apply (cases "sc_ptr' \ sc_ptr") apply head_insufficient_loop_simple apply (clarsimp simp: head_insufficient_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) apply (intro hoare_vcg_conj_lift_pre_fix; (solves wpsimp)?) defer - apply (wpsimp wp: non_overlapping_merge_refills_ordered_disjoint) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum_lower_bound) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (wpsimp wp: non_overlapping_merge_refills_no_overflow) - apply (fastforce dest: head_insufficient_length_at_least_two) + apply (wpsimp wp: merge_nonoverlapping_head_refill_nonempty_refills) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_ordered_disjoint) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum_lower_bound) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_no_overflow) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) done lemma head_insufficient_loop_refills_window: "\\s. pred_map (\cfg. window (scrc_refills cfg) (scrc_period cfg)) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ (sc_ptr' = sc_ptr \ pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr @@ -8761,21 +8810,24 @@ lemma head_insufficient_loop_refills_window: apply (cases "sc_ptr' \ sc_ptr") apply head_insufficient_loop_simple apply (clarsimp simp: head_insufficient_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) apply (intro hoare_vcg_conj_lift_pre_fix; (solves wpsimp)?) defer - apply (wpsimp wp: non_overlapping_merge_refills_ordered_disjoint) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum_lower_bound) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (clarsimp simp: non_overlapping_merge_refills_def refill_pop_head_def merge_refill_def + apply (wpsimp wp: merge_nonoverlapping_head_refill_nonempty_refills) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_ordered_disjoint) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum_lower_bound) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (clarsimp simp: merge_nonoverlapping_head_refill_def refill_pop_head_def merge_refill_def update_sched_context_set_refills_rewrite update_refill_hd_rewrite) apply (wpsimp wp: set_refills_wp get_refills_wp) - apply (frule head_insufficient_length_at_least_two, clarsimp) + apply (frule no_ofailD[OF no_ofail_head_insufficient]) + apply (frule head_insufficient_length) + apply fastforce apply (clarsimp simp: vs_all_heap_simps refills_unat_sum_def obj_at_def window_def) - apply (rename_tac sc n) + apply (rename_tac sc n bool) apply (prop_tac "last (tl (tl (sc_refills sc))) = refill_tl sc") apply (case_tac "sc_refills sc"; simp) apply (rename_tac list) @@ -8783,8 +8835,9 @@ lemma head_insufficient_loop_refills_window: apply (prop_tac "unat (r_time (refill_hd sc)) \ unat (r_time (hd (tl (sc_refills sc))) - r_amount (refill_hd sc))") apply (subst unat_sub) - apply (fastforce simp: ordered_disjoint_def word_le_nat_alt hd_conv_nth nth_tl tl_Nil) - apply (fastforce simp: ordered_disjoint_def hd_conv_nth nth_tl tl_Nil) + apply (fastforce simp: ordered_disjoint_def word_le_nat_alt hd_conv_nth nth_tl tl_Nil + sc_at_ppred_def obj_at_def) + apply (fastforce simp: ordered_disjoint_def hd_conv_nth nth_tl tl_Nil sc_at_ppred_def obj_at_def) apply clarsimp done @@ -8798,7 +8851,7 @@ lemma head_insufficient_loop_hd_r_time: (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s else pred_map (\cfg. unat (r_time (hd (scrc_refills cfg))) \ P (cur_time s)) (sc_refill_cfgs_of s) sc_ptr\ head_insufficient_loop sc_ptr' @@ -8808,42 +8861,42 @@ lemma head_insufficient_loop_hd_r_time: apply (cases "sc_ptr' \ sc_ptr") apply head_insufficient_loop_simple apply (clarsimp simp: head_insufficient_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) apply (intro hoare_vcg_conj_lift_pre_fix; (solves wpsimp)?) defer - apply (wpsimp wp: non_overlapping_merge_refills_ordered_disjoint) - apply (wpsimp wp: non_overlapping_merge_refills_no_overflow) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum_lower_bound) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (wpsimp wp: non_overlapping_merge_refills_nonempty_refills) - apply (clarsimp simp: vs_all_heap_simps obj_at_def) + apply (wpsimp wp: merge_nonoverlapping_head_refill_ordered_disjoint) + apply (wpsimp wp: merge_nonoverlapping_head_refill_no_overflow) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum_lower_bound) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_nonempty_refills) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_nonempty_refills) + apply (clarsimp simp: vs_all_heap_simps obj_at_def sc_at_ppred_def) apply (rename_tac sc n) - apply (rule_tac y="unat (r_time (refill_tl sc))" in order_trans - ; fastforce?) + apply (rule_tac y="unat (r_time (refill_tl sc))" in order_trans; fastforce?) apply (frule_tac k=0 and l="length (sc_refills sc) - 1" - in ordered_disjoint_no_overflow_implies_sorted - ; fastforce?) + in ordered_disjoint_no_overflow_implies_sorted; + fastforce?) apply (simp add: hd_conv_nth last_conv_nth unat_arith_simps(1)) - apply (clarsimp simp: non_overlapping_merge_refills_def refill_pop_head_def merge_refill_def + apply (clarsimp simp: merge_nonoverlapping_head_refill_def refill_pop_head_def merge_refill_def update_sched_context_set_refills_rewrite update_refill_hd_rewrite) apply (wpsimp wp: set_refills_wp get_refills_wp) - apply (frule head_insufficient_length_at_least_two, clarsimp) - apply (clarsimp simp: vs_all_heap_simps obj_at_def) + apply (frule no_ofailD[OF no_ofail_head_insufficient]) + apply (frule head_insufficient_length) + apply fastforce + apply (clarsimp simp: vs_all_heap_simps obj_at_def sc_at_ppred_def) apply (intro conjI impI) apply (subst unat_sub) apply (clarsimp simp: word_le_nat_alt ordered_disjoint_def) apply (drule_tac x=0 in spec) apply clarsimp apply (metis add_leE hd_conv_nth hd_tl_nth numeral_nat(7)) - apply (rename_tac sc n) - apply (rule_tac y="unat (r_time (hd (tl (sc_refills sc))))" in order_trans - ; fastforce?) - apply (frule_tac k=0 and l=1 in ordered_disjoint_no_overflow_implies_sorted - ; fastforce?) + apply (rename_tac sc n bool) + apply (rule_tac y="unat (r_time (hd (tl (sc_refills sc))))" in order_trans; fastforce?) + apply (frule_tac k=0 and l=1 in ordered_disjoint_no_overflow_implies_sorted; fastforce?) apply (metis Nitpick.size_list_simp(2) One_nat_def hd_conv_nth last_conv_nth last_tl length_tl rel_simps(71)) apply (simp add: last_tl tl_Nil) @@ -8852,6 +8905,7 @@ lemma head_insufficient_loop_hd_r_time: lemma head_insufficient_loop_hd_r_time_same: "\\s. pred_map (\cfg. r_time (hd (scrc_refills cfg)) \ P (cur_time s)) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ (sc_ptr' = sc_ptr \ pred_map (\cfg. MIN_BUDGET \ r_amount (hd (scrc_refills cfg))) (sc_refill_cfgs_of s) sc_ptr)\ head_insufficient_loop sc_ptr' @@ -8861,12 +8915,12 @@ lemma head_insufficient_loop_hd_r_time_same: apply (cases "sc_ptr' \ sc_ptr") apply head_insufficient_loop_simple apply (clarsimp simp: head_insufficient_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) apply (rule hoare_weaken_pre) apply (rule hoare_pre_cont) - apply (fastforce dest: head_insufficient_true_imp_insufficient - simp: vs_all_heap_simps refills_unat_sum_def) + apply clarsimp + apply (frule no_ofailD[OF no_ofail_head_insufficient]) + apply (fastforce dest: head_insufficient_true_imp_insufficient simp: vs_all_heap_simps) done lemma head_insufficient_loop_MIN_BUDGET_in_head: @@ -8875,7 +8929,7 @@ lemma head_insufficient_loop_MIN_BUDGET_in_head: (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s else pred_map (\cfg. MIN_BUDGET \ r_amount (hd (scrc_refills cfg))) (sc_refill_cfgs_of s) sc_ptr\ head_insufficient_loop sc_ptr' @@ -8885,33 +8939,34 @@ lemma head_insufficient_loop_MIN_BUDGET_in_head: apply (cases "sc_ptr' \ sc_ptr") apply head_insufficient_loop_simple apply (clarsimp simp: head_insufficient_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) apply (intro hoare_vcg_conj_lift_pre_fix; (solves wpsimp)?) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum_lower_bound) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (wpsimp wp: non_overlapping_merge_refills_nonempty_refills) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum_lower_bound) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_nonempty_refills) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply clarsimp + apply (frule no_ofailD[OF no_ofail_head_insufficient]) apply (fastforce dest: head_insufficient_false_imp_sufficient simp: vs_all_heap_simps refills_unat_sum_def) done -lemma non_overlapping_merge_refills_length_bounded: +lemma merge_nonoverlapping_head_refill_length_bounded: "\\s. pred_map (\cfg. length (scrc_refills cfg) \ scrc_refill_max cfg) (sc_refill_cfgs_of s) sc_ptr - \ (sc_ptr' = sc_ptr \ pred_map (\cfg. 1 < length (scrc_refills cfg)) - (sc_refill_cfgs_of s) sc_ptr)\ - non_overlapping_merge_refills sc_ptr' + \ (sc_ptr' = sc_ptr \ sc_refills_sc_at (\refills. 1 < length refills) sc_ptr s)\ + merge_nonoverlapping_head_refill sc_ptr' \\_ s. pred_map (\cfg. length (scrc_refills cfg) \ scrc_refill_max cfg) (sc_refill_cfgs_of s) sc_ptr\" - apply non_overlapping_merge_refills_simple - apply fastforce - done + apply merge_nonoverlapping_head_refill_simple + by (fastforce simp: vs_all_heap_simps obj_at_kh_kheap_simps) lemma head_insufficient_loop_length_bounded: "\\s. pred_map (\cfg. length (scrc_refills cfg) \ scrc_refill_max cfg) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ (sc_ptr' = sc_ptr \ pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr @@ -8924,17 +8979,18 @@ lemma head_insufficient_loop_length_bounded: apply (cases "sc_ptr' \ sc_ptr") apply head_insufficient_loop_simple apply (clarsimp simp: head_insufficient_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) apply (intro hoare_vcg_conj_lift_pre_fix; (solves wpsimp)?) - defer - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum_lower_bound) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (wpsimp wp: non_overlapping_merge_refills_length_bounded) - apply (fastforce dest: head_insufficient_length_at_least_two - simp: vs_all_heap_simps obj_at_def) + defer + apply (wpsimp wp: merge_nonoverlapping_head_refill_nonempty_refills) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum_lower_bound) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_length_bounded) + apply (frule no_ofailD[OF no_ofail_head_insufficient]) + apply (fastforce dest: head_insufficient_length simp: vs_all_heap_simps obj_at_def) done lemma valid_refills_unbundled: @@ -9220,7 +9276,7 @@ lemma schedule_used_release_time_bounded: apply (wpsimp wp: update_sched_context_wp get_refills_wp simp: schedule_used_defs) apply (clarsimp simp: obj_at_def vs_all_heap_simps) - apply (case_tac "sc_refills sc"; fastforce) + apply (rename_tac sc, case_tac "sc_refills sc"; fastforce) done lemma schedule_used_refills_unat_sum: @@ -9235,7 +9291,6 @@ lemma schedule_used_refills_unat_sum: supply map_map[simp del] apply (wpsimp wp: update_sched_context_wp get_refills_wp simp: schedule_used_def refill_add_tail_def update_refill_tl_def) - apply (rename_tac sc n) apply (clarsimp simp: schedule_used_def vs_all_heap_simps obj_at_def) apply (case_tac "can_merge_refill (refill_tl sc) new") @@ -9342,8 +9397,8 @@ lemma schedule_used_no_overflow: apply (rule_tac left="a # butlast lista" and right=" [\r_time = r_time new - r_amount (last lista), r_amount = r_amount (last lista) + r_amount new\]" - in no_overflow_append - ; fastforce?) + in no_overflow_append; + fastforce?) apply (metis no_overflow_sublist butlast.simps(2) sublist_butlast) apply (clarsimp simp: no_overflow_def can_merge_refill_def) apply (subst unat_add_lem', clarsimp simp: unat_minus_one_word) @@ -9411,8 +9466,8 @@ lemma schedule_used_ordered_disjoint: apply (rule_tac left="a # butlast lista" and right="[\r_time = r_time new - r_amount (last lista), r_amount = r_amount (last lista) + r_amount new\]" - in ordered_disjoint_append - ; fastforce?) + in ordered_disjoint_append; + fastforce?) apply (metis ordered_disjoint_sublist butlast.simps(2) sublist_butlast) apply (clarsimp simp: ordered_disjoint_def can_merge_refill_def) apply (prop_tac "unat (r_time (last lista) + r_amount (last lista)) @@ -9445,7 +9500,7 @@ lemma schedule_used_window: \ (sc_ptr' = sc_ptr \ pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. no_overflow (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ pred_map (\cfg. r_time new \ r_time (scrc_refill_hd cfg) + scrc_period cfg) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. unat (r_time (scrc_refill_hd cfg)) + unat MAX_PERIOD \ unat max_time) @@ -9558,7 +9613,6 @@ lemma head_insufficient_loop_r_time_helper: (sc_refill_cfgs_of s) sc_ptr\" apply (cases "sc_ptr'\sc_ptr") apply head_insufficient_loop_simple - apply (clarsimp simp: vs_all_heap_simps) apply (rule_tac Q'="\_ s. pred_map (\cfg. unat (r_time (scrc_refill_hd cfg)) \ unat max_time - n * unat MAX_PERIOD) (sc_refill_cfgs_of s) sc_ptr" @@ -9566,47 +9620,33 @@ lemma head_insufficient_loop_r_time_helper: apply (clarsimp simp: vs_all_heap_simps) apply fastforce apply (wpsimp wp: head_insufficient_loop_hd_r_time) - apply (fastforce simp: window_def vs_all_heap_simps word_le_nat_alt) + apply (fastforce simp: window_def vs_all_heap_simps word_le_nat_alt sc_at_ppred_def obj_at_def) done -lemma no_ofail_head_time_buffer: - "no_ofail (\s. \sc n. kheap s (cur_sc s) = Some (SchedContext sc n)) (head_time_buffer usage)" - unfolding head_time_buffer_def no_ofail_def - apply (clarsimp simp: obind_def ogets_def read_sched_context_def) - done - -lemma bound_head_time_buffer: - "pred_map \ (scs_of s) (cur_sc s) \ bound (head_time_buffer usage s)" - apply (clarsimp simp: head_time_buffer_def obind_def read_sched_context_def ogets_def - vs_all_heap_simps - split: kernel_object.splits) - done +lemma no_ofail_head_refill_overrun: + "no_ofail (sc_refills_sc_at (\refills. refills \ []) sc_ptr) (head_refill_overrun sc_ptr usage)" + unfolding head_refill_overrun_def + by wpsimp -lemma head_time_buffer_true_imp_buffer: - "pred_map \ (scs_of s) (cur_sc s) - \ the (head_time_buffer usage s) +lemma head_refill_overrun_true_imp_buffer: + "sc_refills_sc_at (\refills. refills \ []) sc_ptr s + \ the (head_refill_overrun sc_ptr usage s) = (pred_map (\cfg. r_time (scrc_refill_hd cfg) < MAX_RELEASE_TIME) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. r_amount (scrc_refill_hd cfg) \ usage ) - (sc_refill_cfgs_of s) (cur_sc s))" - apply (fastforce simp: head_time_buffer_def obind_def read_sched_context_def ogets_def - vs_all_heap_simps MAX_RELEASE_TIME_def - split: kernel_object.splits) - done - -lemma head_time_buffer_true_imp_unat_buffer: - "pred_map \ (scs_of s) (cur_sc s) - \ the (head_time_buffer usage s) + (sc_refill_cfgs_of s) sc_ptr)" + by (fastforce simp: head_refill_overrun_def obind_def read_refill_head_def read_sched_context_def ogets_def + vs_all_heap_simps MAX_RELEASE_TIME_def sc_at_ppred_def obj_at_def + split: kernel_object.splits option.splits) + +lemma head_refill_overrun_true_imp_unat_buffer: + "sc_refills_sc_at (\refills. refills \ []) sc_ptr s + \ the (head_refill_overrun sc_ptr usage s) = (pred_map (\cfg. unat (r_time (scrc_refill_hd cfg)) < unat MAX_RELEASE_TIME) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. r_amount (scrc_refill_hd cfg) \ usage ) - (sc_refill_cfgs_of s) (cur_sc s))" - apply (intro iffI) - apply (fastforce dest!: head_time_buffer_true_imp_buffer[THEN iffD1, where usage1=usage] - simp: vs_all_heap_simps obj_at_kh_kheap_simps word_less_nat_alt) - apply (fastforce simp: head_time_buffer_true_imp_buffer[THEN iffD2, where usage1=usage] - simp: vs_all_heap_simps obj_at_kh_kheap_simps word_less_nat_alt) - done + (sc_refill_cfgs_of s) sc_ptr)" + by (fastforce simp: head_refill_overrun_true_imp_buffer[where usage=usage] word_less_nat_alt) lemma refill_single_sp: "\P\ @@ -9619,78 +9659,72 @@ lemma refill_single_sp: apply (clarsimp simp: obj_at_def) done -lemma handle_overrun_loop_body_nonzero_refills[wp]: - "handle_overrun_loop_body usage - \\s. pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr\" - apply handle_overrun_loop_simple - done +lemma charge_entire_head_refill_nonzero_refills[wp]: + "charge_entire_head_refill sc_ptr' usage \sc_refills_sc_at (\refills. refills \ []) sc_ptr\" + by charge_entire_head_refill_simple -lemma handle_overrun_loop_nonzero_refills: - "handle_overrun_loop usage - \\s. pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr\" - apply (clarsimp simp: handle_overrun_loop_def) - apply (wpsimp wp: whileLoop_valid_inv handle_overrun_loop_body_nonzero_refills) - done +lemma handle_overrun_nonzero_refills: + "handle_overrun sc_ptr usage \sc_refills_sc_at (\refills. refills \ []) sc_ptr\" + by handle_overrun_simple -lemma handle_overrun_loop_body_refills_unat_sum_equals_budget: +lemma charge_entire_head_refill_refills_unat_sum_equals_budget: "\\s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr - \ (sc_ptr = cur_sc s \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr)\ - handle_overrun_loop_body usage + \ (sc_ptr = sc_ptr' \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s)\ + charge_entire_head_refill sc_ptr' usage \\_ s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr\" (is "valid _ _ (\_ s. ?post s)") supply map_map[simp del] - apply (clarsimp simp: handle_overrun_loop_body_def) - apply (rule bind_wp[OF _ gets_sp]) - apply (rule bind_wp[OF _ refill_single_sp]) + apply (clarsimp simp: charge_entire_head_refill_def) + apply (rule bind_wp[OF _ get_refill_head_sp]) apply (rule bind_wp[OF _ get_sched_context_sp]) + apply (rule bind_wp[OF _ refill_single_sp]) apply (rule_tac Q'="\_ s. ?post s" in bind_wp; (solves wpsimp)?) apply (rule hoare_if) apply (wpsimp simp: update_sched_context_set_refills_rewrite update_refill_hd_rewrite wp: set_refills_wp get_refills_wp) apply (clarsimp simp: vs_all_heap_simps obj_at_def refills_unat_sum_def) - apply (rename_tac sc n) + apply (rename_tac sc) apply (case_tac "sc_refills sc"; clarsimp) apply (wpsimp wp: schedule_used_refills_unat_sum) apply (wpsimp simp: refill_pop_head_def update_sched_context_set_refills_rewrite update_refill_hd_rewrite wp: set_refills_wp get_refills_wp) - apply (clarsimp simp: vs_all_heap_simps obj_at_def refills_unat_sum_def) + apply (clarsimp simp: vs_all_heap_simps obj_at_def refills_unat_sum_def sc_at_ppred_def) apply (rename_tac sc n) apply (case_tac "sc_refills sc"; simp) done -lemma handle_overrun_loop_refills_unat_sum_equals_budget: +lemma handle_overrun_refills_unat_sum_equals_budget: "\\s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr - \ (sc_ptr = cur_sc s \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr)\ - handle_overrun_loop usage + \ (sc_ptr = sc_ptr' \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s)\ + handle_overrun sc_ptr' usage \\_ s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr\" (is "\?pre\ _ \_\") - apply (clarsimp simp: handle_overrun_loop_def) + apply (clarsimp simp: handle_overrun_def) apply (rule_tac I="\_. ?pre" in valid_whileLoop; fastforce?) - apply (wpsimp wp: handle_overrun_loop_body_refills_unat_sum_equals_budget - handle_overrun_loop_body_nonzero_refills hoare_vcg_if_lift2 hoare_vcg_imp_lift') + apply (wpsimp wp: charge_entire_head_refill_refills_unat_sum_equals_budget hoare_vcg_imp_lift') done -lemma handle_overrun_loop_body_window: +lemma charge_entire_head_refill_window: "\\s. pred_map (\cfg. window (scrc_refills cfg) (scrc_period cfg)) (sc_refill_cfgs_of s) sc_ptr - \ (sc_ptr = cur_sc s + \ (sc_ptr = sc_ptr' \ pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. no_overflow (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ pred_map (\cfg. unat (r_time (scrc_refill_hd cfg)) + 2 * unat MAX_PERIOD \ unat max_time) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. scrc_period cfg \ MAX_PERIOD) (sc_refill_cfgs_of s) sc_ptr)\ - handle_overrun_loop_body usage + charge_entire_head_refill sc_ptr' usage \\_ s. pred_map (\cfg. window (scrc_refills cfg) (scrc_period cfg)) (sc_refill_cfgs_of s) sc_ptr\" (is "\_\ _ \(\_ s. ?post s)\") - apply (clarsimp simp: handle_overrun_loop_body_def) - apply (rule bind_wp[OF _ gets_sp]) - apply (rule bind_wp[OF _ refill_single_sp]) + apply (clarsimp simp: charge_entire_head_refill_def) + apply (rule bind_wp[OF _ get_refill_head_sp]) apply (rule bind_wp[OF _ get_sched_context_sp]) + apply (rule bind_wp[OF _ refill_single_sp]) apply (rule_tac Q'="\_ s. ?post s" in bind_wp; (solves wpsimp)?) apply (rule hoare_if) apply (wpsimp simp: update_sched_context_set_refills_rewrite update_refill_hd_rewrite @@ -9699,10 +9733,10 @@ lemma handle_overrun_loop_body_window: apply (simp add: less_not_refl2 tail_nonempty_length) apply (wpsimp simp: refill_pop_head_def update_sched_context_set_refills_rewrite wp: schedule_used_window set_refills_wp get_refills_wp) - apply (case_tac "sc_ptr \ cur_sc s") + apply (case_tac "sc_ptr \ sc_ptr'") apply (clarsimp simp: vs_all_heap_simps obj_at_def window_def) - apply (clarsimp simp: vs_all_heap_simps obj_at_def window_def) - apply (rename_tac sc n) + apply (clarsimp simp: vs_all_heap_simps obj_at_def window_def sc_at_ppred_def) + apply (rename_tac sc) apply (prop_tac "unat (r_time (refill_hd sc)) \ unat (r_time (hd (tl (sc_refills sc))))") apply (clarsimp simp: hd_conv_nth hd_tl_nth) apply (subst hd_tl_nth) @@ -9737,12 +9771,12 @@ lemma head_time_buffer_implies_no_overflow: pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr; pred_map (\cfg. window (scrc_refills cfg) (scrc_period cfg)) (sc_refill_cfgs_of s) sc_ptr; pred_map (\cfg. scrc_period cfg \ MAX_PERIOD) (sc_refill_cfgs_of s) sc_ptr; - pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr; + sc_refills_sc_at (\refills. refills \ []) sc_ptr s; pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr; pred_map (\cfg. scrc_budget cfg \ scrc_period cfg) (sc_refill_cfgs_of s) sc_ptr\ \ pred_map (\cfg. no_overflow (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr" - apply (clarsimp simp: vs_all_heap_simps) + apply (clarsimp simp: vs_all_heap_simps sc_at_ppred_def obj_at_def) apply (rename_tac sc n) apply (clarsimp simp: no_overflow_def) apply (clarsimp simp: in_set_conv_nth) @@ -9763,33 +9797,33 @@ lemma head_time_buffer_implies_no_overflow: apply (meson nat_le_Suc_less_imp nat_less_le order_trans unat_bounded_above) done -lemma handle_overrun_loop_body_ordered_disjoint: +lemma charge_entire_head_refill_ordered_disjoint: "\\s. pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr - \ (sc_ptr = cur_sc s + \ (sc_ptr = sc_ptr' \ pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. window (scrc_refills cfg) (scrc_period cfg)) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. unat (r_time (scrc_refill_hd cfg)) + 2 * unat MAX_PERIOD \ unat max_time) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ pred_map (\cfg. scrc_period cfg \ MAX_PERIOD) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. scrc_budget cfg \ scrc_period cfg) (sc_refill_cfgs_of s) sc_ptr)\ - handle_overrun_loop_body usage + charge_entire_head_refill sc_ptr' usage \\_ s. pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr\" (is "valid _ _ (\_ s. ?post s)") - apply (rule_tac P'1="\s. sc_ptr = cur_sc s + apply (rule_tac P'1="\s. sc_ptr = sc_ptr' \ pred_map (\cfg. no_overflow (scrc_refills cfg)) - (sc_refill_cfgs_of s) (cur_sc s)" + (sc_refill_cfgs_of s) sc_ptr" in hoare_pre_add[THEN iffD2, simplified pred_conj_def]) apply (clarsimp split: if_splits) - apply (case_tac "sc_ptr \ cur_sc s") + apply (case_tac "sc_ptr \ sc_ptr'") apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps) apply (fastforce intro: head_time_buffer_implies_no_overflow) - apply (clarsimp simp: handle_overrun_loop_body_def) - apply (rule bind_wp[OF _ gets_sp]) - apply (rule bind_wp[OF _ refill_single_sp]) + apply (clarsimp simp: charge_entire_head_refill_def) + apply (rule bind_wp[OF _ get_refill_head_sp]) apply (rule bind_wp[OF _ get_sched_context_sp]) + apply (rule bind_wp[OF _ refill_single_sp]) apply (rule_tac Q'="\_ s. ?post s" in bind_wp; (solves wpsimp)?) apply (rule hoare_if) apply (wpsimp simp: update_refill_hd_rewrite @@ -9797,7 +9831,7 @@ lemma handle_overrun_loop_body_ordered_disjoint: apply (clarsimp simp: vs_all_heap_simps ordered_disjoint_def obj_at_kh_kheap_simps) apply (wpsimp simp: refill_pop_head_def update_sched_context_set_refills_rewrite wp: schedule_used_ordered_disjoint set_refills_wp get_refills_wp) - apply (case_tac "sc_ptr \ cur_sc s") + apply (case_tac "sc_ptr \ sc_ptr'") apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps) apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps) apply (intro conjI impI) @@ -9807,162 +9841,133 @@ lemma handle_overrun_loop_body_ordered_disjoint: apply (subst unat_add_lem'; fastforce simp: unat_minus_one_word window_def last_tl) done -lemma hoare_vcg_imp_lift_pre_add: - "\ \P and Q\ f \\rv s. R rv s\; f \\s. \ Q s\ \ \ \P\ f \\rv s. Q s \ R rv s\" - apply (rule hoare_weaken_pre) - apply (rule hoare_vcg_imp_lift') - apply fastforce - apply fastforce - apply (clarsimp simp: pred_conj_def valid_def) - done - -lemma handle_overrun_loop_ordered_disjoint: +lemma handle_overrun_ordered_disjoint: "\\s. pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr - \ (sc_ptr = cur_sc s + \ (sc_ptr = sc_ptr' \ pred_map (\cfg. window (scrc_refills cfg) (scrc_period cfg)) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ pred_map (\cfg. scrc_period cfg \ MAX_PERIOD) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. scrc_budget cfg \ scrc_period cfg) (sc_refill_cfgs_of s) sc_ptr)\ - handle_overrun_loop usage + handle_overrun sc_ptr' usage \\_ s. pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr\" (is "valid ?pre _ _") supply if_split[split del] - apply (clarsimp simp: handle_overrun_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_ s. ?pre s"] - ; (fastforce split: if_splits)?) + apply (clarsimp simp: handle_overrun_def) + apply (wpsimp wp: valid_whileLoop[where I="\_ s. ?pre s"]; (fastforce split: if_splits)?) apply (intro hoare_vcg_conj_lift_pre_fix) - apply (wpsimp wp: handle_overrun_loop_body_ordered_disjoint) - apply (fastforce dest: head_time_buffer_true_imp_unat_buffer[THEN iffD1, rotated] - simp: vs_all_heap_simps unat_MAX_RELEASE_TIME) - apply (rule hoare_vcg_imp_lift_pre_add; (solves handle_overrun_loop_body_simple)?) - apply (intro hoare_vcg_conj_lift_pre_fix - ; (solves handle_overrun_loop_body_simple)?) - apply wpsimp - apply (rule_tac f=cur_sc in hoare_lift_Pf2) - apply (wpsimp wp: handle_overrun_loop_body_window) - apply (fastforce intro!: head_time_buffer_implies_no_overflow - dest: head_time_buffer_true_imp_unat_buffer[THEN iffD1, rotated] - simp: vs_all_heap_simps unat_MAX_RELEASE_TIME) - apply wpsimp - apply clarsimp - apply (rule_tac f=cur_sc in hoare_lift_Pf2) - apply (wpsimp wp: handle_overrun_loop_body_refills_unat_sum_equals_budget) - apply wpsimp + apply (wpsimp wp: charge_entire_head_refill_ordered_disjoint) + apply (fastforce dest: head_refill_overrun_true_imp_unat_buffer + simp: vs_all_heap_simps unat_MAX_RELEASE_TIME) + apply (rule hoare_vcg_imp_lift_pre_add; (solves wpsimp)?) + apply (intro hoare_vcg_conj_lift_pre_fix) + apply (wpsimp wp: charge_entire_head_refill_window) + apply (fastforce intro!: head_time_buffer_implies_no_overflow + simp: vs_all_heap_simps unat_MAX_RELEASE_TIME + head_refill_overrun_true_imp_unat_buffer) + apply wpsimp + apply charge_entire_head_refill_simple + apply (wpsimp wp: charge_entire_head_refill_refills_unat_sum_equals_budget) + apply charge_entire_head_refill_simple done -lemma handle_overrun_loop_window: +lemma handle_overrun_window: "\\s. pred_map (\cfg. window (scrc_refills cfg) (scrc_period cfg)) (sc_refill_cfgs_of s) sc_ptr - \ (sc_ptr = cur_sc s + \ (sc_ptr = sc_ptr' \ pred_map (\cfg. window (scrc_refills cfg) (scrc_period cfg)) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ pred_map (\cfg. scrc_period cfg \ MAX_PERIOD) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. scrc_budget cfg \ scrc_period cfg) (sc_refill_cfgs_of s) sc_ptr)\ - handle_overrun_loop usage + handle_overrun sc_ptr' usage \\_ s. pred_map (\cfg. window (scrc_refills cfg) (scrc_period cfg)) (sc_refill_cfgs_of s) sc_ptr\" (is "valid ?pre _ _") supply if_split[split del] - apply (clarsimp simp: handle_overrun_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) + apply (clarsimp simp: handle_overrun_def) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) apply (rule hoare_vcg_conj_lift_pre_fix) - apply (wpsimp wp: handle_overrun_loop_body_window) + apply (wpsimp wp: charge_entire_head_refill_window) apply (fastforce intro!: head_time_buffer_implies_no_overflow - dest!: head_time_buffer_true_imp_unat_buffer[THEN iffD1, rotated] - simp: vs_all_heap_simps unat_MAX_RELEASE_TIME) - apply (rule hoare_vcg_imp_lift_pre_add; (solves handle_overrun_loop_body_simple)?) - apply (intro hoare_vcg_conj_lift_pre_fix - ; (solves handle_overrun_loop_simple)?) - apply (rule hoare_weaken_pre) - apply (rule_tac f=cur_sc in hoare_lift_Pf2) - apply (wpsimp wp: handle_overrun_loop_body_window) + simp: head_refill_overrun_true_imp_unat_buffer vs_all_heap_simps + unat_MAX_RELEASE_TIME) + apply (rule hoare_vcg_imp_lift_pre_add; (solves wpsimp)?) + apply (intro hoare_vcg_conj_lift_pre_fix) + apply (wpsimp wp: charge_entire_head_refill_window) + apply (fastforce intro!: head_time_buffer_implies_no_overflow + simp: head_refill_overrun_true_imp_unat_buffer vs_all_heap_simps + unat_MAX_RELEASE_TIME) + apply (wpsimp wp: charge_entire_head_refill_ordered_disjoint) + apply (fastforce dest: head_refill_overrun_true_imp_unat_buffer[THEN iffD1, rotated] + simp: vs_all_heap_simps MAX_RELEASE_TIME_def MAX_PERIOD_mult unat_sub) apply wpsimp - apply (fastforce intro!: head_time_buffer_implies_no_overflow - dest!: head_time_buffer_true_imp_unat_buffer[THEN iffD1, rotated] - simp: vs_all_heap_simps unat_MAX_RELEASE_TIME) - apply (rule hoare_weaken_pre) - apply (rule_tac f=cur_sc in hoare_lift_Pf2) - apply (wpsimp wp: handle_overrun_loop_body_ordered_disjoint) - apply wpsimp - apply (fastforce dest: head_time_buffer_true_imp_unat_buffer[THEN iffD1, rotated] - simp: vs_all_heap_simps MAX_RELEASE_TIME_def MAX_PERIOD_mult unat_sub) - apply wpsimp - apply (rule_tac f=cur_sc in hoare_lift_Pf2) - apply (wpsimp wp: handle_overrun_loop_body_refills_unat_sum_equals_budget) - apply wpsimp - done - -lemma handle_overrun_loop_head_bound: - "\\s. cur_sc_offset_ready 0 s \ valid_refills (cur_sc s) s \ current_time_bounded s - \ \ round_robin (cur_sc s) s - \ pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) (cur_sc s) - \ pred_map (\cfg. no_overflow (scrc_refills cfg)) (sc_refill_cfgs_of s) (cur_sc s) - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) (cur_sc s) + apply charge_entire_head_refill_simple + apply (wpsimp wp: charge_entire_head_refill_refills_unat_sum_equals_budget) + apply charge_entire_head_refill_simple + done + +lemma handle_overrun_head_bound: + "\\s. pred_map (refill_ready_no_overflow_sc 0 (cur_time s)) (sc_refill_cfgs_of s) sc_ptr + \ valid_refills sc_ptr s \ current_time_bounded s + \ \ round_robin sc_ptr s + \ pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr + \ pred_map (\cfg. no_overflow (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) - (sc_refill_cfgs_of s) (cur_sc s) - \ pred_map (\cfg. scrc_budget cfg \ scrc_period cfg) (sc_refill_cfgs_of s) (cur_sc s) - \ pred_map (\cfg. MIN_REFILLS \ scrc_refill_max cfg) (sc_refill_cfgs_of s) (cur_sc s) - \ cur_sc s \ idle_sc_ptr\ - handle_overrun_loop usage + (sc_refill_cfgs_of s) sc_ptr + \ pred_map (\cfg. scrc_budget cfg \ scrc_period cfg) (sc_refill_cfgs_of s) sc_ptr + \ pred_map (\cfg. MIN_REFILLS \ scrc_refill_max cfg) (sc_refill_cfgs_of s) sc_ptr + \ sc_ptr \ idle_sc_ptr\ + handle_overrun sc_ptr usage \\_ s. pred_map (\cfg. unat (r_time (hd (scrc_refills cfg))) + 4 * unat MAX_PERIOD \ unat max_time) - (sc_refill_cfgs_of s) (cur_sc s)\" + (sc_refill_cfgs_of s) sc_ptr\" (is "valid _ _ (\_ s. ?post s)") - apply (rule_tac P'1="\s. pred_map (\cfg. no_overflow (scrc_refills cfg)) (sc_refill_cfgs_of s) (cur_sc s)" + apply (rule_tac P'1="\s. pred_map (\cfg. no_overflow (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr" in hoare_pre_add[THEN iffD2, simplified pred_conj_def]) apply (blast intro: head_time_buffer_implies_no_overflow) - apply (clarsimp simp: handle_overrun_loop_def) + apply (clarsimp simp: handle_overrun_def) apply (rule_tac I="\_ s. ?post s \ pred_map (\cfg. scrc_period cfg \ MAX_PERIOD) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. window (scrc_refills cfg) (scrc_period cfg)) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) - (sc_refill_cfgs_of s) (cur_sc s) - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. scrc_budget cfg \ scrc_period cfg) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. MIN_REFILLS \ scrc_refill_max cfg) - (sc_refill_cfgs_of s) (cur_sc s)" - in valid_whileLoop - ; fastforce?) - apply (fastforce simp: cur_sc_offset_ready_def vs_all_heap_simps current_time_bounded_def - valid_refills_def round_robin_def) + (sc_refill_cfgs_of s) sc_ptr" + in valid_whileLoop; + fastforce?) + apply (clarsimp simp: cur_sc_offset_ready_def vs_all_heap_simps current_time_bounded_def + valid_refills_def round_robin_def) - apply (intro hoare_vcg_conj_lift_pre_fix - ; (solves handle_overrun_loop_body_simple)?) - defer - apply (rule_tac f=cur_sc in hoare_lift_Pf2) - apply (wpsimp wp: handle_overrun_loop_body_window) + apply (intro hoare_vcg_conj_lift_pre_fix; (solves charge_entire_head_refill_simple)?) + defer + apply (wpsimp wp: charge_entire_head_refill_window) apply (fastforce intro!: head_time_buffer_implies_no_overflow - simp: vs_all_heap_simps) - apply wpsimp - apply (rule_tac f=cur_sc in hoare_lift_Pf2) - apply (wpsimp wp: handle_overrun_loop_body_ordered_disjoint) - apply (fastforce simp: vs_all_heap_simps) - apply wpsimp - apply (rule_tac f=cur_sc in hoare_lift_Pf2) - apply (wpsimp wp: handle_overrun_loop_body_refills_unat_sum_equals_budget) - apply wpsimp + simp: vs_all_heap_simps sc_at_ppred_def obj_at_def) + apply (wpsimp wp: charge_entire_head_refill_ordered_disjoint) + apply (fastforce simp: vs_all_heap_simps sc_at_ppred_def obj_at_def) + apply (wpsimp wp: charge_entire_head_refill_nonzero_refills) + apply (wpsimp wp: charge_entire_head_refill_refills_unat_sum_equals_budget) - apply (rule_tac f=cur_sc in hoare_lift_Pf2; (solves wpsimp)?) - apply (rename_tac sc_ptr) apply (rule_tac P'1="\s. pred_map (\cfg. no_overflow (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr" in hoare_pre_add[THEN iffD2, simplified pred_conj_def]) apply (fastforce dest!: head_time_buffer_implies_no_overflow[rotated 1] simp: vs_all_heap_simps) - apply (clarsimp simp: handle_overrun_loop_body_def) - apply (rule bind_wp[OF _ gets_sp]) - apply (rule bind_wp[OF _ refill_single_sp]) + apply (clarsimp simp: charge_entire_head_refill_def) + apply (rule bind_wp[OF _ get_refill_head_sp]) apply (rule bind_wp[OF _ get_sched_context_sp]) + apply (rule bind_wp[OF _ refill_single_sp]) apply (rule_tac Q'="\_ s. pred_map (\cfg. unat (r_time (hd (scrc_refills cfg))) + 4 * unat MAX_PERIOD \ unat max_time) (sc_refill_cfgs_of s) sc_ptr @@ -9976,7 +9981,7 @@ lemma handle_overrun_loop_head_bound: apply (rule hoare_if) apply (clarsimp simp: update_sched_context_set_refills_rewrite update_refill_hd_rewrite) apply (wpsimp wp: get_refills_wp set_refills_wp) - apply (frule head_time_buffer_true_imp_unat_buffer[THEN iffD1, rotated]) + apply (frule head_refill_overrun_true_imp_unat_buffer[THEN iffD1, rotated]) apply (fastforce simp: vs_all_heap_simps obj_at_kh_kheap_simps) apply (clarsimp simp: obj_at_def vs_all_heap_simps word_le_nat_alt) apply (rename_tac sc n) @@ -9986,76 +9991,74 @@ lemma handle_overrun_loop_head_bound: apply (clarsimp simp: window_def) apply (subst unat_add_lem') apply (clarsimp simp: unat_minus_one_word word_le_nat_alt) - apply (frule head_time_buffer_true_imp_unat_buffer[THEN iffD1, rotated]) + apply (frule head_refill_overrun_true_imp_unat_buffer[THEN iffD1, rotated]) apply (fastforce simp: vs_all_heap_simps obj_at_kh_kheap_simps) apply (fastforce simp: unat_minus_one_word unat_MAX_RELEASE_TIME less_not_refl2 tail_nonempty_length) - apply (intro hoare_vcg_conj_lift_pre_fix; (solves handle_overrun_loop_simple)?) - defer - apply (rule_tac f=cur_sc in hoare_lift_Pf2) - apply (wpsimp wp: schedule_used_window get_refills_wp set_refills_wp update_sched_context_wp - simp: refill_pop_head_def) - apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps pred_map_simps) - apply (rename_tac sc n) - apply (frule head_time_buffer_true_imp_unat_buffer[THEN iffD1, rotated]) - apply (fastforce simp: vs_all_heap_simps obj_at_kh_kheap_simps) - apply (prop_tac "unat (r_time (refill_hd sc)) \ unat (r_time (hd (tl (sc_refills sc))))") - apply (subst hd_tl_nth) - apply (simp add: Suc_lessI) - apply (subst hd_conv_nth) - apply blast - apply (erule (1) ordered_disjoint_no_overflow_implies_sorted) - apply force - apply (simp add: Suc_lessI) - apply simp - - apply (prop_tac "hd (tl (sc_refills sc)) \ set (sc_refills sc)") - apply (metis Nitpick.size_list_simp(2) hd_Cons_tl hd_in_set list.set_intros(2)) - apply (frule (2) ordered_disjoint_last) - apply (intro conjI impI) - apply (clarsimp simp: window_def) - apply (prop_tac "last (tl (sc_refills sc)) = refill_tl sc") - apply (metis Nitpick.size_list_simp(2) last_tl) - apply clarsimp - apply (blast intro: ordered_disjoint_sublist) - apply (blast intro: no_overflow_sublist) - apply (metis Nitpick.size_list_simp(2)) - apply (rule word_add_le_mono1) - apply (clarsimp simp: word_le_nat_alt) - apply (frule (1) ordered_disjoint_last) - apply metis - apply (subst power_two_max_word_fold) - apply (clarsimp simp: unat_max_word word_le_nat_alt) - apply (rule le_imp_less_Suc) - apply (fastforce simp: window_def) - apply (fastforce simp: window_def word_le_nat_alt) - apply wpsimp - + apply (intro hoare_vcg_conj_lift_pre_fix) + defer + apply (wpsimp simp: refill_budget_check_defs update_sched_context_set_refills_rewrite + wp: set_refills_wp get_refills_wp) + apply (clarsimp simp: vs_all_heap_simps sc_at_ppred_def obj_at_def) + apply (wpsimp wp: schedule_used_window get_refills_wp set_refills_wp update_sched_context_wp + simp: refill_pop_head_def) + apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps pred_map_simps) + apply (rename_tac sc) + apply (frule head_refill_overrun_true_imp_unat_buffer[THEN iffD1, rotated]) + apply (fastforce simp: vs_all_heap_simps obj_at_kh_kheap_simps) + apply (prop_tac "unat (r_time (refill_hd sc)) \ unat (r_time (hd (tl (sc_refills sc))))") + apply (subst hd_tl_nth) + apply (simp add: Suc_lessI) + apply (subst hd_conv_nth) + apply blast + apply (erule (1) ordered_disjoint_no_overflow_implies_sorted) + apply force + apply (simp add: Suc_lessI) + apply simp + apply (prop_tac "hd (tl (sc_refills sc)) \ set (sc_refills sc)") + apply (metis Nitpick.size_list_simp(2) hd_Cons_tl hd_in_set list.set_intros(2)) + apply (frule (2) ordered_disjoint_last) + apply (intro conjI impI) + apply (clarsimp simp: window_def) + apply (prop_tac "last (tl (sc_refills sc)) = refill_tl sc") + apply (metis Nitpick.size_list_simp(2) last_tl) + apply clarsimp + apply (blast intro: ordered_disjoint_sublist) + apply (blast intro: no_overflow_sublist) + apply (metis Nitpick.size_list_simp(2)) + apply (rule word_add_le_mono1) + apply (clarsimp simp: word_le_nat_alt) + apply (frule (1) ordered_disjoint_last) + apply metis + apply (subst power_two_max_word_fold) + apply (clarsimp simp: unat_max_word word_le_nat_alt) + apply (rule le_imp_less_Suc) + apply (fastforce simp: window_def) + apply (fastforce simp: window_def word_le_nat_alt) apply (wpsimp wp: get_refills_wp set_refills_wp update_sched_context_wp simp: refill_pop_head_def schedule_used_defs) - apply (prop_tac "pred_map \ (scs_of s) (cur_sc s)") - apply (clarsimp simp: obj_at_def vs_all_heap_simps split: if_splits) - apply (frule_tac usage=r in head_time_buffer_true_imp_unat_buffer) + apply (prop_tac "pred_map \ (scs_of s) sc_ptr") + apply (clarsimp simp: obj_at_def vs_all_heap_simps sc_at_ppred_def split: if_splits) + apply (frule_tac usage=r in head_refill_overrun_true_imp_unat_buffer) apply (clarsimp simp: obj_at_def vs_all_heap_simps split: if_splits) apply (frule (1) ordered_disjoint_last) - apply metis - apply (case_tac "tl (sc_refills scb)") + apply fastforce + apply (rename_tac sc) + apply (case_tac "tl (sc_refills sc)") apply (clarsimp simp: window_def MAX_RELEASE_TIME_def unat_sub) apply (prop_tac "unat (r_time a) + 4 * unat MAX_PERIOD \ unat max_time") - apply (prop_tac "hd (tl (sc_refills scb)) \ set (tl (sc_refills scb))", fastforce) - apply (drule_tac x="hd (tl (sc_refills scb))" in bspec) - apply (simp add: list.set_sel(2)) - apply (frule (2) ordered_disjoint_last) - apply (drule_tac x="hd (tl (sc_refills scb))" in bspec) + apply (prop_tac "hd (tl (sc_refills sc)) \ set (tl (sc_refills sc))", fastforce) + apply (drule_tac x="hd (tl (sc_refills sc))" in bspec) apply (simp add: list.set_sel(2)) + apply (metis list.set_intros(1) list.set_sel(2) tl_Nil) + apply (simp add: list.set_sel(2)) apply (clarsimp simp: window_def MAX_RELEASE_TIME_def unat_sub word_le_nat_alt) apply (fastforce simp: MAX_PERIOD_mult[where n=5]) apply clarsimp apply (subst unat_sub) apply (clarsimp simp: window_def MAX_RELEASE_TIME_def unat_sub word_le_nat_alt) apply (subst unat_add_lem''; fastforce?) - apply (metis MIN_REFILLS_def Nitpick.size_list_simp(2) le_def list.size(4) list_exhaust_size_eq0 - minus_Suc_0_lt mult_is_add.mult_1_left numerals(2)) + apply (metis Suc_length_not_empty hd_Cons_tl le_trans nat_iffs(2) non_empty_tail_length refills_unat_sum_cons refills_unat_sum_length_one) apply (clarsimp simp: window_def MAX_RELEASE_TIME_def unat_sub word_le_nat_alt) by (metis MIN_REFILLS_def One_nat_def length_tl list.size(3) list.size(4) mult_is_add.mult_1_left not_less_eq_eq numerals(2) order_refl) @@ -10088,98 +10091,94 @@ lemma schedule_used_refills_sum: apply (metis add_ac) done -lemma handle_overrun_loop_body_refills_sum: +lemma charge_entire_head_refill_refills_sum: "\\s. pred_map (\cfg. refills_sum (scrc_refills cfg) = scrc_budget cfg) (sc_refill_cfgs_of s) sc_ptr - \ (sc_ptr = cur_sc s \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr)\ - handle_overrun_loop_body r + \ (sc_ptr = sc_ptr' \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr)\ + charge_entire_head_refill sc_ptr' r \\_ s. pred_map (\cfg. refills_sum (scrc_refills cfg) = scrc_budget cfg) (sc_refill_cfgs_of s) sc_ptr\" (is "valid _ _ (\_ s. ?post s)") - apply (clarsimp simp: handle_overrun_loop_body_def) - apply (rule bind_wp[OF _ gets_sp]) - apply (rule bind_wp[OF _ refill_single_sp]) + apply (clarsimp simp: charge_entire_head_refill_def) + apply (rule bind_wp[OF _ get_refill_head_sp]) apply (rule bind_wp[OF _ get_sched_context_sp]) + apply (rule bind_wp[OF _ refill_single_sp]) apply (rule_tac Q'="\_ s. ?post s" in bind_wp; (solves wpsimp)?) apply (rule hoare_if) apply (wpsimp simp: update_sched_context_set_refills_rewrite update_refill_hd_rewrite wp: set_refills_wp get_refills_wp) apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps) - apply (rename_tac sc n) + apply (rename_tac sc) apply (case_tac "sc_refills sc"; clarsimp) apply (wpsimp simp: refill_pop_head_def update_sched_context_set_refills_rewrite schedule_used_defs wp: schedule_used_refills_sum set_refills_wp get_refills_wp) apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps pred_map_simps) - apply (rename_tac sc n) + apply (rename_tac sc) apply (case_tac "sc_refills sc"; clarsimp simp: refills_sum_def add_ac) done -lemma handle_overrun_loop_refills_sum: +lemma handle_overrun_refills_sum: "\\s. pred_map (\cfg. refills_sum (scrc_refills cfg) = scrc_budget cfg) (sc_refill_cfgs_of s) sc_ptr - \ (sc_ptr = cur_sc s \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr)\ - handle_overrun_loop usage + \ (sc_ptr = sc_ptr' \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr)\ + handle_overrun sc_ptr' usage \\_ s. pred_map (\cfg. refills_sum (scrc_refills cfg) = scrc_budget cfg) (sc_refill_cfgs_of s) sc_ptr\" (is "valid ?pre _ _") supply if_split[split del] - apply (clarsimp simp: handle_overrun_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) - apply (intro hoare_vcg_conj_lift_pre_fix; (solves handle_overrun_loop_body_simple)?) - apply (wpsimp wp: handle_overrun_loop_body_refills_sum) + apply (clarsimp simp: handle_overrun_def) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) + apply (intro hoare_vcg_conj_lift_pre_fix) + apply (wpsimp wp: charge_entire_head_refill_refills_sum) + apply charge_entire_head_refill_simple done lemma head_insufficient_loop_round_robin[wp]: "head_insufficient_loop sc_ptr' \\s. P (round_robin sc_ptr s)\" - apply head_insufficient_loop_simple - done + by head_insufficient_loop_simple -lemma handle_overrun_loop_round_robin[wp]: - "handle_overrun_loop sc_ptr' \\s. P (round_robin sc_ptr s)\" - apply handle_overrun_loop_simple - done +lemma handle_overrun_round_robin[wp]: + "handle_overrun sc_ptr usage \\s. P (round_robin sc_ptr s)\" + by handle_overrun_simple -lemma handle_overrun_loop_body_length: +lemma charge_entire_head_refill_length: "\\s. pred_map (\cfg. length (scrc_refills cfg) \ scrc_refill_max cfg) (sc_refill_cfgs_of s) sc_ptr - \ (sc_ptr = cur_sc s \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr)\ - handle_overrun_loop_body usage + \ (sc_ptr = sc_ptr' \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s)\ + charge_entire_head_refill sc_ptr' usage \\_ s. pred_map (\cfg. length (scrc_refills cfg) \ scrc_refill_max cfg) (sc_refill_cfgs_of s) sc_ptr\" - (is "valid _ _ (\_ s. ?post s)") - apply handle_overrun_loop_body_simple + apply (wpsimp simp: refill_budget_check_defs update_sched_context_set_refills_rewrite + wp: set_refills_wp get_refills_wp whileLoop_valid_inv) + apply (clarsimp simp: vs_all_heap_simps sc_at_ppred_def obj_at_def round_robin_def) apply (rename_tac sc n) apply (case_tac "sc_refills sc"; clarsimp) done -lemma handle_overrun_loop_length: +lemma handle_overrun_length: "\\s. pred_map (\cfg. length (scrc_refills cfg) \ scrc_refill_max cfg) (sc_refill_cfgs_of s) sc_ptr - \ (sc_ptr = cur_sc s \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr)\ - handle_overrun_loop usage + \ (sc_ptr = sc_ptr' \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s)\ + handle_overrun sc_ptr' usage \\_ s. pred_map (\cfg. length (scrc_refills cfg) \ scrc_refill_max cfg) (sc_refill_cfgs_of s) sc_ptr\" (is "valid ?pre _ _") supply if_split[split del] - apply (clarsimp simp: handle_overrun_loop_def) + apply (clarsimp simp: handle_overrun_def) apply (rule_tac I="\_. ?pre" in valid_whileLoop; (solves simp)?) - apply (intro hoare_vcg_conj_lift_pre_fix; (solves handle_overrun_loop_body_simple)?) - apply (wpsimp wp: handle_overrun_loop_body_length) + apply (wpsimp wp: charge_entire_head_refill_length + charge_entire_head_refill_nonzero_refills hoare_vcg_imp_lift') done -lemma handle_overrun_loop_body_scs_of_cur_sc[wp]: - "handle_overrun_loop_body usage \\s. pred_map \ (scs_of s) (cur_sc s)\" - apply (clarsimp simp: handle_overrun_loop_body_def schedule_used_defs) - apply (wpsimp simp: update_refill_hd_def refill_pop_head_def refill_single_def refill_size_def - wp: update_sched_context_wp get_refills_wp) - apply (clarsimp simp: vs_all_heap_simps) - done +lemma charge_entire_head_refill_scs_of[wp]: + "charge_entire_head_refill sc_ptr' usage \\s. pred_map \ (scs_of s) sc_ptr\" + by charge_entire_head_refill_simple lemma refill_budget_check_ordered_disjoint_helper: "\ordered_disjoint (sc_refills sc); sc_refills sc \ []; u < r_amount (refill_hd sc); no_overflow (sc_refills sc)\ - \ ordered_disjoint (r_time_update (\t. t + u) (r_amount_update (\m. m - u) (refill_hd sc)) + \ ordered_disjoint (refill_hd sc\r_amount := r_amount (refill_hd sc) - u, + r_time := r_time (refill_hd sc) + u\ # tl (sc_refills sc))" apply (rule_tac left="[r_time_update (\t. t + u) (r_amount_update (\m. m - u) (refill_hd sc))]" and right="tl (sc_refills sc)" @@ -10201,7 +10200,8 @@ lemma refill_budget_check_ordered_disjoint_helper: lemma refill_budget_check_no_overflow_helper: "\no_overflow (sc_refills sc); ordered_disjoint (sc_refills sc); sc_refills sc \ []; u < r_amount (refill_hd sc)\ - \ no_overflow (r_time_update (\t. t + u) (r_amount_update (\m. m - u) (refill_hd sc)) + \ no_overflow (refill_hd sc \r_amount := r_amount (refill_hd sc) - u, + r_time := r_time (refill_hd sc) + u\ # tl (sc_refills sc))" apply (rule_tac left="[r_time_update (\t. t + u) (r_amount_update (\m. m - u) (refill_hd sc))]" and right="tl (sc_refills sc)" @@ -10219,38 +10219,39 @@ lemma refill_budget_check_no_overflow_helper: lemma refill_budget_check_refills_sum_helper: "\\s. all_sp_valid_refills_but_MIN_BUDGET_in_head_no_overflow sc_ptr s \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) - (sc_refill_cfgs_of s) sc_ptr\ - do y \ update_refill_hd csc_ptr (r_time_update (\t. t + u) \ r_amount_update (\m. m - u)); - schedule_used csc_ptr \r_time = r_time (refill_hd sc) + sc_period sc, r_amount = u\ + (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ [] \ hd refills = refill_head) sc_ptr s\ + do y \ update_refill_hd sc_ptr (r_amount_update (\_. r_amount refill_head - u)); + y \ update_refill_hd sc_ptr (r_time_update (\_. r_time refill_head + u)); + schedule_used sc_ptr \r_time = r_time refill_head + sc_period sc, r_amount = u\ od \\_ s. pred_map (\cfg. refills_sum (scrc_refills cfg) = scrc_budget cfg) (sc_refill_cfgs_of s) sc_ptr\" - apply (rule_tac f=cur_sc in hoare_lift_Pf2) - apply (wpsimp wp: schedule_used_refills_sum set_refills_wp get_refills_wp - simp: update_refill_hd_rewrite) - apply (clarsimp simp: obj_at_def vs_all_heap_simps) - apply (case_tac "sc_refills sc"; simp) - apply wpsimp + apply (wpsimp wp: schedule_used_refills_sum set_refills_wp get_refills_wp + simp: update_refill_hd_rewrite) + apply (clarsimp simp: obj_at_def vs_all_heap_simps sc_at_ppred_def) + apply (case_tac "sc_refills sc"; simp) done lemma refill_budget_check_ordered_disjoint_helper': - "r_time (hd refills) < MAX_RELEASE_TIME \ + "r_time refill_head < MAX_RELEASE_TIME \ \\s. all_sp_valid_refills_but_MIN_BUDGET_in_head sc_ptr s + \ sc_refills_sc_at (\refills. refills \ [] \ hd refills = refill_head) sc_ptr s \ pred_map (\cfg. unat (r_time (scrc_refill_hd cfg)) + 4 * unat MAX_PERIOD \ unat max_time) (sc_refill_cfgs_of s) sc_ptr \ (pred_map (\cfg. r_time (scrc_refill_hd cfg) < MAX_RELEASE_TIME) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. u < r_amount (scrc_refill_hd cfg)) (sc_refill_cfgs_of s) sc_ptr) - \ (\sc n. ko_at (SchedContext sc n) sc_ptr s \ refills = sc_refills sc) + \ sc_refills_sc_at (\refills. u < r_amount refill_head) sc_ptr s) \ (\n. ko_at (SchedContext sc n) sc_ptr s)\ - do y \ update_refill_hd csc_ptr (r_time_update (\t. t + u) \ r_amount_update (\m. m - u)); - schedule_used csc_ptr \r_time = r_time (refill_hd sc) + sc_period sc, r_amount = u\ + do y \ update_refill_hd sc_ptr (r_amount_update (\_. r_amount refill_head - u)); + y \ update_refill_hd sc_ptr (r_time_update (\_. r_time refill_head + u)); + schedule_used sc_ptr \r_time = r_time refill_head + sc_period sc, r_amount = u\ od \\_ s. pred_map (\cfg. ordered_disjoint (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr\" apply (wpsimp wp: schedule_used_ordered_disjoint set_refills_wp get_refills_wp simp: update_refill_hd_rewrite) - apply (clarsimp simp: obj_at_def vs_all_heap_simps) + apply (clarsimp simp: sc_at_ppred_def obj_at_def vs_all_heap_simps) apply (intro conjI impI) apply (clarsimp simp: no_overflow_def) apply (subst unat_sub) @@ -10276,8 +10277,9 @@ lemma refill_budget_check_ordered_disjoint_helper': done lemma refill_budget_check_no_overflow': - "r_time (hd refills) < MAX_RELEASE_TIME \ + "r_time refill_head < MAX_RELEASE_TIME \ \\s. all_sp_valid_refills_but_MIN_BUDGET_in_head sc_ptr s + \ sc_refills_sc_at (\refills. refills \ [] \ hd refills = refill_head) sc_ptr s \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. unat (r_time (scrc_refill_hd cfg)) + 4 * unat MAX_PERIOD @@ -10285,16 +10287,16 @@ lemma refill_budget_check_no_overflow': (sc_refill_cfgs_of s) sc_ptr \ (pred_map (\cfg. r_time (scrc_refill_hd cfg) < MAX_RELEASE_TIME) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. u < r_amount (scrc_refill_hd cfg)) (sc_refill_cfgs_of s) sc_ptr) - \ (\sc n. ko_at (SchedContext sc n) sc_ptr s \ refills = sc_refills sc) + \ sc_refills_sc_at (\refills. u < r_amount refill_head) sc_ptr s) \ (\n. ko_at (SchedContext sc n) sc_ptr s)\ - do y \ update_refill_hd csc_ptr (r_time_update (\t. t + u) \ r_amount_update (\m. m - u)); - schedule_used csc_ptr \r_time = r_time (refill_hd sc) + sc_period sc, r_amount = u\ + do y \ update_refill_hd sc_ptr (r_amount_update (\_. r_amount refill_head - u)); + y \ update_refill_hd sc_ptr (r_time_update (\_. r_time refill_head + u)); + schedule_used sc_ptr \r_time = r_time refill_head + sc_period sc, r_amount = u\ od \\_ s. pred_map (\cfg. no_overflow (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr\" apply (wpsimp wp: schedule_used_no_overflow set_refills_wp get_refills_wp simp: update_refill_hd_rewrite) - apply (clarsimp simp: obj_at_def vs_all_heap_simps) + apply (clarsimp simp: sc_at_ppred_def obj_at_def vs_all_heap_simps) apply (intro conjI impI allI) apply (clarsimp simp: no_overflow_def) apply (subst unat_add_subtract_cancel) @@ -10333,23 +10335,24 @@ lemma refill_budget_check_no_overflow': done lemma refill_budget_check_window_helper: - "r_time (hd refills) < MAX_RELEASE_TIME \ + "r_time refill_head < MAX_RELEASE_TIME \ \\s. all_sp_valid_refills_but_MIN_BUDGET_in_head sc_ptr s + \ sc_refills_sc_at (\refills. refills \ [] \ hd refills = refill_head) sc_ptr s \ pred_map (\cfg. unat (r_time (scrc_refill_hd cfg)) + 4 * unat MAX_PERIOD \ unat max_time) (sc_refill_cfgs_of s) sc_ptr \ (pred_map (\cfg. r_time (scrc_refill_hd cfg) < MAX_RELEASE_TIME) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. u < r_amount (scrc_refill_hd cfg)) (sc_refill_cfgs_of s) sc_ptr) - \ (\sc n. ko_at (SchedContext sc n) sc_ptr s \ refills = sc_refills sc) + \ sc_refills_sc_at (\refills. u < r_amount refill_head) sc_ptr s) \ (\n. ko_at (SchedContext sc n) sc_ptr s)\ - do y \ update_refill_hd csc_ptr (r_time_update (\t. t + u) \ r_amount_update (\m. m - u)); - schedule_used csc_ptr \r_time = r_time (refill_hd sc) + sc_period sc, r_amount = u\ + do y \ update_refill_hd sc_ptr (r_amount_update (\_. r_amount refill_head - u)); + y \ update_refill_hd sc_ptr (r_time_update (\_. r_time refill_head + u)); + schedule_used sc_ptr \r_time = r_time refill_head + sc_period sc, r_amount = u\ od \\_ s. pred_map (\cfg. window (scrc_refills cfg) (scrc_period cfg)) (sc_refill_cfgs_of s) sc_ptr\" apply (wpsimp wp: schedule_used_window set_refills_wp get_refills_wp simp: update_refill_hd_rewrite) - apply (clarsimp simp: obj_at_def vs_all_heap_simps) + apply (clarsimp simp: sc_at_ppred_def obj_at_def vs_all_heap_simps) apply (frule_tac refill="refill_hd sc" in refill_budget_check_refill_amount_helper; simp?) apply (intro conjI impI allI) apply (fastforce simp: unat_add_lem'' word_less_nat_alt word_le_nat_alt window_def last_tl) @@ -10360,41 +10363,42 @@ lemma refill_budget_check_window_helper: done lemma refill_budget_check_length_helper: - "\\s. pred_map (\cfg. 0 < length (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr + "\\s. sc_refills_sc_at (\refills. 0 < length refills \ hd refills = refill_head) sc_ptr s \ pred_map (\cfg. length (scrc_refills cfg) \ scrc_refill_max cfg) - (sc_refill_cfgs_of s) sc_ptr \ sc_ptr = csc_ptr\ - do y \ update_refill_hd csc_ptr (r_time_update (\t. t + u) \ r_amount_update (\m. m - u)); - schedule_used csc_ptr \r_time = r_time (refill_hd sc) + sc_period sc, r_amount = u\ + (sc_refill_cfgs_of s) sc_ptr\ + do y \ update_refill_hd sc_ptr (r_amount_update (\_. r_amount refill_head - u)); + y \ update_refill_hd sc_ptr (r_time_update (\_. r_time refill_head + u)); + schedule_used sc_ptr \r_time = r_time refill_head + sc_period sc, r_amount = u\ od \\_ s. pred_map (\cfg. length (scrc_refills cfg) \ scrc_refill_max cfg) (sc_refill_cfgs_of s) sc_ptr\" apply (wpsimp wp: set_refills_wp get_refills_wp simp: vs_all_heap_simps update_sched_context_set_refills_rewrite update_refill_hd_rewrite schedule_used_defs) - apply (clarsimp simp: obj_at_def) - apply (metis Nitpick.size_list_simp(2) Suc_leI Suc_le_mono Suc_length_not_empty diff_Suc_Suc - minus_Suc_0_lt) - done + apply (clarsimp simp: sc_at_ppred_def obj_at_def) + apply (rename_tac sc' n sc''c) + apply (case_tac "sc_refills sc'"; clarsimp) + by (metis Suc_diff_le Suc_length_not_empty diff_Suc_1' diff_Suc_Suc le_Suc_eq) lemma refill_budget_check_refills_unat_sum_helper: - "r_time (hd refills) < MAX_RELEASE_TIME \ + "r_time refill_head < MAX_RELEASE_TIME \ \\s. all_sp_valid_refills_but_MIN_BUDGET_in_head_no_overflow sc_ptr s + \ sc_refills_sc_at (\refills. refills \ [] \ hd refills = refill_head) sc_ptr s \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr \ (pred_map (\cfg. r_time (scrc_refill_hd cfg) < MAX_RELEASE_TIME) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. u < r_amount (scrc_refill_hd cfg)) (sc_refill_cfgs_of s) sc_ptr) - \ (\sc n. ko_at (SchedContext sc n) sc_ptr s \ refills = sc_refills sc)\ - do y \ update_refill_hd csc_ptr (r_time_update (\t. t + u) \ r_amount_update (\m. m - u)); - schedule_used csc_ptr \r_time = r_time (refill_hd sc) + sc_period sc, r_amount = u\ + \ sc_refills_sc_at (\refills. u < r_amount refill_head) sc_ptr s)\ + do y \ update_refill_hd sc_ptr (r_amount_update (\_. r_amount refill_head - u)); + y \ update_refill_hd sc_ptr (r_time_update (\_. r_time refill_head + u)); + schedule_used sc_ptr \r_time = r_time refill_head + sc_period sc, r_amount = u\ od \\_ s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr\" apply (wpsimp wp: schedule_used_refills_unat_sum set_refills_wp get_refills_wp simp: update_refill_hd_rewrite) - apply (clarsimp simp: obj_at_def vs_all_heap_simps) apply (fastforce dest: refill_budget_check_refills_unat_sum_word_helper - simp: refills_unat_sum_def) + simp: refills_unat_sum_def sc_at_ppred_def obj_at_def vs_all_heap_simps) done lemma schedule_used_non_zero_refills: @@ -10409,6 +10413,7 @@ lemma schedule_used_non_zero_refills: (sc_refill_cfgs_of s) sc_ptr\" supply map_map[simp del] apply schedule_used_simple + apply (rename_tac sc) apply (prop_tac " 0 < unat (r_amount (refill_tl sc) + r_amount new)") apply (subst unat_add_lem'') apply (clarsimp simp: refills_unat_sum_append) @@ -10421,84 +10426,91 @@ lemma schedule_used_non_zero_refills: apply (meson in_set_butlastD) done -crunch refill_pop_head - for non_zero_refills[wp]: - "\s. pred_map (\cfg. \refill \ set (scrc_refills cfg). 0 < unat (r_amount refill)) - (sc_refill_cfgs_of s) sc_ptr" - (wp: crunch_wps update_sched_context_wp) +lemma refill_pop_head_non_zero_refills: + "refill_pop_head sc_ptr' + \\s. pred_map (\cfg. \refill \ set (scrc_refills cfg). 0 < unat (r_amount refill)) + (sc_refill_cfgs_of s) sc_ptr\" + unfolding refill_pop_head_def + apply (wpsimp wp: schedule_used_non_zero_refills update_sched_context_wp) + apply (clarsimp simp: sc_at_ppred_def obj_at_def vs_all_heap_simps) + by (metis list.set_sel(2) tl_Nil) -lemma handle_overrun_loop_body_non_zero_refills: - "\\s. if sc_ptr = cur_sc s +lemma charge_entire_head_refill_non_zero_refills: + "\\s. if sc_ptr = sc_ptr' then pred_map (\cfg. \refill \ set (scrc_refills cfg). 0 < unat (r_amount refill)) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr else pred_map (\cfg. \refill \ set (scrc_refills cfg). 0 < unat (r_amount refill)) (sc_refill_cfgs_of s) sc_ptr\ - handle_overrun_loop_body usage + charge_entire_head_refill sc_ptr' usage \\_ s. pred_map (\cfg. \refill \ set (scrc_refills cfg). 0 < unat (r_amount refill)) (sc_refill_cfgs_of s) sc_ptr\" (is "\_\ _ \\_. ?Q\") supply map_map[simp del] list.set_sel[simp add] - apply (clarsimp simp: handle_overrun_loop_body_def) - apply (rule bind_wp[OF _ gets_sp], rename_tac csc_ptr) - apply (rule bind_wp_fwd_skip, wpsimp simp: refill_single_def refill_size_def) - apply (rule bind_wp_fwd_skip, wpsimp) - apply (rule_tac Q'="\_ s. ?Q s \ csc_ptr = cur_sc s" in bind_wp) + apply (clarsimp simp: charge_entire_head_refill_def) + apply (rule bind_wp[OF _ get_refill_head_sp]) + apply (rule bind_wp[OF _ get_sched_context_sp]) + apply (rule bind_wp[OF _ refill_single_sp]) + apply (rule_tac Q'="\_ s. ?Q s" in bind_wp) apply wpsimp apply (rule hoare_if) apply (wpsimp wp: update_sched_context_wp simp: update_refill_hd_def) - apply (clarsimp simp: vs_all_heap_simps obj_at_def) - apply (intro hoare_vcg_conj_lift_pre_fix; (solves wpsimp)?) + apply (clarsimp simp: sc_at_ppred_def vs_all_heap_simps obj_at_def) apply (wpsimp wp: schedule_used_non_zero_refills update_sched_context_wp get_refills_wp simp: refill_pop_head_def) - apply (clarsimp simp: vs_all_heap_simps obj_at_def refills_unat_sum_def) - apply (intro conjI impI; force?) - apply (clarsimp simp: vs_all_heap_simps obj_at_def refills_unat_sum_def) - apply (rename_tac sc n) + apply (clarsimp simp: sc_at_ppred_def vs_all_heap_simps obj_at_def refills_unat_sum_def) + apply (rename_tac sc) apply (case_tac "sc_refills sc"; clarsimp) done lemma handle_overrun_loop_non_zero_refills: "\\s. pred_map (\cfg. \refill \ set (scrc_refills cfg). 0 < unat (r_amount refill)) (sc_refill_cfgs_of s) sc_ptr - \ (sc_ptr = cur_sc s - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr + \ (sc_ptr = sc_ptr' + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr)\ - handle_overrun_loop usage + handle_overrun sc_ptr' usage \\_ s. pred_map (\cfg. \refill\set (scrc_refills cfg). 0 < unat (r_amount refill)) (sc_refill_cfgs_of s) sc_ptr\" (is "valid ?pre _ _") supply if_split[split del] - apply (clarsimp simp: handle_overrun_loop_def) + apply (clarsimp simp: handle_overrun_def) apply (rule_tac I="\_. ?pre" in valid_whileLoop; (solves simp)?) - apply (intro hoare_vcg_conj_lift_pre_fix; (solves handle_overrun_loop_body_simple)?) - apply (wpsimp wp: handle_overrun_loop_body_non_zero_refills) + apply (intro hoare_vcg_conj_lift_pre_fix) + apply (wpsimp wp: charge_entire_head_refill_non_zero_refills) apply (clarsimp split: if_splits) apply (rule hoare_vcg_imp_lift_pre_add; (solves simp)?) - apply (wpsimp wp: handle_overrun_loop_body_refills_unat_sum_equals_budget) - apply (wpsimp wp: handle_overrun_loop_body_non_zero_refills) + apply (wpsimp wp: charge_entire_head_refill_refills_unat_sum_equals_budget) + apply (wpsimp wp: charge_entire_head_refill_non_zero_refills) done -lemma non_overlapping_merge_refills_non_zero_refills: +lemma merge_nonoverlapping_head_refill_non_zero_refills: "\\s. pred_map (\cfg. \refill\set (scrc_refills cfg). 0 < unat (r_amount refill)) (sc_refill_cfgs_of s) sc_ptr \ (sc_ptr' = sc_ptr - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr - \ the (head_insufficient sc_ptr s))\ - non_overlapping_merge_refills sc_ptr' + \ the (refill_head_insufficient sc_ptr s))\ + merge_nonoverlapping_head_refill sc_ptr' \\_ s. pred_map (\cfg. \refill\set (scrc_refills cfg). 0 < unat (r_amount refill)) (sc_refill_cfgs_of s) sc_ptr\" - apply non_overlapping_merge_refills_simple - apply (frule head_insufficient_length_at_least_two[rotated]) - apply (clarsimp simp: vs_all_heap_simps) + apply (cases "sc_ptr' \ sc_ptr") + apply merge_nonoverlapping_head_refill_simple + apply (wpsimp simp: refill_budget_check_defs update_sched_context_set_refills_rewrite + wp: set_refills_wp get_refills_wp) + apply (frule no_ofailD[OF no_ofail_head_insufficient]) + apply (frule head_insufficient_length) + apply (clarsimp simp: sc_at_ppred_def obj_at_def vs_all_heap_simps) + apply (frule head_insufficient_length) + apply clarsimp + apply (clarsimp simp: sc_at_ppred_def obj_at_def vs_all_heap_simps) apply (intro conjI impI) apply (subst unat_add_lem'') apply (case_tac "sc_refills sc"; clarsimp) @@ -10520,7 +10532,7 @@ lemma head_insufficient_loop_non_zero_refills: "\\s. pred_map (\cfg. \refill \ set (scrc_refills cfg). 0 < unat (r_amount refill)) (sc_refill_cfgs_of s) sc_ptr \ (sc_ptr' = sc_ptr - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s \ pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) @@ -10529,37 +10541,42 @@ lemma head_insufficient_loop_non_zero_refills: \\_ s. pred_map (\cfg. \refill \ set (scrc_refills cfg). 0 < unat (r_amount refill)) (sc_refill_cfgs_of s) sc_ptr\" (is "\?P\ _ \_\") + apply (cases "sc_ptr' \ sc_ptr") + apply head_insufficient_loop_simple apply (clarsimp simp: head_insufficient_loop_def) apply (rule_tac I="\_. ?P" in valid_whileLoop; fastforce?) - apply (intro hoare_vcg_conj_lift_pre_fix; (solves non_overlapping_merge_refills_simple)?) - apply (wpsimp wp: non_overlapping_merge_refills_non_zero_refills) + apply (intro hoare_vcg_conj_lift_pre_fix) + apply (wpsimp wp: merge_nonoverlapping_head_refill_non_zero_refills) apply (rule hoare_vcg_imp_lift_pre_add) - apply (intro hoare_vcg_conj_lift_pre_fix; (solves non_overlapping_merge_refills_simple)?) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum_lower_bound) - apply (fastforce dest!: head_insufficient_length_at_least_two[rotated] - simp: vs_all_heap_simps) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum_equals_budget) - apply (fastforce dest!: head_insufficient_length_at_least_two[rotated] - simp: vs_all_heap_simps) - apply (wpsimp wp: non_overlapping_merge_refills_non_zero_refills) + apply (intro hoare_vcg_conj_lift_pre_fix) + apply (wpsimp wp: merge_nonoverlapping_head_refill_nonempty_refills) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum_lower_bound) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length + simp: vs_all_heap_simps) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum_equals_budget) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length + simp: vs_all_heap_simps) + apply wpsimp done lemma refill_budget_check_non_zero_refills_helper: - "\0 < u \ r_time (hd refills) < MAX_RELEASE_TIME\ \ + "\0 < u \ r_time refill_head < MAX_RELEASE_TIME\ \ \\s. all_sp_valid_refills_but_MIN_BUDGET_in_head_no_overflow sc_ptr s + \ sc_refills_sc_at (\refills. refills \ [] \ hd refills = refill_head) sc_ptr s \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr \ (pred_map (\cfg. r_time (scrc_refill_hd cfg) < MAX_RELEASE_TIME) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. u < r_amount (scrc_refill_hd cfg)) (sc_refill_cfgs_of s) sc_ptr) - \ (\sc n. ko_at (SchedContext sc n) csc_ptr s \ refills = sc_refills sc)\ - do y \ update_refill_hd csc_ptr (r_time_update (\t. t + u) \ r_amount_update (\m. m - u)); - schedule_used csc_ptr \r_time = r_time (refill_hd sc) + sc_period sc, r_amount = u\ + \ sc_refills_sc_at (\refills. u < r_amount refill_head) sc_ptr s)\ + do y \ update_refill_hd sc_ptr (r_amount_update (\_. r_amount refill_head - u)); + y \ update_refill_hd sc_ptr (r_time_update (\_. r_time refill_head + u)); + schedule_used sc_ptr \r_time = r_time refill_head + sc_period sc, r_amount = u\ od \\_ s. pred_map (\cfg. \refill \ set (scrc_refills cfg). 0 < unat (r_amount refill)) (sc_refill_cfgs_of s) sc_ptr\" apply (wpsimp wp: schedule_used_non_zero_refills set_refills_wp get_refills_wp simp: update_refill_hd_rewrite) - apply (clarsimp simp: vs_all_heap_simps obj_at_def) + apply (clarsimp simp: sc_at_ppred_def vs_all_heap_simps obj_at_def) apply (intro conjI impI allI) apply (subst unat_sub) apply fastforce @@ -10581,7 +10598,7 @@ lemma refill_budget_check_valid_refills[wp]: apply (rule bind_wp[OF _ is_round_robin_sp], rename_tac robin) apply (rule bind_wp[OF _ assert_sp], clarsimp) - apply (case_tac "sc_ptr \ csc_ptr") + apply (case_tac "csc_ptr \ sc_ptr") apply (rule_tac Q'="\_ s. valid_refills sc_ptr s \ robin = round_robin csc_ptr s \ cur_sc s = csc_ptr" in bind_wp_fwd) @@ -10594,13 +10611,13 @@ lemma refill_budget_check_valid_refills[wp]: schedule_used_defs | fastforce simp: obj_at_def vs_all_heap_simps)+)[1] + apply clarsimp apply (rule_tac P'1="\s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr" in hoare_pre_add[THEN iffD2, simplified pred_conj_def]) apply (clarsimp simp: vs_all_heap_simps) - apply (frule valid_refills_refills_unat_sum_equals_budget - ; (fastforce simp: )?) + apply (frule valid_refills_refills_unat_sum_equals_budget; fastforce?) apply (clarsimp simp: sc_valid_refills_def vs_all_heap_simps) apply (rule_tac Q'="\usage' s. \ round_robin sc_ptr s @@ -10611,86 +10628,87 @@ lemma refill_budget_check_valid_refills[wp]: \ pred_map (\cfg. unat (r_time (hd (scrc_refills cfg))) + 4 * unat MAX_PERIOD \ unat max_time) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ cur_sc s = sc_ptr \ cur_sc s \ idle_sc_ptr \ (pred_map (\cfg. r_time (hd (scrc_refills cfg)) < MAX_RELEASE_TIME) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. usage' < r_amount (hd (scrc_refills cfg))) - (sc_refill_cfgs_of s) sc_ptr)" + \ sc_refills_sc_at (\refills. usage' < r_amount (hd refills)) sc_ptr s)" in bind_wp_fwd) - apply (intro hoare_vcg_conj_lift_pre_fix - ; (solves handle_overrun_loop_simple)?) - apply (wpsimp wp: handle_overrun_loop_refills_sum) + apply (intro hoare_vcg_conj_lift_pre_fix; (solves handle_overrun_simple)?) + apply (wpsimp wp: handle_overrun_refills_sum) apply (fastforce simp: sc_valid_refills_def vs_all_heap_simps) - apply (wpsimp wp: handle_overrun_loop_ordered_disjoint) - apply (fastforce simp: sc_valid_refills_def vs_all_heap_simps) - apply (wpsimp wp: handle_overrun_loop_window) - apply (fastforce simp: sc_valid_refills_def vs_all_heap_simps) - apply (wpsimp wp: handle_overrun_loop_length) - apply (clarsimp simp: sc_valid_refills_def vs_all_heap_simps) + apply (wpsimp wp: handle_overrun_ordered_disjoint) + apply (fastforce simp: sc_valid_refills_def vs_all_heap_simps sc_at_ppred_def obj_at_def) + apply (wpsimp wp: handle_overrun_window) + apply (fastforce simp: sc_valid_refills_def vs_all_heap_simps sc_at_ppred_def obj_at_def) + apply (wpsimp wp: handle_overrun_length) + apply (clarsimp simp: sc_valid_refills_def vs_all_heap_simps sc_at_ppred_def obj_at_def) apply (wpsimp wp: handle_overrun_loop_non_zero_refills) - apply (clarsimp simp: sc_valid_refills_def vs_all_heap_simps) - apply (wpsimp wp: handle_overrun_loop_refills_unat_sum_equals_budget) - apply (clarsimp simp: sc_valid_refills_def vs_all_heap_simps) - apply (wpsimp wp: handle_overrun_loop_head_bound) - apply (fastforce simp: sc_valid_refills_def vs_all_heap_simps) - apply (clarsimp simp: handle_overrun_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_ s. pred_map \ (scs_of s) (cur_sc s) - \ sc_ptr = cur_sc s"]) - apply (fastforce simp: head_time_buffer_true_imp_unat_buffer vs_all_heap_simps word_less_nat_alt - word_le_nat_alt) - apply (clarsimp simp: vs_all_heap_simps) + apply (clarsimp simp: sc_valid_refills_def vs_all_heap_simps sc_at_ppred_def obj_at_def) + apply (wpsimp wp: handle_overrun_refills_unat_sum_equals_budget) + apply (clarsimp simp: sc_valid_refills_def vs_all_heap_simps sc_at_ppred_def obj_at_def) + apply (wpsimp wp: handle_overrun_head_bound) + apply (fastforce simp: sc_valid_refills_def vs_all_heap_simps sc_at_ppred_def obj_at_def) + apply (clarsimp simp: handle_overrun_def) + apply (wpsimp wp: valid_whileLoop[where I="\_ s. sc_refills_sc_at (\refills. refills \ []) sc_ptr s"]) + apply (fastforce simp: head_refill_overrun_true_imp_unat_buffer vs_all_heap_simps word_less_nat_alt + word_le_nat_alt sc_at_ppred_def obj_at_def) + apply (clarsimp simp: sc_at_ppred_def obj_at_def vs_all_heap_simps valid_refills_def) apply (rule_tac P'1="\s. pred_map (\cfg. no_overflow (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr" in hoare_pre_add[THEN iffD2, simplified pred_conj_def]) apply (fastforce intro!: head_time_buffer_implies_no_overflow - simp: vs_all_heap_simps) + simp: vs_all_heap_simps sc_at_ppred_def obj_at_def) - apply (rule bind_wp[OF _ get_refills_sp]) + apply (rule bind_wp[OF _ get_refill_head_sp]) apply (rule_tac Q'="\_ s. \ round_robin sc_ptr s \ all_sp_valid_refills_but_MIN_BUDGET_in_head sc_ptr s \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr - \ sc_ptr = cur_sc s \ cur_sc s \ idle_sc_ptr" + \ cur_sc s = sc_ptr \ cur_sc s \ idle_sc_ptr" in bind_wp_fwd) apply (rule hoare_when_cases, clarsimp simp: vs_all_heap_simps) apply (rule bind_wp[OF _ get_sched_context_sp]) - apply (intro hoare_vcg_conj_lift_pre_fix - ; (solves schedule_used_simple)?) + apply (intro hoare_vcg_conj_lift_pre_fix; (solves schedule_used_simple)?) - apply (find_goal \match conclusion in "\_\ head_insufficient_loop _ \\_. _\" \ \-\\) - apply (rule_tac Q'="\_ s. \ round_robin sc_ptr s \ sp_valid_refills_unbundled sc_ptr s" - in hoare_strengthen_post[rotated]) - apply (fastforce simp: valid_refills_def vs_all_heap_simps) - apply (intro hoare_vcg_conj_lift_pre_fix - ; (solves head_insufficient_loop_simple)?) - apply (wpsimp wp: head_insufficient_loop_refills_sum) - apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt) - apply (wpsimp wp: head_insufficient_loop_ordered_disjoint) - apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt) - apply (wpsimp wp: head_insufficient_loop_no_overflow) - apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt) - apply (wpsimp wp: head_insufficient_loop_refills_window) - apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt) - apply (wpsimp wp: head_insufficient_loop_MIN_BUDGET_in_head) - apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt) - apply (wpsimp wp: head_insufficient_loop_length_bounded) - apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt) - apply (wpsimp wp: head_insufficient_loop_non_zero_refills) - apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt) - apply (wpsimp wp: refill_budget_check_refills_sum_helper) - apply (wpsimp wp: refill_budget_check_ordered_disjoint_helper', fastforce+) - apply (wpsimp wp: refill_budget_check_no_overflow', fastforce+) - apply (wpsimp wp: refill_budget_check_window_helper, fastforce+) - apply (wpsimp wp: refill_budget_check_length_helper) - apply (wpsimp wp: refill_budget_check_non_zero_refills_helper) - apply (wpsimp wp: refill_budget_check_refills_unat_sum_helper, fastforce+) - apply (clarsimp simp: obj_at_def) - apply (wpsimp wp: set_refills_wp get_refills_wp - simp: vs_all_heap_simps update_sched_context_set_refills_rewrite update_refill_hd_rewrite)+ + apply (find_goal \match conclusion in "\_\ head_insufficient_loop _ \\_. _\" \ \-\\) + apply (rule_tac Q'="\_ s. \ round_robin sc_ptr s \ sp_valid_refills_unbundled sc_ptr s" + in hoare_strengthen_post[rotated]) + apply (fastforce simp: valid_refills_def vs_all_heap_simps) + apply (intro hoare_vcg_conj_lift_pre_fix; (solves head_insufficient_loop_simple)?) + apply (wpsimp wp: head_insufficient_loop_refills_sum) + apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt sc_at_ppred_def obj_at_def) + apply (wpsimp wp: head_insufficient_loop_ordered_disjoint) + apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt sc_at_ppred_def obj_at_def) + apply (wpsimp wp: head_insufficient_loop_no_overflow) + apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt sc_at_ppred_def obj_at_def) + apply (wpsimp wp: head_insufficient_loop_refills_window) + apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt sc_at_ppred_def obj_at_def) + apply (wpsimp wp: head_insufficient_loop_MIN_BUDGET_in_head) + apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt sc_at_ppred_def obj_at_def) + apply (wpsimp wp: head_insufficient_loop_length_bounded) + apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt sc_at_ppred_def obj_at_def) + apply (wpsimp wp: head_insufficient_loop_non_zero_refills) + apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt sc_at_ppred_def obj_at_def) + apply (wpsimp wp: refill_budget_check_refills_sum_helper) + apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt sc_at_ppred_def obj_at_def) + apply (wpsimp wp: refill_budget_check_ordered_disjoint_helper') + apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt sc_at_ppred_def obj_at_def) + apply (wpsimp wp: refill_budget_check_no_overflow') + apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt sc_at_ppred_def obj_at_def) + apply (wpsimp wp: refill_budget_check_window_helper) + apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt sc_at_ppred_def obj_at_def) + apply (wpsimp wp: refill_budget_check_length_helper) + apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt sc_at_ppred_def obj_at_def) + apply (wpsimp wp: refill_budget_check_non_zero_refills_helper) + apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt sc_at_ppred_def obj_at_def) + apply (wpsimp wp: refill_budget_check_refills_unat_sum_helper) + apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt sc_at_ppred_def obj_at_def) + apply wpsimp + apply wpsimp done lemma valid_refills_sc_update: @@ -11011,7 +11029,7 @@ lemma sc_refills_update_sc_tcb_sc_at[wp]: done crunch check_domain_time, refill_budget_check, refill_budget_check_round_robin, - refill_unblock_check + refill_unblock_check for sc_tcb_sc_at[wp]: "\s. Q (sc_tcb_sc_at P sc_ptr s)" and ct_in_cur_domain[wp]: "ct_in_cur_domain" and in_cur_domain[wp]: "in_cur_domain t" @@ -11048,20 +11066,11 @@ lemma head_insufficient_loop_valid_sched_action_not: \\_. valid_sched_action\" (is "valid ?pre _ _") apply (clarsimp simp: head_insufficient_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) - apply (intro hoare_vcg_conj_lift_pre_fix) - apply non_overlapping_merge_refills_simple - apply (fastforce simp: valid_sched_action_def weak_valid_sched_action_def vs_all_heap_simps - obj_at_def scheduler_act_not_def is_activatable_def - switch_in_cur_domain_def in_cur_domain_def etcb_at'_def) - apply non_overlapping_merge_refills_simple - apply (fastforce simp: valid_sched_action_def weak_valid_sched_action_def vs_all_heap_simps - obj_at_def scheduler_act_not_def is_activatable_def - switch_in_cur_domain_def in_cur_domain_def etcb_at'_def - split: if_splits) - apply non_overlapping_merge_refills_simple - done + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) + apply merge_nonoverlapping_head_refill_simple + by (fastforce simp: valid_sched_action_def weak_valid_sched_action_def vs_all_heap_simps + obj_at_def scheduler_act_not_def is_activatable_def + switch_in_cur_domain_def in_cur_domain_def etcb_at'_def) lemma refill_head_overlapping_loop_valid_sched_action_not: "\\s. valid_sched_action s \ sc_scheduler_act_not (cur_sc s) s \ cur_sc s = csc_ptr\ @@ -11069,36 +11078,26 @@ lemma refill_head_overlapping_loop_valid_sched_action_not: \\_. valid_sched_action\" (is "valid ?pre _ _") apply (clarsimp simp: refill_head_overlapping_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) - apply (intro hoare_vcg_conj_lift_pre_fix) - apply merge_overlapping_refills_simple - apply (fastforce simp: valid_sched_action_def weak_valid_sched_action_def vs_all_heap_simps - obj_at_def scheduler_act_not_def is_activatable_def - switch_in_cur_domain_def in_cur_domain_def etcb_at'_def) - apply merge_overlapping_refills_simple - apply (fastforce simp: valid_sched_action_def weak_valid_sched_action_def vs_all_heap_simps - obj_at_def scheduler_act_not_def is_activatable_def - switch_in_cur_domain_def in_cur_domain_def etcb_at'_def - split: if_splits) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) apply merge_overlapping_refills_simple - done + by (fastforce simp: valid_sched_action_def weak_valid_sched_action_def vs_all_heap_simps + obj_at_def scheduler_act_not_def is_activatable_def + switch_in_cur_domain_def in_cur_domain_def etcb_at'_def + split: if_splits) lemma handle_overrun_loop_valid_sched_action_not: - "\\s. valid_sched_action s \ sc_scheduler_act_not (cur_sc s) s \ cur_sc s = csc_ptr\ - handle_overrun_loop usage + "\\s. valid_sched_action s \ sc_scheduler_act_not sc_ptr s\ + handle_overrun sc_ptr usage \\_. valid_sched_action\" (is "valid ?pre _ _") - apply (clarsimp simp: handle_overrun_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) - apply handle_overrun_loop_body_simple - apply (fastforce simp: valid_sched_action_def weak_valid_sched_action_def vs_all_heap_simps - obj_at_def scheduler_act_not_def is_activatable_def - switch_in_cur_domain_def in_cur_domain_def etcb_at'_def) - done + apply (clarsimp simp: handle_overrun_def) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) + apply charge_entire_head_refill_simple + by (fastforce simp: valid_sched_action_def weak_valid_sched_action_def vs_all_heap_simps + obj_at_def scheduler_act_not_def is_activatable_def + switch_in_cur_domain_def in_cur_domain_def etcb_at'_def) -crunch head_insufficient_loop, handle_overrun_loop +crunch head_insufficient_loop, handle_overrun for sc_scheduler_act_not_cur_sc[wp]: "\s. sc_scheduler_act_not (cur_sc s) s" (wp: crunch_wps) @@ -11111,7 +11110,6 @@ lemma refill_budget_check_valid_sched_action_act_not: apply (rule bind_wp_fwd_skip, solves wpsimp)+ apply (rule bind_wp_fwd_skip) apply (wpsimp wp: handle_overrun_loop_valid_sched_action_not) - apply fastforce apply (rule bind_wp_fwd_skip, solves wpsimp) apply (rule_tac Q'="\_ s. valid_sched_action s \ sc_scheduler_act_not (cur_sc s) s \ cur_sc s = csc_ptr" @@ -11119,10 +11117,11 @@ lemma refill_budget_check_valid_sched_action_act_not: apply (wpsimp wp: set_refills_wp get_refills_wp simp: update_refill_hd_rewrite update_sched_context_set_refills_rewrite schedule_used_defs) - subgoal by (fastforce simp: valid_sched_action_def weak_valid_sched_action_def vs_all_heap_simps - obj_at_def scheduler_act_not_def is_activatable_def - switch_in_cur_domain_def in_cur_domain_def etcb_at'_def - split: if_splits) + subgoal + by (fastforce simp: valid_sched_action_def weak_valid_sched_action_def vs_all_heap_simps + obj_at_def sc_at_ppred_def scheduler_act_not_def is_activatable_def + switch_in_cur_domain_def in_cur_domain_def etcb_at'_def + split: if_splits) apply (wpsimp wp: head_insufficient_loop_valid_sched_action_not) done @@ -11140,14 +11139,9 @@ lemma head_insufficient_loop_valid_ready_qs_not: \\_. valid_ready_qs\" (is "valid ?pre _ _") apply (clarsimp simp: head_insufficient_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) - apply (intro hoare_vcg_conj_lift_pre_fix) - apply non_overlapping_merge_refills_simple - apply (fastforce simp: valid_ready_qs_def vs_all_heap_simps obj_at_def not_queued_def) - apply non_overlapping_merge_refills_simple - apply (fastforce simp: vs_all_heap_simps obj_at_def split: if_splits) - apply non_overlapping_merge_refills_simple + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) + apply merge_nonoverlapping_head_refill_simple + apply (fastforce simp: valid_ready_qs_def vs_all_heap_simps obj_at_def not_queued_def) done lemma refill_head_overlapping_loop_valid_ready_qs_not: @@ -11167,18 +11161,17 @@ lemma refill_head_overlapping_loop_valid_ready_qs_not: done lemma handle_overrun_loop_valid_ready_qs_not: - "\\s. valid_ready_qs s \ sc_not_in_ready_q (cur_sc s) s \ cur_sc s = csc_ptr\ - handle_overrun_loop usage + "\\s. valid_ready_qs s \ sc_not_in_ready_q sc_ptr s\ + handle_overrun sc_ptr usage \\_. valid_ready_qs\" (is "valid ?pre _ _") - apply (clarsimp simp: handle_overrun_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) - apply handle_overrun_loop_body_simple + apply (clarsimp simp: handle_overrun_def) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) + apply charge_entire_head_refill_simple apply (fastforce simp: valid_ready_qs_def vs_all_heap_simps obj_at_def not_queued_def) done -crunch head_insufficient_loop, handle_overrun_loop +crunch head_insufficient_loop, handle_overrun for sc_not_in_ready_q_cur_sc[wp]: "\s. sc_not_in_ready_q (cur_sc s) s" (wp: crunch_wps) @@ -11192,7 +11185,6 @@ lemma refill_budget_check_valid_ready_qs_not_queued: apply (rule bind_wp_fwd_skip, solves wpsimp)+ apply (rule bind_wp_fwd_skip) apply (wpsimp wp: handle_overrun_loop_valid_ready_qs_not) - apply fastforce apply (rule bind_wp_fwd_skip, solves wpsimp) apply (rule_tac Q'="\_ s. valid_ready_qs s \ sc_not_in_ready_q (cur_sc s) s \ cur_sc s = csc_ptr" @@ -11211,12 +11203,13 @@ lemma is_refill_ready_alt: by (clarsimp simp: refill_ready_def vs_all_heap_simps) method handle_overrun_loop_fail_simple - = (clarsimp simp: handle_overrun_loop_def - , rule valid_whileLoop_cond_fail - , (fastforce simp: head_time_buffer_def obind_def ogets_def read_sched_context_def - vs_all_heap_simps refill_sufficient_def refill_capacity_def - MIN_BUDGET_nonzero sc_valid_refills_def vs_all_heap_simps - split: option.splits kernel_object.splits if_splits)+) + = clarsimp simp: handle_overrun_def, + rule valid_whileLoop_cond_fail, + (fastforce simp: head_refill_overrun_def obind_def ogets_def read_sched_context_def + read_refill_head_def + vs_all_heap_simps refill_sufficient_def refill_capacity_def + MIN_BUDGET_nonzero sc_valid_refills_def sc_at_ppred_def + split: option.splits kernel_object.splits if_splits)+ lemma refill_budget_check_refill_ready_offset_ready_and_sufficient: "\\s. if sc_ptr = cur_sc s @@ -11238,15 +11231,13 @@ lemma refill_budget_check_refill_ready_offset_ready_and_sufficient: \ robin = round_robin csc_ptr s \ cur_sc s = csc_ptr" in bind_wp_fwd) - apply ((wpsimp wp: set_refills_wp get_refills_wp whileLoop_valid_inv - simp: refill_budget_check_defs update_sched_context_set_refills_rewrite - schedule_used_defs - | fastforce simp: obj_at_def vs_all_heap_simps )+)[1] + apply handle_overrun_simple apply ((wpsimp wp: set_refills_wp get_refills_wp whileLoop_valid_inv simp: refill_budget_check_defs update_sched_context_set_refills_rewrite schedule_used_defs | fastforce simp: obj_at_def vs_all_heap_simps)+)[1] + apply clarsimp apply (rule_tac P'1="\s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) (cur_sc s)" @@ -11260,66 +11251,65 @@ lemma refill_budget_check_refill_ready_offset_ready_and_sufficient: \ cur_sc_offset_sufficient usage s \ current_time_bounded s \ pred_map (\cfg. MIN_BUDGET \ r_amount (scrc_refill_hd cfg)) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. scrc_period cfg \ MAX_PERIOD) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. MIN_REFILLS \ scrc_refill_max cfg) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. no_overflow (scrc_refills cfg)) - (sc_refill_cfgs_of s) (cur_sc s) - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr + \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) - (sc_refill_cfgs_of s) (cur_sc s) - \ csc_ptr = cur_sc s \ cur_sc s \ idle_sc_ptr + (sc_refill_cfgs_of s) sc_ptr + \ sc_ptr = cur_sc s \ sc_ptr \ idle_sc_ptr \ (pred_map (\cfg. r_time (hd (scrc_refills cfg)) < MAX_RELEASE_TIME) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. rv < r_amount (hd (scrc_refills cfg))) - (sc_refill_cfgs_of s) (cur_sc s)) + (sc_refill_cfgs_of s) sc_ptr) \ rv = usage" in bind_wp_fwd) apply (intro hoare_vcg_conj_lift_pre_fix; (solves handle_overrun_loop_fail_simple)?) - apply (clarsimp simp: handle_overrun_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_ s. pred_map \ (scs_of s) (cur_sc s)"]) - apply (fastforce simp: head_time_buffer_true_imp_unat_buffer vs_all_heap_simps word_less_nat_alt + apply (clarsimp simp: handle_overrun_def) + apply (wpsimp wp: valid_whileLoop[where I="\_ s. sc_refills_sc_at (\refills. refills \ []) sc_ptr s"]) + apply (fastforce simp: head_refill_overrun_true_imp_unat_buffer vs_all_heap_simps word_less_nat_alt word_le_nat_alt) - apply (clarsimp simp: vs_all_heap_simps) + apply (clarsimp simp: vs_all_heap_simps valid_refills_def sc_at_ppred_def obj_at_def) - apply (rule bind_wp[OF _ get_refills_sp]) + apply (rule bind_wp[OF _ get_refill_head_sp]) apply (rule_tac Q'="\_ s. pred_map (\cfg. r_time (scrc_refill_hd cfg) \ cur_time s) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. MIN_BUDGET \ r_amount (scrc_refill_hd cfg)) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. scrc_period cfg \ MAX_PERIOD) - (sc_refill_cfgs_of s) (cur_sc s) - \ csc_ptr = cur_sc s" + (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s + \ sc_ptr = cur_sc s" in bind_wp) apply (wpsimp wp: head_insufficient_loop_hd_r_time_same) apply (intro hoare_vcg_conj_lift_pre_fix; (solves wpsimp)?) apply (wpsimp wp: update_sched_context_wp set_refills_wp get_refills_wp - simp: schedule_used_defs) - apply (clarsimp simp: vs_all_heap_simps word_le_nat_alt refill_ready_no_overflow_def - current_time_bounded_def obj_at_def) + simp: schedule_used_defs vs_all_heap_simps word_le_nat_alt refill_ready_no_overflow_def + current_time_bounded_def obj_at_def sc_at_ppred_def) apply (subst unat_add_lem''; fastforce?) apply (intro conjI impI) apply (subst unat_add_lem''; fastforce?) apply (subst unat_add_lem''; fastforce?) - apply (rename_tac sc n) + apply (rename_tac sc) apply (case_tac "sc_refills sc"; clarsimp) apply (simp add: MIN_REFILLS_def) apply (subst unat_add_lem''; fastforce?) - apply (clarsimp simp: refills_unat_sum_def unat_sub word_less_nat_alt word_le_nat_alt) + apply (wpsimp wp: update_sched_context_wp set_refills_wp get_refills_wp + simp: schedule_used_defs vs_all_heap_simps current_time_bounded_def) + subgoal + by (fastforce simp: refill_sufficient_def refill_capacity_def + word_less_nat_alt word_le_nat_alt obj_at_def sc_at_ppred_def) apply (wpsimp wp: update_sched_context_wp set_refills_wp get_refills_wp - simp: schedule_used_defs) - apply (clarsimp simp: vs_all_heap_simps current_time_bounded_def obj_at_def) - apply (intro conjI impI) - apply (rule word_le_nat_alt[THEN iffD1]) - apply (clarsimp simp: refill_sufficient_def refill_capacity_def split: if_splits) - apply (clarsimp simp: word_less_nat_alt word_le_nat_alt refill_sufficient_def refill_capacity_def) - apply (wpsimp wp: update_sched_context_wp set_refills_wp get_refills_wp - simp: schedule_used_defs) - apply (clarsimp simp: vs_all_heap_simps obj_at_def) + simp: schedule_used_defs sc_at_ppred_def obj_at_def) + apply (clarsimp simp: vs_all_heap_simps obj_at_def) + apply (wpsimp wp: schedule_used_non_nil) + apply (clarsimp simp: vs_all_heap_simps sc_at_ppred_def obj_at_def) done lemma is_refill_sufficient_0_alt: @@ -11346,10 +11336,7 @@ lemma refill_budget_check_is_refill_sufficient: apply (rule_tac Q'="\_ s. is_refill_sufficient 0 sc_ptr s \ robin = round_robin csc_ptr s \ cur_sc s = csc_ptr" in bind_wp_fwd) - apply ((wpsimp wp: set_refills_wp get_refills_wp whileLoop_valid_inv - simp: refill_budget_check_defs update_sched_context_set_refills_rewrite - schedule_used_defs - | fastforce simp: obj_at_def vs_all_heap_simps )+)[1] + apply handle_overrun_simple apply ((wpsimp wp: set_refills_wp get_refills_wp whileLoop_valid_inv simp: refill_budget_check_defs update_sched_context_set_refills_rewrite schedule_used_defs @@ -11364,43 +11351,49 @@ lemma refill_budget_check_is_refill_sufficient: ; fastforce?) apply (clarsimp simp: sc_valid_refills_def vs_all_heap_simps) + apply clarsimp apply (rule_tac Q'="\usage' s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) - (sc_refill_cfgs_of s) (cur_sc s) - \ valid_refills (cur_sc s) s - \ csc_ptr = cur_sc s \ cur_sc s \ idle_sc_ptr + (sc_refill_cfgs_of s) sc_ptr + \ valid_refills sc_ptr s + \ sc_ptr = cur_sc s \ cur_sc s \ idle_sc_ptr \ \ round_robin (cur_sc s) s \ (pred_map (\cfg. r_time (hd (scrc_refills cfg)) < MAX_RELEASE_TIME) - (sc_refill_cfgs_of s) (cur_sc s) - \ pred_map (\cfg. usage' < r_amount (hd (scrc_refills cfg))) - (sc_refill_cfgs_of s) (cur_sc s))" + (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. usage' < r_amount (hd refills)) + sc_ptr s)" in bind_wp_fwd) apply (intro hoare_vcg_conj_lift_pre_fix; (solves handle_overrun_loop_fail_simple)?) - apply (clarsimp simp: handle_overrun_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_ s. pred_map \ (scs_of s) (cur_sc s)"]) - apply (fastforce simp: head_time_buffer_true_imp_unat_buffer vs_all_heap_simps word_less_nat_alt - word_le_nat_alt) - apply (clarsimp simp: vs_all_heap_simps) - - apply (rule bind_wp[OF _ get_refills_sp]) + apply (clarsimp simp: handle_overrun_def) + apply (wpsimp wp: valid_whileLoop + [where I="\_ s. sc_refills_sc_at (\refills. refills \ []) sc_ptr s"]) + apply (frule_tac usage=r in head_refill_overrun_true_imp_unat_buffer) + apply (clarsimp simp: vs_all_heap_simps word_less_nat_alt word_le_nat_alt + sc_at_ppred_def obj_at_def) + apply (clarsimp simp: valid_refills_def sc_at_ppred_def obj_at_def vs_all_heap_simps) + + apply (rule bind_wp[OF _ get_refill_head_sp]) apply (rule_tac Q'="\_ s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. MIN_BUDGET \ scrc_budget cfg) - (sc_refill_cfgs_of s) (cur_sc s) - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) (cur_sc s) - \ csc_ptr = cur_sc s" + (sc_refill_cfgs_of s) sc_ptr + \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr + \ sc_ptr = cur_sc s" in bind_wp) apply (subst is_refill_sufficient_0_alt) apply (wpsimp wp: head_insufficient_loop_MIN_BUDGET_in_head) - apply (fastforce simp: vs_all_heap_simps word_le_nat_alt) - apply (intro hoare_vcg_conj_lift_pre_fix; (solves head_insufficient_loop_simple)?) - apply (rule_tac f=cur_sc in hoare_lift_Pf2) - apply (wpsimp wp: schedule_used_refills_unat_sum set_refills_wp get_refills_wp - simp: update_refill_hd_rewrite) - apply (fastforce dest!: refill_budget_check_refills_unat_sum_word_helper - simp: refills_unat_sum_def vs_all_heap_simps obj_at_def sc_valid_refills_def) - apply wpsimp + apply (fastforce simp: vs_all_heap_simps word_le_nat_alt sc_at_ppred_def obj_at_def) + apply (intro hoare_vcg_conj_lift_pre_fix) + apply (wpsimp wp: schedule_used_refills_unat_sum set_refills_wp get_refills_wp + simp: update_refill_hd_rewrite) + apply (fastforce dest!: refill_budget_check_refills_unat_sum_word_helper + simp: refills_unat_sum_def vs_all_heap_simps obj_at_def sc_valid_refills_def + sc_at_ppred_def) + apply (wpsimp simp: refill_budget_check_defs update_sched_context_set_refills_rewrite + wp: set_refills_wp get_refills_wp) + apply (clarsimp simp: valid_refills_def vs_all_heap_simps obj_at_def round_robin_def) + apply schedule_used_simple apply wpsimp done @@ -11460,47 +11453,43 @@ lemma update_refill_hd_valid_release_q: sc_heap_of_state_def map_project_def map_join_def sc_refill_cfgs_of_scs_def elim!: opt_mapE) -lemma non_overlapping_merge_refills_cur_sc_not_in_release_q: - "non_overlapping_merge_refills csc_ptr \\s. sc_not_in_release_q (cur_sc s) s\" - apply non_overlapping_merge_refills_simple - apply (fastforce split: if_splits) - done +lemma merge_nonoverlapping_head_refill_cur_sc_not_in_release_q: + "merge_nonoverlapping_head_refill csc_ptr \\s. sc_not_in_release_q (cur_sc s) s\" + by (wpsimp simp: refill_budget_check_defs) + +lemma set_refills_valid_release_q_sc_not_in_release_q: + "\\s. valid_release_q s \ sc_not_in_release_q sc_ptr s\ + set_refills sc_ptr refills + \\_. valid_release_q\" + by (wpsimp wp: set_refills_valid_release_q) lemma head_insufficient_loop_valid_release_q: - "\\s. valid_release_q s \ sc_not_in_release_q (cur_sc s) s \ cur_sc s = csc_ptr\ - head_insufficient_loop csc_ptr + "\\s. valid_release_q s \ sc_not_in_release_q sc_ptr s\ + head_insufficient_loop sc_ptr \\_. valid_release_q\" (is "valid ?pre _ _") apply (clarsimp simp: head_insufficient_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) - apply (intro hoare_vcg_conj_lift_pre_fix) - apply (wpsimp wp: set_refills_valid_release_q get_refills_wp hoare_vcg_all_lift - hoare_vcg_imp_lift' set_refills_wp - simp: non_overlapping_merge_refills_def refill_pop_head_def - update_refill_hd_rewrite update_sched_context_set_refills_rewrite) - apply (fastforce simp: vs_all_heap_simps obj_at_def split: if_splits) - apply non_overlapping_merge_refills_simple - apply (fastforce split: if_splits) - apply non_overlapping_merge_refills_simple + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) + apply (wpsimp simp: refill_budget_check_defs update_sched_context_set_refills_rewrite + wp: set_refills_valid_release_q_sc_not_in_release_q) done lemma handle_overrun_loop_valid_release_q: - "\\s. valid_release_q s \ sc_not_in_release_q (cur_sc s) s\ - handle_overrun_loop usage + "\\s. valid_release_q s \ sc_not_in_release_q sc_ptr s\ + handle_overrun sc_ptr usage \\_. valid_release_q\" (is "valid ?pre _ _") - apply (clarsimp simp: handle_overrun_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) + apply (clarsimp simp: handle_overrun_def) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) apply (wpsimp wp: set_refills_valid_release_q get_refills_wp hoare_vcg_all_lift hoare_vcg_imp_lift' simp: refill_budget_check_defs update_sched_context_set_refills_rewrite schedule_used_defs | wpsimp wp: set_refills_wp)+ done -crunch head_insufficient_loop, handle_overrun_loop - for sc_not_in_release_q_cur_sc[wp]: "\s. sc_not_in_release_q (cur_sc s) s" +crunch head_insufficient_loop, handle_overrun + for sc_not_in_release_q[wp]: "sc_not_in_release_q sc_ptr" + and cur_sc[wp]: "\s. P (cur_sc s)" (wp: crunch_wps) lemma sorted_release_q_sc_not_in_sc_update: @@ -11525,7 +11514,7 @@ lemma valid_release_q_sc_not_in_sc_update: fastforce simp: obj_at_def is_tcb vs_all_heap_simps) lemma schedule_used_valid_release_q: - "\valid_release_q and (\s. sc_not_in_release_q (cur_sc s) s) and (\s. sc_ptr = cur_sc s)\ + "\valid_release_q and sc_not_in_release_q sc_ptr\ schedule_used sc_ptr usage \\_. valid_release_q\" apply (clarsimp simp: schedule_used_defs) @@ -11543,12 +11532,13 @@ lemma refill_budget_check_valid_release_q: update_refill_hd_rewrite) apply (rule bind_wp[OF _ gets_sp], rename_tac csc_ptr) apply (rule bind_wp_fwd_skip, solves wpsimp)+ - apply (rule bind_wp_fwd_skip, wpsimp wp: handle_overrun_loop_valid_release_q) - apply (rule bind_wp[OF _ get_refills_sp]) + apply (rule bind_wp_fwd_skip) + apply ((wpsimp wp: handle_overrun_loop_valid_release_q | wps)+)[1] + apply (rule bind_wp[OF _ get_refill_head_sp]) apply (rule_tac Q'="\_ s. valid_release_q s \ sc_not_in_release_q (cur_sc s) s \ cur_sc s = csc_ptr" in bind_wp_fwd) - apply (wpsimp wp: schedule_used_valid_release_q get_refills_wp set_refills_wp) + apply ((wpsimp wp: schedule_used_valid_release_q get_refills_wp set_refills_wp | wps)+)[1] apply (clarsimp simp: vs_all_heap_simps obj_at_def) apply (drule_tac scp="cur_sc s" in valid_release_q_sc_not_in_sc_update; simp?) apply (clarsimp simp: vs_all_heap_simps obj_at_def is_sc_obj valid_release_q_sc_not_in_sc_update) @@ -11810,34 +11800,34 @@ lemma set_refills_bounded_release_time: apply (wpsimp wp: valid_sched_wp) by (clarsimp simp: vs_all_heap_simps bounded_release_time_def pred_map_simps obj_at_def heap_upd_def split: if_split) - +thm refill_budget_check_def lemma refill_budget_check_schedule_used_r_time_helper: - "r_time (hd refills) < MAX_RELEASE_TIME \ - \\s. all_sp_valid_refills_but_MIN_BUDGET_in_head (cur_sc s) s + "r_time refill_head < MAX_RELEASE_TIME \ + \\s. all_sp_valid_refills_but_MIN_BUDGET_in_head sc_ptr s + \ sc_refills_sc_at (\refills. refills \ [] \ hd refills = refill_head) sc_ptr s \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. unat (r_time (scrc_refill_hd cfg)) + (Suc n) * unat MAX_PERIOD \ unat max_time) - (sc_refill_cfgs_of s) (cur_sc s) - \ cur_sc s = csc_ptr + (sc_refill_cfgs_of s) sc_ptr \ (pred_map (\cfg. r_time (scrc_refill_hd cfg) < MAX_RELEASE_TIME) - (sc_refill_cfgs_of s) (cur_sc s) - \ pred_map (\cfg. u < r_amount (scrc_refill_hd cfg)) (sc_refill_cfgs_of s) (cur_sc s)) - \ (\sc n. ko_at (SchedContext sc n) csc_ptr s \ refills = sc_refills sc) - \ (\n. ko_at (SchedContext sc n) csc_ptr s)\ - do y \ update_refill_hd csc_ptr (r_time_update (\t. t + u) \ r_amount_update (\m. m - u)); - schedule_used csc_ptr \r_time = r_time (refill_hd sc) + sc_period sc, r_amount = u\ - od - \\_ s. pred_map (\cfg. unat (r_time (scrc_refill_hd cfg)) + n * unat MAX_PERIOD \ unat max_time) - (sc_refill_cfgs_of s) (cur_sc s)\" + (sc_refill_cfgs_of s) sc_ptr + \ sc_refills_sc_at (\refills. u < r_amount (hd refills)) sc_ptr s) + \ (\n. ko_at (SchedContext sc n) sc_ptr s)\ + do y \ update_refill_hd sc_ptr (r_amount_update (\_. r_amount refill_head - u)); + y \ update_refill_hd sc_ptr (r_time_update (\_. r_time refill_head + u)); + schedule_used sc_ptr \r_time = r_time refill_head + sc_period sc, r_amount = u\ + od + \\_ s. pred_map (\cfg. unat (r_time (scrc_refill_hd cfg)) + n * unat MAX_PERIOD \ unat max_time) + (sc_refill_cfgs_of s) sc_ptr\" apply (wpsimp wp: update_sched_context_wp get_refills_wp simp: schedule_used_def refill_add_tail_def update_refill_tl_def update_refill_hd_def) - apply (clarsimp simp: vs_all_heap_simps obj_at_def) + apply (clarsimp simp: vs_all_heap_simps obj_at_def sc_at_ppred_def) apply (case_tac "sc_refills sc") apply (clarsimp simp: schedule_used_def vs_all_heap_simps obj_at_def) apply (prop_tac "unat (r_amount a) \ unat (sc_budget sc)") apply (fastforce simp: refills_unat_sum_member_bound refills_unat_sum_cons) - apply (subst unat_sub | subst unat_add_lem'' | linarith + apply (intro conjI impI allI | subst unat_sub | subst unat_add_lem'' | linarith | clarsimp simp: word_le_nat_alt word_less_nat_alt no_overflow_def unat_sub window_def)+ done @@ -11857,10 +11847,7 @@ lemma refill_budget_check_bounded_release_time: apply (rule_tac Q'="\_ s. bounded_release_time sc_ptr s \ robin = round_robin csc_ptr s \ cur_sc s = csc_ptr" in bind_wp_fwd) - apply ((wpsimp wp: set_refills_wp get_refills_wp whileLoop_valid_inv - simp: refill_budget_check_defs update_sched_context_set_refills_rewrite - schedule_used_defs - | fastforce simp: obj_at_def vs_all_heap_simps )+)[1] + apply handle_overrun_simple apply ((wpsimp wp: set_refills_wp get_refills_wp whileLoop_valid_inv simp: refill_budget_check_defs update_sched_context_set_refills_rewrite schedule_used_defs @@ -11871,10 +11858,10 @@ lemma refill_budget_check_bounded_release_time: (sc_refill_cfgs_of s) sc_ptr" in hoare_pre_add[THEN iffD2, simplified pred_conj_def]) apply (clarsimp simp: vs_all_heap_simps) - apply (frule valid_refills_refills_unat_sum_equals_budget - ; (fastforce simp: )?) + apply (frule valid_refills_refills_unat_sum_equals_budget; (fastforce simp: )?) apply (clarsimp simp: sc_valid_refills_def vs_all_heap_simps) + apply clarsimp apply (rule_tac Q'="\usage' s. \ round_robin sc_ptr s \ all_sp_valid_refills_but_MIN_BUDGET_in_head_no_overflow sc_ptr s \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) @@ -11883,66 +11870,73 @@ lemma refill_budget_check_bounded_release_time: \ pred_map (\cfg. unat (r_time (hd (scrc_refills cfg))) + 4 * unat MAX_PERIOD \ unat max_time) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ cur_sc s = sc_ptr \ cur_sc s \ idle_sc_ptr \ (pred_map (\cfg. r_time (hd (scrc_refills cfg)) < MAX_RELEASE_TIME) (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. usage' < r_amount (hd (scrc_refills cfg))) (sc_refill_cfgs_of s) sc_ptr)" in bind_wp_fwd) - apply (intro hoare_vcg_conj_lift_pre_fix - ; (solves handle_overrun_loop_simple)?) - apply (wpsimp wp: handle_overrun_loop_refills_sum) - apply (fastforce simp: sc_valid_refills_def vs_all_heap_simps) - apply (wpsimp wp: handle_overrun_loop_ordered_disjoint) - apply (fastforce simp: sc_valid_refills_def vs_all_heap_simps) - apply (wpsimp wp: handle_overrun_loop_window) - apply (fastforce simp: sc_valid_refills_def vs_all_heap_simps) - apply (wpsimp wp: handle_overrun_loop_length) - apply (clarsimp simp: sc_valid_refills_def vs_all_heap_simps) + apply (intro hoare_vcg_conj_lift_pre_fix; (solves handle_overrun_simple)?) + apply (wpsimp wp: handle_overrun_refills_sum) + apply (fastforce simp: sc_valid_refills_def vs_all_heap_simps sc_at_ppred_def obj_at_def) + apply (wpsimp wp: handle_overrun_ordered_disjoint) + apply (fastforce simp: sc_valid_refills_def vs_all_heap_simps sc_at_ppred_def obj_at_def) + apply (wpsimp wp: handle_overrun_window) + apply (fastforce simp: sc_valid_refills_def vs_all_heap_simps sc_at_ppred_def obj_at_def) + apply (wpsimp wp: handle_overrun_length) + apply (clarsimp simp: sc_valid_refills_def vs_all_heap_simps sc_at_ppred_def obj_at_def) apply (wpsimp wp: handle_overrun_loop_non_zero_refills) - apply (clarsimp simp: sc_valid_refills_def vs_all_heap_simps) - apply (wpsimp wp: handle_overrun_loop_refills_unat_sum_equals_budget) - apply (clarsimp simp: sc_valid_refills_def vs_all_heap_simps) - apply (wpsimp wp: handle_overrun_loop_head_bound) - apply (fastforce simp: sc_valid_refills_def vs_all_heap_simps) - apply (clarsimp simp: handle_overrun_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_ s. pred_map \ (scs_of s) (cur_sc s) - \ sc_ptr = cur_sc s"]) - apply (fastforce simp: head_time_buffer_true_imp_unat_buffer vs_all_heap_simps word_less_nat_alt - word_le_nat_alt is_sc_obj_def) - apply (clarsimp simp: vs_all_heap_simps) + apply (clarsimp simp: sc_valid_refills_def vs_all_heap_simps obj_at_def sc_at_ppred_def) + apply (wpsimp wp: handle_overrun_refills_unat_sum_equals_budget) + apply (clarsimp simp: sc_valid_refills_def vs_all_heap_simps obj_at_def sc_at_ppred_def) + apply (wpsimp wp: handle_overrun_head_bound) + apply (fastforce simp: sc_valid_refills_def vs_all_heap_simps sc_at_ppred_def obj_at_def) + apply (clarsimp simp: handle_overrun_def) + apply (wpsimp wp: valid_whileLoop + [where I="\_ s. sc_refills_sc_at (\refills. refills \ []) sc_ptr s"]) + apply (fastforce simp: head_refill_overrun_true_imp_unat_buffer vs_all_heap_simps + word_less_nat_alt word_le_nat_alt) + apply (clarsimp simp: vs_all_heap_simps sc_at_ppred_def obj_at_def valid_refills_def) apply (rule_tac P'1="\s. pred_map (\cfg. no_overflow (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr" in hoare_pre_add[THEN iffD2, simplified pred_conj_def]) - apply (fastforce intro!: head_time_buffer_implies_no_overflow simp: vs_all_heap_simps) + apply (fastforce intro!: head_time_buffer_implies_no_overflow + simp: vs_all_heap_simps sc_at_ppred_def obj_at_def) - apply (rule bind_wp[OF _ get_refills_sp]) + apply (rule bind_wp[OF _ get_refill_head_sp]) apply (rule_tac Q'="\_ s. \ round_robin (cur_sc s) s \ all_sp_valid_refills_but_MIN_BUDGET_in_head sc_ptr s \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) (sc_refill_cfgs_of s) sc_ptr - \ sc_ptr = cur_sc s \ cur_sc s \ idle_sc_ptr + \ cur_sc s = sc_ptr \ cur_sc s \ idle_sc_ptr \ pred_map (\cfg. unat (r_time (scrc_refill_hd cfg)) + 3 * unat MAX_PERIOD \ unat max_time) - (sc_refill_cfgs_of s) (cur_sc s)" + (sc_refill_cfgs_of s) sc_ptr" in bind_wp_fwd) apply (rule hoare_when_cases) apply ( clarsimp simp: vs_all_heap_simps current_time_bounded_def) apply (rule bind_wp[OF _ get_sched_context_sp]) apply (intro hoare_vcg_conj_lift_pre_fix; (solves schedule_used_simple)?) apply (wpsimp wp: refill_budget_check_refills_sum_helper) - apply (wpsimp wp: refill_budget_check_ordered_disjoint_helper', fastforce+) - apply (wpsimp wp: refill_budget_check_no_overflow', fastforce+) - apply (wpsimp wp: refill_budget_check_window_helper, fastforce+) + apply (clarsimp simp: vs_all_heap_simps sc_at_ppred_def obj_at_def valid_refills_def) + apply (wpsimp wp: refill_budget_check_ordered_disjoint_helper') + apply (clarsimp simp: vs_all_heap_simps sc_at_ppred_def obj_at_def valid_refills_def) + apply (wpsimp wp: refill_budget_check_no_overflow') + apply (clarsimp simp: vs_all_heap_simps sc_at_ppred_def obj_at_def valid_refills_def) + apply (wpsimp wp: refill_budget_check_window_helper) + apply (clarsimp simp: vs_all_heap_simps sc_at_ppred_def obj_at_def valid_refills_def) apply (wpsimp wp: refill_budget_check_length_helper) + apply (clarsimp simp: vs_all_heap_simps sc_at_ppred_def obj_at_def valid_refills_def) apply (wpsimp wp: refill_budget_check_non_zero_refills_helper) - apply (wpsimp wp: refill_budget_check_refills_unat_sum_helper, fastforce+) + apply (clarsimp simp: vs_all_heap_simps sc_at_ppred_def obj_at_def valid_refills_def) + apply (wpsimp wp: refill_budget_check_refills_unat_sum_helper) + apply (clarsimp simp: vs_all_heap_simps sc_at_ppred_def obj_at_def valid_refills_def) apply wpsimp - apply (wpsimp wp: set_refills_wp get_refills_wp - simp: vs_all_heap_simps update_sched_context_set_refills_rewrite update_refill_hd_rewrite) - apply (wpsimp wp: refill_budget_check_schedule_used_r_time_helper, fastforce+) + apply wpsimp + apply (wpsimp wp: refill_budget_check_schedule_used_r_time_helper) + apply (clarsimp simp: vs_all_heap_simps sc_at_ppred_def obj_at_def) apply (rule_tac Q'="\_ s. pred_map (\cfg. unat (r_time (scrc_refill_hd cfg)) + 2 * unat MAX_PERIOD \ unat max_time) (sc_refill_cfgs_of s) sc_ptr" in hoare_strengthen_post[rotated]) @@ -11981,10 +11975,10 @@ lemma cur_sc_bounded_release_helper: by (simp add: word_le_nat_alt) lemma handle_overrun_loop_scrc_cfg[wp]: - "handle_overrun_loop sc_ptr + "handle_overrun sc_ptr usage \\s. Q (pred_map (\cfg. P (scrc_refill_max cfg) (scrc_budget cfg) (scrc_sporadic cfg)) (sc_refill_cfgs_of s) p)\" - by handle_overrun_loop_simple + by handle_overrun_simple lemma head_insufficient_loop_scrc_cfg[wp]: "head_insufficient_loop sc_ptr @@ -12062,16 +12056,16 @@ lemma refill_budget_check_active_reply_scs[wp]: apply (clarsimp simp: refill_budget_check_def) apply (rule bind_wp_fwd_skip, solves wpsimp)+ apply (rule bind_wp_fwd_skip) - apply handle_overrun_loop_simple + apply handle_overrun_simple apply (fastforce simp: active_reply_scs_def active_if_reply_sc_at_def vs_all_heap_simps active_sc_def obj_at_def)+ apply (rule bind_wp_fwd_skip, solves \wpsimp simp: update_refill_hd_rewrite update_sched_context_set_refills_rewrite schedule_used_defs\)+ - apply head_insufficient_loop_simple + apply handle_overrun_simple apply (fastforce simp: active_reply_scs_def active_if_reply_sc_at_def vs_all_heap_simps active_sc_def obj_at_def) - apply head_insufficient_loop_simple + apply handle_overrun_simple apply (fastforce simp: active_reply_scs_def active_if_reply_sc_at_def vs_all_heap_simps active_sc_def obj_at_def) done @@ -16992,24 +16986,23 @@ lemma end_timeslice_released_ipc_queues: unfolding end_timeslice_def by (wpsimp wp: handle_timeout_released_ipc_queues thread_get_wp' simp: is_timeout_fault_def) -lemma head_insufficient_loop_weak_valid_sched_action_not: +lemma head_insufficient_weak_valid_sched_action_not: "\\s. weak_valid_sched_action s \ sc_scheduler_act_not (cur_sc s) s \ cur_sc s = csc_ptr\ head_insufficient_loop csc_ptr \\_. weak_valid_sched_action\" (is "valid ?pre _ _") apply (clarsimp simp: head_insufficient_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) apply (intro hoare_vcg_conj_lift_pre_fix) apply (wpsimp wp: set_refills_wp get_refills_wp - simp: non_overlapping_merge_refills_def refill_pop_head_def update_refill_hd_rewrite - update_sched_context_set_refills_rewrite) + simp: merge_nonoverlapping_head_refill_def refill_pop_head_def + update_refill_hd_rewrite update_sched_context_set_refills_rewrite) apply (fastforce simp: weak_valid_sched_action_def vs_all_heap_simps obj_at_def scheduler_act_not_def) apply (wpsimp wp: set_refills_wp get_refills_wp - simp: non_overlapping_merge_refills_def refill_pop_head_def) + simp: merge_nonoverlapping_head_refill_def refill_pop_head_def) apply (wpsimp wp: set_refills_wp get_refills_wp - simp: non_overlapping_merge_refills_def refill_pop_head_def) + simp: merge_nonoverlapping_head_refill_def refill_pop_head_def) done lemma refill_head_overlapping_loop_weak_valid_sched_action_not: @@ -17030,21 +17023,18 @@ lemma refill_head_overlapping_loop_weak_valid_sched_action_not: simp: merge_overlapping_refills_def refill_pop_head_def)+ done -lemma handle_overrun_loop_weak_valid_sched_action_not: - "\\s. weak_valid_sched_action s \ sc_scheduler_act_not (cur_sc s) s \ cur_sc s = csc_ptr\ - handle_overrun_loop sc_ptr +lemma handle_overrun_weak_valid_sched_action_not: + "\\s. weak_valid_sched_action s \ sc_scheduler_act_not sc_ptr s\ + handle_overrun sc_ptr usage \\_. weak_valid_sched_action\" (is "valid ?pre _ _") - apply (clarsimp simp: handle_overrun_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) + apply (clarsimp simp: handle_overrun_def) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) apply (intro hoare_vcg_conj_lift_pre_fix) - apply handle_overrun_loop_body_simple - apply (fastforce simp: weak_valid_sched_action_def vs_all_heap_simps - obj_at_def scheduler_act_not_def) - apply handle_overrun_loop_body_simple - apply (wpsimp wp: set_refills_wp get_refills_wp - simp: merge_overlapping_refills_def refill_pop_head_def) + apply charge_entire_head_refill_simple + apply (fastforce simp: weak_valid_sched_action_def vs_all_heap_simps + obj_at_def scheduler_act_not_def) + apply charge_entire_head_refill_simple done lemma refill_budget_check_weak_valid_sched_action_act_not: @@ -17055,7 +17045,7 @@ lemma refill_budget_check_weak_valid_sched_action_act_not: apply (clarsimp simp: refill_budget_check_def) apply (rule bind_wp[OF _ gets_sp], rename_tac csc_ptr) apply (rule bind_wp_fwd_skip, solves wpsimp)+ - apply (rule bind_wp_fwd_skip, wp handle_overrun_loop_weak_valid_sched_action_not) + apply (rule bind_wp_fwd_skip, wp handle_overrun_weak_valid_sched_action_not) apply force apply (rule bind_wp_fwd_skip, solves wpsimp) apply (rule_tac Q'="\_ s. weak_valid_sched_action s \ sc_scheduler_act_not (cur_sc s) s @@ -17067,7 +17057,7 @@ lemma refill_budget_check_weak_valid_sched_action_act_not: apply (fastforce simp: valid_sched_action_def weak_valid_sched_action_def vs_all_heap_simps obj_at_def scheduler_act_not_def split: if_splits) - apply (wpsimp wp: head_insufficient_loop_weak_valid_sched_action_not) + apply (wpsimp wp: head_insufficient_weak_valid_sched_action_not) done lemma refill_budget_check_round_robin_weak_valid_sched_action: @@ -17087,11 +17077,9 @@ lemma refill_budget_check_cur_sc_tcb[wp]: apply (clarsimp simp: refill_budget_check_def) apply (rule bind_wp_fwd_skip, solves wpsimp)+ apply (rule bind_wp_fwd_skip) - apply (handle_overrun_loop_simple - ; clarsimp simp: obj_at_def cur_sc_tcb_def sc_at_pred_n_def) + apply (handle_overrun_simple; clarsimp simp: obj_at_def cur_sc_tcb_def sc_at_pred_n_def) apply (rule bind_wp_fwd_skip, solves wpsimp)+ - apply (head_insufficient_loop_simple - ; clarsimp simp: obj_at_def cur_sc_tcb_def sc_at_pred_n_def) + apply (head_insufficient_loop_simple; clarsimp simp: obj_at_def cur_sc_tcb_def sc_at_pred_n_def) done lemma refill_budget_check_active_sc_tcb_at': @@ -17578,19 +17566,21 @@ lemma refill_pop_head_cur_sc_not_blocked: apply (clarsimp simp: cur_sc_not_blocked_def vs_all_heap_simps) done -lemma non_overlapping_merge_refills_released_ipc_queues: +lemma merge_nonoverlapping_head_refill_released_ipc_queues: "\\s. released_ipc_queues s \ cur_sc_not_blocked s \ sc_ptr = cur_sc s\ - non_overlapping_merge_refills sc_ptr + merge_nonoverlapping_head_refill sc_ptr \\_. released_ipc_queues\" - apply (clarsimp simp: non_overlapping_merge_refills_def update_refill_hd_rewrite bind_assoc) + apply (clarsimp simp: merge_nonoverlapping_head_refill_def update_refill_hd_rewrite bind_assoc) apply (rule bind_wp_fwd_skip) apply (intro hoare_vcg_conj_lift_pre_fix) apply (wpsimp wp: refill_pop_head_released_ipc_queues) apply (wpsimp wp: refill_pop_head_cur_sc_not_blocked) apply (clarsimp simp: refill_pop_head_def update_sched_context_set_refills_rewrite bind_assoc) apply (wpsimp wp: set_refills_wp get_refills_wp) - apply (wpsimp wp: set_refills_released_ipc_queues get_refills_wp set_refills_wp) - apply (clarsimp simp: pred_map_ipc_queued_thread_state_iff cur_sc_not_blocked_def obj_at_def) + apply (wpsimp wp: set_refills_released_ipc_queues get_refills_wp + hoare_vcg_imp_lift' hoare_vcg_all_lift + | wpsimp wp: set_refills_wp)+ + apply (clarsimp simp: pred_map_ipc_queued_thread_state_iff cur_sc_not_blocked_def) done lemma head_insufficient_loop_released_ipc_queues: @@ -17599,14 +17589,12 @@ lemma head_insufficient_loop_released_ipc_queues: \\_. released_ipc_queues\" (is "valid ?pre _ _") apply (clarsimp simp: head_insufficient_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) apply (intro hoare_vcg_conj_lift_pre_fix) - apply (wpsimp wp: non_overlapping_merge_refills_released_ipc_queues) - apply (clarsimp simp: non_overlapping_merge_refills_def refill_pop_head_def update_refill_hd_def) + apply (wpsimp wp: merge_nonoverlapping_head_refill_released_ipc_queues) + apply (clarsimp simp: merge_nonoverlapping_head_refill_def refill_pop_head_def update_refill_hd_def) apply (wpsimp wp: set_refills_wp get_refills_wp) apply (clarsimp simp: cur_sc_not_blocked_def vs_all_heap_simps) - apply (clarsimp simp: non_overlapping_merge_refills_def refill_pop_head_def) apply (wpsimp wp: set_refills_wp get_refills_wp) done @@ -17624,43 +17612,28 @@ lemma schedule_used_released_ipc_queues: crunch refill_single for inv[wp]: P -lemma handle_overrun_loop_body_released_ipc_queues: +lemma charge_entire_head_refill_released_ipc_queues: "\\s. released_ipc_queues s \ cur_sc_not_blocked s \ sc_ptr = cur_sc s\ - handle_overrun_loop_body usage + charge_entire_head_refill sc_ptr usage \\_. released_ipc_queues\" (is "valid _ _ (\_ s. ?post s)") - apply (clarsimp simp: handle_overrun_loop_body_def) - apply (rule bind_wp[OF _ gets_sp]) - apply (rule bind_wp_fwd_skip, wpsimp) - apply (rule bind_wp_fwd_skip, wpsimp) - apply (rule_tac Q'="\_ s. ?post s" in bind_wp; (solves wpsimp)?) - apply (rule hoare_if) - apply (wpsimp simp: update_refill_hd_rewrite - wp: set_refills_released_ipc_queues get_refills_wp) + apply (clarsimp simp: charge_entire_head_refill_def) + apply (wpsimp wp: set_refills_released_ipc_queues get_refills_wp + schedule_used_released_ipc_queues + simp: update_sched_context_set_refills_rewrite refill_budget_check_defs + | wpsimp wp: set_refills_wp)+ apply (clarsimp simp: pred_map_ipc_queued_thread_state_iff cur_sc_not_blocked_def) - apply (rule bind_wp_fwd_skip) - apply (clarsimp simp: pred_conj_def) - apply (intro hoare_vcg_conj_lift_pre_fix) - apply (wpsimp wp: refill_pop_head_released_ipc_queues set_refills_released_ipc_queues - get_refills_wp hoare_vcg_all_lift hoare_drop_imps) - apply (wpsimp wp: refill_pop_head_cur_sc_not_blocked) - apply (wpsimp wp: set_refills_wp) - apply (wpsimp wp: set_refills_wp) - apply (wpsimp wp: schedule_used_released_ipc_queues) done lemma handle_overrun_loop_released_ipc_queues: "\\s. released_ipc_queues s \ cur_sc_not_blocked s \ sc_ptr = cur_sc s\ - handle_overrun_loop usage + handle_overrun sc_ptr usage \\_. released_ipc_queues\" (is "valid ?pre _ _") - apply (clarsimp simp: handle_overrun_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] - ; fastforce?) - apply (intro hoare_vcg_conj_lift_pre_fix - ; (solves handle_overrun_loop_body_simple)?) - apply (wpsimp wp: handle_overrun_loop_body_released_ipc_queues) - apply fastforce + apply (clarsimp simp: handle_overrun_def) + apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]; fastforce?) + apply (intro hoare_vcg_conj_lift_pre_fix; (solves charge_entire_head_refill_simple)?) + apply (wpsimp wp: charge_entire_head_refill_released_ipc_queues) done lemma merge_overlapping_refills_released_ipc_queues: @@ -17698,7 +17671,7 @@ lemma refill_head_overlapping_loop_released_ipc_queues: apply (wpsimp wp: set_refills_wp get_refills_wp) done -crunch head_insufficient_loop, handle_overrun_loop +crunch head_insufficient_loop, handle_overrun for cur_sc_not_blocked[wp]: cur_sc_not_blocked (wp: crunch_wps) @@ -17712,20 +17685,20 @@ lemma refill_budget_check_released_ipc_queues: apply (rule bind_wp_fwd_skip) apply (clarsimp simp: pred_conj_def) apply (intro hoare_vcg_conj_lift_pre_fix - ; (solves \wpsimp | handle_overrun_loop_simple\)?) + ; (solves \wpsimp | handle_overrun_simple\)?) apply (wpsimp wp: handle_overrun_loop_released_ipc_queues) - apply fastforce apply (rule bind_wp_fwd_skip, solves wpsimp)+ apply (rule bind_wp_fwd_skip) + apply (clarsimp simp: pred_conj_def) apply (intro hoare_vcg_conj_lift_pre_fix) - apply (unfold update_refill_hd_def) - apply (wpsimp wp: schedule_used_released_ipc_queues update_sched_context_set_refills_rewrite) - apply (clarsimp simp: update_sched_context_set_refills_rewrite) - apply (wpsimp wp: set_refills_released_ipc_queues get_refills_wp - simp: cur_sc_not_blocked_def merge_overlapping_refills_def refill_pop_head_def - obj_at_def)+ - apply (fastforce simp: pred_map_ipc_queued_thread_state_iff) + apply (wpsimp wp: schedule_used_released_ipc_queues set_refills_released_ipc_queues + get_refills_wp hoare_vcg_all_lift hoare_vcg_imp_lift' + simp: update_refill_hd_def update_sched_context_set_refills_rewrite + | wpsimp wp: set_refills_wp)+ + apply (fastforce simp: cur_sc_not_blocked_def vs_all_heap_simps + ipc_queued_thread_state_def2 + split: if_splits) apply (wpsimp wp: set_refills_wp get_refills_wp) apply (clarsimp simp: cur_sc_not_blocked_def vs_all_heap_simps) apply (wpsimp wp: set_refills_wp) @@ -18288,15 +18261,13 @@ lemma refill_budget_check_sc_refill_max_sc_at: apply (clarsimp simp: refill_budget_check_def schedule_used_defs) apply (rule bind_wp_fwd_skip, solves wpsimp)+ apply (rule bind_wp_fwd_skip) - apply (handle_overrun_loop_simple - ; clarsimp simp: sc_at_pred_n_def obj_at_def) + apply (handle_overrun_simple; clarsimp simp: sc_at_pred_n_def obj_at_def) apply (rule bind_wp_fwd_skip, solves wpsimp)+ apply (rule bind_wp_fwd_skip) apply (wpsimp wp: set_refills_wp get_refills_wp whileLoop_valid_inv simp: obj_at_def merge_overlapping_refills_def refill_pop_head_def sc_at_pred_n_def update_sched_context_set_refills_rewrite update_refill_hd_rewrite) - apply (head_insufficient_loop_simple - ; clarsimp simp: sc_at_pred_n_def obj_at_def) + apply (head_insufficient_loop_simple; clarsimp simp: sc_at_pred_n_def obj_at_def) done lemma refill_budget_check_round_robin_sc_refill_max_sc_at: diff --git a/proof/invariant-abstract/Deterministic_AI.thy b/proof/invariant-abstract/Deterministic_AI.thy index 20128220d3..cb84fc0ced 100644 --- a/proof/invariant-abstract/Deterministic_AI.thy +++ b/proof/invariant-abstract/Deterministic_AI.thy @@ -4103,11 +4103,9 @@ crunch set_refills,refill_size, schedule_used for valid_list[wp]: valid_list (simp: schedule_used_defs wp: get_refills_wp) -lemma refill_budget_check_valid_list[wp]: - "refill_budget_check usage \valid_list\" - unfolding refill_budget_check_defs - apply (wpsimp wp: whileLoop_valid_inv | wp (once) hoare_drop_imps)+ - done +crunch refill_budget_check + for valid_list[wp]: valid_list + (simp: crunch_simps wp: crunch_wps) lemma refill_budget_check_round_robin_valid_list[wp]: "\valid_list\ refill_budget_check_round_robin usage \\_.valid_list\" diff --git a/proof/invariant-abstract/RealTime_AI.thy b/proof/invariant-abstract/RealTime_AI.thy index 3a053cc211..fbe61a295d 100644 --- a/proof/invariant-abstract/RealTime_AI.thy +++ b/proof/invariant-abstract/RealTime_AI.thy @@ -1050,6 +1050,16 @@ lemma read_refill_head_wp[wp]: lemmas get_refill_head_wp[wp] = ovalid_gets_the[OF read_refill_head_wp, simplified get_refill_head_def[symmetric]] +lemma read_refill_tail_wp[wp]: + "\\s. \sc n. kheap s sc_ptr = Some (SchedContext sc n) \ Q (refill_tl sc) s\ + read_refill_tail sc_ptr + \Q\" + unfolding read_refill_tail_def + by wpsimp + +lemmas get_refill_tail_wp[wp] = + ovalid_gets_the[OF read_refill_tail_wp, simplified get_refill_tail_def[symmetric]] + lemma read_sc_refill_ready_wp[wp]: "\\s. \sc n. kheap s sc_ptr = Some (SchedContext sc n) \ Q (refill_ready (cur_time s) (refill_hd sc)) s\ @@ -1104,6 +1114,16 @@ lemma no_ofail_read_refill_head[wp]: lemmas no_fail_get_refill_head[wp] = no_ofail_gets_the[OF no_ofail_read_refill_head, simplified get_refill_head_def[symmetric]] +lemma no_ofail_read_refill_tail[wp]: + "no_ofail (\s. sc_refills_sc_at (\refills. refills \ []) scp s) (read_refill_tail scp)" + unfolding read_refill_tail_def + apply wpsimp + apply (clarsimp simp: sc_at_pred_n_def obj_at_def) + done + +lemmas no_fail_get_refill_tail[wp] = + no_ofail_gets_the[OF no_ofail_read_refill_tail, simplified get_refill_tail_def[symmetric]] + lemma no_ofail_read_sc_refill_ready[wp]: "no_ofail (\s. sc_refills_sc_at (\refills. refills \ []) scp s) (read_sc_refill_ready scp)" unfolding read_sc_refill_ready_def @@ -1138,11 +1158,27 @@ lemma refill_pop_head_no_fail[wp]: unfolding refill_pop_head_def by wpsimp -lemma non_overlapping_merge_refills_no_fail[wp]: +lemma update_refill_hd_no_fail[wp]: + "no_fail (\s. sc_refills_sc_at (\refills. refills \ []) sc_ptr s) + (update_refill_hd sc_ptr f)" + unfolding update_refill_hd_def refill_pop_head_def + apply wpsimp + by (clarsimp simp: obj_at_def sc_at_pred_n_def) + +lemma set_refill_hd_no_fail[wp]: "no_fail (\s. sc_refills_sc_at (\refills. refills \ []) sc_ptr s) - (non_overlapping_merge_refills sc_ptr)" - unfolding non_overlapping_merge_refills_def refill_pop_head_def - by (wpsimp simp: update_refill_hd_def wp: update_sched_context_wp) + (set_refill_hd sc_ptr f)" + unfolding set_refill_hd_def refill_pop_head_def + by wpsimp + +lemma merge_nonoverlapping_head_refill_no_fail[wp]: + "no_fail (\s. sc_refills_sc_at (\refills. 1 < length refills) sc_ptr s) + (merge_nonoverlapping_head_refill sc_ptr)" + unfolding merge_nonoverlapping_head_refill_def refill_pop_head_def + apply (wpsimp simp: set_refill_hd_def update_refill_hd_def wp: update_sched_context_wp) + apply (clarsimp simp: obj_at_def sc_at_pred_n_def) + apply (case_tac "sc_refills sc"; clarsimp) + done lemma schedule_used_no_fail[wp]: "no_fail (\s. sc_refills_sc_at (\refills. refills \ []) sc_ptr s) @@ -1162,21 +1198,10 @@ lemma refill_pop_head_nonempty_refills: apply (clarsimp simp: sc_at_pred_n_def obj_at_def neq_Nil_lengthI) done -lemma non_overlapping_merge_refills_nonempty_refills: - "\\s. sc_refills_sc_at (\refills. 1 < length refills) sc_ptr s\ - non_overlapping_merge_refills sc_ptr - \\_ s. sc_refills_sc_at (\refills. refills \ []) sc_ptr s\" - unfolding non_overlapping_merge_refills_def update_refill_hd_def - apply (rule_tac Q'="\_ s. sc_refills_sc_at (\refills. refills \ []) sc_ptr s" in bind_wp) - apply (wpsimp wp: update_sched_context_wp) - apply (clarsimp simp: sc_at_pred_n_def obj_at_def neq_Nil_lengthI) - apply (wpsimp wp: refill_pop_head_nonempty_refills) - done - -lemma handle_overrun_loop_body_no_fail: - "no_fail (\s. sc_refills_sc_at (\refills. refills \ []) (cur_sc s) s) - (handle_overrun_loop_body usage)" - unfolding handle_overrun_loop_body_def +lemma charge_entire_head_refill_no_fail: + "no_fail (\s. sc_refills_sc_at (\refills. refills \ []) sc_ptr s) + (charge_entire_head_refill sc_ptr usage)" + unfolding charge_entire_head_refill_def apply (wpsimp simp: refill_single_def refill_size_def get_refills_def update_refill_hd_def wp: refill_pop_head_no_fail refill_pop_head_nonempty_refills) apply (clarsimp simp: sc_at_pred_n_def obj_at_def Suc_lessI) diff --git a/proof/invariant-abstract/SchedContext_AI.thy b/proof/invariant-abstract/SchedContext_AI.thy index 159d0c9452..3f2ff9724c 100644 --- a/proof/invariant-abstract/SchedContext_AI.thy +++ b/proof/invariant-abstract/SchedContext_AI.thy @@ -91,10 +91,10 @@ lemmas refill_unblock_check_defs refill_head_overlapping_loop_def set_refill_hd_def update_refill_hd_def lemmas refill_budget_check_defs - = refill_budget_check_def non_overlapping_merge_refills_def merge_overlapping_refills_def - refill_pop_head_def head_insufficient_loop_def handle_overrun_loop_def - handle_overrun_loop_body_def set_refill_hd_def update_refill_hd_def refill_single_def - refill_size_def + = refill_budget_check_def merge_nonoverlapping_head_refill_def + refill_pop_head_def head_insufficient_loop_def handle_overrun_def charge_entire_head_refill_def + set_refill_hd_def update_refill_hd_def update_refill_tl_def refill_single_def refill_size_def + schedule_used_def refill_add_tail_def lemma update_sched_context_decompose: "update_sched_context scp (\sc. f (g sc)) @@ -215,7 +215,8 @@ lemma refill_pop_head_valid_objs[wp]: lemma head_insufficient_loop_valid_objs[wp]: "head_insufficient_loop usage \valid_objs\" unfolding head_insufficient_loop_def - by (wpsimp wp: whileLoop_valid_inv simp: non_overlapping_merge_refills_def update_refill_hd_def) + by (wpsimp wp: whileLoop_valid_inv hoare_drop_imps + simp: merge_nonoverlapping_head_refill_def set_refill_hd_def) lemma schedule_used_valid_objs[wp]: "schedule_used sc_ptr new \valid_objs\" @@ -224,19 +225,19 @@ lemma schedule_used_valid_objs[wp]: apply (clarsimp simp: obj_at_def sc_at_ppred_def valid_sched_context_def) done -lemma handle_overrun_loop_valid_objs[wp]: - "handle_overrun_loop usage \valid_objs\" - unfolding handle_overrun_loop_def handle_overrun_loop_body_def refill_single_def refill_size_def +lemma handle_overrun_valid_objs[wp]: + "handle_overrun sc_ptr usage \valid_objs\" + unfolding handle_overrun_def charge_entire_head_refill_def refill_single_def refill_size_def update_refill_hd_def apply (wpsimp wp: whileLoop_valid_inv - simp: valid_sched_context_def) + simp: set_refill_hd_def valid_sched_context_def) done lemma refill_budget_check_valid_objs[wp]: "refill_budget_check usage \valid_objs\" apply (clarsimp simp: refill_budget_check_def) apply (rule bind_wp_fwd_skip, solves wpsimp)+ - apply (wpsimp wp: set_refills_valid_objs simp: update_refill_hd_def) + apply (wpsimp wp: set_refills_valid_objs simp: set_refill_hd_def update_refill_hd_def) done (* FIXME RT: move to Invariants_AI *) @@ -528,7 +529,7 @@ crunch refill_budget_check for ex_nonz_cap_tp[wp]: "\s. ex_nonz_cap_to ptr s" (simp: crunch_simps wp: crunch_wps) -crunch head_insufficient_loop, handle_overrun_loop +crunch head_insufficient_loop, handle_overrun for if_live_then_nonz_cap[wp]: if_live_then_nonz_cap (wp: crunch_wps update_sched_context_iflive_implies) @@ -774,7 +775,7 @@ lemma valid_sched_context_domain_time_update[simp]: "valid_sched_context p (domain_time_update f s) = valid_sched_context p s" by (simp add: valid_sched_context_def valid_bound_obj_def split: option.splits) -crunch head_insufficient_loop, handle_overrun_loop +crunch head_insufficient_loop, handle_overrun for valid_replies_pred[wp]: "valid_replies_pred P" (wp: crunch_wps ignore: update_sched_context) @@ -880,7 +881,7 @@ lemma set_refills_bound_sc_tcb_at_ct[wp]: by (wpsimp simp: set_refills_def update_sched_context_def set_object_def get_object_def pred_tcb_at_def obj_at_def) -crunch handle_overrun_loop, head_insufficient_loop +crunch handle_overrun, head_insufficient_loop for bound_sc_tcb_at_ct[wp]: "\s. bound_sc_tcb_at P (cur_thread s) s" (wp: crunch_wps ignore: update_sched_context) @@ -995,7 +996,7 @@ lemma set_refills_ct_in_state[wp]: "\ ct_in_state t \ set_refills p r \ \rv. ct_in_state t \" by (wpsimp simp: set_refills_def wp: get_sched_context_wp) -crunch head_insufficient_loop, handle_overrun_loop +crunch head_insufficient_loop, handle_overrun for ct_in_state[wp]: "ct_in_state t" (wp: crunch_wps) @@ -1065,13 +1066,13 @@ lemma refill_budget_check_round_robin_sc_obj_at_n[wp]: "refill_budget_check_round_robin usage \\s. Q (sc_obj_at n sc_ptr s)\" by (wpsimp simp: refill_budget_check_round_robin_def update_refill_hd_def update_refill_tl_def) -crunch handle_overrun_loop, head_insufficient_loop +crunch handle_overrun, head_insufficient_loop for sc_obj_at[wp]: "\s. Q (sc_obj_at n sc_ptr s)" (wp: crunch_wps) lemma refill_budget_check_sc_obj_at_n[wp]: "refill_budget_check usage \\s. Q (sc_obj_at n sc_ptr s)\" - by (wpsimp simp: refill_budget_check_def is_round_robin_def) + by (wpsimp simp: refill_budget_check_def is_round_robin_def wp: hoare_drop_imp) lemma commit_time_sc_obj_at[wp]: "commit_time \\s. Q (sc_obj_at n sc_ptr s)\" diff --git a/proof/refine/RISCV64/Ipc_R.thy b/proof/refine/RISCV64/Ipc_R.thy index 049f0e4c57..0550495d57 100644 --- a/proof/refine/RISCV64/Ipc_R.thy +++ b/proof/refine/RISCV64/Ipc_R.thy @@ -4010,15 +4010,6 @@ lemma schedContextResume_corres: opt_map_def opt_pred_def is_active_sc'_def) done -lemma updateRefillHd_valid_objs': - "\valid_objs' and active_sc_at' scPtr\ updateRefillHd scPtr f \\_. valid_objs'\" - apply (clarsimp simp: updateRefillHd_def updateSchedContext_def) - apply wpsimp - apply (frule (1) sc_ko_at_valid_objs_valid_sc') - apply (fastforce simp: valid_sched_context'_def active_sc_at'_def obj_at'_real_def ko_wp_at'_def - valid_sched_context_size'_def objBits_def objBitsKO_def refillSize_def) - done - lemma getCTE_cap_to_refs[wp]: "\\\ getCTE p \\rv s. \r\zobj_refs' (cteCap rv). ex_nonz_cap_to' r s\" apply (rule hoare_strengthen_post [OF getCTE_sp]) diff --git a/proof/refine/RISCV64/KHeap_R.thy b/proof/refine/RISCV64/KHeap_R.thy index 08b9475127..5862388bc1 100644 --- a/proof/refine/RISCV64/KHeap_R.thy +++ b/proof/refine/RISCV64/KHeap_R.thy @@ -4203,10 +4203,26 @@ lemma readRefillHead_wp[wp]: lemmas getRefillHead_wp[wp] = ovalid_gets_the[OF readRefillHead_wp, simplified getRefillHead_def[symmetric]] +lemma readRefillTail_wp[wp]: + "\\s. \sc. scs_of' s scPtr = Some sc \ Q (refillTl sc) s\ + readRefillTail scPtr + \Q\" + unfolding readRefillTail_def readSchedContext_def + apply (wpsimp wp: set_sc'.readObject_wp) + apply (clarsimp simp: opt_map_def obj_at'_def) + done + +lemmas getRefillTail_wp[wp] = + ovalid_gets_the[OF readRefillTail_wp, simplified getRefillTail_def[symmetric]] + lemma getRefillHead_sp: "\P\ getRefillHead scPtr \\rv s. P s \ (\sc. scs_of' s scPtr = Some sc \ refillHd sc = rv)\" by wpsimp +lemma getRefillTail_sp: + "\P\ getRefillTail scPtr \\rv s. P s \ (\sc. scs_of' s scPtr = Some sc \ refillTl sc = rv)\" + by wpsimp + lemma readRefillReady_wp: "\\s. \sc. scs_of' s scp = Some sc \ P (rTime (refillHd sc) \ ksCurTime s) s\ readRefillReady scp diff --git a/proof/refine/RISCV64/SchedContextInv_R.thy b/proof/refine/RISCV64/SchedContextInv_R.thy index 78865be38a..a0d23810d9 100644 --- a/proof/refine/RISCV64/SchedContextInv_R.thy +++ b/proof/refine/RISCV64/SchedContextInv_R.thy @@ -1234,7 +1234,7 @@ crunch commitTime, refillNew, refillUpdate crunch commitTime for ex_nonz_cap_to'[wp]: "ex_nonz_cap_to' ptr" - (wp: crunch_wps simp: crunch_simps) + (wp: crunch_wps hoare_vcg_all_lift simp: crunch_simps) lemma invokeSchedControlConfigureFlags_corres: "sc_ctrl_inv_rel sc_inv sc_inv' \ diff --git a/proof/refine/RISCV64/Schedule_R.thy b/proof/refine/RISCV64/Schedule_R.thy index 81d9ad5583..d680ad5b74 100644 --- a/proof/refine/RISCV64/Schedule_R.thy +++ b/proof/refine/RISCV64/Schedule_R.thy @@ -55,7 +55,7 @@ crunch refillUnblockCheck, refillBudgetCheck, ifCondRefillUnblockCheck, refillBu and valid_dom_schedule'[wp]: valid_dom_schedule' and ksCurSc[wp]: "\s. P (ksCurSc s)" and pspace_in_kernel_mappings'[wp]: pspace_in_kernel_mappings' - (wp: crunch_wps valid_dom_schedule'_lift simp: crunch_simps refillSingle_def) + (wp: crunch_wps hoare_vcg_all_lift valid_dom_schedule'_lift simp: crunch_simps refillSingle_def) crunch commitTime, refillNew, refillUpdate for typ_at'[wp]: "\s. P (typ_at' T p s)" @@ -70,10 +70,10 @@ global_interpretation refillPopHead: typ_at_all_props' "refillPopHead scPtr" global_interpretation updateRefillTl: typ_at_all_props' "updateRefillTl scPtr f" by typ_at_props' -global_interpretation handleOverrunLoopBody: typ_at_all_props' "handleOverrunLoopBody usage" +global_interpretation chargeEntireHeadRefill: typ_at_all_props' "chargeEntireHeadRefill scPtr usage" by typ_at_props' -global_interpretation handleOverrunLoop: typ_at_all_props' "handleOverrunLoop usage" +global_interpretation handleOverrun: typ_at_all_props' "handleOverrun scPtr usage" by typ_at_props' global_interpretation scheduleUsed: typ_at_all_props' "scheduleUsed scPtr new" @@ -88,7 +88,7 @@ global_interpretation mergeOverlappingRefills: typ_at_all_props' "mergeOverlappi global_interpretation setRefillHd: typ_at_all_props' "setRefillHd scPtr v" by typ_at_props' -global_interpretation nonOverlappingMergeRefills: typ_at_all_props' "nonOverlappingMergeRefills scPtr" +global_interpretation mergeNonoverlappingHeadRefill: typ_at_all_props' "mergeNonoverlappingHeadRefill scPtr" by typ_at_props' global_interpretation refillBudgetCheck: typ_at_all_props' "refillBudgetCheck usage" @@ -1976,7 +1976,7 @@ lemma scheduleUsed_if_live_then_nonz_cap'[wp]: "scheduleUsed scPtr refill \if_live_then_nonz_cap'\" unfolding scheduleUsed_def by (wpsimp simp: scheduleUsed_def - wp: updateSchedContext_wp getRefillSize_wp getRefillNext_wp getRefillFull_wp) + wp: updateSchedContext_wp getRefillSize_wp getRefillNext_wp getRefillFull_wp getRefillTail_wp) lemma updateSchedContext_valid_idle': "\valid_idle' and (\s. \sc. idle_sc' sc \ idle_sc' (f sc))\ @@ -2093,22 +2093,16 @@ crunch refillBudgetCheck, refillUnblockCheck for typ_at'[wp]: "\s. P (typ_at' T p s)" and sc_at'_n[wp]: "\s. P (sc_at'_n T p s)" and active_sc_at'[wp]: "active_sc_at' scPtr" - (wp: crunch_wps) + (wp: crunch_wps hoare_vcg_all_lift simp: crunch_simps) -lemma updateRefillHd_valid_objs': - "\\s. valid_objs' s \ sc_at' scPtr s\ - updateRefillHd scPtr f - \\_. valid_objs'\" - apply (clarsimp simp: updateRefillHd_def) - apply (wpsimp wp: set_sc_valid_objs' simp: updateSchedContext_def) - apply (clarsimp simp: is_active_sc'_def valid_sched_context'_def obj_at_simps valid_obj'_def - valid_sched_context_size'_def opt_map_red opt_pred_def active_sc_at'_rewrite - refillSize_def - dest!: sc_ko_at_valid_objs_valid_sc') - done +lemma updateRefillHd_valid_objs'[wp]: + "updateRefillHd scPtr f \valid_objs'\" + unfolding updateRefillHd_def updateSchedContext_def + apply wpsimp + by (fastforce dest: sc_ko_at_valid_objs_valid_sc' simp: valid_sched_context'_def refillSize_def) lemma mergeOverlappingRefills_valid_objs': - "\\s. valid_objs' s \ sc_at' scPtr s \ ((\sc. 1 < refillSize sc) |< scs_of' s) scPtr\ + "\\s. valid_objs' s \ ((\sc. 1 < refillSize sc) |< scs_of' s) scPtr\ mergeOverlappingRefills scPtr \\_. valid_objs'\" unfolding mergeOverlappingRefills_def @@ -2231,17 +2225,6 @@ lemma updateRefillHd_if_live_then_nonz_cap'[wp]: simp: ko_wp_at'_def obj_at'_real_def live_sc'_def) done -crunch mergeOverlappingRefills - for if_live_then_nonz_cap'[wp]: if_live_then_nonz_cap' - -lemma nonOverlappingMergeRefills_if_live_then_nonz_cap'[wp]: - "nonOverlappingMergeRefills scPtr \if_live_then_nonz_cap'\" - apply (clarsimp simp: nonOverlappingMergeRefills_def updateRefillHd_def) - apply (rule bind_wp_fwd_skip, wpsimp wp: getRefillSize_wp)+ - by (wpsimp simp: updateSchedContext_def getRefillSize_wp, - fastforce intro: if_live_then_nonz_capE' - simp: ko_wp_at'_def obj_at'_real_def live_sc'_def)+ - crunch refillHeadOverlappingLoop, headInsufficientLoop for if_live_then_nonz_cap'[wp]: if_live_then_nonz_cap' (wp: crunch_wps) @@ -2251,28 +2234,11 @@ lemma setRefillHd_if_live_then_nonz_cap'[wp]: apply (wpsimp simp: setRefillHd_def) done -crunch handleOverrunLoop +crunch handleOverrun, refillUnblockCheck for if_live_then_nonz_cap'[wp]: if_live_then_nonz_cap' (wp: crunch_wps) -lemma refillUnblockCheck_if_live_then_nonz_cap'[wp]: - "refillUnblockCheck scPtr \if_live_then_nonz_cap'\" - apply (clarsimp simp: refillUnblockCheck_def setReprogramTimer_def refillReady_def - isRoundRobin_def) - apply (wpsimp wp: scActive_wp) - done - -lemma updateSchedContext_scRefills_updateAt_valid_idle'[wp]: - "updateSchedContext scPtr (\sc. scRefills_update (\_. updateAt (scRefillHead sc) (scRefills sc) f) sc) - \valid_idle'\" - unfolding updateSchedContext_def - by (wpsimp simp: valid_idle'_def obj_at'_def) - -lemma nonOverlappingMergeRefills_valid_idle'[wp]: - "nonOverlappingMergeRefills scPtr \valid_idle'\" - by (wpsimp wp: getRefillSize_wp simp: nonOverlappingMergeRefills_def updateRefillHd_def) - -crunch refillHeadOverlappingLoop, headInsufficientLoop, handleOverrunLoop +crunch refillHeadOverlappingLoop, headInsufficientLoop, handleOverrun for valid_idle'[wp]: valid_idle' (wp: crunch_wps) @@ -2285,7 +2251,7 @@ crunch refillUnblockCheck, refillBudgetCheck and ksReadyQueues[wp]: "\s. P (ksReadyQueues s)" and ksReadyQueuesL1Bitmap[wp]: "\s. P (ksReadyQueuesL1Bitmap s)" and ksReadyQueuesL2Bitmap[wp]: "\s. P (ksReadyQueuesL2Bitmap s)" - (wp: crunch_wps valid_replies'_lift) + (wp: crunch_wps hoare_vcg_all_lift valid_replies'_lift hoare_vcg_if_lift2) lemma refillUnblockCheck_invs': "refillUnblockCheck scPtr \invs'\" @@ -2302,65 +2268,13 @@ lemma getRefillSize_sp: apply (clarsimp simp: obj_at'_def) done -lemma nonOverlappingMergeRefills_valid_objs': - "\\s. valid_objs' s \ sc_at' scPtr s\ - nonOverlappingMergeRefills scPtr - \\_. valid_objs'\" - apply (clarsimp simp: nonOverlappingMergeRefills_def updateRefillHd_def) - apply (rule bind_wp[OF _ getRefillSize_sp]) - apply (rule bind_wp[OF _ assert_sp]) - apply (rule_tac Q'="\_ s. valid_objs' s \ sc_at' scPtr s" in bind_wp_fwd) - apply wpsimp - apply (rule bind_wp_fwd_skip) - apply (wpsimp wp: updateSchedContext_valid_objs') - apply (clarsimp simp: valid_sched_context'_def obj_at_simps valid_obj'_def - valid_sched_context_size'_def in_omonad refillSize_def) - apply (wpsimp wp: updateSchedContext_valid_objs') - apply (clarsimp simp: valid_sched_context'_def obj_at_simps valid_obj'_def - valid_sched_context_size'_def in_omonad refillSize_def) - done - -crunch refillHeadOverlappingLoop, headInsufficientLoop, handleOverrunLoop - for active_sc_at'[wp]: "active_sc_at' scPtr" - and ksCurSc[wp]: "\s. P (ksCurSc s)" - (wp: crunch_wps) - -lemma headInsufficientLoop_valid_objs': - "\\s. valid_objs' s \ sc_at' scPtr s\ - headInsufficientLoop scPtr - \\_. valid_objs'\" - (is "valid ?pre _ _") - apply (clarsimp simp: headInsufficientLoop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"] nonOverlappingMergeRefills_valid_objs') - done - -lemma handleOverrunLoop_valid_objs': - "\\s. valid_objs' s \ active_sc_at' (ksCurSc s) s\ - handleOverrunLoop usage - \\_. valid_objs'\" - (is "valid ?pre _ _") - apply (clarsimp simp: handleOverrunLoop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_. ?pre"]) - apply (intro hoare_vcg_conj_lift_pre_fix) - apply (clarsimp simp: handleOverrunLoopBody_def) - apply (wpsimp wp: updateRefillHd_valid_objs' simp: refillSingle_def) - apply (rule_tac f=ksCurSc in hoare_lift_Pf3) - apply wpsimp+ - done - -lemma refillBudgetCheck_valid_objs': - "refillBudgetCheck usage \valid_objs'\" - apply (clarsimp simp: refillBudgetCheck_def isRoundRobin_def refillReady_def getCurSc_def) - apply (rule bind_wp[OF _ gets_sp]) - apply (rule bind_wp[OF _ scActive_sp]) - apply (wpsimp wp: headInsufficientLoop_valid_objs' handleOverrunLoop_valid_objs' - hoare_vcg_all_lift updateRefillHd_valid_objs' hoare_vcg_if_lift2 hoare_drop_imps) - apply (clarsimp simp: active_sc_at'_def obj_at'_def) - done +crunch headInsufficientLoop, handleOverrun, refillBudgetCheck + for valid_objs'[wp]: valid_objs' + (wp: crunch_wps hoare_vcg_all_lift simp: crunch_simps) lemma refillBudgetCheck_valid_mdb'[wp]: "refillBudgetCheck usage \valid_mdb'\" - apply (clarsimp simp: handleOverrunLoop_def valid_mdb'_def) + apply (clarsimp simp: handleOverrun_def valid_mdb'_def) apply (wpsimp wp: scActive_wp) done @@ -2374,36 +2288,14 @@ lemma refillAddTail_list_refs_of_replies'[wp]: by (wpsimp wp: hoare_drop_imps getRefillNext_wp getRefillSize_wp simp: o_def updateSchedContext_def refillAddTail_def updateRefillIndex_def) -crunch scheduleUsed, refillPopHead, handleOverrunLoop, nonOverlappingMergeRefills, - headInsufficientLoop +crunch refillBudgetCheck for list_refs_of_replies'[wp]: "\s. P (list_refs_of_replies' s)" and valid_machine_state'[wp]: valid_machine_state' - (wp: crunch_wps) - -lemma refillBudgetCheck_list_refs_of_replies'[wp]: - "refillBudgetCheck usage \\s. sym_refs (list_refs_of_replies' s)\" - unfolding refillBudgetCheck_def - apply (wpsimp wp: hoare_drop_imps) - apply (simp add: o_def) - done + (wp: crunch_wps hoare_vcg_all_lift hoare_vcg_if_lift2) -lemma refillBudgetCheck_if_live_then_nonz_cap'[wp]: - "refillBudgetCheck uage \if_live_then_nonz_cap'\" - by (wpsimp simp: refillBudgetCheck_def setReprogramTimer_def refillReady_def isRoundRobin_def - wp: hoare_drop_imps) - -lemma refillBudgetCheck_valid_idle'[wp]: - "refillBudgetCheck usage \valid_idle'\" - apply (clarsimp simp: refillBudgetCheck_def isRoundRobin_def refillReady_def - setReprogramTimer_def updateRefillHd_def setRefillHd_def) - apply (rule bind_wp_fwd_skip, solves wpsimp)+ - apply (wpsimp wp: updateSchedContext_valid_idle') - done - -lemma refillBudgetCheck_valid_machine_state'[wp]: - "refillBudgetCheck usage \valid_machine_state'\" - unfolding refillBudgetCheck_def - by (wpsimp wp: hoare_drop_imps) +crunch refillBudgetCheck + for if_live_then_nonz_cap'[wp]: if_live_then_nonz_cap' + (simp: crunch_simps wp: crunch_wps hoare_vcg_all_lift) lemma refillBudgetCheck_invs'[wp]: "refillBudgetCheck usage \invs'\" @@ -3508,7 +3400,7 @@ lemma refillHeadOverlapping_corres_eq: by linarith+ crunch update_refill_hd, refill_pop_head, merge_overlapping_refills, schedule_used, - handle_overrun_loop_body + charge_entire_head_refill for is_active_sc2[wp]: "is_active_sc2 scp" (wp: crunch_wps ignore: update_sched_context simp: crunch_simps update_refill_hd_rewrite update_sched_context_set_refills_rewrite) @@ -3724,20 +3616,18 @@ lemma refillUnblockCheck_corres: apply (rule corres_assert_gen_asm_cross[where P=P' and P'=P' for P', where Q=Q' and Q'=Q' for Q', simplified]) apply (fastforce dest!: active_sc_at'_cross_valid_objs simp: active_sc_at'_def obj_at'_def) - apply simp - apply (rule corres_guard_imp) - apply (rule corres_split_eqr[OF isRoundRobin_corres]) - apply (rule corres_split_eqr[OF refillReady_corres], simp) - apply (rule corres_when, fastforce) - apply (rule corres_split[OF getCurTime_corres]) - apply (rule corres_split[OF updateRefillHd_corres], simp) - apply (clarsimp simp: refill_map_def) - apply (rule corres_split[OF setReprogramTimer_corres]) - apply (rule refillHeadOverlappingLoop_corres) - apply (wpsimp simp: setReprogramTimer_def)+ - apply (wpsimp simp: is_active_sc_rewrite)+ - apply (wpsimp wp: updateRefillHd_valid_objs')+ - apply (wpsimp wp: is_round_robin_wp isRoundRobin_wp)+ + apply simp + apply (rule corres_guard_imp) + apply (rule corres_split_eqr[OF isRoundRobin_corres]) + apply (rule corres_split_eqr[OF refillReady_corres], simp) + apply (rule corres_when, fastforce) + apply (rule corres_split[OF getCurTime_corres]) + apply (rule corres_split[OF updateRefillHd_corres], simp) + apply (clarsimp simp: refill_map_def) + apply (rule corres_split[OF setReprogramTimer_corres]) + apply (rule refillHeadOverlappingLoop_corres) + apply (wpsimp simp: setReprogramTimer_def)+ + apply (wpsimp wp: is_round_robin_wp isRoundRobin_wp)+ apply (clarsimp simp: obj_at_simps opt_map_red opt_pred_def is_active_sc_rewrite sc_at_pred_n_def split: Structures_A.kernel_object.splits) @@ -3951,31 +3841,6 @@ lemma refillBudgetCheckRoundRobin_corres: apply (wpsimp wp: updateRefillHd_valid_objs') done -lemma head_insufficient_length_greater_than_one: - "\the (head_insufficient sc_ptr s); - pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr\ - \ pred_map (\cfg. Suc 0 < length (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr" - apply (frule head_insufficient_true_imp_insufficient) - apply (clarsimp simp: vs_all_heap_simps) - apply (clarsimp simp: vs_all_heap_simps sc_at_ppred_def refills_unat_sum_def word_less_nat_alt) - apply (case_tac "sc_refills y"; fastforce dest!: member_le_sum_list) - done - -lemma update_refill_hd_comp: - "update_refill_hd scPtr (f \ g) - = do update_refill_hd scPtr g; - update_refill_hd scPtr f - od" - apply (clarsimp simp: update_refill_hd_def) - apply (rule box_equals[OF update_sched_context_decompose]; fastforce) - done - -lemma refill_pop_head_is_active_sc[wp]: - "refill_pop_head sc_ptr \is_active_sc sc_ptr'\" - apply (wpsimp simp: refill_pop_head_def wp: update_sched_context_wp) - apply (clarsimp simp: vs_all_heap_simps obj_at_kh_kheap_simps active_sc_def) - done - lemma setSchedContext_is_active_sc_at': "\is_active_sc' scPtr' and K (scPtr' = scPtr \ 0 < scRefillMax sc)\ setSchedContext scPtr sc @@ -4000,38 +3865,29 @@ lemma refillPopHead_is_active_sc_at'[wp]: apply (wpsimp wp: updateSchedContext_is_active_sc_at' getRefillNext_wp) done -lemma nonOverlappingMergeRefills_corres: +lemma mergeNonoverlappingHeadRefill_corres: "sc_ptr = scPtr \ corres dc (pspace_aligned and pspace_distinct and sc_at sc_ptr and is_active_sc sc_ptr and valid_objs - and (\s. pred_map (\cfg. Suc 0 < length (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr)) + and (\s. sc_refills_sc_at (\refills. Suc 0 < length refills) sc_ptr s)) (valid_objs' and valid_refills' sc_ptr) - (non_overlapping_merge_refills sc_ptr) (nonOverlappingMergeRefills scPtr)" - apply (clarsimp simp: non_overlapping_merge_refills_def nonOverlappingMergeRefills_def) + (merge_nonoverlapping_head_refill sc_ptr) (mergeNonoverlappingHeadRefill scPtr)" + apply (clarsimp simp: merge_nonoverlapping_head_refill_def mergeNonoverlappingHeadRefill_def) apply (rule corres_cross[OF sc_at'_cross_rel[where t=scPtr]], simp) - apply (rule corres_symb_exec_r[OF _ getRefillSize_sp, rotated]) - apply (wpsimp wp: getRefillSize_wp) - apply wpsimp apply (rule_tac Q'="is_active_sc' scPtr" in corres_cross_add_guard) apply (fastforce dest: is_active_sc'2_cross) apply (rule_tac Q'="obj_at' (\sc'. Suc 0 < refillSize sc') scPtr" in corres_cross_add_guard) apply (fastforce dest!: length_sc_refills_cross[where P="\l. Suc 0 < l"] - simp: opt_map_red opt_pred_def vs_all_heap_simps obj_at'_def) - apply (rule corres_symb_exec_r[OF _ assert_sp, rotated]) - apply wpsimp - apply (rule no_fail_pre) - apply (rule no_fail_assert) - apply (clarsimp simp: no_fail_def obj_at'_def) - + simp: opt_map_red opt_pred_def vs_all_heap_simps obj_at'_def + sc_at_ppred_def obj_at_def) apply (rule_tac Q="\_. sc_at sc_ptr and is_active_sc sc_ptr" and Q'="\_. valid_refills' scPtr and sc_at' scPtr" - in corres_underlying_split - ; (solves wpsimp)?) - apply (corresKsimp corres: refillPopHead_corres - simp: obj_at_def vs_all_heap_simps pred_map_simps sc_at_ppred_def) - apply (subst update_refill_hd_comp) + in corres_underlying_split; + (solves wpsimp)?) + apply (corres corres: refillPopHead_corres + simp: obj_at_def vs_all_heap_simps pred_map_simps sc_at_ppred_def) apply (rule corres_guard_imp) apply (rule corres_underlying_split[OF updateRefillHd_corres]) apply blast @@ -4045,63 +3901,15 @@ lemma nonOverlappingMergeRefills_corres: apply blast done -lemma head_insufficient_simp: - "sc_at scp s - \ head_insufficient scp s = Some (sc_at_pred (\sc. r_amount (refill_hd sc) < MIN_BUDGET) scp s)" - unfolding head_insufficient_def - by (clarsimp simp: obind_def read_sched_context_def obj_at_def is_sc_obj sc_at_pred_n_def) - -lemma refillHdInsufficient_simp: - "sc_at' scp s - \ refillHdInsufficient scp s - = Some (obj_at' (\sc :: sched_context. rAmount (refillHd sc) < minBudget) scp s)" - unfolding refillHdInsufficient_def - apply (clarsimp simp: obind_def readSchedContext_def split: option.splits) - apply (frule no_ofailD[OF no_ofail_sc_at'_readObject]) - apply (clarsimp simp: obj_at'_def dest!: readObject_misc_ko_at') - done - -lemma head_insufficient_equiv: - "\(s, s') \ state_relation; pspace_aligned s; pspace_distinct s; valid_objs s; - pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) scPtr; valid_refills' scPtr s'\ - \ head_insufficient scPtr s = refillHdInsufficient scPtr s'" - apply (prop_tac "sc_at scPtr s") - apply (fastforce dest: valid_objs_valid_sched_context_size - simp: vs_all_heap_simps obj_at_kh_kheap_simps is_sc_obj_def) - apply (frule state_relation_pspace_relation) - apply (frule sc_at_cross; simp?) - apply (frule state_relation_sc_relation; simp?) - apply (subst head_insufficient_simp; simp?) - apply (subst refillHdInsufficient_simp; simp) - apply (frule refill_hd_relation) - apply (fastforce simp: vs_all_heap_simps valid_refills'_def opt_map_def opt_pred_def obj_at_simps) - apply (clarsimp simp: obj_at_def sc_at_ppred_def obj_at'_def minBudget_def refill_map_def - MIN_BUDGET_def kernelWCETTicks_def opt_map_def vs_all_heap_simps) - done - lemma refill_pop_head_sched_context_at[wp]: "refill_pop_head sc_ptr' \\s. \sc n. kheap s sc_ptr = Some (Structures_A.SchedContext sc n)\" apply (clarsimp simp: refill_pop_head_def) apply (wpsimp wp: update_sched_context_wp) done -lemma non_overlapping_merge_refills_is_active_sc[wp]: - "non_overlapping_merge_refills sc_ptr \is_active_sc sc_ptr'\" - apply (clarsimp simp: non_overlapping_merge_refills_def update_refill_hd_def) - apply (rule bind_wp_fwd_skip, solves wpsimp) - apply (wpsimp wp: update_sched_context_wp) - apply (clarsimp simp: vs_all_heap_simps obj_at_def) - done - -crunch non_overlapping_merge_refills +crunch merge_nonoverlapping_head_refill for valid_objs[wp]: valid_objs -lemma nonOverLappingMergeRefills_valid_refills'[wp]: - "nonOverlappingMergeRefills scPtr \valid_refills' scPtr\" - apply (wpsimp simp: nonOverlappingMergeRefills_def wp: getRefillSize_wp) - apply (clarsimp simp: obj_at'_def) - done - definition head_insufficient_loop_measure where "head_insufficient_loop_measure sc_ptr \ measure (\(_, s). case kheap s sc_ptr of Some (Structures_A.SchedContext sc _) @@ -4111,59 +3919,84 @@ lemma non_overlapping_merge_refills_terminates: "\pred_map (\cfg. refills_unat_sum (scrc_refills cfg) \ unat max_time) (sc_refill_cfgs_of s) sc_ptr; pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) - (sc_refill_cfgs_of s) sc_ptr\ - \ whileLoop_terminates (\_ s. the (head_insufficient sc_ptr s)) - (\_. non_overlapping_merge_refills sc_ptr) r s" - (is "\?P s; ?Q s\ \ _") - apply (rule_tac I="\_. ?P and ?Q" + (sc_refill_cfgs_of s) sc_ptr; + sc_refills_sc_at (\refills. refills \ []) sc_ptr s\ + \ whileLoop_terminates (\_ s. the (refill_head_insufficient sc_ptr s)) + (\_. merge_nonoverlapping_head_refill sc_ptr) r s" + (is "\?P s; ?Q s; ?R s\ \ _") + apply (rule_tac I="\_. ?P and ?Q and ?R" in whileLoop_terminates_inv[where R="head_insufficient_loop_measure sc_ptr"]) apply simp apply (intro hoare_vcg_conj_lift_pre_fix) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum_lower_bound - non_overlapping_merge_refills_refills_unat_sum) - apply (fastforce dest: head_insufficient_length_at_least_two) + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum_lower_bound + merge_nonoverlapping_head_refill_refills_unat_sum + merge_nonoverlapping_head_refill_nonempty_refills) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length) apply (wpsimp wp: update_sched_context_wp - simp: head_insufficient_loop_measure_def non_overlapping_merge_refills_def + simp: head_insufficient_loop_measure_def merge_nonoverlapping_head_refill_def refill_pop_head_def update_refill_hd_def) - apply (fastforce dest: head_insufficient_length_at_least_two - simp: vs_all_heap_simps obj_at_def) + apply (fastforce dest: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length + simp: vs_all_heap_simps obj_at_def sc_at_ppred_def) apply (clarsimp simp: head_insufficient_loop_measure_def) done lemma refills_unat_sum_MIN_BUDGET_implies_non_empty_refills: "pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) (sc_refill_cfgs_of s) sc_ptr - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) sc_ptr" - apply (auto simp: vs_all_heap_simps refills_unat_sum_def MIN_BUDGET_nonzero unat_eq_zero) + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s" + by (auto simp: vs_all_heap_simps refills_unat_sum_def MIN_BUDGET_nonzero unat_eq_zero + sc_at_ppred_def obj_at_def) + +lemma gets_the_refillHdInsufficient_corres: + "sc_ptr = scPtr \ + corres (=) + (valid_objs and pspace_aligned and pspace_distinct + and is_active_sc sc_ptr and sc_refills_sc_at (\refills. refills \ []) sc_ptr) + valid_objs' + (gets_the (refill_head_insufficient sc_ptr)) (gets_the (refillHdInsufficient scPtr))" + apply (simp add: refill_head_insufficient_def refillHdInsufficient_def + flip: get_refill_head_def getRefillHead_def) + apply (corres corres: getRefillHead_corres) + apply (rule corres_return_eq_same) + apply (clarsimp simp: refill_map_def minBudget_def MIN_BUDGET_def kernelWCETTicks_def) + apply wpsimp+ done lemma headInsufficientLoop_corres: - "sc_ptr = scPtr - \ corres dc (pspace_aligned and pspace_distinct and sc_at sc_ptr and is_active_sc sc_ptr - and valid_objs - and (\s. pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) - (sc_refill_cfgs_of s) sc_ptr) - and (\s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) \ unat max_time) - (sc_refill_cfgs_of s) sc_ptr) - and (\s. sc_refills_sc_at (\refills. refills \ []) sc_ptr s)) - (valid_objs' and valid_refills' sc_ptr) - (head_insufficient_loop sc_ptr) - (headInsufficientLoop scPtr)" + "sc_ptr = scPtr \ + corres dc + (pspace_aligned and pspace_distinct and sc_at sc_ptr and is_active_sc sc_ptr + and valid_objs + and (\s. pred_map (\cfg. unat MIN_BUDGET \ refills_unat_sum (scrc_refills cfg)) + (sc_refill_cfgs_of s) sc_ptr) + and (\s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) \ unat max_time) + (sc_refill_cfgs_of s) sc_ptr) + and (\s. sc_refills_sc_at (\refills. refills \ []) sc_ptr s)) + (valid_objs' and valid_refills' sc_ptr) + (head_insufficient_loop sc_ptr) (headInsufficientLoop scPtr)" apply (clarsimp simp: head_insufficient_loop_def headInsufficientLoop_def runReaderT_def) apply (rule_tac Q'="active_sc_at' scPtr" in corres_cross_add_guard) apply (fastforce dest: active_sc_at'_cross) apply (rule corres_whileLoop_abs; simp?) - apply (frule head_insufficient_equiv[where scPtr=scPtr]; simp?) - apply (fastforce intro: refills_unat_sum_MIN_BUDGET_implies_non_empty_refills) - apply (corresKsimp corres: nonOverlappingMergeRefills_corres) - apply (fastforce dest: head_insufficient_length_at_least_two) - apply (wpsimp wp: non_overlapping_merge_refills_refills_unat_sum_lower_bound - non_overlapping_merge_refills_refills_unat_sum - non_overlapping_merge_refills_nonempty_refills - simp: obj_at_kh_kheap_simps) - apply (frule (1) head_insufficient_length_at_least_two) - apply (clarsimp simp: sc_at_pred_n_def obj_at_def vs_all_heap_simps) - apply (wpsimp wp: nonOverlappingMergeRefills_valid_objs') - apply (clarsimp simp: active_sc_at'_def) + apply (rule_tac f="refill_head_insufficient scPtr" + and f'="refillHdInsufficient scPtr" + in gets_the_corres_relation[rotated]) + apply (corres corres: gets_the_refillHdInsufficient_corres) + apply fastforce + apply fastforce + apply fastforce + apply (wpsimp wp: no_ofail_head_insufficient) + apply (corres corres: mergeNonoverlappingHeadRefill_corres) + apply (fastforce dest!: no_ofailD[OF no_ofail_head_insufficient] + head_insufficient_length + simp: vs_all_heap_simps sc_at_ppred_def obj_at_def) + apply fastforce + apply (wpsimp wp: merge_nonoverlapping_head_refill_refills_unat_sum_lower_bound + merge_nonoverlapping_head_refill_refills_unat_sum + merge_nonoverlapping_head_refill_nonempty_refills) + apply (fastforce dest!: no_ofailD[OF no_ofail_head_insufficient] head_insufficient_length + simp: vs_all_heap_simps sc_at_ppred_def obj_at_def) + apply (wpsimp | strengthen valid_objs'_valid_refills' active_sc_at'_imp_is_active_sc')+ + apply (clarsimp simp: active_sc_at'_rewrite) apply (fastforce intro!: non_overlapping_merge_refills_terminates) done @@ -4192,90 +4025,87 @@ lemma getRefillFull_sp: apply (fastforce simp: obj_at'_def refillSize_def) done +(* FIXME RT: move to AInvs *) +lemma get_refill_tail_sp: + "\P\ + get_refill_tail sc_ptr + \\rv s. P s \ (\sc n. ko_at (Structures_A.SchedContext sc n) sc_ptr s \ rv = refill_tl sc)\" + apply (wpsimp wp: get_refill_tail_wp) + apply (clarsimp simp: obj_at_def) + done + +lemma getRefillTail_corres: + "sc_ptr = scPtr \ + corres (\rv rv'. refill_map rv' = rv) + (valid_objs and pspace_aligned and pspace_distinct + and is_active_sc sc_ptr and sc_refills_sc_at (\refills. refills \ []) sc_ptr) + valid_objs' + (get_refill_tail sc_ptr) (getRefillTail scPtr)" + apply (clarsimp simp: get_refill_tail_def getRefillTail_def read_refill_tail_def + readRefillTail_def getSchedContext_def[symmetric] + read_sched_context_get_sched_context + readSchedContext_def getObject_def[symmetric]) + apply (rule stronger_corres_guard_imp) + apply (rule corres_split[OF get_sc_corres]) + apply (rule corres_assert_assume_l) + apply clarsimp + apply (rule refills_tl_equal[symmetric]) + apply simp + apply simp + apply wpsimp + apply wpsimp + apply (clarsimp simp: sc_at_ppred_def obj_at_def is_sc_obj_def) + apply (fastforce intro: valid_objs_valid_sched_context_size) + apply clarsimp + apply (frule (4) active_sc_at'_cross_valid_objs) + by (fastforce dest: valid_objs'_valid_refills' + simp: active_sc_at'_def obj_at'_def is_active_sc'_def in_omonad valid_refills'_def) + lemma scheduleUsed_corres: "\sc_ptr = scPtr; new = refill_map new'\ \ corres dc (sc_at sc_ptr and is_active_sc2 sc_ptr and (\s. sc_refills_sc_at (\refills. 0 < length refills) sc_ptr s) - and pspace_aligned and pspace_distinct) - (valid_refills' scPtr) + and pspace_aligned and pspace_distinct and valid_objs) + valid_objs' (schedule_used sc_ptr new) (scheduleUsed scPtr new')" apply (clarsimp simp: schedule_used_def scheduleUsed_def get_refills_def bind_assoc) apply (rule_tac Q'="sc_at' scPtr" in corres_cross_add_guard) apply (fastforce intro: sc_at_cross) apply (rule_tac Q'="is_active_sc' scPtr" in corres_cross_add_guard) apply (fastforce intro: is_active_sc'_cross) - apply (rule corres_underlying_split[rotated 2, OF get_sched_context_sp get_sc_sp']) - apply (corresKsimp corres: get_sc_corres) - apply (rename_tac sc sc') + apply (rule corres_symb_exec_r[OF _ scActive_sp]; (solves wpsimp)?) apply (rule corres_symb_exec_r[rotated, OF assert_sp]; (solves wpsimp)?) apply wpsimp apply (clarsimp simp: is_active_sc'_def in_omonad obj_at_simps) + apply (rule corres_symb_exec_l[OF _ _ get_sched_context_sp, rotated]) + apply (wpsimp wp: get_sched_context_exs_valid) + apply (clarsimp simp: sc_at_ppred_def obj_at_def) + apply clarsimp + apply wpsimp + apply (clarsimp simp: sc_at_ppred_def obj_at_def) apply (rule corres_symb_exec_l[rotated, OF _ assert_sp]) apply wpsimp apply (fastforce simp: sc_at_pred_n_def obj_at_def) apply wpsimp apply (fastforce simp: sc_at_pred_n_def obj_at_def) - apply (rule_tac F="sc_refills sc \ [] \ sc_valid_refills' sc'" in corres_req) - apply (fastforce simp: valid_refills'_def opt_pred_def in_omonad obj_at'_def - split: option.splits) + apply (rule corres_underlying_split[rotated 2, OF get_refill_tail_sp getRefillTail_sp]) + apply (corres corres: getRefillTail_corres simp: is_active_sc_rewrite) + apply (rename_tac tail tail') apply (rule corres_if_split; (solves simp)?) apply (fastforce simp: can_merge_refill_def refill_map_def dest: refill_hd_relation refills_tl_equal) - apply (corresKsimp corres: updateRefillTl_corres - simp: refill_map_def) + apply (corres corres: updateRefillTl_corres simp: refill_map_def) + apply (fastforce intro!: valid_objs'_valid_refills') apply (rule corres_underlying_split[rotated 2, OF refill_full_sp getRefillFull_sp]) apply (corres corres: getRefillFull_corres) + apply (fastforce intro!: valid_objs'_valid_refills') apply (rule corres_if_split; (solves simp)?) apply (corres corres: refillAddTail_corres) - apply (clarsimp simp: refill_map_def obj_at_simps opt_map_red opt_pred_def) + apply (fastforce dest!: valid_objs'_valid_refills' + simp: obj_at_simps opt_map_red opt_pred_def valid_refills'_def) apply (corres corres: updateRefillTl_corres simp: refill_map_def) - done - -lemma head_time_buffer_simp: - "sc_at (cur_sc s) s - \ head_time_buffer usage s - = Some (sc_at_pred (\sc. r_amount (refill_hd sc) \ usage - \ r_time (refill_hd sc) < MAX_RELEASE_TIME) - (cur_sc s) s)" - unfolding head_time_buffer_def - apply (clarsimp simp: obind_def read_sched_context_def obj_at_def is_sc_obj sc_at_pred_n_def - ogets_def) - done - -lemma headTimeBuffer_simp: - "sc_at' (ksCurSc s) s - \ headTimeBuffer usage s - = Some (obj_at' (\sc :: sched_context. rAmount (refillHd sc) \ usage - \ rTime (refillHd sc) < maxReleaseTime) - (ksCurSc s) s)" - unfolding headTimeBuffer_def - apply (clarsimp simp: obind_def readSchedContext_def split: option.splits) - apply (frule no_ofailD[OF no_ofail_sc_at'_readObject]) - apply (fastforce simp: readObject_def obind_def omonad_defs split_def loadObject_default_def - readCurSc_def obj_at'_def ogets_def - split: option.splits if_split_asm) - done - -lemma head_time_buffer_equiv: - "\(s, s') \ state_relation; pspace_aligned s; pspace_distinct s; valid_objs s; - pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) (cur_sc s); - valid_refills' (ksCurSc s') s'; usage = usage'\ - \ head_time_buffer usage s = headTimeBuffer usage' s'" - apply (prop_tac "sc_at (cur_sc s) s") - apply (fastforce dest: valid_objs_valid_sched_context_size - simp: vs_all_heap_simps obj_at_kh_kheap_simps is_sc_obj_def) - apply (frule state_relation_pspace_relation) - apply (frule sc_at_cross; simp?) - apply (frule state_relation_sc_relation; simp?) - apply (subst head_time_buffer_simp; simp?) - apply (subst headTimeBuffer_simp) - apply (clarsimp simp: state_relation_def) - apply (frule refill_hd_relation) - apply (clarsimp simp: valid_refills'_def obj_at_simps state_relation_def) - apply (clarsimp simp: sc_ko_at_valid_objs_valid_sc' opt_map_def opt_pred_def obj_at_simps) - apply (clarsimp simp: obj_at_def sc_at_ppred_def obj_at'_def state_relation_def - maxReleaseTime_equiv opt_map_def vs_all_heap_simps refill_map_def) + apply (fastforce intro!: valid_objs'_valid_refills') done lemma readRefillSingle_sp: @@ -4296,47 +4126,45 @@ lemma refillSingle_sp: apply (fastforce dest: readRefillSingle_SomeD simp: obj_at'_def) done -lemma handleOverrunLoopBody_corres: - "r = r' \ +lemma chargeEntireHeadRefill_corres: + "\sc_ptr = scPtr; r = r'\ \ corres (=) - (\s. sc_at (cur_sc s) s \ is_active_sc2 (cur_sc s) s - \ pspace_aligned s \ pspace_distinct s \ valid_objs s - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) (cur_sc s)) - (\s'. valid_objs' s' \ valid_refills' (ksCurSc s') s') - (handle_overrun_loop_body r) (handleOverrunLoopBody r')" - apply (rule_tac Q'="\s'. sc_at' (ksCurSc s') s'" in corres_cross_add_guard) - apply (fastforce intro: sc_at_cross simp: state_relation_def) - apply (rule_tac Q'="\s'. is_active_sc' (ksCurSc s') s'" in corres_cross_add_guard) + (sc_at sc_ptr and is_active_sc2 sc_ptr + and pspace_aligned and pspace_distinct and valid_objs + and sc_refills_sc_at (\refills. refills \ []) sc_ptr) + valid_objs' + (charge_entire_head_refill sc_ptr r) (chargeEntireHeadRefill scPtr r')" + apply (rule_tac Q'="\s'. sc_at' scPtr s'" in corres_cross_add_guard) + apply (rule sc_at_cross, fastforce+) + apply (rule_tac Q'="\s'. is_active_sc' scPtr s'" in corres_cross_add_guard) apply (fastforce intro: is_active_sc'_cross simp: state_relation_def) - apply (clarsimp simp: handle_overrun_loop_body_def handleOverrunLoopBody_def) - apply (rule corres_underlying_split[rotated 2, OF gets_sp getCurSc_sp]) - apply (corresKsimp corres: getCurSc_corres) + apply (clarsimp simp: charge_entire_head_refill_def chargeEntireHeadRefill_def) + apply (rule corres_underlying_split[rotated 2, OF get_refill_head_sp getRefillHead_sp]) + apply (corres corres: getRefillHead_corres simp: is_active_sc_rewrite) + apply (rule corres_underlying_split[rotated 2, OF get_sched_context_sp get_sc_sp']) + apply (corres corres: get_sc_corres) + apply (rename_tac sc') apply (rule corres_underlying_split[rotated 2, OF refill_single_sp refillSingle_sp]) apply (corresKsimp corres: refillSingle_corres) - apply (fastforce simp: obj_at_simps valid_refills'_def opt_map_red opt_pred_def) - apply (rule corres_underlying_split[rotated 2, OF get_sched_context_sp get_sc_sp']) - apply (corresKsimp corres: get_sc_corres) - apply (rename_tac sc sc') - apply (rule_tac Q="\_ s. sc_refills sc \ []" - and Q'="\_ _. sc_valid_refills' sc'" - and r'=dc - in corres_underlying_split[rotated]) - apply corresKsimp - apply (fastforce dest: refill_hd_relation simp: refill_map_def) - apply (wpsimp simp: update_refill_hd_def - wp: update_sched_context_wp) - apply (clarsimp simp: vs_all_heap_simps obj_at_def) + apply (fastforce dest!: valid_objs'_valid_refills' + simp: obj_at_simps valid_refills'_def opt_map_red opt_pred_def) + apply (rule corres_underlying_split[rotated, where r'=dc]) + apply (rule corres_return_eq_same) + apply (clarsimp simp: refill_map_def) + apply wpsimp apply wpsimp - apply (clarsimp simp: valid_refills'_def obj_at_simps opt_map_red opt_pred_def) apply (rule corres_if_split; simp?) - apply (corresKsimp corres: updateRefillHd_corres) + apply (corres corres: updateRefillHd_corres) apply (fastforce simp: refill_map_def sc_relation_def) + apply fastforce + apply (fastforce dest!: valid_objs'_valid_refills' + simp: obj_at_simps valid_refills'_def opt_map_red opt_pred_def) apply (rule_tac F="1 < refillSize sc'" in corres_req) apply (frule_tac scp="scPtr" and P="\l. 1 < l" in length_sc_refills_cross) apply (clarsimp simp: state_relation_def) apply simp - apply (fastforce simp: refill_map_def sc_relation_def) - apply (clarsimp simp: opt_map_red opt_pred_def vs_all_heap_simps obj_at'_def Suc_lessI) + apply (fastforce dest!: valid_objs'_valid_refills') + apply (clarsimp simp: opt_map_red opt_pred_def obj_at'_def Suc_lessI obj_at_def sc_at_ppred_def) apply (rule corres_guard_imp) apply (rule corres_split[OF refillPopHead_corres]) apply (rule scheduleUsed_corres) @@ -4346,16 +4174,7 @@ lemma handleOverrunLoopBody_corres: apply wpsimp apply (clarsimp simp: is_active_sc2_def sc_at_ppred_def obj_at_def active_sc_def vs_all_heap_simps Suc_lessI in_omonad) - apply (clarsimp simp: obj_at_simps) - done - -lemma handle_overrun_loop_body_no_fail: - "no_fail (\s. sc_refills_sc_at (\refills. refills \ []) (cur_sc s) s) - (handle_overrun_loop_body usage)" - unfolding handle_overrun_loop_body_def - apply (wpsimp simp: refill_single_def refill_size_def get_refills_def update_refill_hd_def - wp: refill_pop_head_no_fail refill_pop_head_nonempty_refills) - apply (clarsimp simp: sc_at_pred_n_def obj_at_def Suc_lessI) + apply (fastforce dest!: valid_objs'_valid_refills') done lemma updateRefillIndex_is_active_sc'[wp]: @@ -4373,9 +4192,9 @@ lemma scheduleUsed_is_active_sc'[wp]: apply (clarsimp simp: obj_at_simps is_active_sc'_def opt_map_def opt_pred_def) done -lemma handleOverrunLoopBody_is_active_sc'[wp]: - "handleOverrunLoopBody usage \is_active_sc' sc_ptr\" - apply (clarsimp simp: handleOverrunLoopBody_def) +lemma chargeEntireHeadRefill_is_active_sc'[wp]: + "chargeEntireHeadRefill scPtr usage \is_active_sc' sc_ptr\" + apply (clarsimp simp: chargeEntireHeadRefill_def) apply (wpsimp simp: updateRefillHd_def refillSingle_def wp: updateSchedContext_wp) apply (clarsimp simp: obj_at_simps is_active_sc'_def opt_map_def opt_pred_def) @@ -4395,22 +4214,13 @@ lemma scheduleUsed_valid_refills'[wp]: apply (wpsimp wp: getRefillFull_wp) done -crunch handle_overrun_loop_body +crunch charge_entire_head_refill for valid_objs[wp]: valid_objs -crunch handleOverrunLoopBody +crunch chargeEntireHeadRefill for valid_objs'[wp]: valid_objs' (wp: crunch_wps simp: refillSingle_def) -lemma handleOverrunLoopBody_valid_refills'[wp]: - "handleOverrunLoopBody r' \valid_refills' scPtr\" - apply (clarsimp simp: handleOverrunLoopBody_def updateRefillHd_def refillSingle_def) - apply wpsimp - apply (frule readRefillSingle_SomeD) - apply (fastforce simp: valid_refills'_def opt_map_red opt_pred_def obj_at_simps - refillSingle_equiv[THEN arg_cong_Not, symmetric] refillSize_def) - done - lemma schedule_used_length_nonzero[wp]: "\\s. if sc_ptr' = sc_ptr then pred_map \ (scs_of s) sc_ptr @@ -4422,22 +4232,6 @@ lemma schedule_used_length_nonzero[wp]: apply (clarsimp simp: obj_at_def vs_all_heap_simps) done -lemma handle_overrun_loop_body_terminates_wf_helper: - "wf {((r' :: ticks, s' :: 'a state), (r, s)). unat r' < unat r}" - apply (insert wf_inv_image[where r="{(m, n). m < n}" - and f="\(r :: ticks, s :: 'a state). unat r"]) - apply (clarsimp simp: inv_image_def) - apply (prop_tac "wf {(m, n). m < n}") - apply (fastforce intro: wf) - apply (drule meta_mp, simp) - apply (prop_tac "{(x :: ticks \ 'a state, y :: ticks \ 'a state). - (case x of (r, s) \ unat r) - < (case y of (r, s) \ unat r)} - = {((r, s), r', s'). unat r < unat r'}") - apply fastforce - apply fastforce - done - \ \The following method can be used to try to solve Hoare triples in the following way. Note that the method works best when the precondition is not schematic. First, if the postcondition is not a conjunction, it will try to solve the goal with the @@ -4450,68 +4244,85 @@ method wps_conj_solves uses wp simp wps ; (solves \rule hoare_weaken_pre, (wpsimp wp: wp simp: simp | wps wps)+\)?) | (solves \rule hoare_weaken_pre, (wpsimp wp: wp simp: simp | wps wps)+\)? -lemma handle_overrun_loop_body_terminates: - "\sc_at (cur_sc s) s; +lemma charge_entire_head_refill_terminates: + "\sc_at sc_ptr s; pred_map (\cfg. \refill \ set (scrc_refills cfg). 0 < unat (r_amount refill)) - (sc_refill_cfgs_of s) (cur_sc s); - pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) (cur_sc s); + (sc_refill_cfgs_of s) sc_ptr; + sc_refills_sc_at (\refills. refills \ []) sc_ptr s; pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) - (sc_refill_cfgs_of s) (cur_sc s)\ - \ whileLoop_terminates (\r s. the (head_time_buffer r s)) handle_overrun_loop_body usage s" + (sc_refill_cfgs_of s) sc_ptr\ + \ whileLoop_terminates (\r s. the (head_refill_overrun sc_ptr r s)) + (charge_entire_head_refill sc_ptr) usage s" (is "\?P1 s; ?P2 s; ?P3 s; ?P4 s\ \ _") apply (rule_tac R="{((r' :: ticks, s' :: 'a state), (r, s)). unat r' < unat r}" - and I="\_ s'. ?P1 s' \ ?P2 s' \ ?P3 s' \ ?P4 s' \ cur_sc s' = cur_sc s" + and I="\_. ?P1 and ?P2 and ?P3 and ?P4" in whileLoop_terminates_inv) apply simp prefer 2 - apply (fastforce simp: handle_overrun_loop_body_terminates_wf_helper) + apply (prop_tac "{((r', s'), r, s). unat r' < unat r} = measure (\(r, s). unat r)") + apply (clarsimp simp: measure_def inv_image_def split_def) + apply (fastforce dest: sym elim: subst) apply (rename_tac r s') - apply (wps_conj_solves wp: handle_overrun_loop_body_non_zero_refills - handle_overrun_loop_body_refills_unat_sum_equals_budget) - apply (wpsimp simp: handle_overrun_loop_body_def) + apply (wps_conj_solves wp: charge_entire_head_refill_non_zero_refills + charge_entire_head_refill_refills_unat_sum_equals_budget) + apply (wpsimp simp: charge_entire_head_refill_def) apply (rename_tac sc n) + apply (frule_tac usage=r and s=s' in head_refill_overrun_true_imp_buffer) + apply (clarsimp simp: vs_all_heap_simps) apply (subst unat_sub) - apply (prop_tac "sc_at (cur_sc s') s'", simp) - apply (frule_tac usage=r and s=s' in head_time_buffer_simp) - apply (clarsimp simp: sc_at_ppred_def obj_at_def) - apply (clarsimp simp: obj_at_def vs_all_heap_simps) - apply (frule_tac x="refill_hd sc" in bspec, fastforce) - apply (prop_tac "0 < unat r") - apply (prop_tac "sc_at (cur_sc s') s'") - apply (clarsimp simp: obj_at_def is_sc_obj_def) - apply (frule_tac usage=r and s=s' in head_time_buffer_simp) - apply (clarsimp simp: sc_at_ppred_def obj_at_def) - apply (frule_tac x="refill_hd sc" in bspec, fastforce) - apply (fastforce simp: word_le_nat_alt) - apply fastforce + apply clarsimp + apply (drule_tac x="refill_hd sc" in bspec) + apply (fastforce simp: obj_at_def sc_at_ppred_def) + apply (fastforce simp: word_le_nat_alt) done -lemma handleOverrunLoop_corres: - "usage = usage' \ - corres (=) (\s. sc_at (cur_sc s) s \ is_active_sc2 (cur_sc s) s - \ pspace_aligned s \ pspace_distinct s - \ valid_objs s - \ pred_map (\cfg. scrc_refills cfg \ []) (sc_refill_cfgs_of s) (cur_sc s) - \ pred_map (\cfg. \refill\set (scrc_refills cfg). 0 < unat (r_amount refill)) - (sc_refill_cfgs_of s) (cur_sc s) - \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) - (sc_refill_cfgs_of s) (cur_sc s)) - (\s'. valid_objs' s' \ valid_refills' (ksCurSc s') s') - (handle_overrun_loop usage) - (handleOverrunLoop usage')" - apply (rule_tac Q'="\s'. sc_at' (ksCurSc s') s'" in corres_cross_add_guard) - apply (fastforce intro: sc_at_cross simp: state_relation_def) - apply (rule_tac Q'="\s'. is_active_sc' (ksCurSc s') s'" in corres_cross_add_guard) +lemma gets_the_headRefillOverrun_corres: + "\sc_ptr = scPtr; usage = usage'\ \ + corres (=) + (valid_objs and pspace_aligned and pspace_distinct + and is_active_sc sc_ptr and sc_refills_sc_at (\refills. refills \ []) sc_ptr) + valid_objs' + (gets_the (head_refill_overrun sc_ptr usage)) (gets_the (headRefillOverrun scPtr usage'))" + apply (simp add: head_refill_overrun_def headRefillOverrun_def + flip: get_refill_head_def getRefillHead_def) + apply (corres corres: getRefillHead_corres) + apply (rule corres_return_eq_same) + apply (clarsimp simp: refill_map_def maxReleaseTime_equiv) + apply wpsimp+ + done + +lemma handleOverrun_corres: + "\sc_ptr = scPtr; usage = usage'\ \ + corres (=) + (sc_at sc_ptr and is_active_sc2 sc_ptr + and pspace_aligned and pspace_distinct + and valid_objs + and sc_refills_sc_at (\refills. refills \ []) sc_ptr + and (\s. pred_map (\cfg. \refill\set (scrc_refills cfg). 0 < unat (r_amount refill)) + (sc_refill_cfgs_of s) sc_ptr) + and (\s. pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) + (sc_refill_cfgs_of s) sc_ptr)) + valid_objs' + (handle_overrun sc_ptr usage) (handleOverrun scPtr usage')" + apply (rule_tac Q'="\s'. sc_at' scPtr s'" in corres_cross_add_guard) + apply (rule sc_at_cross, fastforce+) + apply (rule_tac Q'="\s'. is_active_sc' scPtr s'" in corres_cross_add_guard) apply (fastforce intro: is_active_sc'_cross simp: state_relation_def) - apply (clarsimp simp: handle_overrun_loop_def handleOverrunLoop_def runReaderT_def) + apply (clarsimp simp: handle_overrun_def handleOverrun_def runReaderT_def) apply (rule corres_whileLoop_abs; simp?) - apply (frule_tac usage=r' in head_time_buffer_equiv; simp?) - apply fastforce - apply (corres corres: handleOverrunLoopBody_corres) - apply (wps_conj_solves wp: handle_overrun_loop_body_non_zero_refills - handle_overrun_loop_body_refills_unat_sum_equals_budget) - apply wps_conj_solves - apply (fastforce intro: handle_overrun_loop_body_terminates) + apply (rule_tac f="head_refill_overrun scPtr r'" + and f'="headRefillOverrun scPtr r'" and sr=state_relation + in gets_the_corres_relation[rotated]) + apply (corres corres: gets_the_headRefillOverrun_corres) + apply (fastforce simp: is_active_sc_rewrite) + apply fastforce + apply fastforce + apply (wpsimp wp: no_ofail_head_refill_overrun) + apply (corres corres: chargeEntireHeadRefill_corres) + apply (wpsimp wp: charge_entire_head_refill_non_zero_refills + charge_entire_head_refill_refills_unat_sum_equals_budget) + apply wpsimp + apply (fastforce intro: charge_entire_head_refill_terminates) done lemma get_refills_exs_valid[wp]: @@ -4522,10 +4333,6 @@ lemma get_refills_exs_valid[wp]: apply simp done -crunch handleOverrunLoop - for valid_refills'[wp]: "valid_refills' scPtr" - (wp: crunch_wps) - lemma update_refill_hd_nonempty_refills[wp]: "update_refill_hd sc_ptr f \\s. sc_refills_sc_at (\refills. refills \ []) sc_ptr' s\" unfolding update_refill_hd_def @@ -4534,17 +4341,17 @@ lemma update_refill_hd_nonempty_refills[wp]: done lemma refillBudgetCheck_corres: - "usage = usage' - \ corres dc ((\s. sc_at (cur_sc s) s \ is_active_sc (cur_sc s) s - \ valid_objs s - \ pspace_aligned s \ pspace_distinct s) - and (\s. \ round_robin (cur_sc s) s \ valid_refills (cur_sc s) s)) - (\s'. valid_objs' s' \ valid_refills' (ksCurSc s') s') - (refill_budget_check usage) - (refillBudgetCheck usage')" + "usage = usage' \ + corres dc + ((\s. sc_at (cur_sc s) s \ is_active_sc (cur_sc s) s + \ valid_objs s + \ pspace_aligned s \ pspace_distinct s) + and (\s. \ round_robin (cur_sc s) s \ valid_refills (cur_sc s) s)) + valid_objs' + (refill_budget_check usage) (refillBudgetCheck usage')" (is "_ \ corres _ (?P and _) _ _ _") apply (rule_tac Q'="\s'. sc_at' (ksCurSc s') s'" in corres_cross_add_guard) - apply (fastforce intro: sc_at_cross simp: state_relation_def) + apply (rule sc_at_cross, (fastforce simp: state_relation_def)+) apply (rule_tac Q'="\s'. is_active_sc' (ksCurSc s') s'" in corres_cross_add_guard) apply (fastforce intro!: is_active_sc'2_cross simp: state_relation_def) @@ -4568,42 +4375,51 @@ lemma refillBudgetCheck_corres: (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. MIN_BUDGET \ scrc_budget cfg) (sc_refill_cfgs_of s) sc_ptr - \ sc_ptr = cur_sc s + \ cur_sc s = sc_ptr \ (pred_map (\cfg. r_time (hd (scrc_refills cfg)) < MAX_RELEASE_TIME) - (sc_refill_cfgs_of s) (cur_sc s) + (sc_refill_cfgs_of s) sc_ptr \ pred_map (\cfg. usage' < r_amount (hd (scrc_refills cfg))) - (sc_refill_cfgs_of s) (cur_sc s)) - \ pred_map (\cfg. scrc_refills cfg \ []) - (sc_refill_cfgs_of s) sc_ptr" - and Q'="\_ s'. valid_objs' s' \ valid_refills' scPtr s' \ active_sc_at' scPtr s' + (sc_refill_cfgs_of s) sc_ptr) + \ sc_refills_sc_at (\refills. refills \ []) sc_ptr s" + and Q'="\_ s'. valid_objs' s' \ active_sc_at' scPtr s' \ scPtr = ksCurSc s'" in corres_underlying_split) - apply (corresKsimp corres: handleOverrunLoop_corres) - apply (fastforce intro: valid_refills_refills_unat_sum_equals_budget - simp: vs_all_heap_simps cfg_valid_refills_def round_robin_def - sp_valid_refills_def is_active_sc_rewrite[symmetric]) - apply (find_goal \match conclusion in "\P\ handle_overrun_loop _ \Q\" for P Q \ -\) + apply (corres corres: handleOverrun_corres) + apply (fastforce intro: valid_refills_refills_unat_sum_equals_budget + simp: vs_all_heap_simps cfg_valid_refills_def round_robin_def + sp_valid_refills_def is_active_sc_rewrite[symmetric] + sc_at_ppred_def obj_at_def) + apply fastforce + apply (find_goal \match conclusion in "\P\ handle_overrun _ usage' \Q\" for P Q \ -\) apply (clarsimp simp: pred_conj_def) - apply (intro hoare_vcg_conj_lift_pre_fix; (solves handle_overrun_loop_simple)?) - apply wps_conj_solves - apply (wpsimp wp: handle_overrun_loop_refills_unat_sum_equals_budget) - apply (fastforce intro: valid_refills_refills_unat_sum_equals_budget - simp: vs_all_heap_simps cfg_valid_refills_def round_robin_def - sp_valid_refills_def) - apply (clarsimp simp: handle_overrun_loop_def) - apply (wpsimp wp: valid_whileLoop[where I="\_ s. pred_map \ (scs_of s) (cur_sc s)"]) - apply (fastforce simp: head_time_buffer_true_imp_unat_buffer vs_all_heap_simps - word_less_nat_alt word_le_nat_alt) - apply (clarsimp simp: vs_all_heap_simps) - apply (find_goal \match conclusion in "\P\ handleOverrunLoop _ \Q\" for P Q \ -\) - apply (wpsimp wp: handleOverrunLoop_valid_objs') - apply (clarsimp simp: active_sc_at'_def obj_at_simps) - - apply (clarsimp simp: get_refills_def) - apply (rule corres_underlying_split[rotated 2, OF get_sched_context_sp get_sc_sp']) - apply (corresKsimp corres: get_sc_corres + apply (intro hoare_vcg_conj_lift_pre_fix; (solves handle_overrun_simple)?) + apply wps_conj_solves + apply (wpsimp wp: handle_overrun_refills_unat_sum_equals_budget) + apply (fastforce intro: valid_refills_refills_unat_sum_equals_budget + simp: vs_all_heap_simps cfg_valid_refills_def round_robin_def + sp_valid_refills_def sc_at_ppred_def obj_at_def) + apply (clarsimp simp: handle_overrun_def) + apply (rule hoare_weaken_pre) + apply (rule_tac I="\_ s. sc_refills_sc_at (\refills. refills \ []) scPtr s" + in valid_whileLoop) + apply assumption + apply (wpsimp wp: charge_entire_head_refill_nonzero_refills) + apply (frule_tac usage=r and sc_ptr=scPtr in head_refill_overrun_true_imp_unat_buffer) + apply (fastforce simp: head_refill_overrun_true_imp_unat_buffer vs_all_heap_simps + word_less_nat_alt word_le_nat_alt) + apply (fastforce simp: vs_all_heap_simps valid_refills_def round_robin_def + sc_at_ppred_def obj_at_def) + apply (find_goal \match conclusion in "\P\ handleOverrun _ usage' \Q\" for P Q \ -\) + apply (wpsimp wp: handleOverrun_valid_objs') + apply (clarsimp simp: active_sc_at'_def obj_at_simps) + apply (wpsimp wp: handle_overrun_nonzero_refills) + apply (fastforce simp: vs_all_heap_simps valid_refills_def round_robin_def + sc_at_ppred_def obj_at_def) + + apply (rule corres_underlying_split[rotated 2, OF get_refill_head_sp getRefillHead_sp]) + apply (corresKsimp corres: getRefillHead_corres simp: state_relation_def active_sc_at'_def obj_at_simps) - apply (rename_tac sc sc') + apply (rename_tac head head') apply (rule_tac Q="\_ s. ?P s \ pred_map (\cfg. refills_unat_sum (scrc_refills cfg) = unat (scrc_budget cfg)) @@ -4611,22 +4427,22 @@ lemma refillBudgetCheck_corres: \ pred_map (\cfg. MIN_BUDGET \ scrc_budget cfg) (sc_refill_cfgs_of s) scPtr \ scPtr = cur_sc s" - and Q'="\_ s'. valid_objs' s' \ valid_refills' scPtr s' \ active_sc_at' scPtr s' - \ scPtr = ksCurSc s'" + and Q'="\_ s'. valid_objs' s' \ active_sc_at' scPtr s' \ scPtr = ksCurSc s'" and r'=dc in corres_underlying_split[rotated]) apply (corresKsimp corres: headInsufficientLoop_corres) using MIN_BUDGET_pos - apply (fastforce simp: vs_all_heap_simps word_le_nat_alt sc_at_pred_n_def obj_at_def - refills_unat_sum_def unat_eq_0) + apply (fastforce intro!: valid_objs'_valid_refills' + simp: vs_all_heap_simps word_le_nat_alt sc_at_pred_n_def obj_at_def + refills_unat_sum_def unat_eq_0 active_sc_at'_def obj_at'_def + is_active_sc'_def in_omonad) apply (intro hoare_vcg_conj_lift_pre_fix; (solves wpsimp)?) - apply schedule_used_simple - apply (clarsimp simp: obj_at_def is_sc_obj_def) - apply schedule_used_simple + apply ((wpsimp | wps)+)[1] + apply ((wpsimp | wps)+)[1] apply (wpsimp wp: schedule_used_refills_unat_sum update_sched_context_wp simp: update_refill_hd_def) apply (clarsimp simp: obj_at_def vs_all_heap_simps refills_unat_sum_cons - refills_unat_sum_append) + refills_unat_sum_append sc_at_ppred_def) apply (subst unat_sub) apply fastforce apply (clarsimp simp: word_less_nat_alt) @@ -4635,29 +4451,13 @@ lemma refillBudgetCheck_corres: apply (case_tac "sc_refills sc"; clarsimp simp: refills_unat_sum_cons) apply schedule_used_simple apply (wpsimp wp: updateRefillHd_valid_objs') - - apply (rule_tac F="refill_hd sc = refill_map (refillHd sc')" in corres_req) - apply clarsimp - apply (rule refill_hd_relation) - apply fastforce - apply (clarsimp simp: vs_all_heap_simps obj_at_def) - apply (fastforce simp: valid_refills'_def obj_at_simps opt_map_red opt_pred_def) - apply (clarsimp simp: maxReleaseTime_equiv) apply (simp add: when_def split del: if_split) apply (rule corres_if_split; (solves simp)?) apply (clarsimp simp: refill_map_def) - apply (rule corres_symb_exec_l[rotated 2, OF get_sched_context_sp]) - apply wpsimp - apply (clarsimp simp: obj_at_def) - apply (find_goal \match conclusion in "\P\ f \\Q\" for P f Q \ -\) - apply (wpsimp wp: get_sched_context_exs_valid) - apply (clarsimp simp: obj_at_def) - apply simp - apply (rename_tac new_sc) - apply (rule_tac F="new_sc=sc" in corres_req) - apply (clarsimp simp: obj_at_def) - apply (subst update_refill_hd_comp) + apply (rule corres_underlying_split[rotated 2, OF get_sched_context_sp get_sc_sp']) + apply (corres corres: get_sc_corres) + apply (clarsimp simp: active_sc_at'_def) apply (clarsimp simp: bind_assoc) apply (rule corres_guard_imp) apply (rule corres_split[OF updateRefillHd_corres]) @@ -4670,12 +4470,11 @@ lemma refillBudgetCheck_corres: apply simp apply (clarsimp simp: refill_map_def sc_relation_def) apply wpsimp - apply wpsimp + apply (wpsimp wp: updateRefillHd_valid_objs') apply wpsimp - apply wpsimp - apply (clarsimp simp: is_active_sc_rewrite sc_at_pred_n_def obj_at_def vs_all_heap_simps - is_active_sc2_def in_omonad active_sc_def) - apply (fastforce simp: active_sc_at'_def is_active_sc'_def obj_at_simps opt_map_red) + apply (wpsimp wp: updateRefillHd_valid_objs') + apply (clarsimp simp: vs_all_heap_simps is_active_sc2_def in_omonad active_sc_def) + apply (fastforce dest: valid_objs'_valid_refills' simp: active_sc_at'_rewrite) done (* schedule_corres *) diff --git a/proof/refine/RISCV64/Syscall_R.thy b/proof/refine/RISCV64/Syscall_R.thy index 4e125b6dc2..55f52effb3 100644 --- a/proof/refine/RISCV64/Syscall_R.thy +++ b/proof/refine/RISCV64/Syscall_R.thy @@ -1030,7 +1030,7 @@ crunch refillNew, refillUpdate, commitTime for pred_tcb_at''[wp]: "\s. Q (pred_tcb_at' proj P tcbPtr s)" and ksCurThread[wp]: "\s. P (ksCurThread s)" and ex_nonz_cap_to'[wp]: "ex_nonz_cap_to' ptr" - (simp: crunch_simps wp: crunch_wps) + (simp: crunch_simps wp: crunch_wps hoare_vcg_all_lift) lemma scSBadge_update_invs'[wp]: "updateSchedContext scPtr (scBadge_update f) \invs'\" @@ -1917,7 +1917,7 @@ crunch refillResetRR, refillBudgetCheck and ex_nonz_cap_to'[wp]: "ex_nonz_cap_to' p" and typ_at'[wp]: "\s. Q (typ_at' P p s)" and sc_at'_n[wp]: "\s. Q (sc_at'_n n p s)" - (wp: crunch_wps) + (wp: crunch_wps hoare_vcg_all_lift simp: crunch_simps) crunch chargeBudget for typ_at'[wp]: "\s. Q (typ_at' P p s)" @@ -2412,7 +2412,7 @@ crunch check_budget crunch checkBudgetRestart for ksInterruptState[wp]: "\s. P (ksInterruptState s)" and ksCurThread[wp]: "\s. P (ksCurThread s)" - (wp: crunch_wps simp: crunch_simps) + (wp: crunch_wps hoare_vcg_all_lift simp: crunch_simps) lemma getCurrentTime_invs'[wp]: "doMachineOp getCurrentTime \invs'\" diff --git a/spec/abstract/KHeap_A.thy b/spec/abstract/KHeap_A.thy index 69f72edcf3..6185272e65 100644 --- a/spec/abstract/KHeap_A.thy +++ b/spec/abstract/KHeap_A.thy @@ -378,6 +378,16 @@ definition read_refill_head :: "obj_ref \ (refill, 'z::state_ext) r_ definition get_refill_head :: "obj_ref \ (refill, 'z::state_ext) s_monad" where "get_refill_head sc_ptr \ gets_the (read_refill_head sc_ptr)" +definition read_refill_tail :: "obj_ref \ (refill, 'z::state_ext) r_monad" where + "read_refill_tail sc_ptr \ do { + sc \ read_sched_context sc_ptr; + oassert (sc_refills sc \ []); + oreturn (refill_tl sc) + }" + +definition get_refill_tail :: "obj_ref \ (refill, 'z::state_ext) s_monad" where + "get_refill_tail sc_ptr \ gets_the (read_refill_tail sc_ptr)" + definition refill_capacity :: "time \ refill \ time" where diff --git a/spec/abstract/SchedContext_A.thy b/spec/abstract/SchedContext_A.thy index 2e09b413d9..37bef80595 100644 --- a/spec/abstract/SchedContext_A.thy +++ b/spec/abstract/SchedContext_A.thy @@ -135,19 +135,18 @@ where maybe_add_empty_tail sc_ptr od" -definition - schedule_used :: "obj_ref \ refill \ (unit, 'z::state_ext) s_monad" -where +definition schedule_used :: "obj_ref \ refill \ (unit, 'z::state_ext) s_monad" where "schedule_used sc_ptr new \ do refills \ get_refills sc_ptr; assert (refills \ []); - if can_merge_refill (last refills) new - then update_refill_tl sc_ptr (r_amount_update (\amount. amount + r_amount new)) + tail \ get_refill_tail sc_ptr; + if can_merge_refill tail new + then update_refill_tl sc_ptr (\r. r\r_amount := r_amount tail + r_amount new\) else do full \ refill_full sc_ptr; if \ full then refill_add_tail sc_ptr new - else do update_refill_tl sc_ptr (\r. r\r_time := r_time new - r_amount r\); - update_refill_tl sc_ptr (\r. r\r_amount := r_amount r + r_amount new\) + else do update_refill_tl sc_ptr (\r. r\r_time := r_time new - r_amount tail\); + update_refill_tl sc_ptr (\r. r\r_amount := r_amount tail + r_amount new\) od od od" @@ -162,81 +161,66 @@ where od" definition - head_insufficient :: "obj_ref \ (bool, 'z::state_ext) r_monad" + refill_head_insufficient :: "obj_ref \ (bool, 'z::state_ext) r_monad" where - "head_insufficient sc_ptr \ do { - sc \ read_sched_context sc_ptr; - oreturn (r_amount (hd (sc_refills sc)) < MIN_BUDGET) + "refill_head_insufficient sc_ptr \ do { + head \ read_refill_head sc_ptr; + oreturn (r_amount head < MIN_BUDGET) }" -definition - non_overlapping_merge_refills :: "obj_ref \ (unit, 'z::state_ext) s_monad" -where - "non_overlapping_merge_refills sc_ptr \ do +definition merge_nonoverlapping_head_refill :: "obj_ref \ (unit, 'z::state_ext) s_monad" where + "merge_nonoverlapping_head_refill sc_ptr \ do old_head \ refill_pop_head sc_ptr; - update_refill_hd sc_ptr (r_time_update (\t. t - r_amount old_head) o (r_amount_update (\m. m + r_amount old_head))) + update_refill_hd sc_ptr (\head. head\r_amount := r_amount head + r_amount old_head\); + update_refill_hd sc_ptr (\head. head\r_time := r_time head - r_amount old_head\) od" -definition - head_insufficient_loop :: "obj_ref \ (unit, 'z::state_ext) s_monad" -where +definition head_insufficient_loop :: "obj_ref \ (unit, 'z::state_ext) s_monad" where "head_insufficient_loop sc_ptr - \ whileLoop (\_ s. the ((head_insufficient sc_ptr) s)) - (\_. non_overlapping_merge_refills sc_ptr) ()" + \ whileLoop (\_ s. the ((refill_head_insufficient sc_ptr) s)) + (\_. merge_nonoverlapping_head_refill sc_ptr) ()" -definition - head_time_buffer :: "ticks \ (bool, 'z::state_ext) r_monad" -where - "head_time_buffer usage \ do { - sc_ptr \ ogets cur_sc; - sc \ read_sched_context sc_ptr; - oreturn (r_amount (refill_hd sc) \ usage \ r_time (refill_hd sc) < MAX_RELEASE_TIME) +definition head_refill_overrun :: "obj_ref \ ticks \ (bool, 'z::state_ext) r_monad" where + "head_refill_overrun sc_ptr usage \ do { + head \ read_refill_head sc_ptr; + oreturn (r_amount head \ usage \ r_time head < MAX_RELEASE_TIME) }" -definition - handle_overrun_loop_body :: "ticks \ (ticks, 'z::state_ext) s_monad" -where - "handle_overrun_loop_body usage \ do - sc_ptr \ gets cur_sc; - single \ refill_single sc_ptr; +definition charge_entire_head_refill :: "obj_ref \ ticks \ (ticks, 'z::state_ext) s_monad" where + "charge_entire_head_refill sc_ptr usage \ do + head \ get_refill_head sc_ptr; sc \ get_sched_context sc_ptr; - - usage' \ return (usage - r_amount (refill_hd sc)); - + single \ refill_single sc_ptr; if single - then update_refill_hd sc_ptr (r_time_update (\t. t + sc_period sc)) + then update_refill_hd sc_ptr (\r. r\r_time := r_time head + sc_period sc\) else do old_head \ refill_pop_head sc_ptr; schedule_used sc_ptr (old_head\r_time := r_time old_head + sc_period sc\) od; - return usage' + return (usage - r_amount head) od" -definition - handle_overrun_loop :: "ticks \ (ticks, 'z::state_ext) s_monad" -where - "handle_overrun_loop usage - \ whileLoop (\usage s. the (head_time_buffer usage s)) - (\usage. handle_overrun_loop_body usage) usage" +definition handle_overrun :: "obj_ref \ ticks \ (ticks, 'z::state_ext) s_monad" where + "handle_overrun sc_ptr usage + \ whileLoop (\usage s. the (head_refill_overrun sc_ptr usage s)) + (\usage. charge_entire_head_refill sc_ptr usage) + usage" -definition - refill_budget_check :: "ticks \ (unit, 'z::state_ext) s_monad" -where +definition refill_budget_check :: "ticks \ (unit, 'z::state_ext) s_monad" where "refill_budget_check usage \ do sc_ptr \ gets cur_sc; robin \ is_round_robin sc_ptr; assert (\robin); - usage' \ handle_overrun_loop usage; + usage' \ handle_overrun sc_ptr usage; - refills \ get_refills sc_ptr; - new_head_amount \ return (r_amount (hd refills)); + head \ get_refill_head sc_ptr; - when (usage' > 0 \ r_time (hd refills) < MAX_RELEASE_TIME) $ do + when (usage' > 0 \ r_time head < MAX_RELEASE_TIME) $ do sc \ get_sched_context sc_ptr; - period \ return (sc_period sc); - used \ return \r_time = r_time (hd (sc_refills sc)) + period, r_amount = usage'\; - update_refill_hd sc_ptr (r_time_update (\t. t + usage') o (r_amount_update (\m. m - usage'))); + used \ return \r_time = r_time head + sc_period sc, r_amount = usage'\; + update_refill_hd sc_ptr (\r. r\r_amount := r_amount head - usage'\); + update_refill_hd sc_ptr (\r. r\r_time := r_time head + usage'\); schedule_used sc_ptr used od; diff --git a/spec/haskell/src/SEL4/Object/SchedContext.lhs b/spec/haskell/src/SEL4/Object/SchedContext.lhs index 776736ccc8..21e0309df6 100644 --- a/spec/haskell/src/SEL4/Object/SchedContext.lhs +++ b/spec/haskell/src/SEL4/Object/SchedContext.lhs @@ -32,7 +32,7 @@ This module uses the C preprocessor to select a target architecture. > schedContextUnbindReply, schedContextSetInactive, unbindFromSC, > schedContextCancelYieldTo, refillAbsoluteMax, schedContextUpdateConsumed, > scReleased, setConsumed, refillResetRR, preemptionPoint, refillHdInsufficient, -> nonOverlappingMergeRefills, headInsufficientLoop, maxReleaseTime, readScActive, scActive +> mergeNonoverlappingHeadRefill, headInsufficientLoop, maxReleaseTime, readScActive, scActive > ) where \begin{impdetails} @@ -304,17 +304,18 @@ This module uses the C preprocessor to select a target architecture. > scheduleUsed :: PPtr SchedContext -> Refill -> Kernel () > scheduleUsed scPtr new = do -> sc <- getSchedContext scPtr -> assert (scRefillMax sc > 0) "scPtr should be active" -> if rTime (refillTl sc) + rAmount (refillTl sc) >= rTime new -> then updateRefillTl scPtr (\last -> last { rAmount = rAmount last + rAmount new }) +> active <- scActive scPtr +> assert active "scPtr should be active" +> tail <- getRefillTail scPtr +> if rTime tail + rAmount tail >= rTime new +> then updateRefillTl scPtr (\last -> last { rAmount = rAmount tail + rAmount new }) > else do > full <- getRefillFull scPtr > if not full > then refillAddTail scPtr new > else do -> updateRefillTl scPtr (\last -> last { rTime = rTime new - rAmount last}) -> updateRefillTl scPtr (\last -> last { rAmount = rAmount last + rAmount new}) +> updateRefillTl scPtr (\last -> last { rTime = rTime new - rAmount tail}) +> updateRefillTl scPtr (\last -> last { rAmount = rAmount tail + rAmount new}) > refillResetRR :: PPtr SchedContext -> Kernel () > refillResetRR scPtr = do @@ -323,19 +324,17 @@ This module uses the C preprocessor to select a target architecture. > refillHdInsufficient :: PPtr SchedContext -> KernelR Bool > refillHdInsufficient scPtr = do -> sc <- readSchedContext scPtr -> return $ rAmount (refillHd sc) < minBudget +> head <- readRefillHead scPtr +> return $ rAmount head < minBudget -> nonOverlappingMergeRefills :: PPtr SchedContext -> Kernel () -> nonOverlappingMergeRefills scPtr = do -> size <- getRefillSize scPtr -> assert (1 < size) "if head is insufficient, there must be at least 2 refills" +> mergeNonoverlappingHeadRefill :: PPtr SchedContext -> Kernel () +> mergeNonoverlappingHeadRefill scPtr = do > old_head <- refillPopHead scPtr > updateRefillHd scPtr $ \head -> head { rAmount = rAmount head + rAmount old_head } > updateRefillHd scPtr $ \head -> head { rTime = rTime head - rAmount old_head} > headInsufficientLoop :: PPtr SchedContext -> Kernel () -> headInsufficientLoop scPtr = whileLoop (const (fromJust . runReaderT (refillHdInsufficient scPtr))) (const (nonOverlappingMergeRefills scPtr)) () +> headInsufficientLoop scPtr = whileLoop (const (fromJust . runReaderT (refillHdInsufficient scPtr))) (const (mergeNonoverlappingHeadRefill scPtr)) () > mergeOverlappingRefills :: PPtr SchedContext -> Kernel () > mergeOverlappingRefills scPtr = do @@ -362,32 +361,27 @@ This module uses the C preprocessor to select a target architecture. > maxReleaseTime :: Word64 > maxReleaseTime = (maxBound :: Word64) - 5 * usToTicks maxPeriodUs -> headTimeBuffer :: Ticks -> KernelR Bool -> headTimeBuffer usage = do -> scPtr <- readCurSc -> sc <- readSchedContext scPtr -> return $ rAmount (refillHd sc) <= usage && rTime (refillHd sc) < maxReleaseTime +> headRefillOverrun :: PPtr SchedContext -> Ticks -> KernelR Bool +> headRefillOverrun scPtr usage = do +> head <- readRefillHead scPtr +> return $ rAmount head <= usage && rTime head < maxReleaseTime -> handleOverrunLoopBody :: Ticks -> Kernel Ticks -> handleOverrunLoopBody usage = do -> scPtr <- getCurSc -> single <- refillSingle scPtr +> chargeEntireHeadRefill :: PPtr SchedContext -> Ticks -> Kernel Ticks +> chargeEntireHeadRefill scPtr usage = do +> head <- getRefillHead scPtr > sc <- getSchedContext scPtr - -> usage' <- return (usage - rAmount (refillHd sc)) - +> single <- refillSingle scPtr > if single -> then updateRefillHd scPtr $ \r -> r { rTime = rTime r + scPeriod sc} +> then updateRefillHd scPtr $ \r -> r { rTime = rTime head + scPeriod sc} > else do > old_head <- refillPopHead scPtr > let new = old_head { rTime = rTime old_head + scPeriod sc} > scheduleUsed scPtr new +> return (usage - rAmount head) -> return usage' - -> handleOverrunLoop :: Ticks -> Kernel Ticks -> handleOverrunLoop usage = -> whileLoop (\usage -> fromJust . runReaderT (headTimeBuffer usage)) (\usage -> handleOverrunLoopBody usage) usage +> handleOverrun :: PPtr SchedContext -> Ticks -> Kernel Ticks +> handleOverrun scPtr usage = +> whileLoop (\usage -> fromJust . runReaderT (headRefillOverrun scPtr usage)) (\usage -> chargeEntireHeadRefill scPtr usage) usage > refillBudgetCheck :: Ticks -> Kernel () > refillBudgetCheck usage = do @@ -399,15 +393,16 @@ This module uses the C preprocessor to select a target architecture. > roundRobin <- isRoundRobin scPtr > assert (not roundRobin) "the current sc should not be round robin when this function is called" -> usage' <- handleOverrunLoop usage +> usage' <- handleOverrun scPtr usage -> sc <- getSchedContext scPtr +> head <- getRefillHead scPtr -> when (usage' > 0 && rTime (refillHd sc) < maxReleaseTime) $ do -> used <- return Refill { rTime = rTime (refillHd sc) + (scPeriod sc), +> when (usage' > 0 && rTime head < maxReleaseTime) $ do +> sc <- getSchedContext scPtr +> used <- return Refill { rTime = rTime head + (scPeriod sc), > rAmount = usage'} -> updateRefillHd scPtr $ \r -> r { rAmount = rAmount r - usage' } -> updateRefillHd scPtr $ \r -> r { rTime = rTime r + usage' } +> updateRefillHd scPtr $ \r -> r { rAmount = rAmount head - usage' } +> updateRefillHd scPtr $ \r -> r { rTime = rTime head + usage' } > scheduleUsed scPtr used Ensure that the rAmount of the head refill is at least minBudget