1(*
2 * Copyright 2014, General Dynamics C4 Systems
3 *
4 * This software may be distributed and modified according to the terms of
5 * the GNU General Public License version 2. Note that NO WARRANTY is provided.
6 * See "LICENSE_GPLv2.txt" for details.
7 *
8 * @TAG(GD_GPL)
9 *)
10
11theory Recycle_C
12imports Delete_C Retype_C
13begin
14
15context kernel_m
16begin
17
18lemma isArchPageCap_ArchObjectCap:
19  "isArchPageCap (ArchObjectCap acap)
20       = isPageCap acap"
21  by (simp add: isArchPageCap_def isPageCap_def)
22
23definition
24  "replicateHider \<equiv> replicate"
25
26lemma collapse_foldl_replicate:
27  "replicate (length xs) v = xs \<Longrightarrow>
28   foldl (@) [] (map (\<lambda>_. xs) ys)
29        = replicateHider (length xs * length ys) v"
30  apply (induct ys rule: rev_induct)
31   apply (simp add: replicateHider_def)
32  apply (simp add: replicateHider_def)
33  apply (subst add.commute, simp add: replicate_add)
34  done
35
36lemma coerce_memset_to_heap_update_user_data:
37  "heap_update_list x (replicateHider 4096 0)
38      = heap_update (Ptr x :: user_data_C ptr)
39             (user_data_C (FCP (\<lambda>_. 0)))"
40  apply (intro ext, simp add: heap_update_def)
41  apply (rule_tac f="\<lambda>xs. heap_update_list x xs a b" for a b in arg_cong)
42  apply (simp add: to_bytes_def size_of_def typ_info_simps user_data_C_tag_def)
43  apply (simp add: ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td align_of_def padup_def
44                   final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def)
45  apply (simp add: typ_info_simps
46                   user_context_C_tag_def thread_state_C_tag_def seL4_Fault_C_tag_def
47                   lookup_fault_C_tag_def update_ti_t_ptr_0s
48                   ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td
49                   ti_typ_combine_empty_ti ti_typ_combine_td
50                   align_of_def padup_def
51                   final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def
52                   align_td_array' size_td_array)
53  apply (simp add: typ_info_array')
54  apply (subst access_ti_list_array)
55     apply simp
56    apply simp
57   apply (simp add: fcp_beta typ_info_word typ_info_ptr word_rsplit_0)
58   apply fastforce
59  apply (simp add: collapse_foldl_replicate)
60  done
61
62lemma clift_foldl_hrs_mem_update:
63  "\<lbrakk> \<forall>x \<in> set xs. hrs_htd s \<Turnstile>\<^sub>t f x;
64     \<And>x s. hrs_htd s \<Turnstile>\<^sub>t f x \<Longrightarrow> clift (hrs_mem_update (heap_update (f x) v) s)
65                                  = g (clift s :: ('a :: c_type) ptr \<rightharpoonup> 'a) x \<rbrakk>
66   \<Longrightarrow>
67   clift (hrs_mem_update (\<lambda>s. foldl (\<lambda>s x. heap_update (f x) v s) s xs) s)
68       = foldl g (clift s :: 'a ptr \<rightharpoonup> 'a) xs"
69  using [[hypsubst_thin]]
70  apply (cases s, clarsimp)
71  apply (induct xs arbitrary: a b)
72   apply (simp add: hrs_mem_update_def)
73  apply (clarsimp simp add: hrs_mem_update_def split_def hrs_htd_def)
74  done
75
76lemma map_to_user_data_aligned:
77  "\<lbrakk> map_to_user_data (ksPSpace s) x = Some y; pspace_aligned' s \<rbrakk>
78        \<Longrightarrow> is_aligned x pageBits"
79  apply (clarsimp simp: map_comp_eq projectKOs split: option.split_asm)
80  apply (drule(1) pspace_alignedD')
81  apply (simp add: objBits_simps)
82  done
83
84lemma help_force_intvl_range_conv:
85  "\<lbrakk> is_aligned (p::word32) n; v = 2 ^ n; n < word_bits \<rbrakk>
86    \<Longrightarrow> {p ..+ v} = {p .. p + 2 ^ n - 1}"
87  by (simp add: intvl_range_conv word_bits_def)
88
89lemma cmap_relation_If_upd:
90  "\<lbrakk> cmap_relation f g ptrfun rel; rel v v'; ptrfun ` S = S'; inj ptrfun \<rbrakk>
91   \<Longrightarrow> cmap_relation (\<lambda>x. if x \<in> S then Some v else f x)
92                     (\<lambda>y. if y \<in> S' then Some v' else g y)
93        ptrfun rel"
94  apply (simp add: cmap_relation_def dom_If_Some)
95  apply (rule context_conjI)
96   apply blast
97  apply clarsimp
98  apply (case_tac "x \<in> S")
99   apply simp
100  apply clarsimp
101  apply (subst if_not_P)
102   apply (clarsimp simp: inj_eq)
103  apply (drule bspec, erule domI)
104  apply simp
105  done
106
107lemma length_replicateHider [simp]:
108  "length (replicateHider n x) = n"
109  by (simp add: replicateHider_def)
110
111lemma coerce_heap_update_to_heap_updates':
112  "n = chunk * m \<Longrightarrow>
113  heap_update_list x (replicateHider n 0)
114  = (\<lambda>s. foldl (\<lambda>s x. heap_update_list x (replicateHider chunk 0) s) s
115    (map (\<lambda>n. x + (of_nat n * of_nat chunk)) [0 ..< m]))"
116  using [[hypsubst_thin]]
117  apply clarsimp
118  apply (induct m arbitrary: x)
119   apply (rule ext, simp)
120   apply (simp add: replicateHider_def)
121  apply (rule ext)
122  apply (simp only: map_upt_unfold map_Suc_upt[symmetric])
123  apply (simp add: replicate_add[folded replicateHider_def]
124                   heap_update_list_concat_unfold
125                   o_def field_simps
126                   length_replicate[folded replicateHider_def])
127  done
128
129lemma h_t_valid_dom_s:
130  "\<lbrakk> h_t_valid htd c_guard p; x = ptr_val (p :: ('a :: mem_type) ptr);
131              n = size_of TYPE ('a) \<rbrakk>
132    \<Longrightarrow> {x ..+ n} \<times> {SIndexVal, SIndexTyp 0} \<subseteq> dom_s htd"
133  apply (clarsimp simp: h_t_valid_def valid_footprint_def Let_def
134                        intvl_def)
135  apply (drule_tac x=k in spec, simp add: size_of_def)
136  apply (clarsimp simp: dom_s_def)
137  apply (drule_tac x=0 in map_leD, simp_all)
138  done
139
140lemma user_data_at_rf_sr_dom_s:
141  "\<lbrakk> typ_at' UserDataT x s; (s, s') \<in> rf_sr \<rbrakk>
142    \<Longrightarrow> {x ..+ 2 ^ pageBits} \<times> {SIndexVal, SIndexTyp 0}
143    \<subseteq> dom_s (hrs_htd (t_hrs_' (globals s')))"
144  apply (drule rf_sr_heap_user_data_relation)
145  apply (drule user_data_at_ko)
146  apply (erule_tac x=x in cmap_relationE1)
147   apply (simp only: heap_to_user_data_def Let_def ko_at_projectKO_opt)
148   apply simp
149  apply (drule h_t_valid_clift)
150  apply (simp add: h_t_valid_dom_s pageBits_def)
151  done
152
153lemma device_data_at_rf_sr_dom_s:
154  "\<lbrakk> typ_at' UserDataDeviceT x s; (s, s') \<in> rf_sr \<rbrakk>
155    \<Longrightarrow> {x ..+ 2 ^ pageBits} \<times> {SIndexVal, SIndexTyp 0}
156    \<subseteq> dom_s (hrs_htd (t_hrs_' (globals s')))"
157  apply (drule rf_sr_heap_device_data_relation)
158  apply (drule device_data_at_ko)
159  apply (erule_tac x=x in cmap_relationE1)
160   apply (simp only: heap_to_device_data_def Let_def ko_at_projectKO_opt)
161   apply simp
162  apply (drule h_t_valid_clift)
163  apply (simp add: h_t_valid_dom_s pageBits_def)
164  done
165
166lemma intvl_2_power_times_decomp:
167  "\<forall>y < 2 ^ (n - m). {x + y * 2 ^ m ..+ 2 ^ m} \<times> S \<subseteq> T
168    \<Longrightarrow> m \<le> n \<Longrightarrow> n < word_bits
169    \<Longrightarrow> {(x :: word32) ..+ 2 ^ n} \<times> S \<subseteq> T"
170  apply (clarsimp simp: intvl_def)
171  apply (drule_tac x="of_nat k >> m" in spec)
172  apply (drule mp)
173   apply (rule shiftr_less_t2n)
174   apply (rule word_of_nat_less)
175   apply (simp add: word_of_nat_less)
176  apply (erule subsetD)
177  apply (clarsimp simp: shiftl_t2n[simplified mult.commute mult.left_commute, symmetric]
178                        shiftr_shiftl1)
179  apply (rule_tac x="unat (of_nat k && mask m :: word32)" in exI)
180  apply (simp add: field_simps word_plus_and_or_coroll2)
181  apply (simp add: word_bits_def unat_less_power and_mask_less')
182  done
183
184lemma flex_user_data_at_rf_sr_dom_s:
185  "\<lbrakk> (\<forall>p<2 ^ (pageBitsForSize sz - pageBits).
186         typ_at' UserDataT (x + p * 2 ^ pageBits) s); (s, s') \<in> rf_sr \<rbrakk>
187    \<Longrightarrow> {x ..+ 2 ^ pageBitsForSize sz} \<times> {SIndexVal, SIndexTyp 0}
188    \<subseteq> dom_s (hrs_htd (t_hrs_' (globals s')))"
189  apply (rule_tac m=pageBits in intvl_2_power_times_decomp,
190         simp_all add: pbfs_atleast_pageBits pbfs_less_wb')
191  apply (erule allEI, clarsimp)
192  apply (drule(1) user_data_at_rf_sr_dom_s)
193  apply (erule subsetD)
194  apply simp
195  done
196
197lemma hrs_mem_update_fold_eq:
198  "hrs_mem_update (fold f xs)
199    = fold (hrs_mem_update o f) xs"
200  apply (rule sym, induct xs)
201   apply (simp add: hrs_mem_update_def)
202  apply (simp add: hrs_mem_update_def fun_eq_iff)
203  done
204
205lemma power_user_page_foldl_zero_ranges:
206  " \<forall>p<2 ^ (pageBitsForSize sz - pageBits).
207      hrs_htd hrs \<Turnstile>\<^sub>t (Ptr (ptr + of_nat p * 0x1000) :: user_data_C ptr)
208    \<Longrightarrow> zero_ranges_are_zero rngs hrs
209    \<Longrightarrow> zero_ranges_are_zero rngs
210        (hrs_mem_update (\<lambda>s. foldl (\<lambda>s x. heap_update (Ptr x) (user_data_C (arr x)) s) s
211            (map (\<lambda>n. ptr + of_nat n * 0x1000) [0..<2 ^ (pageBitsForSize sz - pageBits)]))
212            hrs)"
213  apply (simp add: foldl_conv_fold hrs_mem_update_fold_eq)
214  apply (rule conjunct1)
215  apply (rule fold_invariant[where P="\<lambda>hrs'. zero_ranges_are_zero rngs hrs'
216          \<and> hrs_htd hrs' = hrs_htd hrs"
217      and xs=xs and Q="\<lambda>x. x \<in> set xs" for xs], simp_all)
218  apply (subst zero_ranges_are_zero_update, simp_all)
219  apply clarsimp
220  done
221
222lemma heap_to_device_data_disj_mdf':
223  "\<lbrakk>is_aligned ptr (pageBitsForSize sz); ksPSpace \<sigma> a = Some obj; objBitsKO obj = pageBits; pspace_aligned' \<sigma>;
224  pspace_distinct' \<sigma>; pspace_no_overlap' ptr (pageBitsForSize sz) \<sigma>\<rbrakk>
225\<Longrightarrow> heap_to_device_data (ksPSpace \<sigma>)
226     (\<lambda>x. if x \<in> {ptr..+2 ^ (pageBitsForSize sz)} then 0
227          else underlying_memory (ksMachineState \<sigma>) x)
228     a =
229    heap_to_device_data (ksPSpace \<sigma>) (underlying_memory (ksMachineState \<sigma>)) a"
230  apply (cut_tac heap_to_device_data_disj_mdf[where ptr = ptr
231     and gbits = "pageBitsForSize sz - pageBits" and n = 1
232     and sz = "pageBitsForSize sz",simplified])
233  apply (simp add: pbfs_atleast_pageBits pbfs_less_wb' field_simps| intro range_cover_full )+
234  done
235
236lemma range_cover_nca_neg: "\<And>x p (off :: 10 word).
237  \<lbrakk>(x::word32) < 4; {p..+2 ^pageBits } \<inter> {ptr..ptr + (of_nat n * 2 ^ bits - 1)} = {};
238    range_cover ptr sz bits n\<rbrakk>
239   \<Longrightarrow> p + ucast off * 4 + x \<notin> {ptr..+n * 2 ^ bits}"
240  apply (case_tac "n = 0")
241   apply simp
242  apply (subst range_cover_intvl,simp)
243   apply simp
244  apply (subgoal_tac " p + ucast off * 4 + x \<in>  {p..+2 ^ pageBits}")
245   apply blast
246  apply (clarsimp simp: intvl_def)
247  apply (rule_tac x = "unat off * 4 + unat x" in exI)
248  apply (simp add: ucast_nat_def)
249  apply (rule nat_add_offset_less [where n = 2, simplified])
250    apply (simp add: word_less_nat_alt)
251   apply (rule unat_lt2p)
252  apply (simp add: pageBits_def objBits_simps)
253  done
254
255
256
257lemma clearMemory_PageCap_ccorres:
258  "ccorres dc xfdc (invs' and valid_cap' (ArchObjectCap (PageCap False ptr undefined sz None))
259           and (\<lambda>s. 2 ^ pageBitsForSize sz \<le> gsMaxObjectSize s)
260           and K ({ptr .. ptr + 2 ^ (pageBitsForSize sz) - 1} \<inter> kernel_data_refs = {})
261           )
262      (UNIV \<inter> {s. bits_' s = of_nat (pageBitsForSize sz)}
263            \<inter> {s. ptr___ptr_to_unsigned_long_' s = Ptr ptr})
264      []
265     (doMachineOp (clearMemory ptr (2 ^ pageBitsForSize sz))) (Call clearMemory_'proc)"
266  (is "ccorres dc xfdc ?P ?P' [] ?m ?c")
267  apply (cinit' lift: bits_' ptr___ptr_to_unsigned_long_')
268   apply (rule_tac P="capAligned (ArchObjectCap (PageCap False ptr undefined sz None))"
269                in ccorres_gen_asm)
270   apply (rule ccorres_Guard_Seq)
271   apply (simp add: clearMemory_def)
272   apply (simp add: doMachineOp_bind ef_storeWord)
273   apply (rule ccorres_split_nothrow_novcg_dc)
274      apply (rule_tac P="?P" in ccorres_from_vcg[where P'=UNIV])
275      apply (rule allI, rule conseqPre, vcg)
276      apply (clarsimp simp: valid_cap'_def capAligned_def
277                            is_aligned_no_wrap'[OF _ word32_power_less_1])
278      apply (subgoal_tac "2 \<le> pageBitsForSize sz")
279       prefer 2
280       apply (simp add: pageBitsForSize_def split: vmpage_size.split)
281      apply (rule conjI)
282       apply (erule is_aligned_weaken)
283       apply (clarsimp simp: pageBitsForSize_def split: vmpage_size.splits)
284      apply (rule conjI)
285       apply (rule is_aligned_power2)
286       apply (clarsimp simp: pageBitsForSize_def split: vmpage_size.splits)
287      apply (clarsimp simp: ghost_assertion_size_logic[unfolded o_def])
288      apply (simp add: flex_user_data_at_rf_sr_dom_s)
289      apply (clarsimp simp: field_simps word_size_def mapM_x_storeWord_step)
290      apply (simp add: doMachineOp_def split_def exec_gets)
291      apply (simp add: select_f_def simpler_modify_def bind_def)
292      apply (fold replicateHider_def)[1]
293      apply (subst coerce_heap_update_to_heap_updates'
294                         [where chunk=4096 and m="2 ^ (pageBitsForSize sz - pageBits)"])
295       apply (simp add: pageBitsForSize_def pageBits_def
296                 split: vmpage_size.split)
297      apply (subst coerce_memset_to_heap_update_user_data)
298      apply (subgoal_tac "\<forall>p<2 ^ (pageBitsForSize sz - pageBits).
299                               x \<Turnstile>\<^sub>c (Ptr (ptr + of_nat p * 0x1000) :: user_data_C ptr)")
300       prefer 2
301       apply (erule allfEI[where f=of_nat])
302       apply clarsimp
303       apply (subst(asm) of_nat_power, assumption)
304        apply simp
305        apply (insert pageBitsForSize_32 [of sz])[1]
306        apply (erule order_le_less_trans [rotated])
307        apply simp
308       apply (simp, drule ko_at_projectKO_opt[OF user_data_at_ko])
309       apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def cpspace_relation_def)
310       apply (erule cmap_relationE1, simp(no_asm) add: heap_to_user_data_def Let_def)
311        apply fastforce
312       subgoal by (simp add: pageBits_def typ_heap_simps)
313      apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def)
314      apply (clarsimp simp: cpspace_relation_def typ_heap_simps
315                            clift_foldl_hrs_mem_update foldl_id
316                            carch_state_relation_def
317                            cmachine_state_relation_def
318                            foldl_fun_upd_const[unfolded fun_upd_def]
319                            power_user_page_foldl_zero_ranges
320                            dom_heap_to_device_data)
321      apply (rule conjI[rotated])
322       apply (simp add:pageBitsForSize_mess_multi)
323       apply (rule cmap_relationI)
324        apply (clarsimp simp: dom_heap_to_device_data cmap_relation_def)
325        apply (simp add:cuser_user_data_device_relation_def)
326      apply (subst help_force_intvl_range_conv, assumption)
327        subgoal by (simp add: pageBitsForSize_def split: vmpage_size.split)
328       apply assumption
329      apply (subst heap_to_user_data_update_region)
330       apply (drule map_to_user_data_aligned, clarsimp)
331       apply (rule aligned_range_offset_mem[where m=pageBits], simp_all)[1]
332       apply (rule pbfs_atleast_pageBits)
333      apply (erule cmap_relation_If_upd)
334        apply (clarsimp simp: cuser_user_data_relation_def fcp_beta
335                              order_less_le_trans[OF unat_lt2p])
336        apply (fold word_rsplit_0, simp add: word_rcat_rsplit)[1]
337       apply (rule image_cong[OF _ refl])
338       apply (rule set_eqI, rule iffI)
339        apply (clarsimp simp del: atLeastAtMost_iff)
340        apply (drule map_to_user_data_aligned, clarsimp)
341        apply (simp only: mask_in_range[symmetric])
342        apply (rule_tac x="unat ((xa && mask (pageBitsForSize sz)) >> pageBits)" in image_eqI)
343         apply (simp add: subtract_mask(2)[symmetric])
344         apply (cut_tac w="xa - ptr" and n=pageBits in and_not_mask[symmetric])
345         apply (simp add: shiftl_t2n field_simps pageBits_def)
346         apply (subst is_aligned_neg_mask_eq, simp_all)[1]
347         apply (erule aligned_sub_aligned, simp_all add: word_bits_def)[1]
348         apply (erule is_aligned_weaken)
349         apply (rule pbfs_atleast_pageBits[unfolded pageBits_def])
350        apply simp
351        apply (rule unat_less_power)
352         apply (fold word_bits_def, simp)
353        apply (rule shiftr_less_t2n)
354        apply (simp add: pbfs_atleast_pageBits)
355        apply (rule and_mask_less_size)
356        apply (simp add: word_bits_def word_size)
357       apply (rule IntI)
358        apply (clarsimp simp del: atLeastAtMost_iff)
359        apply (subst aligned_range_offset_mem, assumption, simp_all)[1]
360        apply (rule order_le_less_trans[rotated], erule shiftl_less_t2n [OF of_nat_power],
361               simp_all add: word_bits_def)[1]
362         apply (insert pageBitsForSize_32 [of sz])[1]
363         apply (erule order_le_less_trans [rotated])
364         subgoal by simp
365        subgoal by (simp add: pageBits_def shiftl_t2n field_simps)
366       apply clarsimp
367       apply (drule_tac x="of_nat n" in spec)
368       apply (simp add: of_nat_power[where 'a=32, folded word_bits_def])
369       apply (rule exI)
370       subgoal by (simp add: pageBits_def ko_at_projectKO_opt[OF user_data_at_ko])
371      subgoal by simp
372     apply csymbr
373     apply (ctac add: cleanCacheRange_PoU_ccorres[unfolded dc_def])
374    apply wp
375   apply (simp add: guard_is_UNIV_def unat_of_nat
376                    word_bits_def capAligned_def word_of_nat_less)
377  apply (clarsimp simp: word_bits_def valid_cap'_def
378                        capAligned_def word_of_nat_less)
379  apply (frule is_aligned_addrFromPPtr_n, simp add: pageBitsForSize_def split: vmpage_size.splits)
380  by (clarsimp simp: is_aligned_no_overflow'[where n=12, simplified]
381                        is_aligned_no_overflow'[where n=16, simplified]
382                        is_aligned_no_overflow'[where n=20, simplified]
383                        is_aligned_no_overflow'[where n=24, simplified] pageBits_def
384                        field_simps is_aligned_mask[symmetric] mask_AND_less_0
385                        pageBitsForSize_def split: vmpage_size.splits)
386
387lemma coerce_memset_to_heap_update_asidpool:
388  "heap_update_list x (replicateHider 4096 0)
389      = heap_update (Ptr x :: asid_pool_C ptr)
390             (asid_pool_C (FCP (\<lambda>x. Ptr 0)))"
391  apply (intro ext, simp add: heap_update_def)
392  apply (rule_tac f="\<lambda>xs. heap_update_list x xs a b" for a b in arg_cong)
393  apply (simp add: to_bytes_def size_of_def typ_info_simps asid_pool_C_tag_def)
394  apply (simp add: ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td align_of_def padup_def
395                   final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def)
396  apply (simp add: typ_info_simps
397                   user_context_C_tag_def thread_state_C_tag_def seL4_Fault_C_tag_def
398                   lookup_fault_C_tag_def update_ti_t_ptr_0s
399                   ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td
400                   ti_typ_combine_empty_ti ti_typ_combine_td
401                   align_of_def padup_def
402                   final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def
403                   align_td_array' size_td_array)
404  apply (simp add: typ_info_array')
405  apply (subst access_ti_list_array)
406     apply simp
407    apply simp
408   apply (simp add: fcp_beta typ_info_word typ_info_ptr word_rsplit_0)
409   apply fastforce
410  apply (simp add: collapse_foldl_replicate)
411  done
412
413declare replicate_numeral [simp]
414
415lemma coerce_memset_to_heap_update_pte:
416  "heap_update_list x (replicateHider 8 0)
417      = heap_update (Ptr x :: pte_C ptr)
418             (pte_C.pte_C (FCP (\<lambda>x. 0)))"
419  apply (intro ext, simp add: heap_update_def)
420  apply (rule_tac f="\<lambda>xs. heap_update_list x xs a b" for a b in arg_cong)
421  apply (simp add: to_bytes_def size_of_def typ_info_simps pte_C_tag_def)
422  apply (simp add: ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td align_of_def padup_def
423                   final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def)
424  apply (simp add: typ_info_simps align_td_array' size_td_array)
425  apply (simp add: typ_info_array' typ_info_word word_rsplit_0)
426  apply (simp add: numeral_nat word_rsplit_0)
427  apply (simp add: replicateHider_def)
428  done
429
430lemma coerce_memset_to_heap_update_pde:
431  "heap_update_list x (replicateHider 8 0)
432      = heap_update (Ptr x :: pde_C ptr)
433             (pde_C.pde_C (FCP (\<lambda>x. 0)))"
434  apply (intro ext, simp add: heap_update_def)
435  apply (rule_tac f="\<lambda>xs. heap_update_list x xs a b" for a b in arg_cong)
436  apply (simp add: to_bytes_def size_of_def typ_info_simps pde_C_tag_def)
437  apply (simp add: ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td align_of_def padup_def
438                   final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def)
439  apply (simp add: typ_info_simps align_td_array' size_td_array)
440  apply (simp add: typ_info_array' typ_info_word word_rsplit_0)
441  apply (simp add: numeral_nat word_rsplit_0)
442  apply (simp add: replicateHider_def)
443  done
444
445lemma objBits_eq_by_type:
446  fixes x :: "'a :: pspace_storable" and y :: 'a
447  shows "objBits x = objBits y"
448  apply (simp add: objBits_def)
449  apply (rule objBits_type)
450  apply (simp add: koTypeOf_injectKO)
451  done
452
453lemma mapM_x_store_memset_ccorres_assist:
454  fixes val :: "'a :: pspace_storable"
455  assumes nofail: "\<not> snd (mapM_x (\<lambda>slot. setObject slot val) slots \<sigma>)"
456  assumes slots1: "\<forall>n < length slots. slots ! n = hd slots + (of_nat n << objBits val)"
457  assumes slots2: "n = length slots * (2 ^ objBits val)"
458  assumes ptr: "ptr = hd slots"
459  assumes ko: "\<And>ko :: 'a. updateObject ko = updateObject_default ko"
460              "\<And>ko :: 'a. (1 :: word32) < 2 ^ objBits ko"
461  assumes restr: "set slots \<subseteq> S"
462  assumes worker: "\<And>ptr s s' (ko :: 'a). \<lbrakk> (s, s') \<in> rf_sr; ko_at' ko ptr s; ptr \<in> S \<rbrakk>
463                                \<Longrightarrow> (s \<lparr> ksPSpace := ksPSpace s (ptr \<mapsto> injectKO val)\<rparr>,
464                                     globals_update (t_hrs_'_update (hrs_mem_update
465                                                    (heap_update_list ptr
466                                                    (replicateHider (2 ^ objBits val) (ucast c))))) s') \<in> rf_sr"
467  assumes rf_sr: "(\<sigma>, s) \<in> rf_sr"
468  shows
469  "\<exists>(rv, \<sigma>') \<in> fst (mapM_x (\<lambda>slot. setObject slot val) slots \<sigma>).
470      (\<sigma>', globals_update (t_hrs_'_update (hrs_mem_update
471                          (heap_update_list ptr (replicateHider n c)))) s) \<in> rf_sr"
472  unfolding slots2 ptr using rf_sr slots1 nofail restr
473proof (induct slots arbitrary: s \<sigma>)
474  case Nil
475  show ?case
476    using Nil.prems
477    apply (simp add: mapM_x_def sequence_x_def return_def replicateHider_def)
478    apply (simp add: rf_sr_def hrs_mem_update_def cstate_relation_def Let_def
479                     carch_state_relation_def cmachine_state_relation_def
480                     h_t_valid_clift_Some_iff)
481    done
482next
483  case (Cons x xs tPre sPre)
484
485  note nofail_bind = Cons.prems(3)[unfolded mapM_x_Cons K_bind_def]
486
487  have obj_at: "obj_at' (\<lambda>x :: 'a. True) x sPre"
488    using not_snd_bindI1[OF nofail_bind]
489    apply (subst(asm) setObject_obj_at_pre, simp_all add: ko snd_bind)
490    apply (clarsimp simp: stateAssert_def exec_get return_def)
491    apply (simp add: koTypeOf_injectKO typ_at_to_obj_at')
492    done
493
494  note in_setObject = setObject_eq[OF _ _ objBits_eq_by_type obj_at,
495                                   where ko=val, simplified ko, simplified]
496
497  note nofail_mapM = not_snd_bindI2[OF nofail_bind, OF in_setObject]
498
499  have hd_xs: "xs \<noteq> [] \<Longrightarrow> hd xs = x + (2 ^ objBits val)"
500    using Cons.prems(2)[rule_format, where n=1]
501    by (simp add: hd_conv_nth)
502
503  show ?case
504    using obj_at_ko_at'[OF obj_at] Cons.prems(4)
505    apply (clarsimp simp add: mapM_x_Cons bind_def split_def)
506    apply (rule rev_bexI, rule in_setObject)
507    apply (cut_tac Cons.hyps[OF _ _ nofail_mapM])
508       defer
509       apply (rule worker, rule Cons.prems, assumption+)
510      apply clarsimp
511      apply (case_tac "xs = []", simp_all)[1]
512      apply (insert Cons.prems, simp)[1]
513      apply (frule_tac x="Suc n" in spec)
514      apply (simp add: hd_xs shiftl_t2n field_simps)
515     apply assumption
516    apply clarsimp
517    apply (rule rev_bexI, assumption)
518    apply (simp add: o_def)
519    apply (case_tac "xs = []")
520     apply (simp add: hrs_mem_update_def split_def replicateHider_def)
521    apply (subst(asm) heap_update_list_concat_fold_hrs_mem)
522     apply (simp add: hd_xs replicateHider_def)
523    apply (simp add: replicateHider_def replicate_add)
524    done
525qed
526
527lemma invalidateTLBByASID_ccorres:
528  "ccorres dc xfdc (valid_pde_mappings' and (\<lambda>_. asid \<le> mask asid_bits))
529                   (UNIV \<inter> {s. asid_' s = asid}) []
530       (invalidateTLBByASID asid) (Call invalidateTLBByASID_'proc)"
531  apply (cinit lift: asid_')
532   apply (ctac(no_vcg) add: loadHWASID_ccorres)
533    apply csymbr
534    apply (simp add: case_option_If2 del: Collect_const)
535    apply (rule ccorres_if_cond_throws2[where Q=\<top> and Q'=\<top>])
536       apply (clarsimp simp: pde_stored_asid_def to_bool_def split: if_split)
537      apply (rule ccorres_return_void_C[unfolded dc_def])
538     apply (simp add: dc_def[symmetric])
539     apply csymbr
540     apply (ctac add: invalidateTranslationASID_ccorres)
541    apply vcg
542   apply (wp hoare_drop_imps)
543  apply (clarsimp simp: pde_stored_asid_def to_bool_def)
544  done
545
546(* FIXME: also in VSpace_C *)
547lemma ignoreFailure_liftM:
548  "ignoreFailure = liftM (\<lambda>v. ())"
549  apply (rule ext)+
550  apply (simp add: ignoreFailure_def liftM_def
551                   catch_def)
552  apply (rule bind_apply_cong[OF refl])
553  apply (simp split: sum.split)
554  done
555
556end
557
558lemma option_to_0_user_mem':
559  "option_to_0 \<circ> user_mem' as =(\<lambda>x. if x \<in> {y. \<not> pointerInUserData y as} then 0
560  else underlying_memory (ksMachineState as) x) "
561  apply (rule ext)
562  apply (simp add:user_mem'_def option_to_0_def split:if_splits)
563  done
564
565lemma heap_to_user_data_in_user_mem'[simp]:
566  "\<lbrakk>pspace_aligned' as;pspace_distinct' as\<rbrakk> \<Longrightarrow> heap_to_user_data (ksPSpace as) (option_to_0 \<circ> user_mem' as) =
567  heap_to_user_data (ksPSpace as)(underlying_memory (ksMachineState as))"
568  apply (rule ext)+
569  apply (clarsimp simp: heap_to_user_data_def option_map_def
570                 split: option.splits)
571  apply (subst option_to_0_user_mem')
572  apply (subst map_option_byte_to_word_heap)
573   apply (clarsimp simp: projectKO_opt_user_data map_comp_def
574                  split: option.split_asm kernel_object.split_asm)
575   apply (frule(1) pspace_alignedD')
576   apply (frule(1) pspace_distinctD')
577   apply (subgoal_tac "x + ucast off * 4 + xa  && ~~ mask pageBits = x" )
578    apply (clarsimp simp: pointerInUserData_def typ_at'_def ko_wp_at'_def)
579   apply (simp add: ARM_HYP.pageBits_def)
580   apply (subst mask_lower_twice2[where n = 2 and m = 12,simplified,symmetric])
581   apply (subst is_aligned_add_helper[THEN conjunct2,where n1 = 2])
582     apply (erule aligned_add_aligned)
583      apply (simp add: is_aligned_mult_triv2[where n = 2,simplified])
584     apply  (clarsimp simp: objBits_simps ARM_HYP.pageBits_def)
585    apply simp
586   apply (rule is_aligned_add_helper[THEN conjunct2])
587    apply (simp add: ARM_HYP.pageBits_def objBits_simps)
588   apply (rule word_less_power_trans2[where k = 2,simplified])
589     apply (rule less_le_trans[OF ucast_less])
590      apply simp+
591  done
592
593context begin interpretation Arch . (*FIXME: arch_split*)
594
595crunch pde_mappings'[wp]: invalidateTLBByASID "valid_pde_mappings'"
596crunch ksArchState[wp]: invalidateTLBByASID "\<lambda>s. P (ksArchState s)"
597
598crunch gsMaxObjectSize[wp]: invalidateTLBByASID "\<lambda>s. P (gsMaxObjectSize s)"
599crunch gsMaxObjectSize[wp]: deleteASIDPool "\<lambda>s. P (gsMaxObjectSize s)"
600  (ignore: setObject getObject wp: crunch_wps getObject_inv loadObject_default_inv
601     simp: crunch_simps)
602end
603
604context kernel_m begin
605
606lemma page_table_at_rf_sr_dom_s:
607  "\<lbrakk> page_table_at' x s; (s, s') \<in> rf_sr \<rbrakk>
608    \<Longrightarrow> {x ..+ 2 ^ ptBits} \<times> {SIndexVal, SIndexTyp 0}
609    \<subseteq> dom_s (hrs_htd (t_hrs_' (globals s')))"
610  apply (rule_tac m=3 in intvl_2_power_times_decomp,
611         simp_all add: shiftl_t2n field_simps table_bits_defs
612                       word_bits_def)
613  apply (clarsimp simp: page_table_at'_def intvl_def)
614  apply (drule spec, drule(1) mp)
615  apply (simp add: typ_at_to_obj_at_arches)
616  apply (drule obj_at_ko_at', clarsimp)
617  apply (erule cmap_relationE1[OF rf_sr_cpte_relation])
618   apply (erule ko_at_projectKO_opt)
619  apply (drule h_t_valid_clift)
620  apply (drule h_t_valid_dom_s[OF _ refl refl])
621  apply (erule subsetD)
622  apply (auto simp add: intvl_def shiftl_t2n)[1]
623  done
624
625lemma page_directory_at_rf_sr_dom_s:
626  "\<lbrakk> page_directory_at' x s; (s, s') \<in> rf_sr \<rbrakk>
627    \<Longrightarrow> {x ..+ 2 ^ pdBits} \<times> {SIndexVal, SIndexTyp 0}
628    \<subseteq> dom_s (hrs_htd (t_hrs_' (globals s')))"
629  apply (rule_tac m=3 in intvl_2_power_times_decomp,
630         simp_all add: shiftl_t2n field_simps table_bits_defs
631                       word_bits_def)
632  apply (clarsimp simp: page_directory_at'_def intvl_def)
633  apply (drule spec, drule(1) mp)
634  apply (simp add: typ_at_to_obj_at_arches)
635  apply (drule obj_at_ko_at', clarsimp)
636  apply (erule cmap_relationE1[OF rf_sr_cpde_relation])
637   apply (erule ko_at_projectKO_opt)
638  apply (drule h_t_valid_clift)
639  apply (drule h_t_valid_dom_s[OF _ refl refl])
640  apply (erule subsetD)
641  apply (auto simp add: intvl_def shiftl_t2n)[1]
642  done
643
644lemma clearMemory_setObject_PTE_ccorres:
645  "ccorres dc xfdc (page_table_at' ptr
646                and (\<lambda>s. 2 ^ ptBits \<le> gsMaxObjectSize s)
647                and (\<lambda>_. is_aligned ptr ptBits \<and> ptr \<noteq> 0 \<and> pstart = addrFromPPtr ptr))
648            (UNIV \<inter> {s. ptr___ptr_to_unsigned_long_' s = Ptr ptr} \<inter> {s. bits_' s = of_nat ptBits}) []
649       (do x \<leftarrow> mapM_x (\<lambda>a. setObject a ARM_HYP_H.InvalidPTE)
650                       [ptr , ptr + 2 ^ objBits ARM_HYP_H.InvalidPTE .e. ptr + 2 ^ ptBits - 1];
651           doMachineOp (cleanCacheRange_PoU ptr (ptr + 2 ^ ptBits - 1) pstart)
652        od)
653       (Call clearMemory_'proc)"
654  apply (rule ccorres_gen_asm)+
655  apply (cinit' lift: ptr___ptr_to_unsigned_long_' bits_')
656   apply (rule ccorres_Guard_Seq)
657   apply (rule ccorres_split_nothrow_novcg_dc)
658      apply (rule_tac P="page_table_at' ptr and (\<lambda>s. 2 ^ ptBits \<le> gsMaxObjectSize s)"
659               in ccorres_from_vcg_nofail[where P'=UNIV])
660      apply (rule allI, rule conseqPre, vcg)
661      apply clarsimp
662      apply (subst ghost_assertion_size_logic[unfolded o_def])
663        apply (simp add: table_bits_defs)
664       apply simp
665      apply (clarsimp simp: replicateHider_def[symmetric])
666      apply (frule is_aligned_no_overflow')
667      apply (intro conjI)
668          apply (clarsimp simp add: table_bits_defs
669                          cong: StateSpace.state.fold_congs globals.fold_congs)
670         apply (erule is_aligned_weaken, simp add: table_bits_defs)
671        apply (clarsimp simp: is_aligned_def table_bits_defs)
672       subgoal
673         apply (subst unat_of_nat32, simp add: table_bits_defs word_bits_def)
674         apply (subst unat_power_lower32', simp add: table_bits_defs word_bits_def)
675         apply (erule (1) page_table_at_rf_sr_dom_s)
676         done
677      apply (clarsimp simp add: table_bits_defs
678                      cong: StateSpace.state.fold_congs globals.fold_congs)
679      apply (simp add: upto_enum_step_def objBits_simps table_bits_defs
680                       field_simps linorder_not_less[symmetric] archObjSize_def
681                       upto_enum_word split_def)
682      apply (erule mapM_x_store_memset_ccorres_assist
683                      [unfolded split_def, OF _ _ _ _ _ _ subset_refl],
684             simp_all add: shiftl_t2n hd_map objBits_simps archObjSize_def table_bits_defs)[1]
685      apply (rule cmap_relationE1, erule rf_sr_cpte_relation, erule ko_at_projectKO_opt)
686(*lemma coerce_memset_to_heap_update_pte: FIXME REMOVE
687  "heap_update_list x (replicateHider 8 0)
688      = heap_update (Ptr x :: pte_C ptr)
689             (pte_C.pte_C (FCP (\<lambda>x. 0)))" *)
690      apply (subst coerce_memset_to_heap_update_pte)
691      apply (clarsimp simp: rf_sr_def Let_def cstate_relation_def typ_heap_simps)
692      apply (rule conjI)
693       apply (simp add: cpspace_relation_def typ_heap_simps update_pte_map_tos
694                        update_pte_map_to_ptes carray_map_relation_upd_triv)
695       apply (rule cmap_relation_updI, simp_all)[1]
696       apply (simp add: cpte_relation_def Let_def pte_lift_def
697                        fcp_beta pte_get_tag_def pte_tag_defs)
698       apply (simp add: carch_state_relation_def cmachine_state_relation_def
699                        typ_heap_simps update_pte_map_tos)
700      apply csymbr
701     apply (rule ccorres_Guard)
702     apply (ctac add: cleanCacheRange_PoU_ccorres)
703    apply (wp mapM_x_wp' setObject_ksPSpace_only updateObject_default_inv | simp)+
704   apply (clarsimp simp: guard_is_UNIV_def table_bits_defs)
705  apply (clarsimp simp: ptBits_def pageBits_def)
706  apply (frule is_aligned_addrFromPPtr_n, simp add: table_bits_defs)
707  apply (clarsimp simp: is_aligned_no_overflow'[where n=10, simplified] table_bits_defs
708                        field_simps is_aligned_mask[symmetric] mask_AND_less_0)
709  apply (erule is_aligned_no_wrap', simp)
710  done
711
712lemma ccorres_make_xfdc:
713  "ccorresG rf_sr \<Gamma> r xf P P' h a c \<Longrightarrow> ccorresG rf_sr \<Gamma> dc xfdc P P' h a c"
714  apply (erule ccorres_rel_imp)
715  apply simp
716  done
717
718lemma ccorres_if_True_False_simps:
719  "ccorres r xf P P' hs a (IF True THEN c ELSE c' FI) = ccorres r xf P P' hs a c"
720  "ccorres r xf P P' hs a (IF False THEN c ELSE c' FI) = ccorres r xf P P' hs a c'"
721  "ccorres r xf P P' hs a (IF True THEN c ELSE c' FI ;; d) = ccorres r xf P P' hs a (c ;; d)"
722  "ccorres r xf P P' hs a (IF False THEN c ELSE c' FI ;; d) = ccorres r xf P P' hs a (c' ;; d)"
723  by (simp_all add: ccorres_cond_iffs ccorres_seq_simps)
724
725lemmas cap_tag_values =
726  cap_untyped_cap_def
727  cap_endpoint_cap_def
728  cap_notification_cap_def
729  cap_reply_cap_def
730  cap_cnode_cap_def
731  cap_thread_cap_def
732  cap_irq_handler_cap_def
733  cap_null_cap_def
734  cap_irq_control_cap_def
735  cap_zombie_cap_def
736  cap_small_frame_cap_def
737  cap_frame_cap_def
738  cap_page_table_cap_def
739  cap_page_directory_cap_def
740  cap_asid_pool_cap_def
741  cap_vcpu_cap_def
742
743lemma ccorres_return_C_seq:
744  "\<lbrakk>\<And>s f. xf (global_exn_var_'_update f (xfu (\<lambda>_. v s) s)) = v s; \<And>s f. globals (xfu f s) = globals s; wfhandlers hs\<rbrakk>
745  \<Longrightarrow> ccorres_underlying rf_sr \<Gamma> r rvxf arrel xf (\<lambda>_. True) {s. arrel rv (v s)} hs (return rv) (return_C xfu v ;; d)"
746  apply (rule ccorres_guard_imp)
747  apply (rule ccorres_split_throws, rule ccorres_return_C, simp+)
748  apply vcg
749  apply simp_all
750  done
751
752lemma ccap_relation_VPIsDevice:
753  "\<lbrakk>isPageCap cp; ccap_relation (ArchObjectCap cp) cap  \<rbrakk> \<Longrightarrow>
754   (generic_frame_cap_get_capFIsDevice_CL (cap_lift cap) = 0) = (\<not> (capVPIsDevice cp))"
755   by (clarsimp elim!:ccap_relationE
756     simp : isPageCap_def generic_frame_cap_get_capFIsDevice_CL_def cap_to_H_def
757            Let_def to_bool_def
758     split: arch_capability.split_asm cap_CL.split_asm if_split_asm)
759
760
761lemma ccap_relation_get_capZombiePtr_CL:
762  "\<lbrakk> ccap_relation cap cap'; isZombie cap; capAligned cap \<rbrakk>
763      \<Longrightarrow> get_capZombiePtr_CL (cap_zombie_cap_lift cap') = capZombiePtr cap"
764  apply (simp only: cap_get_tag_isCap[symmetric])
765  apply (drule(1) cap_get_tag_to_H)
766  apply (clarsimp simp: get_capZombiePtr_CL_def get_capZombieBits_CL_def Let_def split: if_split)
767  apply (subst less_mask_eq)
768   apply (clarsimp simp add: capAligned_def objBits_simps word_bits_conv)
769   apply unat_arith
770  apply simp
771  done
772
773
774lemma modify_gets_helper:
775  "do y \<leftarrow> modify (ksPSpace_update (\<lambda>_. ps)); ps' \<leftarrow> gets ksPSpace; f ps' od
776      = do y \<leftarrow> modify (ksPSpace_update (\<lambda>_. ps)); f ps od"
777  by (simp add: bind_def simpler_modify_def simpler_gets_def)
778
779lemma snd_lookupAround2_update:
780  "ps y \<noteq> None \<Longrightarrow>
781    snd (lookupAround2 x (ps (y \<mapsto> v'))) = snd (lookupAround2 x ps)"
782  apply (clarsimp simp: lookupAround2_def lookupAround_def Let_def
783                        dom_fun_upd2
784              simp del: dom_fun_upd cong: if_cong option.case_cong)
785  apply (clarsimp split: option.split if_split  cong: if_cong)
786  apply auto
787  done
788
789lemma double_setEndpoint:
790  "do y \<leftarrow> setEndpoint epptr v1; setEndpoint epptr v2 od
791       = setEndpoint epptr v2"
792  apply (simp add: setEndpoint_def setObject_def bind_assoc split_def
793                   modify_gets_helper)
794  apply (simp add: updateObject_default_def bind_assoc objBits_simps)
795  apply (rule ext)
796  apply (rule bind_apply_cong, rule refl)+
797  apply (clarsimp simp add: in_monad projectKOs magnitudeCheck_assert
798                            snd_lookupAround2_update)
799  apply (simp add: lookupAround2_known1 assert_opt_def projectKO_def projectKO_opt_ep
800                    alignCheck_assert)
801  apply (simp add: bind_def simpler_modify_def)
802  done
803
804lemma filterM_setEndpoint_adjustment:
805  "\<lbrakk> \<And>v. do setEndpoint epptr IdleEP; body v od
806          = do v' \<leftarrow> body v; setEndpoint epptr IdleEP; return v' od \<rbrakk>
807    \<Longrightarrow>
808   (do q' \<leftarrow> filterM body q; setEndpoint epptr (f q') od)
809    = (do setEndpoint epptr IdleEP; q' \<leftarrow> filterM body q; setEndpoint epptr (f q') od)"
810  apply (rule sym)
811  apply (induct q arbitrary: f)
812   apply (simp add: double_setEndpoint)
813  apply (simp add: bind_assoc)
814  apply (subst bind_assoc[symmetric], simp, simp add: bind_assoc)
815  done
816
817lemma ccorres_inst_voodoo:
818  "\<forall>x. ccorres r xf (P x) (P' x) hs (h x) (c x)
819     \<Longrightarrow> \<forall>x. ccorres r xf (P x) (P' x) hs (h x) (c x)"
820  by simp
821
822lemma cpspace_relation_ep_update_ep2:
823  "\<lbrakk> ko_at' (ep :: endpoint) epptr s;
824      cmap_relation (map_to_eps (ksPSpace s))
825           (cslift t) ep_Ptr (cendpoint_relation (cslift t));
826      cendpoint_relation (cslift t') ep' endpoint;
827      (cslift t' :: tcb_C ptr \<rightharpoonup> tcb_C) = cslift t \<rbrakk>
828     \<Longrightarrow> cmap_relation (map_to_eps (ksPSpace s(epptr \<mapsto> KOEndpoint ep')))
829          (cslift t(ep_Ptr epptr \<mapsto> endpoint))
830          ep_Ptr (cendpoint_relation (cslift t'))"
831  apply (rule cmap_relationE1, assumption, erule ko_at_projectKO_opt)
832  apply (rule_tac P="\<lambda>a. cmap_relation a b c d" for b c d in rsubst,
833                   erule cmap_relation_upd_relI, assumption+)
834    apply simp+
835  apply (rule ext, simp add: map_comp_def projectKO_opt_ep split: if_split)
836  done
837
838end
839
840context begin interpretation Arch . (*FIXME: arch_split*)
841lemma setObject_tcb_ep_obj_at'[wp]:
842  "\<lbrace>obj_at' (P :: endpoint \<Rightarrow> bool) ptr\<rbrace> setObject ptr' (tcb :: tcb) \<lbrace>\<lambda>rv. obj_at' P ptr\<rbrace>"
843  apply (rule obj_at_setObject2, simp_all)
844  apply (clarsimp simp: updateObject_default_def in_monad)
845  done
846end
847
848crunch ep_obj_at'[wp]: setThreadState "obj_at' (P :: endpoint \<Rightarrow> bool) ptr"
849  (ignore: getObject setObject simp: unless_def)
850
851context kernel_m begin
852
853lemma ccorres_abstract_h_val:
854  "(\<And>rv. P rv \<Longrightarrow> ccorres r xf G (G' rv) hs a c) \<Longrightarrow>
855   ccorres r xf G ({s. P (h_val (hrs_mem (t_hrs_' (globals s))) p)
856            \<longrightarrow> s \<in> G' (h_val (hrs_mem (t_hrs_' (globals s)))
857            p)}
858   \<inter> {s. P (h_val (hrs_mem (t_hrs_' (globals s))) p)}) hs a c"
859   apply (rule ccorres_tmp_lift1 [where P = P])
860   apply (clarsimp simp: Collect_conj_eq [symmetric])
861   apply (fastforce intro: ccorres_guard_imp)
862   done
863
864lemma ccorres_subst_basic_helper:
865  "\<lbrakk> \<And>s s'. \<lbrakk> P s; s' \<in> P'; (s, s') \<in> rf_sr \<rbrakk> \<Longrightarrow> f s' = f' s';
866     \<And>s s'. \<lbrakk> P s; s' \<in> P'; (s, s') \<in> rf_sr \<rbrakk> \<Longrightarrow> (s, f' s') \<in> rf_sr;
867     \<And>s'. xf' (f' s') = v; \<And>rv' t t'. ceqv \<Gamma> xf' rv' t t' c (c' rv');
868        ccorres rrel xf Q Q' hs a (c' v) \<rbrakk>
869       \<Longrightarrow> ccorres rrel xf (P and Q) {s. s \<in> P' \<and> f' s \<in> Q'} hs a (Basic f ;; c)"
870  apply (rule ccorres_guard_imp2)
871   apply (rule ccorres_add_return)
872   apply (rule ccorres_split_nothrow[where xf'=xf' and r'="\<lambda>rv rv'. rv' = v"])
873       apply (rule ccorres_from_vcg[where P=P and P'=P'])
874       apply (rule allI, rule conseqPre, vcg)
875       apply (clarsimp simp: return_def)
876      apply assumption
877     apply simp
878    apply wp
879   apply vcg
880  apply clarsimp
881  done
882
883lemma ctcb_relation_blocking_ipc_badge:
884  "\<lbrakk> ctcb_relation tcb ctcb; isBlockedOnSend (tcbState tcb) \<rbrakk> \<Longrightarrow>
885      tsType_CL (thread_state_lift (tcbState_C ctcb)) = scast ThreadState_BlockedOnSend"
886  "\<lbrakk> ctcb_relation tcb ctcb;
887        tsType_CL (thread_state_lift (tcbState_C ctcb)) = scast ThreadState_BlockedOnSend \<rbrakk>
888      \<Longrightarrow> blockingIPCBadge (tcbState tcb)
889             = blockingIPCBadge_CL (thread_state_lift (tcbState_C ctcb))"
890   apply (clarsimp simp add: ctcb_relation_def)
891   apply (simp add: isBlockedOnSend_def split: Structures_H.thread_state.split_asm)
892   apply (clarsimp simp: cthread_state_relation_def)
893  apply (clarsimp simp add: ctcb_relation_def cthread_state_relation_def)
894  apply (cases "tcbState tcb", simp_all add: "StrictC'_thread_state_defs")
895  done
896
897lemma cendpoint_relation_q_cong:
898  "\<lbrakk> \<And>t rf. (t, rf) \<in> ep_q_refs_of' ep \<Longrightarrow> hp (tcb_ptr_to_ctcb_ptr t) = hp' (tcb_ptr_to_ctcb_ptr t) \<rbrakk>
899      \<Longrightarrow> cendpoint_relation hp ep ep' = cendpoint_relation hp' ep ep'"
900  apply (cases ep, simp_all add: cendpoint_relation_def Let_def)
901   apply (rule conj_cong [OF refl])
902   apply (rule tcb_queue_relation'_cong[OF refl refl refl])
903   apply clarsimp
904  apply (rule conj_cong [OF refl])
905  apply (rule tcb_queue_relation'_cong[OF refl refl refl])
906  apply clarsimp
907  done
908
909lemma cnotification_relation_q_cong:
910  "\<lbrakk>\<And>t rf. (t, rf) \<in> ntfn_q_refs_of' (ntfnObj ntfn) \<Longrightarrow>  hp (tcb_ptr_to_ctcb_ptr t) = hp' (tcb_ptr_to_ctcb_ptr t)\<rbrakk>
911      \<Longrightarrow>  cnotification_relation hp ntfn ntfn' = cnotification_relation hp' ntfn ntfn'"
912  apply (cases "ntfnObj ntfn", simp_all add: cnotification_relation_def Let_def)
913  apply (auto intro: iffD1[OF tcb_queue_relation'_cong[OF refl refl refl]])
914  done
915
916lemma tcbSchedEnqueue_ep_at:
917  "\<lbrace>obj_at' (P :: endpoint \<Rightarrow> bool) ep\<rbrace>
918      tcbSchedEnqueue t
919   \<lbrace>\<lambda>rv. obj_at' P ep\<rbrace>"
920  including no_pre
921  apply (simp add: tcbSchedEnqueue_def unless_def null_def)
922  apply (wp threadGet_wp, clarsimp, wp+)
923  apply (clarsimp split: if_split, wp)
924  done
925
926lemma ccorres_duplicate_guard:
927  "ccorres r xf (P and P) Q hs f f' \<Longrightarrow> ccorres r xf P Q hs f f'"
928  by (erule ccorres_guard_imp, auto)
929
930
931lemma ep_q_refs'_no_NTFNBound[simp]:
932  "(x, NTFNBound) \<notin> ep_q_refs_of' ep"
933  by (auto simp: ep_q_refs_of'_def split: endpoint.splits)
934
935
936lemma ntfn_q_refs'_no_NTFNBound[simp]:
937  "(x, NTFNBound) \<notin> ntfn_q_refs_of' ntfn"
938  by (auto simp: ntfn_q_refs_of'_def split: ntfn.splits)
939
940lemma cancelBadgedSends_ccorres:
941  "ccorres dc xfdc (invs' and ep_at' ptr)
942              (UNIV \<inter> {s. epptr_' s = Ptr ptr} \<inter> {s. badge_' s = bdg}) []
943       (cancelBadgedSends ptr bdg) (Call cancelBadgedSends_'proc)"
944  apply (cinit lift: epptr_' badge_' simp: whileAnno_def)
945   apply (simp add: list_case_return2
946              cong: list.case_cong Structures_H.endpoint.case_cong call_ignore_cong
947               del: Collect_const)
948   apply (rule ccorres_pre_getEndpoint)
949   apply (rule_tac R="ko_at' rv ptr" and xf'="ret__unsigned_'"
950               and val="case rv of RecvEP q \<Rightarrow> scast EPState_Recv | IdleEP \<Rightarrow> scast EPState_Idle
951                                | SendEP q \<Rightarrow> scast EPState_Send"
952               in ccorres_symb_exec_r_known_rv_UNIV[where R'=UNIV])
953      apply vcg
954      apply clarsimp
955      apply (erule cmap_relationE1 [OF cmap_relation_ep], erule ko_at_projectKO_opt)
956      apply (clarsimp simp: typ_heap_simps cendpoint_relation_def Let_def
957                     split: Structures_H.endpoint.split_asm)
958     apply ceqv
959    apply wpc
960      apply (simp add: dc_def[symmetric] ccorres_cond_iffs)
961      apply (rule ccorres_return_Skip)
962     apply (simp add: dc_def[symmetric] ccorres_cond_iffs)
963     apply (rule ccorres_return_Skip)
964    apply (rename_tac list)
965    apply (simp add: Collect_True Collect_False endpoint_state_defs
966                     ccorres_cond_iffs dc_def[symmetric]
967                del: Collect_const cong: call_ignore_cong)
968    apply (rule ccorres_rhs_assoc)+
969    apply (csymbr, csymbr)
970    apply (drule_tac s = rv in sym, simp only:)
971    apply (rule_tac P="ko_at' rv ptr and invs'" in ccorres_cross_over_guard)
972    apply (rule ccorres_symb_exec_r)
973      apply (rule ccorres_rhs_assoc2, rule ccorres_rhs_assoc2)
974      apply (rule ccorres_split_nothrow[where r'=dc and xf'=xfdc, OF _ ceqv_refl])
975         apply (rule_tac P="ko_at' rv ptr"
976                    in ccorres_from_vcg[where P'=UNIV])
977         apply (rule allI, rule conseqPre, vcg)
978         apply clarsimp
979         apply (rule cmap_relationE1[OF cmap_relation_ep], assumption)
980          apply (erule ko_at_projectKO_opt)
981         apply (clarsimp simp: typ_heap_simps setEndpoint_def)
982         apply (rule rev_bexI)
983          apply (rule setObject_eq; simp add: objBits_simps')[1]
984         apply (clarsimp simp: rf_sr_def cstate_relation_def
985                               Let_def carch_state_relation_def
986                               cmachine_state_relation_def
987                               )
988         apply (clarsimp simp: cpspace_relation_def
989                               update_ep_map_tos
990                               typ_heap_simps')
991         apply (erule(1) cpspace_relation_ep_update_ep2)
992          apply (simp add: cendpoint_relation_def endpoint_state_defs)
993         subgoal by simp
994        apply (rule ccorres_symb_exec_r)
995          apply (rule_tac xs=list in filterM_voodoo)
996          apply (rule_tac P="\<lambda>xs s. (\<forall>x \<in> set xs \<union> set list.
997                   st_tcb_at' (\<lambda>st. isBlockedOnSend st \<and> blockingObject st = ptr) x s)
998                              \<and> distinct (xs @ list) \<and> ko_at' IdleEP ptr s
999                              \<and> (\<forall>p. \<forall>x \<in> set (xs @ list). \<forall>rf. (x, rf) \<notin> {r \<in> state_refs_of' s p. snd r \<noteq> NTFNBound})
1000                              \<and> valid_queues s \<and> pspace_aligned' s \<and> pspace_distinct' s
1001                              \<and> sch_act_wf (ksSchedulerAction s) s \<and> valid_objs' s"
1002                     and P'="\<lambda>xs. {s. ep_queue_relation' (cslift s) (xs @ list)
1003                                         (head_C (queue_' s)) (end_C (queue_' s))}
1004                                \<inter> {s. thread_' s = (case list of [] \<Rightarrow> tcb_Ptr 0
1005                                                       | x # xs \<Rightarrow> tcb_ptr_to_ctcb_ptr x)}"
1006                      in ccorres_inst_voodoo)
1007          apply (induct_tac list)
1008           apply (rule allI)
1009           apply (rule iffD1 [OF ccorres_expand_while_iff_Seq])
1010           apply (rule ccorres_tmp_lift2 [OF _ _ refl])
1011            apply ceqv
1012           apply (simp add: ccorres_cond_iffs)
1013           apply (rule ccorres_rhs_assoc2)
1014           apply (rule ccorres_duplicate_guard, rule ccorres_split_nothrow_novcg_dc)
1015              apply (rule ccorres_from_vcg, rule allI, rule conseqPre, vcg)
1016              apply clarsimp
1017              apply (drule obj_at_ko_at', clarsimp)
1018              apply (rule cmap_relationE1[OF cmap_relation_ep], assumption)
1019               apply (erule ko_at_projectKO_opt)
1020              apply (clarsimp simp: typ_heap_simps tcb_queue_relation'_def)
1021              apply (case_tac x)
1022               apply (clarsimp simp: setEndpoint_def)
1023               apply (rule rev_bexI, rule setObject_eq,
1024                      (simp add: objBits_simps')+)
1025               apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
1026                 carch_state_relation_def
1027                 cmachine_state_relation_def
1028                 cpspace_relation_def typ_heap_simps'
1029                 update_ep_map_tos)
1030               apply (erule(1) cpspace_relation_ep_update_ep2)
1031                subgoal by (simp add: cendpoint_relation_def Let_def)
1032               subgoal by simp
1033              apply (clarsimp simp: tcb_at_not_NULL[OF pred_tcb_at']
1034                                    setEndpoint_def)
1035              apply (rule rev_bexI, rule setObject_eq,
1036                      (simp add: objBits_simps')+)
1037              apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
1038                                     carch_state_relation_def
1039                                     cmachine_state_relation_def
1040                                     cpspace_relation_def typ_heap_simps'
1041                                     update_ep_map_tos)
1042              apply (erule(1) cpspace_relation_ep_update_ep2)
1043               apply (simp add: cendpoint_relation_def Let_def)
1044               apply (subgoal_tac "tcb_at' (last (a # list)) \<sigma> \<and> tcb_at' a \<sigma>")
1045                apply (clarsimp simp: is_aligned_neg_mask_weaken[
1046                                        OF is_aligned_tcb_ptr_to_ctcb_ptr[where P=\<top>]])
1047                subgoal by (simp add: tcb_queue_relation'_def EPState_Send_def mask_def)
1048               subgoal by (auto split: if_split)
1049              subgoal by simp
1050             apply (ctac add: rescheduleRequired_ccorres[unfolded dc_def])
1051            apply (rule hoare_pre, wp weak_sch_act_wf_lift_linear set_ep_valid_objs')
1052            apply (clarsimp simp: weak_sch_act_wf_def sch_act_wf_def)
1053            apply (fastforce simp: valid_ep'_def pred_tcb_at' split: list.splits)
1054           apply (simp add: guard_is_UNIV_def)
1055          apply (rule allI)
1056          apply (rename_tac a lista x)
1057          apply (rule iffD1 [OF ccorres_expand_while_iff_Seq])
1058          apply (rule ccorres_init_tmp_lift2, ceqv)
1059          apply (rule ccorres_guard_imp2)
1060           apply (simp add: bind_assoc dc_def[symmetric]
1061                       del: Collect_const)
1062           apply (rule ccorres_cond_true)
1063           apply (rule ccorres_rhs_assoc)+
1064           apply (rule ccorres_pre_threadGet[where f=tcbState, folded getThreadState_def])
1065           apply (rule ccorres_move_c_guard_tcb)
1066           apply csymbr
1067           apply (rule ccorres_abstract_cleanup)
1068           apply csymbr
1069           apply (rule ccorres_move_c_guard_tcb)
1070           apply (rule_tac P=\<top>
1071                      and P'="{s. ep_queue_relation' (cslift s) (x @ a # lista)
1072                                        (head_C (queue_' s)) (end_C (queue_' s))}"
1073                      and f'="\<lambda>s. s \<lparr> next___ptr_to_struct_tcb_C_' := (case lista of [] \<Rightarrow> tcb_Ptr 0
1074                                              | v # vs \<Rightarrow> tcb_ptr_to_ctcb_ptr v) \<rparr>"
1075                      and xf'="next___ptr_to_struct_tcb_C_'"
1076                           in ccorres_subst_basic_helper)
1077               apply (thin_tac "\<forall>x. P x" for P)
1078               apply (rule myvars.fold_congs, (rule refl)+)
1079               apply (clarsimp simp: tcb_queue_relation'_def use_tcb_queue_relation2
1080                                     tcb_queue_relation2_concat)
1081               apply (clarsimp simp: typ_heap_simps split: list.split)
1082              subgoal by (simp add: rf_sr_def)
1083             apply simp
1084            apply ceqv
1085           apply (rule_tac P="ret__unsigned=blockingIPCBadge rva" in ccorres_gen_asm2)
1086           apply (rule ccorres_if_bind, rule ccorres_if_lhs)
1087            apply (simp add: bind_assoc dc_def[symmetric])
1088            apply (rule ccorres_rhs_assoc)+
1089            apply (ctac add: setThreadState_ccorres)
1090              apply (ctac add: tcbSchedEnqueue_ccorres)
1091                apply (rule_tac P="\<lambda>s. \<forall>t \<in> set (x @ a # lista). tcb_at' t s"
1092                             in ccorres_cross_over_guard)
1093                apply (rule ccorres_add_return, rule ccorres_split_nothrow[OF _ ceqv_refl])
1094                   apply (rule_tac rrel=dc and xf=xfdc
1095                               and P="\<lambda>s. (\<forall>t \<in> set (x @ a # lista). tcb_at' t s)
1096                                          \<and> (\<forall>p. \<forall>t \<in> set (x @ a # lista). \<forall>rf. (t, rf) \<notin> {r \<in> state_refs_of' s p. snd r \<noteq> NTFNBound})
1097                                          \<and> valid_queues s \<and> distinct (x @ a # lista)
1098                                          \<and> pspace_aligned' s \<and> pspace_distinct' s"
1099                              and P'="{s. ep_queue_relation' (cslift s) (x @ a # lista)
1100                                           (head_C (queue_' s)) (end_C (queue_' s))}"
1101                               in ccorres_from_vcg)
1102                   apply (thin_tac "\<forall>x. P x" for P)
1103                   apply (rule allI, rule conseqPre, vcg)
1104                   apply (clarsimp simp: ball_Un)
1105                   apply (rule exI, rule conjI)
1106                    apply (rule exI, erule conjI)
1107                    apply (intro conjI[rotated])
1108                    apply (assumption)
1109                    apply (fold_subgoals (prefix))[3]
1110                    subgoal premises prems using prems by (fastforce intro: pred_tcb_at')+
1111                   apply (clarsimp simp: return_def rf_sr_def cstate_relation_def Let_def)
1112                   apply (rule conjI)
1113                    apply (clarsimp simp: cpspace_relation_def)
1114                    apply (rule conjI, erule ctcb_relation_null_queue_ptrs)
1115                     apply (rule null_ep_queue)
1116                     subgoal by (simp add: o_def)
1117                    apply (rule conjI)
1118                     apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
1119                     apply clarsimp
1120                     apply (rule cendpoint_relation_q_cong)
1121                     apply (rule sym, erule restrict_map_eqI)
1122                     apply (clarsimp simp: image_iff)
1123                     apply (drule(2) map_to_ko_atI2)
1124                     apply (drule ko_at_state_refs_ofD')
1125                     apply clarsimp
1126                     apply (drule_tac x=p in spec)
1127                     subgoal by fastforce
1128
1129                    apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
1130                    apply clarsimp
1131                    apply (drule(2) map_to_ko_atI2, drule ko_at_state_refs_ofD')
1132
1133                    apply (rule cnotification_relation_q_cong)
1134                    apply (rule sym, erule restrict_map_eqI)
1135                    apply (clarsimp simp: image_iff)
1136                    apply (drule_tac x=p in spec)
1137                    subgoal by fastforce
1138                   apply (rule conjI)
1139                    apply (erule cready_queues_relation_not_queue_ptrs,
1140                           auto dest: null_ep_schedD[unfolded o_def] simp: o_def)[1]
1141                   apply (simp add: carch_state_relation_def
1142                                    cmachine_state_relation_def
1143                                    h_t_valid_clift_Some_iff)
1144                  apply (rule ccorres_symb_exec_r2)
1145                    apply (erule spec)
1146                   apply vcg
1147                  apply (vcg spec=modifies)
1148                 apply wp
1149                apply simp
1150                apply vcg
1151               apply (wp hoare_vcg_const_Ball_lift tcbSchedEnqueue_ep_at
1152                         sch_act_wf_lift)
1153              apply simp
1154              apply (vcg exspec=tcbSchedEnqueue_cslift_spec)
1155             apply (wp hoare_vcg_const_Ball_lift sts_st_tcb_at'_cases
1156                       sts_sch_act sts_valid_queues setThreadState_oa_queued)
1157            apply (vcg exspec=setThreadState_cslift_spec)
1158           apply (simp add: ccorres_cond_iffs dc_def[symmetric])
1159           apply (rule ccorres_symb_exec_r2)
1160             apply (drule_tac x="x @ [a]" in spec, simp add: dc_def[symmetric])
1161            apply vcg
1162           apply (vcg spec=modifies)
1163          apply (thin_tac "\<forall>x. P x" for P)
1164          apply (clarsimp simp: pred_tcb_at' ball_Un)
1165          apply (rule conjI)
1166           apply (clarsimp split: if_split)
1167           subgoal by (fastforce simp: valid_tcb_state'_def valid_objs'_maxDomain
1168                                  valid_objs'_maxPriority dest: pred_tcb_at')
1169          apply (clarsimp simp: tcb_at_not_NULL [OF pred_tcb_at'])
1170          apply (clarsimp simp: typ_heap_simps st_tcb_at'_def)
1171          apply (drule(1) obj_at_cslift_tcb)
1172          apply (clarsimp simp: ctcb_relation_blocking_ipc_badge)
1173          apply (rule conjI, simp add: "StrictC'_thread_state_defs" mask_def)
1174          apply (rule conjI)
1175           apply clarsimp
1176           apply (frule rf_sr_cscheduler_relation)
1177           apply (clarsimp simp: cscheduler_action_relation_def st_tcb_at'_def
1178                          split: scheduler_action.split_asm)
1179           apply (rename_tac word)
1180           apply (frule_tac x=word in tcbSchedEnqueue_cslift_precond_discharge)
1181              apply simp
1182             subgoal by clarsimp
1183            subgoal by clarsimp
1184           subgoal by clarsimp
1185          apply clarsimp
1186          apply (rule conjI)
1187           apply (frule(3) tcbSchedEnqueue_cslift_precond_discharge)
1188           subgoal by clarsimp
1189          apply clarsimp
1190          apply (rule context_conjI)
1191           apply (clarsimp simp: tcb_queue_relation'_def)
1192           apply (erule iffD2[OF ep_queue_relation_shift[rule_format], rotated -1])
1193           subgoal by simp
1194          apply (rule_tac x="x @ a # lista" in exI)
1195          apply (clarsimp simp: ball_Un)
1196          apply (rule conjI, fastforce)
1197          subgoal by (clarsimp simp: remove1_append)
1198         apply vcg
1199        apply (rule conseqPre, vcg)
1200        apply clarsimp
1201       apply (wp hoare_vcg_const_Ball_lift)
1202       apply (wp obj_at_setObject3[where 'a=endpoint, folded setEndpoint_def])
1203         apply (simp add: objBits_simps')+
1204       apply (wp set_ep_valid_objs')
1205      apply vcg
1206     apply vcg
1207    apply (rule conseqPre, vcg)
1208    apply clarsimp
1209   apply (clarsimp simp: guard_is_UNIV_def)
1210   apply (erule cmap_relationE1[OF cmap_relation_ep], erule ko_at_projectKO_opt)
1211   apply (clarsimp simp: typ_heap_simps)
1212   apply (clarsimp simp: cendpoint_relation_def Let_def)
1213   subgoal by (clarsimp simp: tcb_queue_relation'_def neq_Nil_conv
1214                  split: if_split_asm)
1215  apply clarsimp
1216  apply (frule ko_at_valid_objs', clarsimp)
1217   apply (simp add: projectKOs)
1218  apply (clarsimp simp: valid_obj'_def valid_ep'_def)
1219  apply (frule sym_refs_ko_atD', clarsimp)
1220  apply (clarsimp simp: st_tcb_at_refs_of_rev')
1221  apply (rule conjI)
1222   subgoal by (auto simp: isBlockedOnSend_def elim!: pred_tcb'_weakenE)
1223  apply (rule conjI)
1224   apply (clarsimp split: if_split)
1225   apply (drule sym_refsD, clarsimp)
1226   apply (drule(1) bspec)+
1227   by (auto simp: obj_at'_def projectKOs state_refs_of'_def pred_tcb_at'_def tcb_bound_refs'_def
1228              dest!: symreftype_inverse')
1229
1230
1231lemma tcb_ptr_to_ctcb_ptr_force_fold:
1232  "x + 2 ^ ctcb_size_bits = ptr_val (tcb_ptr_to_ctcb_ptr x)"
1233  by (simp add: tcb_ptr_to_ctcb_ptr_def ctcb_offset_def)
1234
1235
1236lemma coerce_memset_to_heap_update:
1237  "heap_update_list x (replicateHider (size_of (TYPE (tcb_C))) 0)
1238      = heap_update (tcb_Ptr x)
1239             (tcb_C (arch_tcb_C (user_context_C (FCP (\<lambda>x. 0))) NULL)
1240                    (thread_state_C (FCP (\<lambda>x. 0)))
1241                    (NULL)
1242                    (seL4_Fault_C (FCP (\<lambda>x. 0)))
1243                    (lookup_fault_C (FCP (\<lambda>x. 0)))
1244                      0 0 0 0 0 0 NULL NULL NULL NULL)"
1245  apply (intro ext, simp add: heap_update_def)
1246  apply (rule_tac f="\<lambda>xs. heap_update_list x xs a b" for a b in arg_cong)
1247  apply (simp add: to_bytes_def size_of_def typ_info_simps tcb_C_tag_def)
1248  apply (simp add: ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td align_of_def padup_def
1249                   final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def)
1250  apply (simp add: typ_info_simps
1251                   user_context_C_tag_def thread_state_C_tag_def seL4_Fault_C_tag_def
1252                   lookup_fault_C_tag_def update_ti_t_ptr_0s arch_tcb_C_tag_def
1253                   ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td
1254                   ti_typ_combine_empty_ti ti_typ_combine_td
1255                   align_of_def padup_def
1256                   final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def
1257                   align_td_array' size_td_array)
1258  apply (simp add: typ_info_array')
1259  apply (simp add: typ_info_word word_rsplit_0 upt_conv_Cons)
1260  apply (simp add: typ_info_word typ_info_ptr word_rsplit_0
1261                   replicateHider_def)
1262  done
1263
1264lemma isArchObjectCap_capBits:
1265  "isArchObjectCap cap \<Longrightarrow> capBits cap = acapBits (capCap cap)"
1266  by (clarsimp simp: isCap_simps)
1267
1268declare Kernel_C.tcb_C_size [simp del]
1269
1270lemma cte_lift_ccte_relation:
1271  "cte_lift cte' = Some ctel'
1272    \<Longrightarrow> c_valid_cte cte'
1273    \<Longrightarrow> ccte_relation (cte_to_H ctel') cte'"
1274  by (simp add: ccte_relation_def)
1275
1276lemma updateFreeIndex_ccorres:
1277  "\<forall>s. \<Gamma> \<turnstile> ({s} \<inter> {s. \<exists>cte cte'. cslift s (cte_Ptr srcSlot) = Some cte'
1278               \<and> cteCap cte = cap' \<and> ccte_relation cte cte'})
1279          c
1280        {t. \<exists>cap. cap_untyped_cap_lift cap = (cap_untyped_cap_lift
1281                    (cte_C.cap_C (the (cslift s (cte_Ptr srcSlot)))))
1282                        \<lparr> cap_untyped_cap_CL.capFreeIndex_CL := ((of_nat idx') >> 4) \<rparr>
1283                \<and> cap_get_tag cap = scast cap_untyped_cap
1284                \<and> t_hrs_' (globals t) = hrs_mem_update (heap_update (cte_Ptr srcSlot)
1285                    (cte_C.cap_C_update (\<lambda>_. cap) (the (cslift s (cte_Ptr srcSlot)))))
1286                    (t_hrs_' (globals s))
1287                \<and> t may_only_modify_globals s in [t_hrs]
1288        }
1289    \<Longrightarrow> ccorres dc xfdc
1290           (valid_objs' and cte_wp_at' (\<lambda>cte. isUntypedCap (cteCap cte)
1291               \<and> cap' = (cteCap cte)) srcSlot
1292           and untyped_ranges_zero'
1293           and (\<lambda>_. is_aligned (of_nat idx' :: word32) 4 \<and> idx' \<le> 2 ^ (capBlockSize cap')))
1294           {s. \<not> capIsDevice cap'
1295               \<longrightarrow> region_actually_is_zero_bytes (capPtr cap' + of_nat idx') (capFreeIndex cap' - idx') s} hs
1296           (updateFreeIndex srcSlot idx') c"
1297  (is "_ \<Longrightarrow> ccorres dc xfdc (valid_objs' and ?cte_wp_at' and _ and _) ?P' hs ?a c")
1298  apply (rule ccorres_gen_asm)
1299  apply (simp add: updateFreeIndex_def getSlotCap_def updateCap_def)
1300  apply (rule ccorres_guard_imp2)
1301   apply (rule ccorres_split_noop_lhs, rule_tac cap'=cap' in updateTrackedFreeIndex_noop_ccorres)
1302    apply (rule ccorres_pre_getCTE)+
1303    apply (rename_tac cte cte2)
1304    apply (rule_tac P = "\<lambda>s. ?cte_wp_at' s \<and> cte2 = cte \<and> cte_wp_at' ((=) cte) srcSlot s"
1305              and P'="{s. \<exists>cte cte'. cslift s (cte_Ptr srcSlot) = Some cte'
1306               \<and> cteCap cte = cap' \<and> ccte_relation cte cte'} \<inter> ?P'" in ccorres_from_vcg)
1307    apply (rule allI, rule HoarePartial.conseq_exploit_pre, clarify)
1308    apply (drule_tac x=s in spec, rule conseqPre, erule conseqPost)
1309      defer
1310      apply clarsimp
1311     apply clarsimp
1312    apply (simp add: cte_wp_at_ctes_of)
1313    apply wp
1314   apply (clarsimp simp: isCap_simps cte_wp_at_ctes_of)
1315   apply (frule(1) rf_sr_ctes_of_clift)
1316   apply clarsimp
1317   apply (frule(1) cte_lift_ccte_relation)
1318   apply (rule exI, intro conjI[rotated], assumption, simp_all)[1]
1319  apply (clarsimp simp: cte_wp_at_ctes_of)
1320  apply (erule(1) rf_sr_ctes_of_cliftE)
1321  apply (frule(1) rf_sr_ctes_of_clift)
1322  apply clarsimp
1323  apply (subgoal_tac "ccap_relation (capFreeIndex_update (\<lambda>_. idx')
1324        (cteCap (the (ctes_of \<sigma> srcSlot)))) cap")
1325   apply (rule fst_setCTE [OF ctes_of_cte_at], assumption)
1326   apply (erule bexI [rotated])
1327   apply (clarsimp simp add: rf_sr_def cstate_relation_def Let_def
1328      cvariable_array_map_const_add_map_option[where f="tcb_no_ctes_proj"]
1329      isCap_simps)
1330   apply (simp add:cpspace_relation_def)
1331   apply (clarsimp simp:typ_heap_simps' modify_map_def mex_def meq_def)
1332   apply (rule conjI)
1333    apply (rule cpspace_cte_relation_upd_capI, assumption+)
1334   apply (rule conjI)
1335    apply (rule setCTE_tcb_case, assumption+)
1336   apply (case_tac s', clarsimp)
1337   subgoal by (simp add: carch_state_relation_def cmachine_state_relation_def
1338                   typ_heap_simps')
1339
1340  apply (clarsimp simp: isCap_simps)
1341  apply (drule(1) cte_lift_ccte_relation,
1342    drule ccte_relation_ccap_relation)
1343  apply (simp add: cte_to_H_def)
1344  apply (frule cap_get_tag_isCap_unfolded_H_cap)
1345  apply (clarsimp simp: ccap_relation_def cap_lift_untyped_cap
1346                        cap_to_H_simps cap_untyped_cap_lift_def
1347                        is_aligned_shiftr_shiftl
1348                 dest!: ccte_relation_ccap_relation)
1349  apply (rule unat_of_nat_eq unat_of_nat_eq[symmetric],
1350    erule order_le_less_trans,
1351    rule power_strict_increasing, simp_all)
1352  apply (rule unat_less_helper, rule order_le_less_trans[OF word_and_le1], simp add: mask_def)
1353  done
1354
1355end
1356
1357(* FIXME: Move *)
1358lemma ccap_relation_isDeviceCap:
1359 "\<lbrakk>ccap_relation cp cap; isUntypedCap cp
1360  \<rbrakk> \<Longrightarrow> to_bool (capIsDevice_CL (cap_untyped_cap_lift cap)) =  (capIsDevice cp)"
1361  apply (frule cap_get_tag_UntypedCap)
1362  apply (simp add:cap_get_tag_isCap )
1363  done
1364
1365lemma ccap_relation_isDeviceCap2:
1366 "\<lbrakk>ccap_relation cp cap; isUntypedCap cp
1367  \<rbrakk> \<Longrightarrow> (capIsDevice_CL (cap_untyped_cap_lift cap) = 0) = (\<not> (capIsDevice cp))"
1368  apply (frule cap_get_tag_UntypedCap)
1369  apply (simp add:cap_get_tag_isCap to_bool_def)
1370  done
1371end
1372