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 4 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: replicateHider_def)
427  done
428
429lemma coerce_memset_to_heap_update_pde:
430  "heap_update_list x (replicateHider 4 0)
431      = heap_update (Ptr x :: pde_C ptr)
432             (pde_C.pde_C (FCP (\<lambda>x. 0)))"
433  apply (intro ext, simp add: heap_update_def)
434  apply (rule_tac f="\<lambda>xs. heap_update_list x xs a b" for a b in arg_cong)
435  apply (simp add: to_bytes_def size_of_def typ_info_simps pde_C_tag_def)
436  apply (simp add: ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td align_of_def padup_def
437                   final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def)
438  apply (simp add: typ_info_simps align_td_array' size_td_array)
439  apply (simp add: typ_info_array' typ_info_word word_rsplit_0)
440  apply (simp add: replicateHider_def)
441  done
442
443lemma objBits_eq_by_type:
444  fixes x :: "'a :: pspace_storable" and y :: 'a
445  shows "objBits x = objBits y"
446  apply (simp add: objBits_def)
447  apply (rule objBits_type)
448  apply (simp add: koTypeOf_injectKO)
449  done
450
451lemma mapM_x_store_memset_ccorres_assist:
452  fixes val :: "'a :: pspace_storable"
453  assumes nofail: "\<not> snd (mapM_x (\<lambda>slot. setObject slot val) slots \<sigma>)"
454  assumes slots1: "\<forall>n < length slots. slots ! n = hd slots + (of_nat n << objBits val)"
455  assumes slots2: "n = length slots * (2 ^ objBits val)"
456  assumes ptr: "ptr = hd slots"
457  assumes ko: "\<And>ko :: 'a. updateObject ko = updateObject_default ko"
458              "\<And>ko :: 'a. (1 :: word32) < 2 ^ objBits ko"
459  assumes restr: "set slots \<subseteq> S"
460  assumes worker: "\<And>ptr s s' (ko :: 'a). \<lbrakk> (s, s') \<in> rf_sr; ko_at' ko ptr s; ptr \<in> S \<rbrakk>
461                                \<Longrightarrow> (s \<lparr> ksPSpace := ksPSpace s (ptr \<mapsto> injectKO val)\<rparr>,
462                                     globals_update (t_hrs_'_update (hrs_mem_update
463                                                    (heap_update_list ptr
464                                                    (replicateHider (2 ^ objBits val) (ucast c))))) s') \<in> rf_sr"
465  assumes rf_sr: "(\<sigma>, s) \<in> rf_sr"
466  shows
467  "\<exists>(rv, \<sigma>') \<in> fst (mapM_x (\<lambda>slot. setObject slot val) slots \<sigma>).
468      (\<sigma>', globals_update (t_hrs_'_update (hrs_mem_update
469                          (heap_update_list ptr (replicateHider n c)))) s) \<in> rf_sr"
470  unfolding slots2 ptr using rf_sr slots1 nofail restr
471proof (induct slots arbitrary: s \<sigma>)
472  case Nil
473  show ?case
474    using Nil.prems
475    apply (simp add: mapM_x_def sequence_x_def return_def replicateHider_def)
476    apply (simp add: rf_sr_def hrs_mem_update_def cstate_relation_def Let_def
477                     carch_state_relation_def cmachine_state_relation_def
478                     h_t_valid_clift_Some_iff)
479    done
480next
481  case (Cons x xs tPre sPre)
482
483  note nofail_bind = Cons.prems(3)[unfolded mapM_x_Cons K_bind_def]
484
485  have obj_at: "obj_at' (\<lambda>x :: 'a. True) x sPre"
486    using not_snd_bindI1[OF nofail_bind]
487    apply (subst(asm) setObject_obj_at_pre, simp_all add: ko snd_bind)
488    apply (clarsimp simp: stateAssert_def exec_get return_def)
489    apply (simp add: koTypeOf_injectKO typ_at_to_obj_at')
490    done
491
492  note in_setObject = setObject_eq[OF _ _ objBits_eq_by_type obj_at,
493                                   where ko=val, simplified ko, simplified]
494
495  note nofail_mapM = not_snd_bindI2[OF nofail_bind, OF in_setObject]
496
497  have hd_xs: "xs \<noteq> [] \<Longrightarrow> hd xs = x + (2 ^ objBits val)"
498    using Cons.prems(2)[rule_format, where n=1]
499    by (simp add: hd_conv_nth)
500
501  show ?case
502    using obj_at_ko_at'[OF obj_at] Cons.prems(4)
503    apply (clarsimp simp add: mapM_x_Cons bind_def split_def)
504    apply (rule rev_bexI, rule in_setObject)
505    apply (cut_tac Cons.hyps[OF _ _ nofail_mapM])
506       defer
507       apply (rule worker, rule Cons.prems, assumption+)
508      apply clarsimp
509      apply (case_tac "xs = []", simp_all)[1]
510      apply (insert Cons.prems, simp)[1]
511      apply (frule_tac x="Suc n" in spec)
512      apply (simp add: hd_xs shiftl_t2n field_simps)
513     apply assumption
514    apply clarsimp
515    apply (rule rev_bexI, assumption)
516    apply (simp add: o_def)
517    apply (case_tac "xs = []")
518     apply (simp add: hrs_mem_update_def split_def replicateHider_def)
519    apply (subst(asm) heap_update_list_concat_fold_hrs_mem)
520     apply (simp add: hd_xs replicateHider_def)
521    apply (simp add: replicateHider_def replicate_add)
522    done
523qed
524
525lemma invalidateTLBByASID_ccorres:
526  "ccorres dc xfdc (valid_pde_mappings' and (\<lambda>_. asid \<le> mask asid_bits))
527                   (UNIV \<inter> {s. asid_' s = asid}) []
528       (invalidateTLBByASID asid) (Call invalidateTLBByASID_'proc)"
529  apply (cinit lift: asid_')
530   apply (ctac(no_vcg) add: loadHWASID_ccorres)
531    apply csymbr
532    apply (simp add: case_option_If2 del: Collect_const)
533    apply (rule ccorres_if_cond_throws2[where Q=\<top> and Q'=\<top>])
534       apply (clarsimp simp: pde_stored_asid_def to_bool_def split: if_split)
535      apply (rule ccorres_return_void_C[unfolded dc_def])
536     apply (simp add: dc_def[symmetric])
537     apply csymbr
538     apply (ctac add: invalidateTranslationASID_ccorres)
539    apply vcg
540   apply (wp hoare_drop_imps)
541  apply (clarsimp simp: pde_stored_asid_def to_bool_def)
542  done
543
544(* FIXME: also in VSpace_C *)
545lemma ignoreFailure_liftM:
546  "ignoreFailure = liftM (\<lambda>v. ())"
547  apply (rule ext)+
548  apply (simp add: ignoreFailure_def liftM_def
549                   catch_def)
550  apply (rule bind_apply_cong[OF refl])
551  apply (simp split: sum.split)
552  done
553
554end
555
556lemma option_to_0_user_mem':
557  "option_to_0 \<circ> user_mem' as =(\<lambda>x. if x \<in> {y. \<not> pointerInUserData y as} then 0
558  else underlying_memory (ksMachineState as) x) "
559  apply (rule ext)
560  apply (simp add:user_mem'_def option_to_0_def split:if_splits)
561  done
562
563lemma heap_to_user_data_in_user_mem'[simp]:
564  "\<lbrakk>pspace_aligned' as;pspace_distinct' as\<rbrakk> \<Longrightarrow> heap_to_user_data (ksPSpace as) (option_to_0 \<circ> user_mem' as) =
565  heap_to_user_data (ksPSpace as)(underlying_memory (ksMachineState as))"
566  apply (rule ext)+
567  apply (clarsimp simp: heap_to_user_data_def option_map_def
568                 split: option.splits)
569  apply (subst option_to_0_user_mem')
570  apply (subst map_option_byte_to_word_heap)
571   apply (clarsimp simp: projectKO_opt_user_data map_comp_def
572                  split: option.split_asm kernel_object.split_asm)
573   apply (frule(1) pspace_alignedD')
574   apply (frule(1) pspace_distinctD')
575   apply (subgoal_tac "x + ucast off * 4 + xa  && ~~ mask pageBits = x" )
576    apply (clarsimp simp: pointerInUserData_def typ_at'_def ko_wp_at'_def)
577   apply (simp add: ARM.pageBits_def)
578   apply (subst mask_lower_twice2[where n = 2 and m = 12,simplified,symmetric])
579   apply (subst is_aligned_add_helper[THEN conjunct2,where n1 = 2])
580     apply (erule aligned_add_aligned)
581      apply (simp add: is_aligned_mult_triv2[where n = 2,simplified])
582     apply  (clarsimp simp: objBits_simps ARM.pageBits_def)
583    apply simp
584   apply (rule is_aligned_add_helper[THEN conjunct2])
585    apply (simp add: ARM.pageBits_def objBits_simps)
586   apply (rule word_less_power_trans2[where k = 2,simplified])
587     apply (rule less_le_trans[OF ucast_less])
588      apply simp+
589  done
590
591context begin interpretation Arch . (*FIXME: arch_split*)
592
593crunch pde_mappings'[wp]: invalidateTLBByASID "valid_pde_mappings'"
594crunch ksArchState[wp]: invalidateTLBByASID "\<lambda>s. P (ksArchState s)"
595
596crunch gsMaxObjectSize[wp]: invalidateTLBByASID "\<lambda>s. P (gsMaxObjectSize s)"
597crunch gsMaxObjectSize[wp]: deleteASIDPool "\<lambda>s. P (gsMaxObjectSize s)"
598  (ignore: setObject getObject wp: crunch_wps getObject_inv loadObject_default_inv
599     simp: crunch_simps)
600end
601
602context kernel_m begin
603
604lemma page_table_at_rf_sr_dom_s:
605  "\<lbrakk> page_table_at' x s; (s, s') \<in> rf_sr \<rbrakk>
606    \<Longrightarrow> {x ..+ 2 ^ ptBits} \<times> {SIndexVal, SIndexTyp 0}
607    \<subseteq> dom_s (hrs_htd (t_hrs_' (globals s')))"
608  apply (rule_tac m=2 in intvl_2_power_times_decomp,
609         simp_all add: shiftl_t2n field_simps ptBits_def pageBits_def
610                       word_bits_def pteBits_def)
611  apply (clarsimp simp: page_table_at'_def intvl_def)
612  apply (drule spec, drule(1) mp)
613  apply (simp add: typ_at_to_obj_at_arches)
614  apply (drule obj_at_ko_at', clarsimp)
615  apply (erule cmap_relationE1[OF rf_sr_cpte_relation])
616   apply (erule ko_at_projectKO_opt)
617  apply (drule h_t_valid_clift)
618  apply (drule h_t_valid_dom_s[OF _ refl refl])
619  apply (erule subsetD)
620  apply (auto simp add: intvl_def shiftl_t2n)[1]
621  done
622
623lemma page_directory_at_rf_sr_dom_s:
624  "\<lbrakk> page_directory_at' x s; (s, s') \<in> rf_sr \<rbrakk>
625    \<Longrightarrow> {x ..+ 2 ^ pdBits} \<times> {SIndexVal, SIndexTyp 0}
626    \<subseteq> dom_s (hrs_htd (t_hrs_' (globals s')))"
627  apply (rule_tac m=2 in intvl_2_power_times_decomp,
628         simp_all add: shiftl_t2n field_simps pdBits_def pageBits_def
629                       word_bits_def pdeBits_def)
630  apply (clarsimp simp: page_directory_at'_def intvl_def)
631  apply (drule spec, drule(1) mp)
632  apply (simp add: typ_at_to_obj_at_arches)
633  apply (drule obj_at_ko_at', clarsimp)
634  apply (erule cmap_relationE1[OF rf_sr_cpde_relation])
635   apply (erule ko_at_projectKO_opt)
636  apply (drule h_t_valid_clift)
637  apply (drule h_t_valid_dom_s[OF _ refl refl])
638  apply (erule subsetD)
639  apply (auto simp add: intvl_def shiftl_t2n)[1]
640  done
641
642lemma clearMemory_setObject_PTE_ccorres:
643  "ccorres dc xfdc (page_table_at' ptr
644                and (\<lambda>s. 2 ^ ptBits \<le> gsMaxObjectSize s)
645                and (\<lambda>_. is_aligned ptr ptBits \<and> ptr \<noteq> 0 \<and> pstart = addrFromPPtr ptr))
646            (UNIV \<inter> {s. ptr___ptr_to_unsigned_long_' s = Ptr ptr} \<inter> {s. bits_' s = of_nat ptBits}) []
647       (do x \<leftarrow> mapM_x (\<lambda>a. setObject a ARM_H.InvalidPTE)
648                       [ptr , ptr + 2 ^ pteBits .e. ptr + 2 ^ ptBits - 1];
649           doMachineOp (cleanCacheRange_PoU ptr (ptr + 2 ^ ptBits - 1) pstart)
650        od)
651       (Call clearMemory_'proc)"
652  apply (rule ccorres_gen_asm)+
653  apply (cinit' lift: ptr___ptr_to_unsigned_long_' bits_')
654   apply (rule ccorres_Guard_Seq)
655   apply (rule ccorres_split_nothrow_novcg_dc)
656      apply (rule_tac P="page_table_at' ptr and (\<lambda>s. 2 ^ ptBits \<le> gsMaxObjectSize s)"
657               in ccorres_from_vcg_nofail[where P'=UNIV])
658      apply (rule allI, rule conseqPre, vcg)
659      apply clarsimp
660      apply (subst ghost_assertion_size_logic[unfolded o_def])
661        apply (simp add: ptBits_def pageBits_def pteBits_def)
662       apply simp
663      apply (clarsimp simp: replicateHider_def[symmetric])
664      apply (frule is_aligned_no_overflow')
665      apply (intro conjI)
666          apply (clarsimp simp add: ptBits_def pageBits_def pteBits_def
667                          cong: StateSpace.state.fold_congs globals.fold_congs)
668         apply (erule is_aligned_weaken, simp add: ptBits_def pageBits_def)
669        apply (clarsimp simp: is_aligned_def ptBits_def pageBits_def pteBits_def)
670       apply (simp add: unat_of_nat32 order_less_le_trans[OF pt_bits_stuff(2)]
671                        word_bits_def page_table_at_rf_sr_dom_s)
672      apply (clarsimp simp add: ptBits_def pageBits_def pteBits_def
673                      cong: StateSpace.state.fold_congs globals.fold_congs)
674      apply (simp add: upto_enum_step_def objBits_simps ptBits_def pageBits_def
675                       field_simps linorder_not_less[symmetric] archObjSize_def
676                       upto_enum_word split_def pteBits_def)
677      apply (erule mapM_x_store_memset_ccorres_assist
678                      [unfolded split_def, OF _ _ _ _ _ _ subset_refl],
679             simp_all add: shiftl_t2n hd_map objBits_simps archObjSize_def pteBits_def)[1]
680      apply (rule cmap_relationE1, erule rf_sr_cpte_relation, erule ko_at_projectKO_opt)
681      apply (subst coerce_memset_to_heap_update_pte)
682      apply (clarsimp simp: rf_sr_def Let_def cstate_relation_def typ_heap_simps)
683      apply (rule conjI)
684       apply (simp add: cpspace_relation_def typ_heap_simps update_pte_map_tos
685                        update_pte_map_to_ptes carray_map_relation_upd_triv)
686       apply (rule cmap_relation_updI, simp_all)[1]
687       apply (simp add: cpte_relation_def Let_def pte_lift_def
688                        fcp_beta pte_get_tag_def pte_tag_defs)
689       apply (simp add: carch_state_relation_def cmachine_state_relation_def
690                        typ_heap_simps update_pte_map_tos)
691      apply csymbr
692     apply (rule ccorres_Guard)
693     apply (ctac add: cleanCacheRange_PoU_ccorres)
694    apply (wp mapM_x_wp' setObject_ksPSpace_only updateObject_default_inv | simp)+
695   apply (clarsimp simp: guard_is_UNIV_def ptBits_def pageBits_def pteBits_def)
696  apply (clarsimp simp: ptBits_def pageBits_def pteBits_def)
697  apply (frule is_aligned_addrFromPPtr_n, simp)
698  apply (clarsimp simp: is_aligned_no_overflow'[where n=10, simplified] pageBits_def
699                        field_simps is_aligned_mask[symmetric] mask_AND_less_0)
700  done
701
702lemma ccorres_make_xfdc:
703  "ccorresG rf_sr \<Gamma> r xf P P' h a c \<Longrightarrow> ccorresG rf_sr \<Gamma> dc xfdc P P' h a c"
704  apply (erule ccorres_rel_imp)
705  apply simp
706  done
707
708lemma ccorres_if_True_False_simps:
709  "ccorres r xf P P' hs a (IF True THEN c ELSE c' FI) = ccorres r xf P P' hs a c"
710  "ccorres r xf P P' hs a (IF False THEN c ELSE c' FI) = ccorres r xf P P' hs a c'"
711  "ccorres r xf P P' hs a (IF True THEN c ELSE c' FI ;; d) = ccorres r xf P P' hs a (c ;; d)"
712  "ccorres r xf P P' hs a (IF False THEN c ELSE c' FI ;; d) = ccorres r xf P P' hs a (c' ;; d)"
713  by (simp_all add: ccorres_cond_iffs ccorres_seq_simps)
714
715lemmas cap_tag_values =
716  cap_untyped_cap_def
717  cap_endpoint_cap_def
718  cap_notification_cap_def
719  cap_reply_cap_def
720  cap_cnode_cap_def
721  cap_thread_cap_def
722  cap_irq_handler_cap_def
723  cap_null_cap_def
724  cap_irq_control_cap_def
725  cap_zombie_cap_def
726  cap_small_frame_cap_def
727  cap_frame_cap_def
728  cap_page_table_cap_def
729  cap_page_directory_cap_def
730  cap_asid_pool_cap_def
731
732lemma ccorres_return_C_seq:
733  "\<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>
734  \<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)"
735  apply (rule ccorres_guard_imp)
736  apply (rule ccorres_split_throws, rule ccorres_return_C, simp+)
737  apply vcg
738  apply simp_all
739  done
740
741lemma ccap_relation_VPIsDevice:
742  "\<lbrakk>isPageCap cp; ccap_relation (ArchObjectCap cp) cap  \<rbrakk> \<Longrightarrow>
743   (generic_frame_cap_get_capFIsDevice_CL (cap_lift cap) = 0) = (\<not> (capVPIsDevice cp))"
744   by (clarsimp elim!:ccap_relationE
745     simp : isPageCap_def generic_frame_cap_get_capFIsDevice_CL_def cap_to_H_def
746            Let_def to_bool_def
747     split: arch_capability.split_asm cap_CL.split_asm if_split_asm)
748
749
750lemma ccap_relation_get_capZombiePtr_CL:
751  "\<lbrakk> ccap_relation cap cap'; isZombie cap; capAligned cap \<rbrakk>
752      \<Longrightarrow> get_capZombiePtr_CL (cap_zombie_cap_lift cap') = capZombiePtr cap"
753  apply (simp only: cap_get_tag_isCap[symmetric])
754  apply (drule(1) cap_get_tag_to_H)
755  apply (clarsimp simp: get_capZombiePtr_CL_def get_capZombieBits_CL_def Let_def split: if_split)
756  apply (subst less_mask_eq)
757   apply (clarsimp simp add: capAligned_def objBits_simps word_bits_conv)
758   apply unat_arith
759  apply simp
760  done
761
762lemma modify_gets_helper:
763  "do y \<leftarrow> modify (ksPSpace_update (\<lambda>_. ps)); ps' \<leftarrow> gets ksPSpace; f ps' od
764      = do y \<leftarrow> modify (ksPSpace_update (\<lambda>_. ps)); f ps od"
765  by (simp add: bind_def simpler_modify_def simpler_gets_def)
766
767lemma snd_lookupAround2_update:
768  "ps y \<noteq> None \<Longrightarrow>
769    snd (lookupAround2 x (ps (y \<mapsto> v'))) = snd (lookupAround2 x ps)"
770  apply (clarsimp simp: lookupAround2_def lookupAround_def Let_def
771                        dom_fun_upd2
772              simp del: dom_fun_upd cong: if_cong option.case_cong)
773  apply (clarsimp split: option.split if_split  cong: if_cong)
774  apply auto
775  done
776
777lemma double_setEndpoint:
778  "do y \<leftarrow> setEndpoint epptr v1; setEndpoint epptr v2 od
779       = setEndpoint epptr v2"
780  apply (simp add: setEndpoint_def setObject_def bind_assoc split_def
781                   modify_gets_helper)
782  apply (simp add: updateObject_default_def bind_assoc objBits_simps)
783  apply (rule ext)
784  apply (rule bind_apply_cong, rule refl)+
785  apply (clarsimp simp add: in_monad projectKOs magnitudeCheck_assert
786                            snd_lookupAround2_update)
787  apply (simp add: lookupAround2_known1 assert_opt_def projectKO_def projectKO_opt_ep
788                    alignCheck_assert)
789  apply (simp add: bind_def simpler_modify_def)
790  done
791
792lemma filterM_setEndpoint_adjustment:
793  "\<lbrakk> \<And>v. do setEndpoint epptr IdleEP; body v od
794          = do v' \<leftarrow> body v; setEndpoint epptr IdleEP; return v' od \<rbrakk>
795    \<Longrightarrow>
796   (do q' \<leftarrow> filterM body q; setEndpoint epptr (f q') od)
797    = (do setEndpoint epptr IdleEP; q' \<leftarrow> filterM body q; setEndpoint epptr (f q') od)"
798  apply (rule sym)
799  apply (induct q arbitrary: f)
800   apply (simp add: double_setEndpoint)
801  apply (simp add: bind_assoc)
802  apply (subst bind_assoc[symmetric], simp, simp add: bind_assoc)
803  done
804
805lemma ccorres_inst_voodoo:
806  "\<forall>x. ccorres r xf (P x) (P' x) hs (h x) (c x)
807     \<Longrightarrow> \<forall>x. ccorres r xf (P x) (P' x) hs (h x) (c x)"
808  by simp
809
810lemma cpspace_relation_ep_update_ep2:
811  "\<lbrakk> ko_at' (ep :: endpoint) epptr s;
812      cmap_relation (map_to_eps (ksPSpace s))
813           (cslift t) ep_Ptr (cendpoint_relation (cslift t));
814      cendpoint_relation (cslift t') ep' endpoint;
815      (cslift t' :: tcb_C ptr \<rightharpoonup> tcb_C) = cslift t \<rbrakk>
816     \<Longrightarrow> cmap_relation (map_to_eps (ksPSpace s(epptr \<mapsto> KOEndpoint ep')))
817          (cslift t(ep_Ptr epptr \<mapsto> endpoint))
818          ep_Ptr (cendpoint_relation (cslift t'))"
819  apply (rule cmap_relationE1, assumption, erule ko_at_projectKO_opt)
820  apply (rule_tac P="\<lambda>a. cmap_relation a b c d" for b c d in rsubst,
821                   erule cmap_relation_upd_relI, assumption+)
822    apply simp+
823  apply (rule ext, simp add: map_comp_def projectKO_opt_ep split: if_split)
824  done
825
826end
827
828context begin interpretation Arch . (*FIXME: arch_split*)
829lemma setObject_tcb_ep_obj_at'[wp]:
830  "\<lbrace>obj_at' (P :: endpoint \<Rightarrow> bool) ptr\<rbrace> setObject ptr' (tcb :: tcb) \<lbrace>\<lambda>rv. obj_at' P ptr\<rbrace>"
831  apply (rule obj_at_setObject2, simp_all)
832  apply (clarsimp simp: updateObject_default_def in_monad)
833  done
834end
835
836crunch ep_obj_at'[wp]: setThreadState "obj_at' (P :: endpoint \<Rightarrow> bool) ptr"
837  (ignore: getObject setObject simp: unless_def)
838
839context kernel_m begin
840
841lemma ccorres_abstract_h_val:
842  "(\<And>rv. P rv \<Longrightarrow> ccorres r xf G (G' rv) hs a c) \<Longrightarrow>
843   ccorres r xf G ({s. P (h_val (hrs_mem (t_hrs_' (globals s))) p)
844            \<longrightarrow> s \<in> G' (h_val (hrs_mem (t_hrs_' (globals s)))
845            p)}
846   \<inter> {s. P (h_val (hrs_mem (t_hrs_' (globals s))) p)}) hs a c"
847   apply (rule ccorres_tmp_lift1 [where P = P])
848   apply (clarsimp simp: Collect_conj_eq [symmetric])
849   apply (fastforce intro: ccorres_guard_imp)
850   done
851
852lemma ccorres_subst_basic_helper:
853  "\<lbrakk> \<And>s s'. \<lbrakk> P s; s' \<in> P'; (s, s') \<in> rf_sr \<rbrakk> \<Longrightarrow> f s' = f' s';
854     \<And>s s'. \<lbrakk> P s; s' \<in> P'; (s, s') \<in> rf_sr \<rbrakk> \<Longrightarrow> (s, f' s') \<in> rf_sr;
855     \<And>s'. xf' (f' s') = v; \<And>rv' t t'. ceqv \<Gamma> xf' rv' t t' c (c' rv');
856        ccorres rrel xf Q Q' hs a (c' v) \<rbrakk>
857       \<Longrightarrow> ccorres rrel xf (P and Q) {s. s \<in> P' \<and> f' s \<in> Q'} hs a (Basic f ;; c)"
858  apply (rule ccorres_guard_imp2)
859   apply (rule ccorres_add_return)
860   apply (rule ccorres_split_nothrow[where xf'=xf' and r'="\<lambda>rv rv'. rv' = v"])
861       apply (rule ccorres_from_vcg[where P=P and P'=P'])
862       apply (rule allI, rule conseqPre, vcg)
863       apply (clarsimp simp: return_def)
864      apply assumption
865     apply simp
866    apply wp
867   apply vcg
868  apply clarsimp
869  done
870
871lemma ctcb_relation_blocking_ipc_badge:
872  "\<lbrakk> ctcb_relation tcb ctcb; isBlockedOnSend (tcbState tcb) \<rbrakk> \<Longrightarrow>
873      tsType_CL (thread_state_lift (tcbState_C ctcb)) = scast ThreadState_BlockedOnSend"
874  "\<lbrakk> ctcb_relation tcb ctcb;
875        tsType_CL (thread_state_lift (tcbState_C ctcb)) = scast ThreadState_BlockedOnSend \<rbrakk>
876      \<Longrightarrow> blockingIPCBadge (tcbState tcb)
877             = blockingIPCBadge_CL (thread_state_lift (tcbState_C ctcb))"
878   apply (clarsimp simp add: ctcb_relation_def)
879   apply (simp add: isBlockedOnSend_def split: Structures_H.thread_state.split_asm)
880   apply (clarsimp simp: cthread_state_relation_def)
881  apply (clarsimp simp add: ctcb_relation_def cthread_state_relation_def)
882  apply (cases "tcbState tcb", simp_all add: "StrictC'_thread_state_defs")
883  done
884
885lemma cendpoint_relation_q_cong:
886  "\<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>
887      \<Longrightarrow> cendpoint_relation hp ep ep' = cendpoint_relation hp' ep ep'"
888  apply (cases ep, simp_all add: cendpoint_relation_def Let_def)
889   apply (rule conj_cong [OF refl])
890   apply (rule tcb_queue_relation'_cong[OF refl refl refl])
891   apply clarsimp
892  apply (rule conj_cong [OF refl])
893  apply (rule tcb_queue_relation'_cong[OF refl refl refl])
894  apply clarsimp
895  done
896
897lemma cnotification_relation_q_cong:
898  "\<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>
899      \<Longrightarrow>  cnotification_relation hp ntfn ntfn' = cnotification_relation hp' ntfn ntfn'"
900  apply (cases "ntfnObj ntfn", simp_all add: cnotification_relation_def Let_def)
901  apply (auto intro: iffD1[OF tcb_queue_relation'_cong[OF refl refl refl]])
902  done
903
904lemma tcbSchedEnqueue_ep_at:
905  "\<lbrace>obj_at' (P :: endpoint \<Rightarrow> bool) ep\<rbrace>
906      tcbSchedEnqueue t
907   \<lbrace>\<lambda>rv. obj_at' P ep\<rbrace>"
908  including no_pre
909  apply (simp add: tcbSchedEnqueue_def unless_def null_def)
910  apply (wp threadGet_wp, clarsimp, wp+)
911  apply (clarsimp split: if_split, wp)
912  done
913
914lemma ccorres_duplicate_guard:
915  "ccorres r xf (P and P) Q hs f f' \<Longrightarrow> ccorres r xf P Q hs f f'"
916  by (erule ccorres_guard_imp, auto)
917
918
919lemma ep_q_refs'_no_NTFNBound[simp]:
920  "(x, NTFNBound) \<notin> ep_q_refs_of' ep"
921  by (auto simp: ep_q_refs_of'_def split: endpoint.splits)
922
923
924lemma ntfn_q_refs'_no_NTFNBound[simp]:
925  "(x, NTFNBound) \<notin> ntfn_q_refs_of' ntfn"
926  by (auto simp: ntfn_q_refs_of'_def split: ntfn.splits)
927
928lemma cancelBadgedSends_ccorres:
929  "ccorres dc xfdc (invs' and ep_at' ptr)
930              (UNIV \<inter> {s. epptr_' s = Ptr ptr} \<inter> {s. badge_' s = bdg}) []
931       (cancelBadgedSends ptr bdg) (Call cancelBadgedSends_'proc)"
932  apply (cinit lift: epptr_' badge_' simp: whileAnno_def)
933   apply (simp add: list_case_return2
934              cong: list.case_cong Structures_H.endpoint.case_cong call_ignore_cong
935               del: Collect_const)
936   apply (rule ccorres_pre_getEndpoint)
937   apply (rule_tac R="ko_at' rv ptr" and xf'="ret__unsigned_'"
938               and val="case rv of RecvEP q \<Rightarrow> scast EPState_Recv | IdleEP \<Rightarrow> scast EPState_Idle
939                                | SendEP q \<Rightarrow> scast EPState_Send"
940               in ccorres_symb_exec_r_known_rv_UNIV[where R'=UNIV])
941      apply vcg
942      apply clarsimp
943      apply (erule cmap_relationE1 [OF cmap_relation_ep], erule ko_at_projectKO_opt)
944      apply (clarsimp simp: typ_heap_simps cendpoint_relation_def Let_def
945                     split: Structures_H.endpoint.split_asm)
946     apply ceqv
947    apply wpc
948      apply (simp add: dc_def[symmetric] ccorres_cond_iffs)
949      apply (rule ccorres_return_Skip)
950     apply (simp add: dc_def[symmetric] ccorres_cond_iffs)
951     apply (rule ccorres_return_Skip)
952    apply (rename_tac list)
953    apply (simp add: Collect_True Collect_False endpoint_state_defs
954                     ccorres_cond_iffs dc_def[symmetric]
955                del: Collect_const cong: call_ignore_cong)
956    apply (rule ccorres_rhs_assoc)+
957    apply (csymbr, csymbr)
958    apply (drule_tac s = rv in sym, simp only:)
959    apply (rule_tac P="ko_at' rv ptr and invs'" in ccorres_cross_over_guard)
960    apply (rule ccorres_symb_exec_r)
961      apply (rule ccorres_rhs_assoc2, rule ccorres_rhs_assoc2)
962      apply (rule ccorres_split_nothrow[where r'=dc and xf'=xfdc, OF _ ceqv_refl])
963         apply (rule_tac P="ko_at' rv ptr"
964                    in ccorres_from_vcg[where P'=UNIV])
965         apply (rule allI, rule conseqPre, vcg)
966         apply clarsimp
967         apply (rule cmap_relationE1[OF cmap_relation_ep], assumption)
968          apply (erule ko_at_projectKO_opt)
969         apply (clarsimp simp: typ_heap_simps setEndpoint_def)
970         apply (rule rev_bexI)
971          apply (rule setObject_eq; simp add: objBits_simps')[1]
972         apply (clarsimp simp: rf_sr_def cstate_relation_def
973                               Let_def carch_state_relation_def
974                               cmachine_state_relation_def
975                               )
976         apply (clarsimp simp: cpspace_relation_def
977                               update_ep_map_tos
978                               typ_heap_simps')
979         apply (erule(1) cpspace_relation_ep_update_ep2)
980          apply (simp add: cendpoint_relation_def endpoint_state_defs)
981         subgoal by simp
982        apply (rule ccorres_symb_exec_r)
983          apply (rule_tac xs=list in filterM_voodoo)
984          apply (rule_tac P="\<lambda>xs s. (\<forall>x \<in> set xs \<union> set list.
985                   st_tcb_at' (\<lambda>st. isBlockedOnSend st \<and> blockingObject st = ptr) x s)
986                              \<and> distinct (xs @ list) \<and> ko_at' IdleEP ptr s
987                              \<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})
988                              \<and> valid_queues s \<and> pspace_aligned' s \<and> pspace_distinct' s
989                              \<and> sch_act_wf (ksSchedulerAction s) s \<and> valid_objs' s"
990                     and P'="\<lambda>xs. {s. ep_queue_relation' (cslift s) (xs @ list)
991                                         (head_C (queue_' s)) (end_C (queue_' s))}
992                                \<inter> {s. thread_' s = (case list of [] \<Rightarrow> tcb_Ptr 0
993                                                       | x # xs \<Rightarrow> tcb_ptr_to_ctcb_ptr x)}"
994                      in ccorres_inst_voodoo)
995          apply (induct_tac list)
996           apply (rule allI)
997           apply (rule iffD1 [OF ccorres_expand_while_iff_Seq])
998           apply (rule ccorres_tmp_lift2 [OF _ _ refl])
999            apply ceqv
1000           apply (simp add: ccorres_cond_iffs)
1001           apply (rule ccorres_rhs_assoc2)
1002           apply (rule ccorres_duplicate_guard, rule ccorres_split_nothrow_novcg_dc)
1003              apply (rule ccorres_from_vcg, rule allI, rule conseqPre, vcg)
1004              apply clarsimp
1005              apply (drule obj_at_ko_at', clarsimp)
1006              apply (rule cmap_relationE1[OF cmap_relation_ep], assumption)
1007               apply (erule ko_at_projectKO_opt)
1008              apply (clarsimp simp: typ_heap_simps tcb_queue_relation'_def)
1009              apply (case_tac x)
1010               apply (clarsimp simp: setEndpoint_def)
1011               apply (rule rev_bexI, rule setObject_eq,
1012                      (simp add: objBits_simps')+)
1013               apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
1014                 carch_state_relation_def
1015                 cmachine_state_relation_def
1016                 cpspace_relation_def typ_heap_simps'
1017                 update_ep_map_tos)
1018               apply (erule(1) cpspace_relation_ep_update_ep2)
1019                subgoal by (simp add: cendpoint_relation_def Let_def)
1020               subgoal by simp
1021              apply (clarsimp simp: tcb_at_not_NULL[OF pred_tcb_at']
1022                                    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               apply (simp add: cendpoint_relation_def Let_def)
1032               apply (subgoal_tac "tcb_at' (last (a # list)) \<sigma> \<and> tcb_at' a \<sigma>")
1033                apply (clarsimp simp: is_aligned_neg_mask_weaken[
1034                                        OF is_aligned_tcb_ptr_to_ctcb_ptr[where P=\<top>]])
1035                subgoal by (simp add: tcb_queue_relation'_def EPState_Send_def mask_def)
1036               subgoal by (auto split: if_split)
1037              subgoal by simp
1038             apply (ctac add: rescheduleRequired_ccorres[unfolded dc_def])
1039            apply (rule hoare_pre, wp weak_sch_act_wf_lift_linear set_ep_valid_objs')
1040            apply (clarsimp simp: weak_sch_act_wf_def sch_act_wf_def)
1041            apply (fastforce simp: valid_ep'_def pred_tcb_at' split: list.splits)
1042           apply (simp add: guard_is_UNIV_def)
1043          apply (rule allI)
1044          apply (rename_tac a lista x)
1045          apply (rule iffD1 [OF ccorres_expand_while_iff_Seq])
1046          apply (rule ccorres_init_tmp_lift2, ceqv)
1047          apply (rule ccorres_guard_imp2)
1048           apply (simp add: bind_assoc dc_def[symmetric]
1049                       del: Collect_const)
1050           apply (rule ccorres_cond_true)
1051           apply (rule ccorres_rhs_assoc)+
1052           apply (rule ccorres_pre_threadGet[where f=tcbState, folded getThreadState_def])
1053           apply (rule ccorres_move_c_guard_tcb)
1054           apply csymbr
1055           apply (rule ccorres_abstract_cleanup)
1056           apply csymbr
1057           apply (rule ccorres_move_c_guard_tcb)
1058           apply (rule_tac P=\<top>
1059                      and P'="{s. ep_queue_relation' (cslift s) (x @ a # lista)
1060                                        (head_C (queue_' s)) (end_C (queue_' s))}"
1061                      and f'="\<lambda>s. s \<lparr> next___ptr_to_struct_tcb_C_' := (case lista of [] \<Rightarrow> tcb_Ptr 0
1062                                              | v # vs \<Rightarrow> tcb_ptr_to_ctcb_ptr v) \<rparr>"
1063                      and xf'="next___ptr_to_struct_tcb_C_'"
1064                           in ccorres_subst_basic_helper)
1065               apply (thin_tac "\<forall>x. P x" for P)
1066               apply (rule myvars.fold_congs, (rule refl)+)
1067               apply (clarsimp simp: tcb_queue_relation'_def use_tcb_queue_relation2
1068                                     tcb_queue_relation2_concat)
1069               apply (clarsimp simp: typ_heap_simps split: list.split)
1070              subgoal by (simp add: rf_sr_def)
1071             apply simp
1072            apply ceqv
1073           apply (rule_tac P="ret__unsigned=blockingIPCBadge rva" in ccorres_gen_asm2)
1074           apply (rule ccorres_if_bind, rule ccorres_if_lhs)
1075            apply (simp add: bind_assoc dc_def[symmetric])
1076            apply (rule ccorres_rhs_assoc)+
1077            apply (ctac add: setThreadState_ccorres)
1078              apply (ctac add: tcbSchedEnqueue_ccorres)
1079                apply (rule_tac P="\<lambda>s. \<forall>t \<in> set (x @ a # lista). tcb_at' t s"
1080                             in ccorres_cross_over_guard)
1081                apply (rule ccorres_add_return, rule ccorres_split_nothrow[OF _ ceqv_refl])
1082                   apply (rule_tac rrel=dc and xf=xfdc
1083                               and P="\<lambda>s. (\<forall>t \<in> set (x @ a # lista). tcb_at' t s)
1084                                          \<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})
1085                                          \<and> valid_queues s \<and> distinct (x @ a # lista)
1086                                          \<and> pspace_aligned' s \<and> pspace_distinct' s"
1087                              and P'="{s. ep_queue_relation' (cslift s) (x @ a # lista)
1088                                           (head_C (queue_' s)) (end_C (queue_' s))}"
1089                               in ccorres_from_vcg)
1090                   apply (thin_tac "\<forall>x. P x" for P)
1091                   apply (rule allI, rule conseqPre, vcg)
1092                   apply (clarsimp simp: ball_Un)
1093                   apply (rule exI, rule conjI)
1094                    apply (rule exI, erule conjI)
1095                    apply (intro conjI[rotated])
1096                    apply (assumption)
1097                    apply (fold_subgoals (prefix))[3]
1098                    subgoal premises prems using prems by (fastforce intro: pred_tcb_at')+
1099                   apply (clarsimp simp: return_def rf_sr_def cstate_relation_def Let_def)
1100                   apply (rule conjI)
1101                    apply (clarsimp simp: cpspace_relation_def)
1102                    apply (rule conjI, erule ctcb_relation_null_queue_ptrs)
1103                     apply (rule null_ep_queue)
1104                     subgoal by (simp add: o_def)
1105                    apply (rule conjI)
1106                     apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
1107                     apply clarsimp
1108                     apply (rule cendpoint_relation_q_cong)
1109                     apply (rule sym, erule restrict_map_eqI)
1110                     apply (clarsimp simp: image_iff)
1111                     apply (drule(2) map_to_ko_atI2)
1112                     apply (drule ko_at_state_refs_ofD')
1113                     apply clarsimp
1114                     apply (drule_tac x=p in spec)
1115                     subgoal by fastforce
1116                    apply (erule iffD1 [OF cmap_relation_cong, OF refl refl, rotated -1])
1117                    apply clarsimp
1118                    apply (drule(2) map_to_ko_atI2, drule ko_at_state_refs_ofD')
1119
1120                    apply (rule cnotification_relation_q_cong)
1121                    apply (rule sym, erule restrict_map_eqI)
1122                    apply (clarsimp simp: image_iff)
1123                    apply (drule_tac x=p in spec)
1124                    subgoal by fastforce
1125                   apply (rule conjI)
1126                    apply (erule cready_queues_relation_not_queue_ptrs,
1127                           auto dest: null_ep_schedD[unfolded o_def] simp: o_def)[1]
1128                   apply (simp add: carch_state_relation_def
1129                                    cmachine_state_relation_def
1130                                    h_t_valid_clift_Some_iff)
1131                  apply (rule ccorres_symb_exec_r2)
1132                    apply (erule spec)
1133                   apply vcg
1134                  apply (vcg spec=modifies)
1135                 apply wp
1136                apply simp
1137                apply vcg
1138               apply (wp hoare_vcg_const_Ball_lift tcbSchedEnqueue_ep_at
1139                         sch_act_wf_lift)
1140              apply simp
1141              apply (vcg exspec=tcbSchedEnqueue_cslift_spec)
1142             apply (wp hoare_vcg_const_Ball_lift sts_st_tcb_at'_cases
1143                       sts_sch_act sts_valid_queues setThreadState_oa_queued)
1144            apply (vcg exspec=setThreadState_cslift_spec)
1145           apply (simp add: ccorres_cond_iffs dc_def[symmetric])
1146           apply (rule ccorres_symb_exec_r2)
1147             apply (drule_tac x="x @ [a]" in spec, simp add: dc_def[symmetric])
1148            apply vcg
1149           apply (vcg spec=modifies)
1150          apply (thin_tac "\<forall>x. P x" for P)
1151          apply (clarsimp simp: pred_tcb_at' ball_Un)
1152          apply (rule conjI)
1153           apply (clarsimp split: if_split)
1154           subgoal by (fastforce simp: valid_tcb_state'_def valid_objs'_maxDomain
1155                                  valid_objs'_maxPriority dest: pred_tcb_at')
1156          apply (clarsimp simp: tcb_at_not_NULL [OF pred_tcb_at'])
1157          apply (clarsimp simp: typ_heap_simps st_tcb_at'_def)
1158          apply (drule(1) obj_at_cslift_tcb)
1159          apply (clarsimp simp: ctcb_relation_blocking_ipc_badge)
1160          apply (rule conjI, simp add: "StrictC'_thread_state_defs" mask_def)
1161          apply (rule conjI)
1162           apply clarsimp
1163           apply (frule rf_sr_cscheduler_relation)
1164           apply (clarsimp simp: cscheduler_action_relation_def st_tcb_at'_def
1165                          split: scheduler_action.split_asm)
1166           apply (rename_tac word)
1167           apply (frule_tac x=word in tcbSchedEnqueue_cslift_precond_discharge)
1168              apply simp
1169             subgoal by clarsimp
1170            subgoal by clarsimp
1171           subgoal by clarsimp
1172          apply clarsimp
1173          apply (rule conjI)
1174           apply (frule(3) tcbSchedEnqueue_cslift_precond_discharge)
1175           subgoal by clarsimp
1176          apply clarsimp
1177          apply (rule context_conjI)
1178           apply (clarsimp simp: tcb_queue_relation'_def)
1179           apply (erule iffD2[OF ep_queue_relation_shift[rule_format], rotated -1])
1180           subgoal by simp
1181          apply (rule_tac x="x @ a # lista" in exI)
1182          apply (clarsimp simp: ball_Un)
1183          apply (rule conjI, fastforce)
1184          subgoal by (clarsimp simp: remove1_append)
1185         apply vcg
1186        apply (rule conseqPre, vcg)
1187        apply clarsimp
1188       apply (wp hoare_vcg_const_Ball_lift)
1189       apply (wp obj_at_setObject3[where 'a=endpoint, folded setEndpoint_def])
1190         apply (simp add: objBits_simps')+
1191       apply (wp set_ep_valid_objs')
1192      apply vcg
1193     apply vcg
1194    apply (rule conseqPre, vcg)
1195    apply clarsimp
1196   apply (clarsimp simp: guard_is_UNIV_def)
1197   apply (erule cmap_relationE1[OF cmap_relation_ep], erule ko_at_projectKO_opt)
1198   apply (clarsimp simp: typ_heap_simps)
1199   apply (clarsimp simp: cendpoint_relation_def Let_def)
1200   subgoal by (clarsimp simp: tcb_queue_relation'_def neq_Nil_conv
1201                  split: if_split_asm)
1202  apply clarsimp
1203  apply (frule ko_at_valid_objs', clarsimp)
1204   apply (simp add: projectKOs)
1205  apply (clarsimp simp: valid_obj'_def valid_ep'_def)
1206  apply (frule sym_refs_ko_atD', clarsimp)
1207  apply (clarsimp simp: st_tcb_at_refs_of_rev')
1208  apply (rule conjI)
1209   subgoal by (auto simp: isBlockedOnSend_def elim!: pred_tcb'_weakenE)
1210  apply (rule conjI)
1211   apply (clarsimp split: if_split)
1212   apply (drule sym_refsD, clarsimp)
1213   apply (drule(1) bspec)+
1214   by (auto simp: obj_at'_def projectKOs state_refs_of'_def pred_tcb_at'_def tcb_bound_refs'_def
1215              dest!: symreftype_inverse')
1216
1217
1218lemma tcb_ptr_to_ctcb_ptr_force_fold:
1219  "x + 2 ^ ctcb_size_bits = ptr_val (tcb_ptr_to_ctcb_ptr x)"
1220  by (simp add: tcb_ptr_to_ctcb_ptr_def ctcb_offset_def)
1221
1222
1223lemma coerce_memset_to_heap_update:
1224  "heap_update_list x (replicateHider (size_of (TYPE (tcb_C))) 0)
1225      = heap_update (tcb_Ptr x)
1226             (tcb_C (arch_tcb_C (user_context_C (FCP (\<lambda>x. 0))))
1227                    (thread_state_C (FCP (\<lambda>x. 0)))
1228                    (NULL)
1229                    (seL4_Fault_C (FCP (\<lambda>x. 0)))
1230                    (lookup_fault_C (FCP (\<lambda>x. 0)))
1231                      0 0 0 0 0 0 NULL NULL NULL NULL)"
1232  apply (intro ext, simp add: heap_update_def)
1233  apply (rule_tac f="\<lambda>xs. heap_update_list x xs a b" for a b in arg_cong)
1234  apply (simp add: to_bytes_def size_of_def typ_info_simps tcb_C_tag_def)
1235  apply (simp add: ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td align_of_def padup_def
1236                   final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def)
1237  apply (simp add: typ_info_simps
1238                   user_context_C_tag_def thread_state_C_tag_def seL4_Fault_C_tag_def
1239                   lookup_fault_C_tag_def update_ti_t_ptr_0s arch_tcb_C_tag_def
1240                   ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td
1241                   ti_typ_combine_empty_ti ti_typ_combine_td
1242                   align_of_def padup_def
1243                   final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def
1244                   align_td_array' size_td_array)
1245  apply (simp add: typ_info_array')
1246  apply (simp add: typ_info_word word_rsplit_0 upt_conv_Cons)
1247  apply (simp add: typ_info_word typ_info_ptr word_rsplit_0
1248                   replicateHider_def)
1249  done
1250
1251lemma isArchObjectCap_capBits:
1252  "isArchObjectCap cap \<Longrightarrow> capBits cap = acapBits (capCap cap)"
1253  by (clarsimp simp: isCap_simps)
1254
1255declare Kernel_C.tcb_C_size [simp del]
1256
1257lemma cte_lift_ccte_relation:
1258  "cte_lift cte' = Some ctel'
1259    \<Longrightarrow> c_valid_cte cte'
1260    \<Longrightarrow> ccte_relation (cte_to_H ctel') cte'"
1261  by (simp add: ccte_relation_def)
1262
1263lemma updateFreeIndex_ccorres:
1264  "\<forall>s. \<Gamma> \<turnstile> ({s} \<inter> {s. \<exists>cte cte'. cslift s (cte_Ptr srcSlot) = Some cte'
1265               \<and> cteCap cte = cap' \<and> ccte_relation cte cte'})
1266          c
1267        {t. \<exists>cap. cap_untyped_cap_lift cap = (cap_untyped_cap_lift
1268                    (cte_C.cap_C (the (cslift s (cte_Ptr srcSlot)))))
1269                        \<lparr> cap_untyped_cap_CL.capFreeIndex_CL := ((of_nat idx') >> 4) \<rparr>
1270                \<and> cap_get_tag cap = scast cap_untyped_cap
1271                \<and> t_hrs_' (globals t) = hrs_mem_update (heap_update (cte_Ptr srcSlot)
1272                    (cte_C.cap_C_update (\<lambda>_. cap) (the (cslift s (cte_Ptr srcSlot)))))
1273                    (t_hrs_' (globals s))
1274                \<and> t may_only_modify_globals s in [t_hrs]
1275        }
1276    \<Longrightarrow> ccorres dc xfdc
1277           (valid_objs' and cte_wp_at' (\<lambda>cte. isUntypedCap (cteCap cte)
1278               \<and> cap' = (cteCap cte)) srcSlot
1279           and untyped_ranges_zero'
1280           and (\<lambda>_. is_aligned (of_nat idx' :: word32) 4 \<and> idx' \<le> 2 ^ (capBlockSize cap')))
1281           {s. \<not> capIsDevice cap'
1282               \<longrightarrow> region_actually_is_zero_bytes (capPtr cap' + of_nat idx') (capFreeIndex cap' - idx') s} hs
1283           (updateFreeIndex srcSlot idx') c"
1284  (is "_ \<Longrightarrow> ccorres dc xfdc (valid_objs' and ?cte_wp_at' and _ and _) ?P' hs ?a c")
1285  apply (rule ccorres_gen_asm)
1286  apply (simp add: updateFreeIndex_def getSlotCap_def updateCap_def)
1287  apply (rule ccorres_guard_imp2)
1288   apply (rule ccorres_split_noop_lhs, rule_tac cap'=cap' in updateTrackedFreeIndex_noop_ccorres)
1289    apply (rule ccorres_pre_getCTE)+
1290    apply (rename_tac cte cte2)
1291    apply (rule_tac P = "\<lambda>s. ?cte_wp_at' s \<and> cte2 = cte \<and> cte_wp_at' ((=) cte) srcSlot s"
1292              and P'="{s. \<exists>cte cte'. cslift s (cte_Ptr srcSlot) = Some cte'
1293               \<and> cteCap cte = cap' \<and> ccte_relation cte cte'} \<inter> ?P'" in ccorres_from_vcg)
1294    apply (rule allI, rule HoarePartial.conseq_exploit_pre, clarify)
1295    apply (drule_tac x=s in spec, rule conseqPre, erule conseqPost)
1296      defer
1297      apply clarsimp
1298     apply clarsimp
1299    apply (simp add: cte_wp_at_ctes_of)
1300    apply wp
1301   apply (clarsimp simp: isCap_simps cte_wp_at_ctes_of)
1302   apply (frule(1) rf_sr_ctes_of_clift)
1303   apply clarsimp
1304   apply (frule(1) cte_lift_ccte_relation)
1305   apply (rule exI, intro conjI[rotated], assumption, simp_all)[1]
1306  apply (clarsimp simp: cte_wp_at_ctes_of)
1307  apply (erule(1) rf_sr_ctes_of_cliftE)
1308  apply (frule(1) rf_sr_ctes_of_clift)
1309  apply clarsimp
1310  apply (subgoal_tac "ccap_relation (capFreeIndex_update (\<lambda>_. idx')
1311        (cteCap (the (ctes_of \<sigma> srcSlot)))) cap")
1312   apply (rule fst_setCTE [OF ctes_of_cte_at], assumption)
1313   apply (erule bexI [rotated])
1314   apply (clarsimp simp add: rf_sr_def cstate_relation_def Let_def
1315      cvariable_array_map_const_add_map_option[where f="tcb_no_ctes_proj"]
1316      isCap_simps)
1317   apply (simp add:cpspace_relation_def)
1318   apply (clarsimp simp:typ_heap_simps' modify_map_def mex_def meq_def)
1319   apply (rule conjI)
1320    apply (rule cpspace_cte_relation_upd_capI, assumption+)
1321   apply (rule conjI)
1322    apply (rule setCTE_tcb_case, assumption+)
1323   apply (case_tac s', clarsimp)
1324   subgoal by (simp add: carch_state_relation_def cmachine_state_relation_def
1325                   typ_heap_simps')
1326
1327  apply (clarsimp simp: isCap_simps)
1328  apply (drule(1) cte_lift_ccte_relation,
1329    drule ccte_relation_ccap_relation)
1330  apply (simp add: cte_to_H_def)
1331  apply (frule cap_get_tag_isCap_unfolded_H_cap)
1332  apply (clarsimp simp: ccap_relation_def cap_lift_untyped_cap
1333                        cap_to_H_simps cap_untyped_cap_lift_def
1334                        is_aligned_shiftr_shiftl
1335                 dest!: ccte_relation_ccap_relation)
1336  apply (rule unat_of_nat_eq unat_of_nat_eq[symmetric],
1337    erule order_le_less_trans,
1338    rule power_strict_increasing, simp_all)
1339  apply (rule unat_less_helper, rule order_le_less_trans[OF word_and_le1], simp add: mask_def)
1340  done
1341
1342end
1343
1344(* FIXME: Move *)
1345lemma ccap_relation_isDeviceCap:
1346 "\<lbrakk>ccap_relation cp cap; isUntypedCap cp
1347  \<rbrakk> \<Longrightarrow> to_bool (capIsDevice_CL (cap_untyped_cap_lift cap)) =  (capIsDevice cp)"
1348  apply (frule cap_get_tag_UntypedCap)
1349  apply (simp add:cap_get_tag_isCap )
1350  done
1351
1352lemma ccap_relation_isDeviceCap2:
1353 "\<lbrakk>ccap_relation cp cap; isUntypedCap cp
1354  \<rbrakk> \<Longrightarrow> (capIsDevice_CL (cap_untyped_cap_lift cap) = 0) = (\<not> (capIsDevice cp))"
1355  apply (frule cap_get_tag_UntypedCap)
1356  apply (simp add:cap_get_tag_isCap to_bool_def)
1357  done
1358end
1359