1243730Srwatson(*
2243730Srwatson * Copyright 2014, General Dynamics C4 Systems
3243730Srwatson *
4243730Srwatson * This software may be distributed and modified according to the terms of
5243730Srwatson * the GNU General Public License version 2. Note that NO WARRANTY is provided.
6243730Srwatson * See "LICENSE_GPLv2.txt" for details.
7243730Srwatson *
8243730Srwatson * @TAG(GD_GPL)
9243730Srwatson *)
10243730Srwatson
11243730Srwatsontheory Retype_C
12243730Srwatsonimports
13243730Srwatson  Detype_C
14243730Srwatson  CSpace_All
15243730Srwatson  StoreWord_C
16243730Srwatsonbegin
17243730Srwatson
18243730Srwatsondeclare word_neq_0_conv [simp del]
19243730Srwatson
20243730Srwatsoninstance cte_C :: array_outer_max_size
21243730Srwatson  by intro_classes simp
22243730Srwatson
23243730Srwatsoninstance virq_C :: array_inner_packed
24243730Srwatson  by intro_classes simp
25243730Srwatson
26243730Srwatsoncontext begin interpretation Arch . (*FIXME: arch_split*)
27243730Srwatson
28243730Srwatson
29243730Srwatsonlemma map_option_byte_to_word_heap:
30243730Srwatson  assumes disj: "\<And>(off :: 10 word) x. x<4 \<Longrightarrow> p + ucast off * 4 + x \<notin> S "
31243730Srwatson  shows "byte_to_word_heap (\<lambda>x. if x \<in> S then 0 else mem x) p
32243730Srwatson        = byte_to_word_heap mem p"
33243730Srwatson  by (clarsimp simp: option_map_def  byte_to_word_heap_def[abs_def]
34243730Srwatson                     Let_def disj disj[where x = 0,simplified]
35243730Srwatson              split: option.splits)
36243730Srwatson
37243730Srwatsontext {* Generalise the different kinds of retypes to allow more general proofs
38243730Srwatsonabout what they might change. *}
39243730Srwatsondefinition
40243730Srwatson  ptr_retyps_gen :: "nat \<Rightarrow> ('a :: c_type) ptr \<Rightarrow> bool \<Rightarrow> heap_typ_desc \<Rightarrow> heap_typ_desc"
41243730Srwatsonwhere
42243730Srwatson  "ptr_retyps_gen n p mk_array
43243730Srwatson    = (if mk_array then ptr_arr_retyps n p else ptr_retyps n p)"
44243730Srwatson
45243730Srwatsonlemma ptr_retyp_gen_one:
46243730Srwatson  "ptr_retyps_gen 1 p False = ptr_retyp p"
47243730Srwatson  unfolding ptr_retyps_gen_def
48243730Srwatson  by (rule ext, clarsimp)
49243730Srwatson
50243730Srwatsonend
51243730Srwatson
52243730Srwatsoncontext kernel_m
53243730Srwatsonbegin
54243730Srwatson
55243730Srwatson(* Ensure that the given region of memory does not contain any typed memory. *)
56243730Srwatsondefinition
57243730Srwatson  region_is_typeless :: "word32 \<Rightarrow> nat \<Rightarrow> ('a globals_scheme, 'b) StateSpace.state_scheme \<Rightarrow> bool"
58243730Srwatsonwhere
59243730Srwatson  "region_is_typeless ptr sz s \<equiv>
60243730Srwatson      \<forall>z\<in>{ptr ..+ sz}. snd (snd (t_hrs_' (globals s)) z) = Map.empty"
61243730Srwatson
62243730Srwatsonlemma c_guard_word8:
63243730Srwatson  "c_guard (p :: word8 ptr) = (ptr_val p \<noteq> 0)"
64243730Srwatson  unfolding c_guard_def ptr_aligned_def c_null_guard_def
65243730Srwatson  apply simp
66243730Srwatson  apply (rule iffI)
67243730Srwatson   apply (drule intvlD)
68243730Srwatson   apply clarsimp
69243730Srwatson  apply simp
70243730Srwatson  apply (rule intvl_self)
71243730Srwatson  apply simp
72243730Srwatson  done
73243730Srwatson
74243730Srwatsonlemma
75243730Srwatson  "(x \<in> {x ..+ n}) = (n \<noteq> 0)"
76243730Srwatson  apply (rule iffI)
77243730Srwatson   apply (drule intvlD)
78243730Srwatson   apply clarsimp
79243730Srwatson  apply (rule intvl_self)
80243730Srwatson  apply simp
81243730Srwatson  done
82243730Srwatson
83243730Srwatsonlemma heap_update_list_append3:
84243730Srwatson    "\<lbrakk> s' = s + of_nat (length xs) \<rbrakk> \<Longrightarrow> heap_update_list s (xs @ ys) H = heap_update_list s' ys (heap_update_list s xs H)"
85243730Srwatson  apply simp
86243730Srwatson  apply (subst heap_update_list_append [symmetric])
87243730Srwatson  apply clarsimp
88243730Srwatson  done
89243730Srwatson
90243730Srwatsonlemma ptr_aligned_word32:
91243730Srwatson  "\<lbrakk> is_aligned p 2  \<rbrakk> \<Longrightarrow> ptr_aligned ((Ptr p) :: word32 ptr)"
92243730Srwatson  apply (clarsimp simp: is_aligned_def ptr_aligned_def)
93243730Srwatson  done
94243730Srwatson
95243730Srwatsonlemma c_guard_word32:
96243730Srwatson  "\<lbrakk> is_aligned (ptr_val p) 2; p \<noteq> NULL  \<rbrakk> \<Longrightarrow> c_guard (p :: (word32 ptr))"
97243730Srwatson  apply (clarsimp simp: c_guard_def)
98243730Srwatson  apply (rule conjI)
99243730Srwatson   apply (case_tac p, clarsimp simp: ptr_aligned_word32)
100243730Srwatson  apply (case_tac p, simp add: c_null_guard_def)
101243730Srwatson  apply (subst intvl_aligned_bottom_eq [where n=2 and bits=2], auto simp: word_bits_def)
102243730Srwatson  done
103243730Srwatson
104243730Srwatsonlemma is_aligned_and_not_zero: "\<lbrakk> is_aligned n k; n \<noteq> 0 \<rbrakk> \<Longrightarrow> 2^k \<le> n"
105243730Srwatson  apply (metis aligned_small_is_0 word_not_le)
106243730Srwatson  done
107243730Srwatson
108243730Srwatsonlemma replicate_append [rule_format]: "\<forall>xs. replicate n x @ (x # xs) = replicate (n + 1) x @ xs"
109243730Srwatson  apply (induct n)
110243730Srwatson   apply clarsimp
111243730Srwatson  apply clarsimp
112243730Srwatson  done
113243730Srwatson
114243730Srwatsonlemmas unat_add_simple =
115243730Srwatson       iffD1 [OF unat_add_lem [where 'a = 32, folded word_bits_def]]
116243730Srwatson
117243730Srwatsonlemma replicate_append_list [rule_format]:
118243730Srwatson  "\<forall>n. set L \<subseteq> {0::word8} \<longrightarrow> (replicate n 0 @ L = replicate (n + length L) 0)"
119243730Srwatson  apply (rule rev_induct)
120243730Srwatson   apply clarsimp
121243730Srwatson  apply (rule allI)
122243730Srwatson  apply (erule_tac x="n+1" in allE)
123243730Srwatson  apply clarsimp
124243730Srwatson  apply (subst append_assoc[symmetric])
125243730Srwatson  apply clarsimp
126243730Srwatson  apply (subgoal_tac "\<And>n. (replicate n 0 @ [0]) = (0 # replicate n (0 :: word8))")
127243730Srwatson   apply clarsimp
128243730Srwatson  apply (induct_tac na)
129243730Srwatson   apply clarsimp
130243730Srwatson  apply clarsimp
131243730Srwatson  done
132243730Srwatson
133243730Srwatsonlemma heap_update_list_replicate:
134243730Srwatson  "\<lbrakk> set L = {0}; n' = n + length L \<rbrakk> \<Longrightarrow>  heap_update_list s ((replicate n 0) @ L) H = heap_update_list s (replicate n' 0) H"
135243730Srwatson  apply (subst replicate_append_list)
136243730Srwatson   apply clarsimp
137243730Srwatson  apply clarsimp
138243730Srwatson  done
139243730Srwatson
140243730Srwatsonlemma heap_update_word32_is_heap_update_list:
141243730Srwatson  "heap_update p (x :: word32) = heap_update_list (ptr_val p) (to_bytes x a)"
142243730Srwatson  apply (rule ext)+
143243730Srwatson  apply (clarsimp simp: heap_update_def)
144243730Srwatson  apply (clarsimp simp: to_bytes_def typ_info_word)
145243730Srwatson  done
146243730Srwatson
147243730Srwatsonlemma to_bytes_word32_0:
148243730Srwatson  "to_bytes (0 :: word32) xs = [0, 0, 0, 0 :: word8]"
149243730Srwatson  apply (simp add: to_bytes_def typ_info_word word_rsplit_same word_rsplit_0)
150243730Srwatson  done
151243730Srwatson
152243730Srwatson
153243730Srwatsonlemma globals_list_distinct_subset:
154243730Srwatson  "\<lbrakk> globals_list_distinct D symtab xs; D' \<subseteq> D \<rbrakk>
155243730Srwatson    \<Longrightarrow> globals_list_distinct D' symtab xs"
156243730Srwatson  by (simp add: globals_list_distinct_def disjoint_subset)
157243730Srwatson
158243730Srwatsonlemma fst_s_footprint:
159243730Srwatson  "(fst ` s_footprint p) = {ptr_val (p :: 'a ptr)
160243730Srwatson        ..+ size_of TYPE('a :: c_type)}"
161243730Srwatson  apply (simp add: s_footprint_def s_footprint_untyped_def)
162243730Srwatson  apply (auto simp: intvl_def size_of_def image_def)
163243730Srwatson  done
164243730Srwatson
165243730Srwatsonlemma memzero_spec:
166243730Srwatson  "\<forall>s. \<Gamma> \<turnstile> \<lbrace>s. ptr_val \<acute>s \<noteq> 0 \<and> ptr_val \<acute>s \<le> ptr_val \<acute>s + (\<acute>n - 1)
167243730Srwatson         \<and> (is_aligned (ptr_val \<acute>s) 2) \<and> (is_aligned (\<acute>n) 2)
168243730Srwatson         \<and> {ptr_val \<acute>s ..+ unat \<acute>n} \<times> {SIndexVal, SIndexTyp 0} \<subseteq> dom_s (hrs_htd \<acute>t_hrs)
169243730Srwatson         \<and> gs_get_assn cap_get_capSizeBits_'proc \<acute>ghost'state \<in> insert 0 {\<acute>n ..}\<rbrace>
170243730Srwatson    Call memzero_'proc {t.
171243730Srwatson     t_hrs_' (globals t) = hrs_mem_update (heap_update_list (ptr_val (s_' s))
172243730Srwatson                                            (replicate (unat (n_' s)) (ucast (0)))) (t_hrs_' (globals s))}"
173243730Srwatson  apply (hoare_rule HoarePartial.ProcNoRec1)
174243730Srwatson  apply (clarsimp simp: whileAnno_def)
175243730Srwatson  apply (rule_tac I1="{t. (ptr_val (s_' s) \<le> ptr_val (s_' s) + ((n_' s) - 1) \<and> ptr_val (s_' s) \<noteq> 0) \<and>
176243730Srwatson                             ptr_val (s_' s) + (n_' s - n_' t) = ptr_val (p_' t) \<and>
177243730Srwatson                             n_' t \<le> n_' s \<and>
178243730Srwatson                             (is_aligned (n_' t) 2) \<and>
179243730Srwatson                             (is_aligned (n_' s) 2) \<and>
180243730Srwatson                             (is_aligned (ptr_val (s_' t)) 2) \<and>
181243730Srwatson                             (is_aligned (ptr_val (s_' s)) 2) \<and>
182243730Srwatson                             (is_aligned (ptr_val (p_' t)) 2) \<and>
183243730Srwatson                             {ptr_val (p_' t) ..+ unat (n_' t)} \<times> {SIndexVal, SIndexTyp 0}
184243730Srwatson                                 \<subseteq> dom_s (hrs_htd (t_hrs_' (globals t))) \<and>
185243730Srwatson                             globals t = (globals s)\<lparr> t_hrs_' :=
186243730Srwatson                             hrs_mem_update (heap_update_list (ptr_val (s_' s))
187243730Srwatson                                               (replicate (unat (n_' s - n_' t)) 0))
188243730Srwatson                                                      (t_hrs_' (globals s))\<rparr> }"
189243730Srwatson            and V1=undefined in subst [OF whileAnno_def])
190243730Srwatson  apply vcg
191243730Srwatson    apply (clarsimp simp add: hrs_mem_update_def)
192243730Srwatson
193243730Srwatson   apply clarsimp
194243730Srwatson   apply (case_tac s, case_tac p)
195243730Srwatson
196243730Srwatson   apply (subgoal_tac "4 \<le> unat na")
197243730Srwatson    apply (intro conjI)
198243730Srwatson           apply (simp add: ptr_safe_def s_footprint_def s_footprint_untyped_def
199243730Srwatson                            typ_uinfo_t_def typ_info_word)
200243730Srwatson           apply (erule order_trans[rotated])
201243730Srwatson            apply (auto intro!: intvlI)[1]
202243730Srwatson          apply (subst c_guard_word32, simp_all)[1]
203243730Srwatson          apply (clarsimp simp: field_simps)
204243730Srwatson          apply (metis le_minus' minus_one_helper5 olen_add_eqv diff_self word_le_0_iff word_le_less_eq)
205243730Srwatson         apply (clarsimp simp: field_simps)
206243730Srwatson        apply (frule is_aligned_and_not_zero)
207243730Srwatson         apply clarsimp
208243730Srwatson        apply (rule word_le_imp_diff_le, auto)[1]
209243730Srwatson       apply clarsimp
210243730Srwatson       apply (rule aligned_sub_aligned [where n=2], simp_all add: is_aligned_def word_bits_def)[1]
211243730Srwatson      apply clarsimp
212243730Srwatson      apply (rule is_aligned_add, simp_all add: is_aligned_def word_bits_def)[1]
213243730Srwatson     apply (erule order_trans[rotated])
214243730Srwatson     apply (clarsimp simp: subset_iff)
215243730Srwatson     apply (erule subsetD[OF intvl_sub_offset, rotated])
216243730Srwatson     apply (simp add: unat_sub word_le_nat_alt)
217243730Srwatson    apply (clarsimp simp: word_bits_def hrs_mem_update_def)
218243730Srwatson    apply (subst heap_update_word32_is_heap_update_list [where a="[]"])
219243730Srwatson    apply (subst heap_update_list_append3[symmetric])
220243730Srwatson     apply clarsimp
221243730Srwatson    apply (subst to_bytes_word32_0)
222243730Srwatson    apply (rule heap_update_list_replicate)
223243730Srwatson     apply clarsimp
224243730Srwatson    apply (rule_tac s="unat ((n - na) + 4)" in trans)
225243730Srwatson     apply (simp add: field_simps)
226243730Srwatson    apply (subst Word.unat_plus_simple[THEN iffD1])
227243730Srwatson     apply (rule is_aligned_no_overflow''[where n=2, simplified])
228243730Srwatson      apply (erule(1) aligned_sub_aligned, simp)
229243730Srwatson     apply (clarsimp simp: field_simps)
230243730Srwatson     apply (frule_tac x=n in is_aligned_no_overflow'', simp)
231243730Srwatson     apply simp
232243730Srwatson    apply simp
233243730Srwatson   apply (rule dvd_imp_le)
234243730Srwatson    apply (simp add: is_aligned_def)
235243730Srwatson   apply (simp add: unat_eq_0[symmetric])
236243730Srwatson  apply clarsimp
237243730Srwatson  done
238243730Srwatson
239243730Srwatson
240243730Srwatsonlemma memset_spec:
241243730Srwatson  "\<forall>s. \<Gamma> \<turnstile> \<lbrace>s. ptr_val \<acute>s \<noteq> 0 \<and> ptr_val \<acute>s \<le> ptr_val \<acute>s + (\<acute>n - 1)
242243730Srwatson         \<and> {ptr_val \<acute>s ..+ unat \<acute>n} \<times> {SIndexVal, SIndexTyp 0} \<subseteq> dom_s (hrs_htd \<acute>t_hrs)
243243730Srwatson         \<and> gs_get_assn cap_get_capSizeBits_'proc \<acute>ghost'state \<in> insert 0 {\<acute>n ..}\<rbrace>
244243730Srwatson    Call memset_'proc
245243730Srwatson   {t. t_hrs_' (globals t) = hrs_mem_update (heap_update_list (ptr_val (s_' s))
246243730Srwatson                                            (replicate (unat (n_' s)) (ucast (c_' s)))) (t_hrs_' (globals s))}"
247243730Srwatson  apply (hoare_rule HoarePartial.ProcNoRec1)
248243730Srwatson  apply (clarsimp simp: whileAnno_def)
249243730Srwatson  apply (rule_tac I1="{t. (ptr_val (s_' s) \<le> ptr_val (s_' s) + ((n_' s) - 1) \<and> ptr_val (s_' s) \<noteq> 0) \<and>
250243730Srwatson                             c_' t = c_' s \<and>
251243730Srwatson                             ptr_val (s_' s) + (n_' s - n_' t) = ptr_val (p_' t) \<and>
252243730Srwatson                             n_' t \<le> n_' s \<and>
253243730Srwatson                             {ptr_val (p_' t) ..+ unat (n_' t)} \<times> {SIndexVal, SIndexTyp 0}
254243730Srwatson                                \<subseteq> dom_s (hrs_htd (t_hrs_' (globals t))) \<and>
255243730Srwatson                             globals t = (globals s)\<lparr> t_hrs_' :=
256243730Srwatson                             hrs_mem_update (heap_update_list (ptr_val (s_' s))
257243730Srwatson                                               (replicate (unat (n_' s - n_' t)) (ucast (c_' t))))
258243730Srwatson                                                      (t_hrs_' (globals s))\<rparr>}"
259243730Srwatson            and V1=undefined in subst [OF whileAnno_def])
260243730Srwatson  apply vcg
261243730Srwatson    apply (clarsimp simp add: hrs_mem_update_def split: if_split_asm)
262243730Srwatson    apply (subst (asm) word_mod_2p_is_mask [where n=2, simplified], simp)
263243730Srwatson    apply (subst (asm) word_mod_2p_is_mask [where n=2, simplified], simp)
264243730Srwatson    apply (rule conjI)
265243730Srwatson     apply (rule is_aligned_and_2_to_k, clarsimp simp: mask_def)
266243730Srwatson    apply (rule is_aligned_and_2_to_k, clarsimp simp: mask_def)
267243730Srwatson   apply clarsimp
268243730Srwatson   apply (intro conjI)
269243730Srwatson        apply (simp add: ptr_safe_def s_footprint_def s_footprint_untyped_def
270243730Srwatson                         typ_uinfo_t_def typ_info_word)
271243730Srwatson        apply (erule order_trans[rotated])
272243730Srwatson        apply (auto simp: intvl_self unat_gt_0 intro!: intvlI)[1]
273243730Srwatson       apply (simp add: c_guard_word8)
274243730Srwatson       apply (erule subst)
275243730Srwatson       apply (subst lt1_neq0 [symmetric])
276243730Srwatson       apply (rule order_trans)
277243730Srwatson        apply (subst lt1_neq0, assumption)
278243730Srwatson       apply (erule word_random)
279243730Srwatson       apply (rule word_le_minus_mono_right)
280243730Srwatson         apply (simp add: lt1_neq0)
281243730Srwatson        apply assumption
282243730Srwatson       apply (erule order_trans [rotated])
283243730Srwatson       apply (simp add: lt1_neq0)
284243730Srwatson      apply (case_tac p, simp add: CTypesDefs.ptr_add_def unat_minus_one field_simps)
285243730Srwatson     apply (metis word_must_wrap word_not_simps(1) linear)
286243730Srwatson    apply (erule order_trans[rotated])
287243730Srwatson    apply (clarsimp simp: ptr_val_case split: ptr.splits)
288243730Srwatson    apply (erule subsetD[OF intvl_sub_offset, rotated])
289243730Srwatson    apply (simp add: unat_sub word_le_nat_alt word_less_nat_alt)
290243730Srwatson   apply (clarsimp simp: ptr_val_case unat_minus_one hrs_mem_update_def split: ptr.splits)
291243730Srwatson   apply (subgoal_tac "unat (n - (na - 1)) = Suc (unat (n - na))")
292243730Srwatson    apply (erule ssubst, subst replicate_Suc_append)
293243730Srwatson    apply (subst heap_update_list_append)
294243730Srwatson    apply (simp add: heap_update_word8)
295243730Srwatson   apply (subst unatSuc [symmetric])
296243730Srwatson    apply (subst add.commute)
297243730Srwatson    apply (metis word_neq_0_conv word_sub_plus_one_nonzero)
298243730Srwatson   apply (simp add: field_simps)
299243730Srwatson  apply (clarsimp)
300243730Srwatson  apply (metis diff_0_right word_gt_0)
301243730Srwatson  done
302243730Srwatson
303243730Srwatson
304243730Srwatsondeclare snd_get[simp]
305243730Srwatson
306243730Srwatsondeclare snd_gets[simp]
307243730Srwatson
308243730Srwatsonlemma snd_when_aligneError[simp]:
309243730Srwatson  shows "(snd ((when P (alignError sz)) s)) = P"
310243730Srwatson  by (simp add: when_def alignError_def fail_def split: if_split)
311243730Srwatson
312243730Srwatsonlemma snd_unless_aligneError[simp]:
313243730Srwatson  shows "(snd ((unless P (alignError sz)) s)) = (\<not> P)"
314243730Srwatson  by (simp add: unless_def)
315243730Srwatson
316243730Srwatsonlemma lift_t_retyp_heap_same:
317243730Srwatson  fixes p :: "'a :: mem_type ptr"
318243730Srwatson  assumes gp: "g p"
319243730Srwatson  shows "lift_t g (hp, ptr_retyp p td) p = Some (from_bytes (heap_list hp (size_of TYPE('a)) (ptr_val p)))"
320243730Srwatson  apply (simp add: lift_t_def lift_typ_heap_if s_valid_def hrs_htd_def)
321243730Srwatson  apply (subst ptr_retyp_h_t_valid)
322243730Srwatson   apply (rule gp)
323243730Srwatson  apply simp
324243730Srwatson  apply (subst heap_list_s_heap_list_dom)
325243730Srwatson  apply (clarsimp simp: s_footprint_intvl)
326243730Srwatson  apply simp
327243730Srwatson  done
328243730Srwatson
329243730Srwatsonlemma lift_t_retyp_heap_same_rep0:
330243730Srwatson  fixes p :: "'a :: mem_type ptr"
331243730Srwatson  assumes gp: "g p"
332243730Srwatson  shows "lift_t g (heap_update_list (ptr_val p) (replicate (size_of TYPE('a)) 0) hp, ptr_retyp p td) p =
333243730Srwatson  Some (from_bytes (replicate (size_of TYPE('a)) 0))"
334243730Srwatson  apply (subst lift_t_retyp_heap_same)
335243730Srwatson   apply (rule gp)
336243730Srwatson  apply (subst heap_list_update [where v = "replicate (size_of TYPE('a)) 0", simplified])
337243730Srwatson  apply (rule order_less_imp_le)
338243730Srwatson  apply simp
339243730Srwatson  apply simp
340243730Srwatson  done
341243730Srwatson
342243730Srwatsonlemma lift_t_retyp_heap_other2:
343243730Srwatson  fixes p :: "'a :: mem_type ptr" and p' :: "'b :: mem_type ptr"
344243730Srwatson  assumes orth: "{ptr_val p..+size_of TYPE('a)} \<inter> {ptr_val p'..+size_of TYPE('b)} = {}"
345243730Srwatson  shows "lift_t g (hp, ptr_retyp p td) p' = lift_t g (hp, td) p'"
346243730Srwatson  apply (simp add: lift_t_def lift_typ_heap_if s_valid_def hrs_htd_def ptr_retyp_disjoint_iff [OF orth])
347243730Srwatson  apply (cases "td, g \<Turnstile>\<^sub>t p'")
348243730Srwatson   apply simp
349243730Srwatson   apply (simp add: h_t_valid_taut heap_list_s_heap_list heap_list_update_disjoint_same
350243730Srwatson     ptr_retyp_disjoint_iff orth)
351243730Srwatson  apply (simp add: h_t_valid_taut heap_list_s_heap_list heap_list_update_disjoint_same
352243730Srwatson    ptr_retyp_disjoint_iff orth)
353243730Srwatson  done
354243730Srwatson
355243730Srwatsonlemma dom_s_SindexValD:
356243730Srwatson  "(x, SIndexVal) \<in> dom_s td \<Longrightarrow> fst (td x)"
357243730Srwatson  unfolding dom_s_def by clarsimp
358243730Srwatson
359243730Srwatsonlemma typ_slice_t_self_nth:
360243730Srwatson  "\<exists>n < length (typ_slice_t td m). \<exists>b. typ_slice_t td m ! n = (td, b)"
361243730Srwatson  using typ_slice_t_self [where td = td and m = m]
362243730Srwatson  by (fastforce simp add: in_set_conv_nth)
363243730Srwatson
364243730Srwatsonlemma ptr_retyp_other_cleared_region:
365243730Srwatson  fixes p :: "'a :: mem_type ptr" and p' :: "'b :: mem_type ptr"
366243730Srwatson  assumes  ht: "ptr_retyp p td, g \<Turnstile>\<^sub>t p'"
367243730Srwatson  and   tdisj: "typ_uinfo_t TYPE('a) \<bottom>\<^sub>t typ_uinfo_t TYPE('b :: mem_type)"
368243730Srwatson  and   clear: "\<forall>x \<in> {ptr_val p ..+ size_of TYPE('a)}. \<forall>n b. snd (td x) n \<noteq> Some (typ_uinfo_t TYPE('b), b)"
369243730Srwatson  shows "{ptr_val p'..+ size_of TYPE('b)} \<inter> {ptr_val p ..+ size_of TYPE('a)} = {}"
370243730Srwatsonproof (rule classical)
371243730Srwatson  assume asm: "{ptr_val p'..+ size_of TYPE('b)} \<inter> {ptr_val p ..+ size_of TYPE('a)} \<noteq> {}"
372243730Srwatson  then obtain mv where mvp: "mv \<in> {ptr_val p..+size_of TYPE('a)}"
373243730Srwatson    and mvp': "mv \<in> {ptr_val p'..+size_of TYPE('b)}"
374243730Srwatson      by blast
375243730Srwatson
376243730Srwatson  then obtain k' where mv: "mv = ptr_val p' + of_nat k'" and klt: "k' < size_td (typ_info_t TYPE('b))"
377243730Srwatson    by (clarsimp dest!: intvlD simp: size_of_def typ_uinfo_size)
378243730Srwatson
379243730Srwatson  let ?mv = "ptr_val p' + of_nat k'"
380243730Srwatson
381243730Srwatson  obtain n b where nl: "n < length (typ_slice_t (typ_uinfo_t TYPE('b)) k')"
382243730Srwatson    and tseq: "typ_slice_t (typ_uinfo_t TYPE('b)) k' ! n = (typ_uinfo_t TYPE('b), b)"
383243730Srwatson    using typ_slice_t_self_nth [where td = "typ_uinfo_t TYPE('b)" and m = k']
384243730Srwatson    by clarsimp
385
386  with ht have "snd (ptr_retyp p td ?mv) n = Some (typ_uinfo_t TYPE('b), b)"
387    unfolding h_t_valid_def
388    apply -
389    apply (clarsimp simp: valid_footprint_def Let_def)
390    apply (drule spec, drule mp [OF _ klt])
391    apply (clarsimp simp: map_le_def)
392    apply (drule bspec)
393    apply simp
394    apply simp
395    done
396
397  moreover {
398    assume "snd (ptr_retyp p empty_htd ?mv) n = Some (typ_uinfo_t TYPE('b), b)"
399    hence "(typ_uinfo_t TYPE('b)) \<in> fst ` set (typ_slice_t (typ_uinfo_t TYPE('a))
400                                                 (unat (ptr_val p' + of_nat k' - ptr_val p)))"
401      using asm mv mvp
402      apply -
403      apply (rule_tac x = "(typ_uinfo_t TYPE('b), b)" in image_eqI)
404       apply simp
405      apply (fastforce simp add: ptr_retyp_footprint list_map_eq in_set_conv_nth split: if_split_asm)
406      done
407
408    with typ_slice_set have "(typ_uinfo_t TYPE('b)) \<in> fst ` td_set (typ_uinfo_t TYPE('a)) 0"
409      by (rule subsetD)
410
411    hence False using tdisj by (clarsimp simp: tag_disj_def typ_tag_le_def)
412  } ultimately show ?thesis using mvp mvp' mv unfolding h_t_valid_def valid_footprint_def
413    apply -
414    apply (subst (asm) ptr_retyp_d_eq_snd)
415    apply (auto simp add: map_add_Some_iff clear)
416    done
417qed
418
419lemma h_t_valid_not_empty:
420  fixes p :: "'a :: c_type ptr"
421  shows "\<lbrakk> d,g \<Turnstile>\<^sub>t p; x \<in> {ptr_val p..+size_of TYPE('a)} \<rbrakk> \<Longrightarrow> snd (d x) \<noteq> Map.empty"
422  apply (drule intvlD)
423  apply (clarsimp simp: h_t_valid_def size_of_def)
424  apply (drule valid_footprintD)
425   apply (simp add: typ_uinfo_size)
426  apply clarsimp
427  done
428
429lemma ptr_retyps_out:
430  fixes p :: "'a :: mem_type ptr"
431  shows "x \<notin> {ptr_val p..+n * size_of TYPE('a)} \<Longrightarrow> ptr_retyps n p td x = td x"
432proof (induct n arbitrary: p)
433  case 0 thus ?case by simp
434next
435  case (Suc m)
436
437  have ih: "ptr_retyps m (CTypesDefs.ptr_add p 1) td x = td x"
438  proof (rule Suc.hyps)
439    from Suc.prems show "x \<notin> {ptr_val (CTypesDefs.ptr_add p 1)..+m * size_of TYPE('a)}"
440      apply (rule contrapos_nn)
441      apply (erule subsetD [rotated])
442      apply (simp add: CTypesDefs.ptr_add_def)
443      apply (rule intvl_sub_offset)
444      apply (simp add: unat_of_nat)
445      done
446  qed
447
448  from Suc.prems have "x \<notin> {ptr_val p..+size_of TYPE('a)}"
449    apply (rule contrapos_nn)
450    apply (erule subsetD [rotated])
451    apply (rule intvl_start_le)
452    apply simp
453    done
454
455  thus ?case
456    by (simp add: ptr_retyp_d ih)
457qed
458
459lemma image_add_intvl:
460  "((+) x) ` {p ..+ n} = {p + x ..+ n}"
461  by (auto simp add: intvl_def)
462
463lemma intvl_sum:
464  "{p..+ i + j}
465    = {p ..+ i} \<union> {(p :: ('a :: len) word) + of_nat i ..+ j}"
466  apply (simp add: intvl_def, safe)
467    apply clarsimp
468    apply (case_tac "k < i")
469     apply auto[1]
470    apply (drule_tac x="k - i" in spec)
471    apply simp
472   apply fastforce
473  apply (rule_tac x="k + i" in exI)
474  apply simp
475  done
476
477lemma intvl_Suc_right:
478  "{p ..+ Suc n} = {p} \<union> {(p :: ('a :: len) word) + 1 ..+ n}"
479  apply (simp add: intvl_sum[where p=p and i=1 and j=n, simplified])
480  apply (auto dest: intvl_Suc simp: intvl_self)
481  done
482
483lemma htd_update_list_same2:
484  "x \<notin> {p ..+ length xs} \<Longrightarrow>
485    htd_update_list p xs htd x = htd x"
486  by (induct xs arbitrary: p htd, simp_all add: intvl_Suc_right)
487
488lemma ptr_retyps_gen_out:
489  fixes p :: "'a :: mem_type ptr"
490  shows "x \<notin> {ptr_val p..+n * size_of TYPE('a)} \<Longrightarrow> ptr_retyps_gen n p arr td x = td x"
491  apply (simp add: ptr_retyps_gen_def ptr_retyps_out split: if_split)
492  apply (clarsimp simp: ptr_arr_retyps_def htd_update_list_same2)
493  done
494
495lemma h_t_valid_intvl_htd_contains_uinfo_t:
496  "h_t_valid d g (p :: ('a :: c_type) ptr) \<Longrightarrow> x \<in> {ptr_val p ..+ size_of TYPE('a)} \<Longrightarrow>
497    (\<exists>n. snd (d x) n \<noteq> None \<and> fst (the (snd (d x) n)) = typ_uinfo_t TYPE ('a))"
498  apply (clarsimp simp: h_t_valid_def valid_footprint_def Let_def intvl_def size_of_def)
499  apply (drule spec, drule(1) mp)
500  apply (cut_tac m=k in typ_slice_t_self[where td="typ_uinfo_t TYPE ('a)"])
501  apply (clarsimp simp: in_set_conv_nth)
502  apply (drule_tac x=i in map_leD)
503   apply simp
504  apply fastforce
505  done
506
507lemma list_map_override_comono:
508  "list_map xs  \<subseteq>\<^sub>m m ++ list_map ys
509    \<Longrightarrow> xs \<le> ys \<or> ys \<le> xs"
510  apply (simp add: map_le_def list_map_eq map_add_def)
511  apply (cases "length xs \<le> length ys")
512   apply (simp add: prefix_eq_nth)
513  apply (simp split: if_split_asm add: prefix_eq_nth)
514  done
515
516lemma list_map_plus_le_not_tag_disj:
517  "list_map (typ_slice_t td y) \<subseteq>\<^sub>m m ++ list_map (typ_slice_t td' y')
518    \<Longrightarrow> \<not> td \<bottom>\<^sub>t td'"
519  apply (drule list_map_override_comono)
520  apply (auto dest: typ_slice_sub)
521  done
522
523lemma htd_update_list_not_tag_disj:
524  "list_map (typ_slice_t td y)
525        \<subseteq>\<^sub>m snd (htd_update_list p xs htd x)
526    \<Longrightarrow> x \<in> {p ..+ length xs}
527    \<Longrightarrow> y < size_td td
528    \<Longrightarrow> length xs < addr_card
529    \<Longrightarrow> set xs \<subseteq> list_map ` typ_slice_t td' ` {..< size_td td'}
530    \<Longrightarrow> \<not> td \<bottom>\<^sub>t td'"
531  apply (induct xs arbitrary: p htd)
532   apply simp
533  apply (clarsimp simp: intvl_Suc_right)
534  apply (erule disjE)
535   apply clarsimp
536   apply (subst(asm) htd_update_list_same2,
537     rule intvl_Suc_nmem'[where n="Suc m" for m, simplified])
538    apply (simp add: addr_card_def card_word)
539   apply (simp add: list_map_plus_le_not_tag_disj)
540  apply blast
541  done
542
543(* Sigh *)
544lemma td_set_offset_ind:
545  "\<forall>j. td_set t (Suc j) = (apsnd Suc :: ('a typ_desc \<times> nat) \<Rightarrow> _) ` td_set t j"
546  "\<forall>j. td_set_struct ts (Suc j) = (apsnd Suc :: ('a typ_desc \<times> nat) \<Rightarrow> _) ` td_set_struct ts j"
547  "\<forall>j. td_set_list xs (Suc j) = (apsnd Suc :: ('a typ_desc \<times> nat) \<Rightarrow> _) ` td_set_list xs j"
548  "\<forall>j. td_set_pair x (Suc j) = (apsnd Suc :: ('a typ_desc \<times> nat) \<Rightarrow> _) ` td_set_pair x j"
549  apply (induct t and ts and xs and x)
550  apply (simp_all add: image_Un)
551  done
552
553lemma td_set_offset:
554  "(td, i) \<in> td_set td' j \<Longrightarrow> (td, i - j) \<in> td_set td' 0"
555  by (induct j arbitrary: i, auto simp: td_set_offset_ind)
556
557lemma typ_le_uinfo_array_tag_n_m:
558  "0 < n \<Longrightarrow> td \<le> uinfo_array_tag_n_m TYPE('a :: c_type) n m
559    = (td \<le> typ_uinfo_t TYPE('a) \<or> td = uinfo_array_tag_n_m TYPE('a) n m)"
560proof -
561  have ind: "\<And>xs cs. \<forall>n'. td_set_list (map (\<lambda>i. DTPair (typ_uinfo_t TYPE('a)) (cs i)) xs) n'
562    \<subseteq> (fst ` (\<Union>i. td_set (typ_uinfo_t TYPE('a)) i)) \<times> UNIV"
563    apply (induct_tac xs)
564     apply clarsimp
565    apply clarsimp
566    apply (fastforce intro: image_eqI[rotated])
567    done
568  assume "0 < n"
569  thus ?thesis
570    apply (simp add: uinfo_array_tag_n_m_def typ_tag_le_def upt_conv_Cons)
571    apply (auto dest!: ind[rule_format, THEN subsetD], (blast dest: td_set_offset)+)
572    done
573qed
574
575lemma h_t_array_valid_retyp:
576  "0 < n \<Longrightarrow> n * size_of TYPE('a) < addr_card
577    \<Longrightarrow> h_t_array_valid (ptr_arr_retyps n p htd) (p :: ('a :: wf_type) ptr) n"
578  apply (clarsimp simp: ptr_arr_retyps_def h_t_array_valid_def
579                        valid_footprint_def)
580  apply (simp add: htd_update_list_index intvlI mult.commute)
581  apply (simp add: addr_card_wb unat_of_nat32)
582  done
583
584lemma valid_call_Spec_eq_subset:
585"\<Gamma>' procname = Some (Spec R)
586\<Longrightarrow> (\<forall>x. \<Gamma>'\<Turnstile>\<^bsub>/NF\<^esub> (P x) Call procname (Q x),(A x))
587  = ((\<forall>x. P x \<subseteq> fst ` R) \<and> (R \<subseteq> (\<Inter>x. (- P x) \<times> UNIV \<union> UNIV \<times> Q x)))"
588  apply (safe, simp_all)
589    apply (clarsimp simp: HoarePartialDef.valid_def)
590    apply (rule ccontr)
591    apply (elim allE, subst(asm) imageI, assumption)
592    apply (drule mp, erule exec.Call, rule exec.SpecStuck)
593     apply (auto simp: image_def)[2]
594   apply (clarsimp simp: HoarePartialDef.valid_def)
595   apply (elim allE, drule mp, erule exec.Call, erule exec.Spec)
596   apply auto[1]
597  apply (clarsimp simp: HoarePartialDef.valid_def)
598  apply (erule exec_Normal_elim_cases, simp_all)
599  apply (erule exec_Normal_elim_cases, auto simp: image_def)
600   apply fastforce
601  apply (thin_tac "R \<subseteq> _", fastforce)
602  done
603
604lemma field_of_t_refl:
605  "field_of_t p p' = (p = p')"
606  apply (safe, simp_all add: field_of_t_def field_of_self)
607  apply (simp add: field_of_def)
608  apply (drule td_set_size_lte)
609  apply (simp add: unat_eq_0)
610  done
611
612lemma typ_slice_list_array:
613  "x < size_td td * n
614    \<Longrightarrow> typ_slice_list (map (\<lambda>i. DTPair td (nm i)) [0..<n]) x
615        = typ_slice_t td (x mod size_td td)"
616proof (induct n arbitrary: x nm)
617  case 0 thus ?case by simp
618next
619  case (Suc n)
620  from Suc.prems show ?case
621    apply (simp add: upt_conv_Cons map_Suc_upt[symmetric]
622                del: upt.simps)
623    apply (split if_split, intro conjI impI)
624     apply auto[1]
625    apply (simp add: o_def)
626    apply (subst Suc.hyps)
627     apply arith
628    apply (metis mod_geq)
629    done
630qed
631
632lemma h_t_array_valid_field:
633  "h_t_array_valid htd (p :: ('a :: wf_type) ptr) n
634    \<Longrightarrow> k < n
635    \<Longrightarrow> gd (p +\<^sub>p int k)
636    \<Longrightarrow> h_t_valid htd gd (p +\<^sub>p int k)"
637  apply (clarsimp simp: h_t_array_valid_def h_t_valid_def valid_footprint_def
638                        size_of_def[symmetric, where t="TYPE('a)"])
639  apply (drule_tac x="k * size_of TYPE('a) + y" in spec)
640  apply (drule mp)
641   apply (frule_tac k="size_of TYPE('a)" in mult_le_mono1[where j=n, OF Suc_leI])
642   apply (simp add: mult.commute)
643  apply (clarsimp simp: ptr_add_def add.assoc)
644  apply (erule map_le_trans[rotated])
645  apply (clarsimp simp: uinfo_array_tag_n_m_def)
646  apply (subst typ_slice_list_array)
647   apply (frule_tac k="size_of TYPE('a)" in mult_le_mono1[where j=n, OF Suc_leI])
648   apply (simp add: mult.commute size_of_def)
649  apply (simp add: size_of_def list_map_mono)
650  done
651
652lemma h_t_valid_ptr_retyps_gen:
653  assumes sz: "nptrs * size_of TYPE('a :: mem_type) < addr_card"
654    and gd: "gd p'"
655  shows
656  "(p' \<in> ((+\<^sub>p) (Ptr p :: 'a ptr) \<circ> int) ` {k. k < nptrs})
657    \<Longrightarrow> h_t_valid (ptr_retyps_gen nptrs (Ptr p :: 'a ptr) arr htd) gd p'"
658  using gd sz
659  apply (cases arr, simp_all add: ptr_retyps_gen_def)
660   apply (cases "nptrs = 0")
661    apply simp
662   apply (cut_tac h_t_array_valid_retyp[where p="Ptr p" and htd=htd, OF _ sz], simp_all)
663   apply clarsimp
664   apply (drule_tac k=x in h_t_array_valid_field, simp_all)
665  apply (induct nptrs arbitrary: p htd)
666   apply simp
667  apply clarsimp
668  apply (case_tac x, simp_all add: ptr_retyp_h_t_valid)
669  apply (rule ptr_retyp_disjoint)
670   apply (elim meta_allE, erule meta_mp, rule image_eqI[rotated], simp)
671   apply (simp add: field_simps)
672  apply simp
673  apply (cut_tac p=p and z="size_of TYPE('a)"
674    and k="Suc nat * size_of TYPE('a)" in init_intvl_disj)
675   apply (erule order_le_less_trans[rotated])
676   apply (simp del: mult_Suc)
677  apply (simp add: field_simps Int_ac)
678  apply (erule disjoint_subset[rotated] disjoint_subset2[rotated])
679  apply (rule intvl_start_le, simp)
680  done
681
682lemma ptr_retyps_gen_not_tag_disj:
683  "x \<in> {p ..+ n * size_of TYPE('a :: mem_type)}
684    \<Longrightarrow> list_map (typ_slice_t td y)
685        \<subseteq>\<^sub>m snd (ptr_retyps_gen n (Ptr p :: 'a ptr) arr htd x)
686    \<Longrightarrow> y < size_td td
687    \<Longrightarrow> n * size_of TYPE('a) < addr_card
688    \<Longrightarrow> 0 < n
689    \<Longrightarrow> \<not> td \<bottom>\<^sub>t typ_uinfo_t TYPE('a)"
690  apply (simp add: ptr_retyps_gen_def ptr_arr_retyps_def
691            split: if_split_asm)
692   apply (drule_tac td'="uinfo_array_tag_n_m TYPE('a) n n"
693     in htd_update_list_not_tag_disj, simp+)
694    apply (clarsimp simp: mult.commute)
695   apply (clarsimp simp: tag_disj_def)
696   apply (erule disjE)
697    apply (metis order_refl typ_le_uinfo_array_tag_n_m)
698   apply (erule notE, erule order_trans[rotated])
699   apply (simp add: typ_le_uinfo_array_tag_n_m)
700  apply clarsimp
701  apply (induct n arbitrary: p htd, simp_all)
702  apply (case_tac "x \<in> {p ..+ size_of TYPE('a)}")
703   apply (simp add: intvl_sum ptr_retyp_def)
704   apply (drule_tac td'="typ_uinfo_t TYPE('a)"
705     in htd_update_list_not_tag_disj, simp+)
706    apply (clarsimp simp add: typ_slices_def size_of_def)
707   apply simp
708  apply (simp add: intvl_sum)
709  apply (case_tac "n = 0")
710   apply simp
711  apply (simp add: ptr_retyps_out[where n=1, simplified])
712  apply blast
713  done
714
715lemma ptr_retyps_gen_valid_footprint:
716  assumes cleared: "region_is_bytes' p (n * size_of TYPE('a)) htd"
717    and distinct: "td \<bottom>\<^sub>t typ_uinfo_t TYPE('a)"
718    and not_byte: "td \<noteq> typ_uinfo_t TYPE(word8)"
719    and sz: "n * size_of TYPE('a) < addr_card"
720  shows
721  "valid_footprint (ptr_retyps_gen n (Ptr p :: 'a :: mem_type ptr) arr htd) p' td
722    = (valid_footprint htd p' td)"
723  apply (cases "n = 0")
724   apply (simp add: ptr_retyps_gen_def ptr_arr_retyps_def split: if_split)
725  apply (simp add: valid_footprint_def Let_def)
726  apply (intro conj_cong refl, rule all_cong)
727  apply (case_tac "p' + of_nat y \<in> {p ..+ n * size_of TYPE('a)}")
728   apply (simp_all add: ptr_retyps_gen_out)
729  apply (rule iffI; clarsimp)
730   apply (frule(1) ptr_retyps_gen_not_tag_disj, (simp add: sz)+)
731   apply (simp add: distinct)
732  apply (cut_tac m=y in typ_slice_t_self[where td=td])
733  apply (clarsimp simp: in_set_conv_nth)
734  apply (drule_tac x=i in map_leD)
735   apply simp
736  apply (simp add: cleared[unfolded region_is_bytes'_def] not_byte)
737  done
738
739lemma list_map_length_is_None [simp]:
740  "list_map xs (length xs) = None"
741  apply (induct xs)
742   apply (simp add: list_map_def)
743  apply (simp add: list_map_def)
744  done
745
746lemma list_map_append_one:
747  "list_map (xs @ [x]) = [length xs \<mapsto> x] ++ list_map xs"
748  by (simp add: list_map_def)
749
750lemma ptr_retyp_same_cleared_region:
751  fixes p :: "'a :: mem_type ptr" and p' :: "'a :: mem_type ptr"
752  assumes  ht: "ptr_retyp p td, g \<Turnstile>\<^sub>t p'"
753  shows "p = p' \<or> {ptr_val p..+ size_of TYPE('a)} \<inter> {ptr_val p' ..+ size_of TYPE('a)} = {}"
754  using ht
755  by (simp add: h_t_valid_ptr_retyp_eq[where p=p and p'=p'] field_of_t_refl
756         split: if_split_asm)
757
758lemma h_t_valid_ptr_retyp_inside_eq:
759  fixes p :: "'a :: mem_type ptr" and p' :: "'a :: mem_type ptr"
760  assumes inside: "ptr_val p' \<in> {ptr_val p ..+ size_of TYPE('a)}"
761  and         ht: "ptr_retyp p td, g \<Turnstile>\<^sub>t p'"
762  shows   "p = p'"
763  using ptr_retyp_same_cleared_region[OF ht] inside mem_type_self[where p=p']
764  by blast
765
766lemma ptr_add_orth:
767  fixes p :: "'a :: mem_type ptr"
768  assumes lt: "Suc n * size_of TYPE('a) < 2 ^ word_bits"
769  shows "{ptr_val p..+size_of TYPE('a)} \<inter> {ptr_val (CTypesDefs.ptr_add p 1)..+n * size_of TYPE('a)} = {}"
770  using lt
771  apply -
772  apply (rule disjointI)
773  apply clarsimp
774  apply (drule intvlD)+
775  apply (clarsimp simp: CTypesDefs.ptr_add_def)
776  apply (simp only: Abs_fnat_hom_add)
777  apply (drule unat_cong)
778  apply (simp only: unat_of_nat)
779  apply (unfold word_bits_len_of)
780   apply (simp add: addr_card_wb [symmetric])
781  done
782
783lemma dom_lift_t_heap_update:
784  "dom (lift_t g (hrs_mem_update v hp)) = dom (lift_t g hp)"
785  by (clarsimp simp add: lift_t_def lift_typ_heap_if s_valid_def hrs_htd_def hrs_mem_update_def split_def dom_def
786    intro!: Collect_cong split: if_split)
787
788lemma h_t_valid_ptr_retyps_gen_same:
789  assumes guard: "\<forall>n' < nptrs. gd (CTypesDefs.ptr_add (Ptr p :: 'a ptr) (of_nat n'))"
790  assumes cleared: "region_is_bytes' p (nptrs * size_of TYPE('a :: mem_type)) htd"
791  and not_byte: "typ_uinfo_t TYPE('a) \<noteq> typ_uinfo_t TYPE(word8)"
792  assumes sz: "nptrs * size_of TYPE('a) < addr_card"
793  shows
794  "h_t_valid (ptr_retyps_gen nptrs (Ptr p :: 'a ptr) arr htd) gd p'
795    = ((p' \<in> ((+\<^sub>p) (Ptr p :: 'a ptr) \<circ> int) ` {k. k < nptrs}) \<or> h_t_valid htd gd p')"
796  (is "h_t_valid ?htd' gd p' = (p' \<in> ?S \<or> h_t_valid htd gd p')")
797proof (cases "{ptr_val p' ..+ size_of TYPE('a)} \<inter> {p ..+ nptrs * size_of TYPE('a)} = {}")
798  case True
799
800  from True have notin:
801    "p' \<notin> ?S"
802    apply clarsimp
803    apply (drule_tac x="p + of_nat (x * size_of TYPE('a))" in eqset_imp_iff)
804    apply (simp only: Int_iff empty_iff simp_thms)
805    apply (subst(asm) intvlI, simp)
806    apply (simp add: intvl_self)
807    done
808
809  from True have same: "\<forall>y < size_of TYPE('a). ?htd' (ptr_val p' + of_nat y)
810        = htd (ptr_val p' + of_nat y)"
811    apply clarsimp
812    apply (rule ptr_retyps_gen_out)
813    apply simp
814    apply (blast intro: intvlI)
815    done
816
817  show ?thesis
818    by (clarsimp simp: h_t_valid_def valid_footprint_def Let_def
819                       notin same size_of_def[symmetric, where t="TYPE('a)"])
820next
821  case False
822
823  from False have nvalid: "\<not> h_t_valid htd gd p'"
824    apply (clarsimp simp: h_t_valid_def valid_footprint_def set_eq_iff
825                          Let_def size_of_def[symmetric, where t="TYPE('a)"]
826                          intvl_def[where x="(ptr_val p', a)" for a])
827    apply (drule cleared[unfolded region_is_bytes'_def, THEN bspec])
828    apply (drule spec, drule(1) mp, clarsimp)
829    apply (cut_tac m=k in typ_slice_t_self[where td="typ_uinfo_t TYPE ('a)"])
830    apply (clarsimp simp: in_set_conv_nth)
831    apply (drule_tac x=i in map_leD, simp_all)
832    apply (simp add: not_byte)
833    done
834
835  have mod_split: "\<And>k. k < nptrs * size_of TYPE('a)
836    \<Longrightarrow> \<exists>quot rem. k = quot * size_of TYPE('a) + rem \<and> rem < size_of TYPE('a) \<and> quot < nptrs"
837    apply (intro exI conjI, rule div_mult_mod_eq[symmetric])
838     apply simp
839    apply (simp add: Word_Miscellaneous.td_gal_lt)
840    done
841
842  have gd: "\<And>p'. p' \<in> ?S \<Longrightarrow> gd p'"
843    using guard by auto
844
845  note htv = h_t_valid_ptr_retyps_gen[where gd=gd, OF sz gd]
846
847  show ?thesis using False
848    apply (simp add: nvalid)
849    apply (rule iffI, simp_all add: htv)
850    apply (clarsimp simp: set_eq_iff intvl_def[where x="(p, a)" for a])
851    apply (drule mod_split, clarsimp)
852    apply (frule_tac htv[OF imageI, simplified])
853     apply fastforce
854    apply (rule ccontr)
855    apply (drule(1) h_t_valid_neq_disjoint)
856      apply simp
857     apply (clarsimp simp: field_of_t_refl)
858    apply (simp add: set_eq_iff)
859    apply (drule spec, drule(1) mp)
860    apply (subst(asm) add.assoc[symmetric], subst(asm) intvlI, assumption)
861    apply simp
862    done
863qed
864
865lemma clift_ptr_retyps_gen_memset_same:
866  assumes guard: "\<forall>n' < n. c_guard (CTypesDefs.ptr_add (Ptr p :: 'a :: mem_type ptr) (of_nat n'))"
867  assumes cleared: "region_is_bytes' p (n * size_of TYPE('a :: mem_type)) (hrs_htd hrs)"
868    and not_byte: "typ_uinfo_t TYPE('a :: mem_type) \<noteq> typ_uinfo_t TYPE(word8)"
869  and nb: "nb = n * size_of TYPE ('a)"
870  and sz: "n * size_of TYPE('a) < 2 ^ word_bits"
871  shows "(clift (hrs_htd_update (ptr_retyps_gen n (Ptr p :: 'a :: mem_type ptr) arr)
872              (hrs_mem_update (heap_update_list p (replicate nb 0))
873               hrs)) :: 'a :: mem_type typ_heap)
874         = (\<lambda>y. if y \<in> (CTypesDefs.ptr_add (Ptr p :: 'a :: mem_type ptr) o of_nat) ` {k. k < n}
875                then Some (from_bytes (replicate (size_of TYPE('a  :: mem_type)) 0)) else clift hrs y)"
876  using sz
877  apply (simp add: nb liftt_if[folded hrs_mem_def hrs_htd_def]
878                   hrs_htd_update hrs_mem_update
879                   h_t_valid_ptr_retyps_gen_same[OF guard cleared not_byte]
880                   addr_card_wb)
881  apply (rule ext, rename_tac p')
882  apply (case_tac "p' \<in> ((+\<^sub>p) (Ptr p) \<circ> int) ` {k. k < n}")
883   apply (clarsimp simp: h_val_def)
884   apply (simp only: Word.Abs_fnat_hom_mult hrs_mem_update)
885   apply (frule_tac k="size_of TYPE('a)" in mult_le_mono1[where j=n, OF Suc_leI])
886   apply (subst heap_list_update_list)
887    apply (simp add: addr_card_def card_word word_bits_def)
888   apply simp
889  apply (clarsimp split: if_split)
890  apply (simp add: h_val_def)
891  apply (subst heap_list_update_disjoint_same, simp_all)
892  apply (simp add: region_is_bytes_disjoint[OF cleared not_byte])
893  done
894
895lemma clift_ptr_retyps_gen_prev_memset_same:
896  assumes guard: "\<forall>n' < n. c_guard (CTypesDefs.ptr_add (Ptr p :: 'a :: mem_type ptr) (of_nat n'))"
897  assumes cleared: "region_is_bytes' p (n * size_of TYPE('a :: mem_type)) (hrs_htd hrs)"
898    and not_byte: "typ_uinfo_t TYPE('a :: mem_type) \<noteq> typ_uinfo_t TYPE(word8)"
899  and nb: "nb = n * size_of TYPE ('a)"
900  and sz: "n * size_of TYPE('a) < 2 ^ word_bits"
901  and rep0:  "heap_list (hrs_mem hrs) nb p = replicate nb 0"
902  shows "(clift (hrs_htd_update (ptr_retyps_gen n (Ptr p :: 'a :: mem_type ptr) arr) hrs) :: 'a :: mem_type typ_heap)
903         = (\<lambda>y. if y \<in> (CTypesDefs.ptr_add (Ptr p :: 'a :: mem_type ptr) o of_nat) ` {k. k < n}
904                then Some (from_bytes (replicate (size_of TYPE('a  :: mem_type)) 0)) else clift hrs y)"
905  using rep0
906  apply (subst clift_ptr_retyps_gen_memset_same[symmetric, OF guard cleared not_byte nb sz])
907  apply (rule arg_cong[where f=clift])
908  apply (rule_tac f="hrs_htd_update f" for f in arg_cong)
909  apply (cases hrs, simp add: hrs_mem_update_def)
910  apply (simp add: heap_update_list_id hrs_mem_def)
911  done
912
913lemma clift_ptr_retyps_gen_other:
914  assumes cleared: "region_is_bytes' (ptr_val p) (nptrs * size_of TYPE('a :: mem_type)) (hrs_htd hrs)"
915  and sz: "nptrs * size_of TYPE('a) < 2 ^ word_bits"
916  and other: "typ_uinfo_t TYPE('b)  \<bottom>\<^sub>t typ_uinfo_t TYPE('a)"
917  and not_byte: "typ_uinfo_t TYPE('b :: mem_type) \<noteq> typ_uinfo_t TYPE(word8)"
918  shows "(clift (hrs_htd_update (ptr_retyps_gen nptrs (p :: 'a ptr) arr) hrs) :: 'b :: mem_type typ_heap)
919         = clift hrs"
920  using sz cleared
921  apply (cases p)
922  apply (simp add: liftt_if[folded hrs_mem_def hrs_htd_def]
923                   h_t_valid_def hrs_htd_update
924                   ptr_retyps_gen_valid_footprint[simplified addr_card_wb, OF _ other not_byte sz])
925  done
926
927lemma clift_heap_list_update_no_heap_other:
928  assumes cleared: "region_is_bytes' p (length xs) (hrs_htd hrs)"
929  and not_byte: "typ_uinfo_t TYPE('a :: c_type) \<noteq> typ_uinfo_t TYPE(word8)"
930  shows "clift (hrs_mem_update (heap_update_list p xs) hrs) = (clift hrs :: 'a typ_heap)"
931  apply (clarsimp simp: liftt_if[folded hrs_mem_def hrs_htd_def] hrs_mem_update
932                        fun_eq_iff h_val_def split: if_split)
933  apply (subst heap_list_update_disjoint_same, simp_all)
934  apply (clarsimp simp: set_eq_iff h_t_valid_def valid_footprint_def Let_def
935                 dest!: intvlD[where n="size_of TYPE('a)"])
936  apply (drule_tac x="of_nat k" in spec, clarsimp simp: size_of_def)
937  apply (cut_tac m=k in typ_slice_t_self[where td="typ_uinfo_t TYPE('a)"])
938  apply (clarsimp simp: in_set_conv_nth)
939  apply (drule_tac x=i in map_leD, simp)
940  apply (simp add: cleared[unfolded region_is_bytes'_def] not_byte size_of_def)
941  done
942
943lemma add_is_injective_ring:
944  "inj ((+) (x :: 'a :: ring))"
945  by (rule inj_onI, clarsimp)
946
947(* assumes that y & elements are n-aligned but not that the compound
948   interval is aligned to a higher power of two. needed for cte arrays. *)
949lemma ptr_span_disjoint_ptr_set_span:
950  fixes y :: "('a :: mem_type) ptr"
951  assumes align: "is_aligned p n"
952  and size_of: "size_of TYPE('a) = 2 ^ n"
953  and al: "is_aligned (ptr_val y) n"
954  and card: "b * 2 ^ n < addr_card"
955  and b: "b \<noteq> 0"
956  shows "y \<notin> ((+\<^sub>p) (Ptr p) \<circ> int) ` {k. k < b}
957    \<longrightarrow> ptr_span y \<inter> {p ..+ b * 2 ^ n} = {}"
958proof -
959  from card b have word_bits: "n < word_bits"
960    using power_increasing[where n=word_bits and N=n and a=2]
961    apply (simp add: word_bits_def addr_card)
962    apply (rule ccontr, simp)
963    apply (cases b, simp_all)
964    apply (drule(1) order_less_le_trans)
965    apply simp
966    done
967
968  note al_sub = aligned_sub_aligned_simple[OF al align]
969
970  have yuck: "of_nat b * 2 ^ n \<noteq> (0 :: word32)"
971    using of_nat_neq_0[where k="b * 2 ^ n" and 'a=32] b card
972    by (clarsimp simp: addr_card_def card_word)
973
974  show ?thesis
975    apply (clarsimp simp add: size_of)
976    apply (rule inj_image_eq_iff[OF add_is_injective_ring[where x="- p"], THEN iffD1])
977    apply (subst image_Int[OF add_is_injective_ring])
978    apply (simp add: image_add_intvl upto_intvl_eq al_sub)
979    apply (subst upto_intvl_eq', simp, simp add: b)
980     apply (cut_tac card, simp add: addr_card_def card_word)
981    apply safe
982    apply (simp only: mask_in_range[symmetric] al_sub)
983    apply simp
984    apply (drule_tac f="(+) p" in arg_cong, simp)
985    apply (erule notE, rule_tac x="unat (x >> n)" in image_eqI)
986     apply (simp add: size_of)
987     apply (cases y, clarsimp simp: and_not_mask shiftl_t2n)
988    apply (simp add: shiftr_div_2n')
989    apply (rule Word_Miscellaneous.td_gal_lt[THEN iffD1], simp)
990    apply (drule minus_one_helper5[OF yuck])
991    apply (rule unat_less_helper, simp)
992    done
993qed
994
995lemma ptr_retyp_to_array:
996  "ptr_retyps_gen 1 (p :: (('a :: wf_type)['b :: finite]) ptr) False
997    = ptr_retyps_gen CARD('b) (ptr_coerce p :: 'a ptr) True"
998  by (intro ext, simp add: ptr_retyps_gen_def ptr_arr_retyps_to_retyp)
999
1000lemma projectKO_opt_retyp_other:
1001  assumes cover: "range_cover ptr sz (objBitsKO ko) n"
1002  assumes pal: "pspace_aligned' \<sigma>"
1003  assumes pno: "pspace_no_overlap' ptr sz \<sigma>"
1004  and  ko_def: "ko \<equiv> x"
1005  and  pko: "\<forall>v. (projectKO_opt x :: ('a :: pre_storable) option) \<noteq> Some v"
1006  shows "projectKO_opt \<circ>\<^sub>m
1007    (\<lambda>x. if x \<in> set (new_cap_addrs n ptr ko) then Some ko else ksPSpace \<sigma> x)
1008  = (projectKO_opt \<circ>\<^sub>m (ksPSpace \<sigma>) :: word32 \<Rightarrow> ('a :: pre_storable) option)" (is "?LHS = ?RHS")
1009proof (rule ext)
1010  fix x
1011  show "?LHS x = ?RHS x"
1012  proof (cases "x \<in> set (new_cap_addrs n ptr ko)")
1013    case False
1014      thus ?thesis by (simp add: map_comp_def)
1015  next
1016    case True
1017      hence "ksPSpace \<sigma> x = None"
1018        apply -
1019        apply (cut_tac no_overlap_new_cap_addrs_disjoint [OF cover pal pno])
1020          apply (rule ccontr)
1021          apply (clarsimp,drule domI[where a = x])
1022          apply blast
1023        done
1024      thus ?thesis using True pko ko_def by simp
1025  qed
1026qed
1027
1028lemma pspace_aligned_to_C:
1029  fixes v :: "'a :: pre_storable"
1030  assumes pal: "pspace_aligned' s"
1031  and    cmap: "cmap_relation (projectKO_opt \<circ>\<^sub>m (ksPSpace s) :: word32 \<rightharpoonup> 'a)
1032                              (cslift x :: 'b :: mem_type typ_heap) Ptr rel"
1033  and     pko: "projectKO_opt ko = Some v"
1034  and   pkorl: "\<And>ko' (v' :: 'a).  projectKO_opt ko' = Some v' \<Longrightarrow> objBitsKO ko = objBitsKO ko'"
1035  shows  "\<forall>x\<in>dom (cslift x :: 'b :: mem_type typ_heap). is_aligned (ptr_val x) (objBitsKO ko)"
1036  (is "\<forall>x\<in>dom ?CS. is_aligned (ptr_val x) (objBitsKO ko)")
1037proof
1038  fix z
1039  assume "z \<in> dom ?CS"
1040  hence "z \<in> Ptr ` dom (projectKO_opt \<circ>\<^sub>m (ksPSpace s) :: word32 \<rightharpoonup> 'a)" using cmap
1041    by (simp add: cmap_relation_def)
1042  hence pvz: "ptr_val z \<in> dom (projectKO_opt \<circ>\<^sub>m (ksPSpace s) :: word32 \<rightharpoonup> 'a)"
1043    by clarsimp
1044  then obtain v' :: 'a where "projectKO_opt (the (ksPSpace s (ptr_val z))) = Some v'"
1045    and pvz: "ptr_val z \<in> dom (ksPSpace s)"
1046    apply -
1047    apply (frule map_comp_subset_domD)
1048    apply (clarsimp simp: dom_def)
1049    done
1050
1051  thus "is_aligned (ptr_val z) (objBitsKO ko)" using pal
1052    unfolding pspace_aligned'_def
1053    apply -
1054    apply (drule (1) bspec)
1055    apply (simp add: pkorl)
1056    done
1057qed
1058
1059lemma pspace_aligned_to_C_cte:
1060  fixes v :: "cte"
1061  assumes pal: "pspace_aligned' s"
1062  and    cmap: "cmap_relation (ctes_of s) (cslift x :: cte_C typ_heap) Ptr ccte_relation"
1063  and     pko: "projectKO_opt ko = Some v"
1064  shows  "\<forall>x\<in>dom (cslift x :: cte_C typ_heap). is_aligned (ptr_val x) (objBitsKO ko)"
1065  (is "\<forall>x\<in>dom ?CS. is_aligned (ptr_val x) (objBitsKO ko)")
1066proof
1067  fix z
1068  assume "z \<in> dom ?CS"
1069  hence "z \<in> Ptr ` dom (ctes_of s)" using cmap
1070    by (simp add: cmap_relation_def)
1071  hence pvz: "ptr_val z \<in> dom (ctes_of s)"
1072    by clarsimp
1073  thus "is_aligned (ptr_val z) (objBitsKO ko)" using pal pko
1074    unfolding pspace_aligned'_def
1075    apply -
1076    apply clarsimp
1077    apply (drule ctes_of_is_aligned)
1078    apply (cases ko, simp_all add: projectKOs)
1079    apply (simp add: objBits_simps)
1080    done
1081qed
1082
1083lemma pspace_aligned_to_C_tcb:
1084  fixes v :: "tcb"
1085  assumes pal: "pspace_aligned' s"
1086  and    cmap: "cpspace_tcb_relation (ksPSpace s) (t_hrs_' (globals x))"
1087  shows  "\<forall>x\<in>dom (cslift x :: tcb_C typ_heap). is_aligned (ptr_val x) ctcb_size_bits"
1088  (is "\<forall>x\<in>dom ?CS. is_aligned (ptr_val x) ctcb_size_bits")
1089proof
1090  fix z
1091  assume "z \<in> dom ?CS"
1092  hence "z \<in> tcb_ptr_to_ctcb_ptr ` dom (map_to_tcbs (ksPSpace s))" using cmap
1093    by (simp add: cmap_relation_def)
1094  hence pvz: "ctcb_ptr_to_tcb_ptr z \<in> dom (map_to_tcbs (ksPSpace s))"
1095    by clarsimp
1096  then obtain v' :: tcb where "projectKO_opt (the (ksPSpace s (ctcb_ptr_to_tcb_ptr z))) = Some v'"
1097    and pvz: "ctcb_ptr_to_tcb_ptr z \<in> dom (ksPSpace s)"
1098    apply -
1099    apply (frule map_comp_subset_domD)
1100    apply (clarsimp simp: dom_def)
1101    done
1102
1103  thus "is_aligned (ptr_val z) ctcb_size_bits" using pal
1104    unfolding pspace_aligned'_def
1105    apply -
1106    apply (drule (1) bspec)
1107    apply (clarsimp simp add: projectKOs objBits_simps)
1108    apply (erule ctcb_ptr_to_tcb_ptr_aligned)
1109    done
1110qed
1111
1112lemma ptr_add_to_new_cap_addrs:
1113  assumes size_of_m: "size_of TYPE('a :: mem_type) = 2 ^ objBitsKO ko"
1114  shows "(CTypesDefs.ptr_add (Ptr ptr :: 'a :: mem_type ptr) \<circ> of_nat) ` {k. k < n}
1115   = Ptr ` set (new_cap_addrs n ptr ko)"
1116  unfolding new_cap_addrs_def
1117  apply (simp add: comp_def image_image shiftl_t2n size_of_m field_simps)
1118  apply (clarsimp simp: atLeastLessThan_def lessThan_def)
1119  done
1120
1121lemma cmap_relation_retype:
1122  assumes cm: "cmap_relation mp mp' Ptr rel"
1123  and   rel: "rel (makeObject :: 'a :: pspace_storable) ko'"
1124  shows "cmap_relation
1125        (\<lambda>x. if x \<in> addrs then Some (makeObject :: 'a :: pspace_storable) else mp x)
1126        (\<lambda>y. if y \<in> Ptr ` addrs then Some ko' else mp' y)
1127        Ptr rel"
1128  using cm rel
1129  apply -
1130  apply (rule cmap_relationI)
1131   apply (simp add: dom_if cmap_relation_def image_Un)
1132  apply (case_tac "x \<in> addrs")
1133   apply simp
1134  apply simp
1135  apply (subst (asm) if_not_P)
1136   apply clarsimp
1137  apply (erule (2) cmap_relation_relI)
1138  done
1139
1140lemma update_ti_t_word32_0s:
1141  "update_ti_t (typ_info_t TYPE(word32)) [0,0,0,0] X = 0"
1142  "word_rcat [0, 0, 0, (0 :: word8)] = (0 :: word32)"
1143  by (simp_all add: typ_info_word word_rcat_def bin_rcat_def)
1144
1145lemma is_aligned_ptr_aligned:
1146  fixes p :: "'a :: c_type ptr"
1147  assumes al: "is_aligned (ptr_val p) n"
1148  and  alignof: "align_of TYPE('a) = 2 ^ n"
1149  shows "ptr_aligned p"
1150  using al unfolding is_aligned_def ptr_aligned_def
1151  by (simp add: alignof)
1152
1153lemma is_aligned_c_guard:
1154  "is_aligned (ptr_val p) n
1155    \<Longrightarrow> ptr_val p \<noteq> 0
1156    \<Longrightarrow> align_of TYPE('a) = 2 ^ m
1157    \<Longrightarrow> size_of TYPE('a) \<le> 2 ^ n
1158    \<Longrightarrow> m \<le> n
1159    \<Longrightarrow> c_guard (p :: ('a :: c_type) ptr)"
1160  apply (clarsimp simp: c_guard_def c_null_guard_def)
1161  apply (rule conjI)
1162   apply (rule is_aligned_ptr_aligned, erule(1) is_aligned_weaken, simp)
1163  apply (erule is_aligned_get_word_bits, simp_all)
1164  apply (rule intvl_nowrap[where x=0, simplified], simp)
1165  apply (erule is_aligned_no_wrap_le, simp+)
1166  done
1167
1168lemma retype_guard_helper:
1169  assumes cover: "range_cover p sz (objBitsKO ko) n"
1170  and ptr0: "p \<noteq> 0"
1171  and szo: "size_of TYPE('a :: c_type) = 2 ^ objBitsKO ko"
1172  and lt2: "m \<le> objBitsKO ko"
1173  and ala: "align_of TYPE('a :: c_type) = 2 ^ m"
1174  shows "\<forall>b < n. c_guard (CTypesDefs.ptr_add (Ptr p :: 'a ptr) (of_nat b))"
1175proof (rule allI, rule impI)
1176  fix b :: nat
1177  assume nv: "b < n"
1178  let ?p = "(Ptr p :: 'a ptr)"
1179
1180  have "of_nat b * of_nat (size_of TYPE('a)) = (of_nat (b * 2 ^ objBitsKO ko) :: word32)"
1181    by (simp add: szo)
1182
1183  also have "\<dots> < (2 :: word32) ^ sz" using nv cover
1184    apply simp
1185    apply (rule word_less_power_trans_ofnat)
1186      apply (erule less_le_trans)
1187      apply (erule range_cover.range_cover_n_le(2))
1188    apply (erule range_cover.sz)+
1189    done
1190
1191  finally have ofn: "of_nat b * of_nat (size_of TYPE('a)) < (2 :: word32) ^ sz" .
1192
1193  have le: "p \<le> p + of_nat b * 2 ^ objBitsKO ko"
1194    using ofn szo nv
1195    apply -
1196    apply (cases b,clarsimp+)
1197    apply (cut_tac n = nat in range_cover_ptr_le)
1198     apply (rule range_cover_le[OF cover])
1199      apply simp
1200     apply (simp add:ptr0)
1201    apply (simp add:shiftl_t2n field_simps)
1202    done
1203
1204  show "c_guard (CTypesDefs.ptr_add ?p (of_nat b))"
1205    apply (rule is_aligned_c_guard[OF _ _ ala _ lt2])
1206      apply (simp add: szo)
1207      apply (rule is_aligned_add)
1208       apply (rule range_cover.aligned, rule cover)
1209      apply (rule is_aligned_mult_triv2)
1210     apply (simp add: szo neq_0_no_wrap[OF le ptr0])
1211    apply (simp add: szo)
1212    done
1213qed
1214
1215lemma retype_guard_helper2:
1216  assumes cover: "range_cover p sz (objBitsKO ko) n"
1217  and ptr0: "p \<noteq> 0"
1218  and szo: "size_of TYPE('a :: c_type) = 2 ^ objBitsKO ko"
1219  and ala: "align_of TYPE('a :: c_type) \<in> set (map (\<lambda>x. 2 ^ x) [0 ..< Suc (objBitsKO ko)])"
1220  shows "\<forall>b < n. c_guard (CTypesDefs.ptr_add (Ptr p :: 'a ptr) (of_nat b))"
1221  using ala retype_guard_helper[OF cover ptr0 szo]
1222  by (clarsimp simp del: upt.simps)
1223
1224(* When we are retyping, CTEs in the system do not change,
1225 * unless we happen to be retyping into a CNode or a TCB,
1226 * in which case new CTEs only pop up in the new object. *)
1227lemma retype_ctes_helper:
1228  assumes pal: "pspace_aligned' s"
1229  and    pdst: "pspace_distinct' s"
1230  and     pno: "pspace_no_overlap' ptr sz s"
1231  and      al: "is_aligned ptr (objBitsKO ko)"
1232  and      sz: "objBitsKO ko \<le> sz"
1233  and     szb: "sz < word_bits"
1234  and     mko: "makeObjectKO dev tp = Some ko"
1235  and      rc: "range_cover ptr sz (objBitsKO ko) n"
1236  shows  "map_to_ctes (\<lambda>xa. if xa \<in> set (new_cap_addrs n ptr ko) then Some ko else ksPSpace s xa) =
1237   (\<lambda>x. if tp = Inr (APIObjectType ArchTypes_H.apiobject_type.CapTableObject) \<and> x \<in> set (new_cap_addrs n ptr ko) \<or>
1238           tp = Inr (APIObjectType ArchTypes_H.apiobject_type.TCBObject) \<and>
1239           x && ~~ mask tcbBlockSizeBits \<in> set (new_cap_addrs n ptr ko) \<and> x && mask tcbBlockSizeBits \<in> dom tcb_cte_cases
1240        then Some (CTE capability.NullCap nullMDBNode) else ctes_of s x)"
1241  using mko pal pdst
1242proof (rule ctes_of_retype)
1243  show "pspace_aligned' (s\<lparr>ksPSpace := \<lambda>xa. if xa \<in> set (new_cap_addrs n ptr ko) then Some ko else ksPSpace s xa\<rparr>)"
1244    using pal pdst pno szb al sz rc
1245    apply -
1246    apply (rule retype_aligned_distinct'', simp_all)
1247    done
1248
1249  show "pspace_distinct' (s\<lparr>ksPSpace := \<lambda>xa. if xa \<in> set (new_cap_addrs n ptr ko) then Some ko else ksPSpace s xa\<rparr>)"
1250    using pal pdst pno szb al sz rc
1251    apply -
1252    apply (rule retype_aligned_distinct'', simp_all)
1253    done
1254
1255  show "\<forall>x\<in>set (new_cap_addrs n ptr ko). is_aligned x (objBitsKO ko)"
1256    using al szb
1257    apply -
1258    apply (rule new_cap_addrs_aligned, simp_all)
1259    done
1260
1261  show "\<forall>x\<in>set (new_cap_addrs n ptr ko). ksPSpace s x = None"
1262    using al szb pno pal rc sz
1263    apply -
1264    apply (drule(1) pspace_no_overlap_disjoint')
1265    apply (frule new_cap_addrs_subset)
1266    apply (clarsimp simp: Word_Lib.ptr_add_def field_simps)
1267    apply fastforce
1268    done
1269qed
1270
1271lemma ptr_retyps_htd_safe:
1272  "\<lbrakk> htd_safe D htd;
1273    {ptr_val ptr ..+ n * size_of TYPE('a :: mem_type)}
1274        \<subseteq> D \<rbrakk>
1275   \<Longrightarrow> htd_safe D (ptr_retyps_gen n (ptr :: 'a ptr) arr htd)"
1276  apply (clarsimp simp: htd_safe_def)
1277  apply (case_tac "a \<in> {ptr_val ptr..+n * size_of TYPE('a)}")
1278   apply blast
1279  apply (case_tac "(a, b) \<in> dom_s htd")
1280   apply blast
1281  apply (clarsimp simp: dom_s_def ptr_retyps_gen_out)
1282  done
1283
1284lemma ptr_retyps_htd_safe_neg:
1285  "\<lbrakk> htd_safe (- D) htd;
1286    {ptr_val ptr ..+ n * size_of TYPE('a :: mem_type)}
1287        \<inter> D = {} \<rbrakk>
1288   \<Longrightarrow> htd_safe (- D) (ptr_retyps_gen n (ptr :: 'a ptr) arr htd)"
1289  using ptr_retyps_htd_safe by blast
1290
1291lemma region_is_bytes_subset:
1292  "region_is_bytes' ptr sz htd
1293    \<Longrightarrow> {ptr' ..+ sz'} \<subseteq> {ptr ..+ sz}
1294    \<Longrightarrow> region_is_bytes' ptr' sz' htd"
1295  by (auto simp: region_is_bytes'_def)
1296
1297lemma (in range_cover) strong_times_32:
1298  "len_of TYPE('a) = len_of TYPE(32) \<Longrightarrow> n * 2 ^ sbit < 2 ^ word_bits"
1299  apply (simp add: nat_mult_power_less_eq)
1300  apply (rule order_less_le_trans, rule string)
1301  apply (simp add: word_bits_def)
1302  done
1303
1304(* Helper for use in the many proofs below. *)
1305lemma cslift_ptr_retyp_other_inst:
1306  assumes   bytes: "region_is_bytes' p (n * (2 ^ bits)) (hrs_htd hp)"
1307  and       cover: "range_cover p sz bits n"
1308  and          sz: "region_sz = n * size_of TYPE('a :: mem_type)"
1309  and         sz2: "size_of TYPE('a :: mem_type) = 2 ^ bits"
1310  and       tdisj: "typ_uinfo_t TYPE('b) \<bottom>\<^sub>t typ_uinfo_t TYPE('a)"
1311  and    not_byte: "typ_uinfo_t TYPE('b :: mem_type) \<noteq> typ_uinfo_t TYPE(word8)"
1312  shows "(clift (hrs_htd_update (ptr_retyps_gen n (Ptr p :: 'a :: mem_type ptr) arr)
1313               hp) :: 'b :: mem_type typ_heap)
1314         = clift hp"
1315  using bytes
1316  apply (subst clift_ptr_retyps_gen_other[OF _ _ tdisj not_byte], simp_all)
1317   apply (simp add: sz2)
1318  apply (simp add: sz2 range_cover.strong_times_32[OF cover])
1319  done
1320
1321(* Helper for use in the many proofs below. *)
1322lemma cslift_ptr_retyp_memset_other_inst:
1323  assumes   bytes: "region_is_bytes p (n * (2 ^ bits)) x"
1324  and       cover: "range_cover p sz bits n"
1325  and          sz: "region_sz = n * size_of TYPE('a :: mem_type)"
1326  and         sz2: "size_of TYPE('a :: mem_type) = 2 ^ bits"
1327  and       tdisj: "typ_uinfo_t TYPE('b) \<bottom>\<^sub>t typ_uinfo_t TYPE('a)"
1328  and    not_byte: "typ_uinfo_t TYPE('b :: mem_type) \<noteq> typ_uinfo_t TYPE(word8)"
1329  shows "(clift (hrs_htd_update (ptr_retyps_gen n (Ptr p :: 'a :: mem_type ptr) arr)
1330              (hrs_mem_update (heap_update_list p (replicate (region_sz) 0))
1331               (t_hrs_' (globals x)))) :: 'b :: mem_type typ_heap)
1332         = cslift x"
1333  using bytes
1334  apply (subst cslift_ptr_retyp_other_inst[OF _ cover sz sz2 tdisj not_byte])
1335   apply simp
1336  apply (rule clift_heap_list_update_no_heap_other[OF _ not_byte])
1337  apply (simp add: hrs_htd_def sz sz2)
1338  done
1339
1340lemma ptr_retyps_one:
1341  "ptr_retyps (Suc 0) = ptr_retyp"
1342  apply (rule ext)+
1343  apply simp
1344  done
1345
1346lemma uinfo_array_tag_n_m_not_le_typ_name:
1347  "typ_name (typ_info_t TYPE('b)) @ ''_array_'' @ nat_to_bin_string m
1348      \<notin> td_names (typ_info_t TYPE('a))
1349    \<Longrightarrow> \<not> uinfo_array_tag_n_m TYPE('b :: c_type) n m \<le> typ_uinfo_t TYPE('a :: c_type)"
1350  apply (clarsimp simp: typ_tag_le_def typ_uinfo_t_def)
1351  apply (drule td_set_td_names)
1352   apply (clarsimp simp: uinfo_array_tag_n_m_def typ_uinfo_t_def)
1353   apply (drule arg_cong[where f="\<lambda>xs. set ''r'' \<subseteq> set xs"], simp)
1354  apply (simp add: uinfo_array_tag_n_m_def typ_uinfo_t_def)
1355  done
1356
1357lemma tag_not_le_via_td_name:
1358  "typ_name (typ_info_t TYPE('a)) \<notin> td_names (typ_info_t TYPE('b))
1359    \<Longrightarrow> typ_name (typ_info_t TYPE('a)) \<noteq> pad_typ_name
1360    \<Longrightarrow> \<not> typ_uinfo_t TYPE('a :: c_type) \<le> typ_uinfo_t TYPE ('b :: c_type)"
1361  apply (clarsimp simp: typ_tag_le_def typ_uinfo_t_def)
1362  apply (drule td_set_td_names, simp+)
1363  done
1364
1365lemma in_set_list_map:
1366  "x \<in> set xs \<Longrightarrow> \<exists>n. [n \<mapsto> x] \<subseteq>\<^sub>m list_map xs"
1367  apply (clarsimp simp: in_set_conv_nth)
1368  apply (rule_tac x=i in exI)
1369  apply (simp add: map_le_def)
1370  done
1371
1372lemma h_t_valid_eq_array_valid:
1373  "h_t_valid htd gd (p :: (('a :: wf_type)['b :: finite]) ptr)
1374    = (gd p \<and> h_t_array_valid htd (ptr_coerce p :: 'a ptr) CARD('b))"
1375  by (auto simp: h_t_array_valid_def h_t_valid_def
1376                 typ_uinfo_array_tag_n_m_eq)
1377
1378lemma h_t_array_valid_ptr_retyps_gen:
1379  assumes sz2: "size_of TYPE('a :: mem_type) = sz"
1380  assumes bytes: "region_is_bytes' (ptr_val p) (n * sz) htd"
1381  shows "h_t_array_valid htd p' n'
1382    \<Longrightarrow> h_t_array_valid (ptr_retyps_gen n (p :: 'a :: mem_type ptr) arr htd) p' n'"
1383  apply (clarsimp simp: h_t_array_valid_def valid_footprint_def)
1384  apply (drule spec, drule(1) mp, clarsimp)
1385  apply (case_tac "ptr_val p' + of_nat y \<in> {ptr_val p ..+ n * size_of TYPE('a)}")
1386   apply (cut_tac s="uinfo_array_tag_n_m TYPE('b) n' n'" and n=y in ladder_set_self)
1387   apply (clarsimp dest!: in_set_list_map)
1388   apply (drule(1) map_le_trans)
1389   apply (simp add: map_le_def)
1390   apply (subst(asm) bytes[unfolded region_is_bytes'_def, rule_format, symmetric])
1391     apply (simp add: sz2)
1392    apply (simp add: uinfo_array_tag_n_m_def typ_uinfo_t_def typ_info_word)
1393   apply simp
1394  apply (simp add: ptr_retyps_gen_out)
1395  done
1396
1397lemma cvariable_array_ptr_retyps:
1398  assumes sz2: "size_of TYPE('a :: mem_type) = sz"
1399  assumes bytes: "region_is_bytes' (ptr_val p) (n * sz) htd"
1400  shows "cvariable_array_map_relation m ns ptrfun htd
1401    \<Longrightarrow> cvariable_array_map_relation m ns (ptrfun :: _ \<Rightarrow> ('b :: mem_type) ptr)
1402            (ptr_retyps_gen n (p :: 'a :: mem_type ptr) arr htd)"
1403  by (clarsimp simp: cvariable_array_map_relation_def
1404                     h_t_array_valid_ptr_retyps_gen[OF sz2 bytes])
1405
1406lemma cvariable_array_ptr_upd:
1407  assumes at: "h_t_array_valid htd (ptrfun x) (ns y)"
1408  shows "cvariable_array_map_relation m ns ptrfun htd
1409    \<Longrightarrow> cvariable_array_map_relation (m(x \<mapsto> y))
1410        ns (ptrfun :: _ \<Rightarrow> ('b :: mem_type) ptr) htd"
1411  by (clarsimp simp: cvariable_array_map_relation_def at
1412              split: if_split)
1413
1414lemma clift_eq_h_t_valid_eq:
1415  "clift hp = (clift hp' :: ('a :: c_type) ptr \<Rightarrow> _)
1416    \<Longrightarrow> (h_t_valid (hrs_htd hp) c_guard :: 'a ptr \<Rightarrow> _)
1417        = h_t_valid (hrs_htd hp') c_guard"
1418  by (rule ext, simp add: h_t_valid_clift_Some_iff)
1419
1420lemma region_is_bytes_typ_region_bytes:
1421  "{ptr ..+ len} \<le> {ptr' ..+ 2 ^ bits}
1422    \<Longrightarrow> region_is_bytes' ptr len (typ_region_bytes ptr' bits htd)"
1423  apply (clarsimp simp: region_is_bytes'_def typ_region_bytes_def hrs_htd_update)
1424  apply (simp add: subsetD split: if_split_asm)
1425  done
1426
1427lemma region_actually_is_bytes_retyp_disjoint:
1428  "{ptr ..+ sz} \<inter> {ptr_val (p :: 'a ptr)..+n * size_of TYPE('a :: mem_type)} = {}
1429    \<Longrightarrow> region_actually_is_bytes' ptr sz htd
1430    \<Longrightarrow> region_actually_is_bytes' ptr sz (ptr_retyps_gen n p arr htd)"
1431  apply (clarsimp simp: region_actually_is_bytes'_def del: impI)
1432  apply (subst ptr_retyps_gen_out)
1433   apply blast
1434  apply simp
1435  done
1436
1437lemma intvl_plus_unat_eq:
1438  "p \<le> p + x - 1 \<Longrightarrow> x \<noteq> 0
1439    \<Longrightarrow> {p ..+ unat x} = {p .. p + x - 1}"
1440  apply (subst upto_intvl_eq', simp_all add: unat_eq_0 field_simps)
1441  apply (rule order_less_imp_le, simp)
1442  done
1443
1444lemma zero_ranges_ptr_retyps:
1445  "zero_ranges_are_zero (gsUntypedZeroRanges s) hrs
1446    \<Longrightarrow> caps_overlap_reserved' {ptr_val (p :: 'a ptr) ..+ n * size_of TYPE ('a :: mem_type)} s
1447    \<Longrightarrow> untyped_ranges_zero' s
1448    \<Longrightarrow> valid_objs' s
1449    \<Longrightarrow> zero_ranges_are_zero (gsUntypedZeroRanges s)
1450       (hrs_htd_update (ptr_retyps_gen n p arr) hrs)"
1451  apply (clarsimp simp: zero_ranges_are_zero_def untyped_ranges_zero_inv_def
1452                        hrs_htd_update)
1453  apply (drule(1) bspec, clarsimp)
1454  apply (rule region_actually_is_bytes_retyp_disjoint, simp_all)
1455  apply (clarsimp simp: map_comp_Some_iff cteCaps_of_def
1456                 elim!: ranE)
1457  apply (frule(1) ctes_of_valid')
1458  apply (simp add: caps_overlap_reserved'_def,
1459      drule bspec, erule ranI)
1460  apply (frule(1) untypedZeroRange_to_usableCapRange)
1461  apply (clarsimp simp: isCap_simps untypedZeroRange_def
1462                        getFreeRef_def max_free_index_def
1463                 split: if_split_asm)
1464  apply (erule disjoint_subset[rotated])
1465  apply (subst intvl_plus_unat_eq)
1466    apply clarsimp
1467   apply clarsimp
1468   apply (clarsimp simp: word_unat.Rep_inject[symmetric]
1469                         valid_cap_simps' capAligned_def
1470                         unat_of_nat
1471               simp del: word_unat.Rep_inject)
1472  apply clarsimp
1473  done
1474
1475abbreviation
1476  "ret_zero ptr sz
1477    \<equiv> valid_objs' and untyped_ranges_zero' and caps_overlap_reserved' {ptr ..+ sz}"
1478
1479lemma createObjects_ccorres_ep:
1480  defines "ko \<equiv> (KOEndpoint (makeObject :: endpoint))"
1481  shows "\<forall>\<sigma> x. (\<sigma>, x) \<in> rf_sr
1482  \<and> ptr \<noteq> 0
1483  \<and> pspace_aligned' \<sigma> \<and> pspace_distinct' \<sigma>
1484  \<and> pspace_no_overlap' ptr sz \<sigma>
1485  \<and> ret_zero ptr (n * (2 ^ objBitsKO ko)) \<sigma>
1486  \<and> region_is_zero_bytes ptr (n * (2 ^ objBitsKO ko)) x
1487  \<and> range_cover ptr sz (objBitsKO ko) n
1488  \<and> {ptr ..+ n * (2 ^ objBitsKO ko)} \<inter> kernel_data_refs = {}
1489  \<longrightarrow>
1490  (\<sigma>\<lparr>ksPSpace := foldr (\<lambda>addr. data_map_insert addr ko) (new_cap_addrs n ptr ko) (ksPSpace \<sigma>)\<rparr>,
1491   x\<lparr>globals := globals x
1492                 \<lparr>t_hrs_' := hrs_htd_update (ptr_retyps_gen n (Ptr ptr :: endpoint_C ptr) False)
1493                   (t_hrs_' (globals x))\<rparr>\<rparr>) \<in> rf_sr"
1494  (is "\<forall>\<sigma> x. ?P \<sigma> x \<longrightarrow>
1495    (\<sigma>\<lparr>ksPSpace := ?ks \<sigma>\<rparr>, x\<lparr>globals := globals x\<lparr>t_hrs_' := ?ks' x\<rparr>\<rparr>) \<in> rf_sr")
1496proof (intro impI allI)
1497  fix \<sigma> x
1498  let ?thesis = "(\<sigma>\<lparr>ksPSpace := ?ks \<sigma>\<rparr>, x\<lparr>globals := globals x\<lparr>t_hrs_' := ?ks' x\<rparr>\<rparr>) \<in> rf_sr"
1499  let ?ks = "?ks \<sigma>"
1500  let ?ks' = "?ks' x"
1501  let ?ptr = "Ptr ptr :: endpoint_C ptr"
1502
1503  assume "?P \<sigma> x"
1504  hence rf: "(\<sigma>, x) \<in> rf_sr"
1505    and cover: "range_cover ptr sz (objBitsKO ko) n"
1506    and al: "is_aligned ptr (objBitsKO ko)" and ptr0: "ptr \<noteq> 0"
1507    and sz: "objBitsKO ko \<le> sz"
1508    and szb: "sz < word_bits"
1509    and pal: "pspace_aligned' \<sigma>" and pdst: "pspace_distinct' \<sigma>"
1510    and pno: "pspace_no_overlap' ptr sz \<sigma>"
1511    and rzo: "ret_zero ptr (n * (2 ^ objBitsKO ko)) \<sigma>"
1512    and empty: "region_is_bytes ptr (n * (2 ^ objBitsKO ko)) x"
1513    and zero: "heap_list_is_zero (hrs_mem (t_hrs_' (globals x))) ptr (n * (2 ^ objBitsKO ko))"
1514    and rc: "range_cover ptr sz (objBitsKO ko) n"
1515    and kdr: "{ptr..+n * (2 ^ objBitsKO ko)} \<inter> kernel_data_refs = {}"
1516    by (clarsimp simp:range_cover_def[where 'a=32, folded word_bits_def])+
1517
1518  (* obj specific *)
1519  have mko: "\<And>dev. makeObjectKO dev (Inr (APIObjectType ArchTypes_H.apiobject_type.EndpointObject)) = Some ko"
1520    by (simp add: ko_def makeObjectKO_def)
1521
1522  have relrl:
1523    "cendpoint_relation (cslift x) makeObject (from_bytes (replicate (size_of TYPE(endpoint_C)) 0))"
1524    unfolding cendpoint_relation_def
1525    apply (simp add: Let_def makeObject_endpoint size_of_def endpoint_lift_def)
1526    apply (simp add: from_bytes_def)
1527    apply (simp add: typ_info_simps endpoint_C_tag_def endpoint_lift_def
1528      size_td_lt_final_pad size_td_lt_ti_typ_pad_combine Let_def size_of_def)
1529    apply (simp add: final_pad_def Let_def size_td_lt_ti_typ_pad_combine Let_def
1530      size_of_def padup_def align_td_array' size_td_array update_ti_adjust_ti
1531      ti_typ_pad_combine_def Let_def ti_typ_combine_def empty_typ_info_def)
1532    apply (simp add: typ_info_array array_tag_def eval_nat_numeral)
1533    apply (simp add: array_tag_n_eq)
1534    apply (simp add: final_pad_def Let_def size_td_lt_ti_typ_pad_combine
1535      size_of_def padup_def align_td_array' size_td_array update_ti_adjust_ti
1536      ti_typ_pad_combine_def ti_typ_combine_def empty_typ_info_def)
1537    apply (simp add: EPState_Idle_def update_ti_t_word32_0s)
1538    done
1539
1540  (* /obj specific *)
1541
1542  (* s/obj/obj'/ *)
1543  have szo: "size_of TYPE(endpoint_C) = 2 ^ objBitsKO ko"
1544    by (simp add: size_of_def objBits_simps' ko_def)
1545  have szo': "n * (2 ^ objBitsKO ko) = n * size_of TYPE(endpoint_C)"
1546    by (metis szo)
1547
1548  note rl' = cslift_ptr_retyp_other_inst[OF empty cover[simplified] szo' szo]
1549
1550  note rl = projectKO_opt_retyp_other [OF rc pal pno ko_def]
1551  note cterl = retype_ctes_helper [OF pal pdst pno al sz szb mko rc, simplified]
1552  note ht_rl = clift_eq_h_t_valid_eq[OF rl', OF tag_disj_via_td_name, simplified]
1553    uinfo_array_tag_n_m_not_le_typ_name
1554
1555 have guard:
1556    "\<forall>b < n. c_guard (CTypesDefs.ptr_add ?ptr (of_nat b))"
1557    apply (rule retype_guard_helper [where m = 2, OF cover ptr0 szo])
1558    apply (simp add: ko_def objBits_simps')
1559    apply (simp add: align_of_def)
1560    done
1561
1562  from rf have "cpspace_relation (ksPSpace \<sigma>) (underlying_memory (ksMachineState \<sigma>)) (t_hrs_' (globals x))"
1563    unfolding rf_sr_def cstate_relation_def by (simp add: Let_def)
1564  hence "cpspace_relation ?ks  (underlying_memory (ksMachineState \<sigma>)) ?ks'"
1565    unfolding cpspace_relation_def
1566    apply -
1567    apply (clarsimp simp: rl' cterl tag_disj_via_td_name foldr_upd_app_if [folded data_map_insert_def]
1568      heap_to_user_data_def cte_C_size heap_to_device_data_def)
1569    apply (subst clift_ptr_retyps_gen_prev_memset_same[OF guard _ _ szo' _ zero],
1570      simp_all only: szo empty, simp_all)
1571     apply (rule range_cover.strong_times_32[OF cover refl])
1572    apply (simp add: ptr_add_to_new_cap_addrs [OF szo] ht_rl)
1573    apply (simp add: rl projectKO_opt_retyp_same projectKOs)
1574    apply (simp add: ko_def projectKO_opt_retyp_same projectKOs cong: if_cong)
1575    apply (erule cmap_relation_retype)
1576    apply (rule relrl[simplified szo ko_def])
1577    done
1578
1579  thus ?thesis using rf empty kdr rzo
1580  apply (simp add: rf_sr_def cstate_relation_def Let_def rl'
1581                   tag_disj_via_td_name)
1582  apply (simp add: carch_state_relation_def cmachine_state_relation_def)
1583  apply (simp add: rl' cterl tag_disj_via_td_name h_t_valid_clift_Some_iff)
1584  apply (clarsimp simp: hrs_htd_update ptr_retyps_htd_safe_neg szo
1585                        kernel_data_refs_domain_eq_rotate
1586                        ht_rl foldr_upd_app_if [folded data_map_insert_def]
1587                        rl projectKOs cvariable_array_ptr_retyps[OF szo]
1588                        zero_ranges_ptr_retyps
1589              simp del: endpoint_C_size)
1590  done
1591qed
1592
1593lemma createObjects_ccorres_ntfn:
1594  defines "ko \<equiv> (KONotification (makeObject :: Structures_H.notification))"
1595  shows "\<forall>\<sigma> x. (\<sigma>, x) \<in> rf_sr \<and> ptr \<noteq> 0
1596  \<and> pspace_aligned' \<sigma> \<and> pspace_distinct' \<sigma>
1597  \<and> pspace_no_overlap' ptr sz \<sigma>
1598  \<and> ret_zero ptr (n * (2 ^ objBitsKO ko)) \<sigma>
1599  \<and> region_is_zero_bytes ptr (n * 2 ^ objBitsKO ko) x
1600  \<and> range_cover ptr sz (objBitsKO ko) n
1601  \<and> {ptr ..+ n * (2 ^ objBitsKO ko)} \<inter> kernel_data_refs = {}
1602  \<longrightarrow>
1603  (\<sigma>\<lparr>ksPSpace := foldr (\<lambda>addr. data_map_insert addr ko) (new_cap_addrs n ptr ko) (ksPSpace \<sigma>)\<rparr>,
1604   x\<lparr>globals := globals x
1605                 \<lparr>t_hrs_' := hrs_htd_update (ptr_retyps_gen n (Ptr ptr :: notification_C ptr) False)
1606                      (t_hrs_' (globals x))\<rparr>\<rparr>) \<in> rf_sr"
1607  (is "\<forall>\<sigma> x. ?P \<sigma> x \<longrightarrow>
1608    (\<sigma>\<lparr>ksPSpace := ?ks \<sigma>\<rparr>, x\<lparr>globals := globals x\<lparr>t_hrs_' := ?ks' x\<rparr>\<rparr>) \<in> rf_sr")
1609
1610proof (intro impI allI)
1611  fix \<sigma> x
1612  let ?thesis = "(\<sigma>\<lparr>ksPSpace := ?ks \<sigma>\<rparr>, x\<lparr>globals := globals x\<lparr>t_hrs_' := ?ks' x\<rparr>\<rparr>) \<in> rf_sr"
1613  let ?ks = "?ks \<sigma>"
1614  let ?ks' = "?ks' x"
1615  let ?ptr = "Ptr ptr :: notification_C ptr"
1616
1617  assume "?P \<sigma> x"
1618  hence rf: "(\<sigma>, x) \<in> rf_sr"
1619    and cover: "range_cover ptr sz (objBitsKO ko) n"
1620    and al: "is_aligned ptr (objBitsKO ko)" and ptr0: "ptr \<noteq> 0"
1621    and sz: "objBitsKO ko \<le> sz"
1622    and szb: "sz < word_bits"
1623    and pal: "pspace_aligned' \<sigma>" and pdst: "pspace_distinct' \<sigma>"
1624    and pno: "pspace_no_overlap' ptr sz \<sigma>"
1625    and rzo: "ret_zero ptr (n * (2 ^ objBitsKO ko)) \<sigma>"
1626    and empty: "region_is_bytes ptr (n * (2 ^ objBitsKO ko)) x"
1627    and zero: "heap_list_is_zero (hrs_mem (t_hrs_' (globals x))) ptr (n * (2 ^ objBitsKO ko))"
1628    and rc: "range_cover ptr sz (objBitsKO ko) n"
1629    and kdr: "{ptr..+n * 2 ^ objBitsKO ko} \<inter> kernel_data_refs = {}"
1630    by (clarsimp simp:range_cover_def[where 'a=32, folded word_bits_def])+
1631
1632  (* obj specific *)
1633  have mko: "\<And> dev. makeObjectKO dev (Inr (APIObjectType ArchTypes_H.apiobject_type.NotificationObject)) = Some ko" by (simp add: ko_def makeObjectKO_def)
1634
1635  have relrl:
1636    "cnotification_relation (cslift x) makeObject (from_bytes (replicate (size_of TYPE(notification_C)) 0))"
1637    unfolding cnotification_relation_def
1638    apply (simp add: Let_def makeObject_notification size_of_def notification_lift_def)
1639    apply (simp add: from_bytes_def)
1640    apply (simp add: typ_info_simps notification_C_tag_def notification_lift_def
1641      size_td_lt_final_pad size_td_lt_ti_typ_pad_combine Let_def size_of_def)
1642    apply (simp add: final_pad_def Let_def size_td_lt_ti_typ_pad_combine Let_def
1643      size_of_def padup_def align_td_array' size_td_array update_ti_adjust_ti
1644      ti_typ_pad_combine_def Let_def ti_typ_combine_def empty_typ_info_def)
1645    apply (simp add: typ_info_array array_tag_def eval_nat_numeral)
1646    apply (simp add: array_tag_n.simps)
1647    apply (simp add: final_pad_def Let_def size_td_lt_ti_typ_pad_combine Let_def
1648      size_of_def padup_def align_td_array' size_td_array update_ti_adjust_ti
1649      ti_typ_pad_combine_def Let_def ti_typ_combine_def empty_typ_info_def)
1650    apply (simp add: update_ti_t_word32_0s NtfnState_Idle_def option_to_ctcb_ptr_def)
1651    done
1652
1653  (* /obj specific *)
1654
1655  (* s/obj/obj'/ *)
1656  have szo: "size_of TYPE(notification_C) = 2 ^ objBitsKO ko"
1657    by (simp add: size_of_def objBits_simps' ko_def)
1658  have szo': "n * (2 ^ objBitsKO ko) = n * size_of TYPE(notification_C)" using sz
1659    apply (subst szo)
1660    apply (simp add: power_add [symmetric])
1661    done
1662
1663  note rl' = cslift_ptr_retyp_other_inst[OF empty cover[simplified] szo' szo]
1664
1665  (* rest is generic *)
1666  note rl = projectKO_opt_retyp_other [OF rc pal pno ko_def]
1667  note cterl = retype_ctes_helper [OF pal pdst pno al sz szb mko rc, simplified]
1668  note ht_rl = clift_eq_h_t_valid_eq[OF rl', OF tag_disj_via_td_name, simplified]
1669    uinfo_array_tag_n_m_not_le_typ_name
1670
1671  have guard:
1672    "\<forall>b<n. c_guard (CTypesDefs.ptr_add ?ptr (of_nat b))"
1673    apply (rule retype_guard_helper[where m=2, OF cover ptr0 szo])
1674    apply (simp add: ko_def objBits_simps' align_of_def)+
1675    done
1676
1677  from rf have "cpspace_relation (ksPSpace \<sigma>) (underlying_memory (ksMachineState \<sigma>)) (t_hrs_' (globals x))"
1678    unfolding rf_sr_def cstate_relation_def by (simp add: Let_def)
1679  hence "cpspace_relation ?ks  (underlying_memory (ksMachineState \<sigma>)) ?ks'"
1680    unfolding cpspace_relation_def
1681    apply -
1682    apply (clarsimp simp: rl' cterl tag_disj_via_td_name foldr_upd_app_if [folded data_map_insert_def]
1683      heap_to_user_data_def cte_C_size)
1684    apply (subst clift_ptr_retyps_gen_prev_memset_same[OF guard _ _ szo' _ zero],
1685      simp_all only: szo empty, simp_all)
1686     apply (rule range_cover.strong_times_32[OF cover refl])
1687    apply (simp add: ptr_add_to_new_cap_addrs [OF szo] ht_rl)
1688    apply (simp add: rl projectKO_opt_retyp_same projectKOs)
1689    apply (simp add: ko_def projectKO_opt_retyp_same projectKOs cong: if_cong)
1690    apply (erule cmap_relation_retype)
1691    apply (rule relrl[simplified szo ko_def])
1692    done
1693
1694  thus ?thesis using rf empty kdr rzo
1695    apply (simp add: rf_sr_def cstate_relation_def Let_def rl' tag_disj_via_td_name)
1696    apply (simp add: carch_state_relation_def cmachine_state_relation_def)
1697    apply (simp add: rl' cterl tag_disj_via_td_name h_t_valid_clift_Some_iff )
1698    apply (clarsimp simp: hrs_htd_update ptr_retyps_htd_safe_neg szo
1699                          kernel_data_refs_domain_eq_rotate
1700                          ht_rl foldr_upd_app_if [folded data_map_insert_def]
1701                          rl projectKOs cvariable_array_ptr_retyps[OF szo]
1702                          zero_ranges_ptr_retyps
1703                simp del: notification_C_size)
1704    done
1705qed
1706
1707
1708lemma ccte_relation_makeObject:
1709  notes option.case_cong_weak [cong]
1710  shows "ccte_relation makeObject (from_bytes (replicate (size_of TYPE(cte_C)) 0))"
1711  apply (simp add: Let_def makeObject_cte size_of_def ccte_relation_def map_option_Some_eq2)
1712  apply (simp add: from_bytes_def)
1713  apply (simp add: typ_info_simps cte_C_tag_def  cte_lift_def
1714    size_td_lt_final_pad size_td_lt_ti_typ_pad_combine Let_def size_of_def)
1715  apply (simp add: final_pad_def Let_def size_td_lt_ti_typ_pad_combine
1716    size_of_def padup_def align_td_array' size_td_array update_ti_adjust_ti
1717    ti_typ_pad_combine_def ti_typ_combine_def empty_typ_info_def align_of_def
1718    typ_info_simps cap_C_tag_def mdb_node_C_tag_def split: option.splits)
1719  apply (simp add: typ_info_array array_tag_def eval_nat_numeral array_tag_n.simps)
1720  apply (simp add: final_pad_def Let_def size_td_lt_ti_typ_pad_combine
1721    size_of_def padup_def align_td_array' size_td_array update_ti_adjust_ti
1722    ti_typ_pad_combine_def ti_typ_combine_def empty_typ_info_def update_ti_t_word32_0s)
1723  apply (simp add: cap_lift_def Let_def cap_get_tag_def cap_tag_defs cte_to_H_def cap_to_H_def mdb_node_to_H_def
1724    mdb_node_lift_def nullMDBNode_def c_valid_cte_def)
1725  done
1726
1727lemma ccte_relation_nullCap:
1728  notes option.case_cong_weak [cong]
1729  shows "ccte_relation (CTE NullCap (MDB 0 0 False False)) (from_bytes (replicate (size_of TYPE(cte_C)) 0))"
1730  apply (simp add: Let_def makeObject_cte size_of_def ccte_relation_def map_option_Some_eq2)
1731  apply (simp add: from_bytes_def)
1732  apply (simp add: typ_info_simps cte_C_tag_def  cte_lift_def
1733    size_td_lt_final_pad size_td_lt_ti_typ_pad_combine Let_def size_of_def)
1734  apply (simp add: final_pad_def Let_def size_td_lt_ti_typ_pad_combine
1735    size_of_def padup_def align_td_array' size_td_array update_ti_adjust_ti
1736    ti_typ_pad_combine_def ti_typ_combine_def empty_typ_info_def align_of_def
1737    typ_info_simps cap_C_tag_def mdb_node_C_tag_def split: option.splits)
1738  apply (simp add: typ_info_array array_tag_def eval_nat_numeral array_tag_n.simps)
1739  apply (simp add: final_pad_def Let_def size_td_lt_ti_typ_pad_combine
1740    size_of_def padup_def align_td_array' size_td_array update_ti_adjust_ti
1741    ti_typ_pad_combine_def ti_typ_combine_def empty_typ_info_def update_ti_t_word32_0s)
1742  apply (simp add: cap_lift_def Let_def cap_get_tag_def cap_tag_defs cte_to_H_def cap_to_H_def mdb_node_to_H_def
1743    mdb_node_lift_def nullMDBNode_def c_valid_cte_def)
1744  done
1745
1746lemma createObjects_ccorres_cte:
1747  defines "ko \<equiv> (KOCTE (makeObject :: cte))"
1748  shows "\<forall>\<sigma> x. (\<sigma>, x) \<in> rf_sr  \<and> ptr \<noteq> 0
1749  \<and> pspace_aligned' \<sigma> \<and> pspace_distinct' \<sigma>
1750  \<and> pspace_no_overlap' ptr sz \<sigma>
1751  \<and> ret_zero ptr (n * 2 ^ objBitsKO ko) \<sigma>
1752  \<and> region_is_zero_bytes ptr (n * 2 ^ objBitsKO ko) x
1753  \<and> range_cover ptr sz (objBitsKO ko) n
1754  \<and> {ptr ..+ n * (2 ^ objBitsKO ko)} \<inter> kernel_data_refs = {}
1755   \<longrightarrow>
1756  (\<sigma>\<lparr>ksPSpace := foldr (\<lambda>addr. data_map_insert addr ko) (new_cap_addrs n ptr ko) (ksPSpace \<sigma>)\<rparr>,
1757   x\<lparr>globals := globals x
1758                 \<lparr>t_hrs_' := hrs_htd_update (ptr_retyps_gen n (Ptr ptr :: cte_C ptr) True)
1759                       (t_hrs_' (globals x))\<rparr>\<rparr>) \<in> rf_sr"
1760  (is "\<forall>\<sigma> x. ?P \<sigma> x \<longrightarrow>
1761    (\<sigma>\<lparr>ksPSpace := ?ks \<sigma>\<rparr>, x\<lparr>globals := globals x\<lparr>t_hrs_' := ?ks' x\<rparr>\<rparr>) \<in> rf_sr")
1762proof (intro impI allI)
1763  fix \<sigma> x
1764  let ?thesis = "(\<sigma>\<lparr>ksPSpace := ?ks \<sigma>\<rparr>, x\<lparr>globals := globals x\<lparr>t_hrs_' := ?ks' x\<rparr>\<rparr>) \<in> rf_sr"
1765  let ?ks = "?ks \<sigma>"
1766  let ?ks' = "?ks' x"
1767  let ?ptr = "Ptr ptr :: cte_C ptr"
1768
1769  assume "?P \<sigma> x"
1770  hence rf: "(\<sigma>, x) \<in> rf_sr"
1771    and cover: "range_cover ptr sz (objBitsKO ko) n"
1772    and al: "is_aligned ptr (objBitsKO ko)" and ptr0: "ptr \<noteq> 0"
1773    and sz: "objBitsKO ko \<le> sz"
1774    and szb: "sz < word_bits"
1775    and pal: "pspace_aligned' \<sigma>" and pdst: "pspace_distinct' \<sigma>"
1776    and pno: "pspace_no_overlap' ptr sz \<sigma>"
1777    and rzo: "ret_zero ptr (n * 2 ^ objBitsKO ko) \<sigma>"
1778    and empty: "region_is_bytes ptr (n * (2 ^ objBitsKO ko)) x"
1779    and zero: "heap_list_is_zero (hrs_mem (t_hrs_' (globals x))) ptr (n * (2 ^ objBitsKO ko))"
1780    and rc: "range_cover ptr sz (objBitsKO ko) n"
1781    and kdr: "{ptr..+n * 2 ^ objBitsKO ko} \<inter> kernel_data_refs = {}"
1782    by (clarsimp simp:range_cover_def[where 'a=32, folded word_bits_def])+
1783
1784  (* obj specific *)
1785  have mko: "\<And>dev. makeObjectKO dev (Inr (APIObjectType  ArchTypes_H.apiobject_type.CapTableObject)) = Some ko"
1786    by (simp add: ko_def makeObjectKO_def)
1787
1788  note relrl = ccte_relation_makeObject
1789
1790  (* /obj specific *)
1791
1792  (* s/obj/obj'/ *)
1793  have szo: "size_of TYPE(cte_C) = 2 ^ objBitsKO ko"
1794    by (simp add: size_of_def objBits_simps' ko_def)
1795  have szo': "n * 2 ^ objBitsKO ko = n * size_of TYPE(cte_C)" using sz
1796    apply (subst szo)
1797    apply (simp add: power_add [symmetric])
1798    done
1799
1800  note rl' = cslift_ptr_retyp_other_inst[OF empty cover szo' szo]
1801
1802  (* rest is generic *)
1803  note rl = projectKO_opt_retyp_other [OF rc pal pno ko_def]
1804  note cterl = retype_ctes_helper [OF pal pdst pno al sz szb mko rc, simplified]
1805  note ht_rl = clift_eq_h_t_valid_eq[OF rl', OF tag_disj_via_td_name, simplified]
1806    uinfo_array_tag_n_m_not_le_typ_name
1807
1808  have guard:
1809    "\<forall>b< n. c_guard (CTypesDefs.ptr_add ?ptr (of_nat b))"
1810    apply (rule retype_guard_helper[where m=2, OF cover ptr0 szo])
1811    apply (simp add: ko_def objBits_simps' align_of_def)+
1812    done
1813
1814  note irq = h_t_valid_eq_array_valid[where 'a=cte_C]
1815    h_t_array_valid_ptr_retyps_gen[where p="Ptr ptr", simplified, OF szo empty]
1816
1817  with rf have irq: "h_t_valid (hrs_htd ?ks') c_guard
1818      (ptr_coerce (intStateIRQNode_' (globals x)) :: (cte_C[256]) ptr)"
1819    apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def)
1820    apply (simp add: hrs_htd_update h_t_valid_eq_array_valid)
1821    apply (simp add: h_t_array_valid_ptr_retyps_gen[OF szo] empty)
1822    done
1823
1824  note if_cong[cong] (* needed by some of the [simplified]'s below. *)
1825  from rf have "cpspace_relation (ksPSpace \<sigma>) (underlying_memory (ksMachineState \<sigma>)) (t_hrs_' (globals x))"
1826    unfolding rf_sr_def cstate_relation_def by (simp add: Let_def)
1827  hence "cpspace_relation ?ks (underlying_memory (ksMachineState \<sigma>)) ?ks'"
1828    unfolding cpspace_relation_def
1829    apply -
1830    apply (clarsimp simp: rl' cterl tag_disj_via_td_name foldr_upd_app_if [folded data_map_insert_def])
1831    apply (subst clift_ptr_retyps_gen_prev_memset_same[OF guard _ _ szo' _ zero],
1832      simp_all only: szo empty, simp_all)
1833     apply (rule range_cover.strong_times_32[OF cover refl])
1834    apply (simp add: ptr_add_to_new_cap_addrs [OF szo] ht_rl)
1835    apply (simp add: rl projectKO_opt_retyp_same projectKOs)
1836    apply (simp add: ko_def projectKO_opt_retyp_same projectKOs cong: if_cong)
1837    apply (subst makeObject_cte[symmetric])
1838    apply (erule cmap_relation_retype)
1839    apply (rule relrl[simplified szo ko_def])
1840    done
1841
1842  thus ?thesis using rf empty kdr irq rzo
1843    apply (simp add: rf_sr_def cstate_relation_def Let_def rl' tag_disj_via_td_name)
1844    apply (simp add: carch_state_relation_def cmachine_state_relation_def)
1845    apply (simp add: rl' cterl tag_disj_via_td_name h_t_valid_clift_Some_iff)
1846    apply (clarsimp simp: hrs_htd_update ptr_retyps_htd_safe_neg szo
1847                          kernel_data_refs_domain_eq_rotate
1848                          rl foldr_upd_app_if [folded data_map_insert_def] projectKOs
1849                          zero_ranges_ptr_retyps
1850                          ht_rl cvariable_array_ptr_retyps[OF szo])
1851    done
1852qed
1853
1854lemma h_t_valid_ptr_retyps_gen_disjoint:
1855  "\<lbrakk> d \<Turnstile>\<^sub>t p; {ptr_val p..+ size_of TYPE('b)} \<inter> {ptr_val ptr..+n * size_of TYPE('a)} = {} \<rbrakk> \<Longrightarrow>
1856  ptr_retyps_gen n (ptr::'a::mem_type ptr) arr d \<Turnstile>\<^sub>t (p::'b::mem_type ptr)"
1857  apply (clarsimp simp: h_t_valid_def valid_footprint_def Let_def)
1858  apply (drule spec, drule (1) mp)
1859  apply (subgoal_tac "ptr_val p + of_nat y \<notin> {ptr_val ptr..+n * size_of TYPE('a)}")
1860   apply (simp add: ptr_retyps_gen_out)
1861  apply clarsimp
1862  apply (drule intvlD)
1863  apply (clarsimp simp: disjoint_iff_not_equal )
1864  apply (drule_tac x = "ptr_val p + of_nat y" in bspec)
1865   apply (rule intvlI)
1866   apply (simp add: size_of_def)
1867  apply (drule_tac x = "ptr_val ptr + of_nat k" in bspec)
1868   apply (erule intvlI)
1869  apply simp
1870  done
1871
1872
1873lemma range_cover_intvl:
1874assumes cover: "range_cover (ptr :: 'a :: len word) sz us n"
1875assumes not0 : "n \<noteq> 0"
1876shows "{ptr..+n * 2 ^ us} = {ptr..ptr + (of_nat n * 2 ^ us - 1)}"
1877  proof
1878    have not0' : "(0 :: 'a word) < of_nat n * (2 :: 'a word) ^ us"
1879      using range_cover_not_zero_shift[OF _ cover,where gbits = "us"]
1880     apply (simp add:not0 shiftl_t2n field_simps)
1881     apply unat_arith
1882     done
1883
1884    show "{ptr..+n * 2 ^ us} \<subseteq> {ptr..ptr + (of_nat n* 2 ^ us - 1)}"
1885     using not0 not0'
1886     apply (clarsimp simp:intvl_def)
1887     apply (intro conjI)
1888      apply (rule word_plus_mono_right2[rotated,where b = "of_nat n * 2^us - 1"])
1889       apply (subst le_m1_iff_lt[THEN iffD1])
1890        apply (simp add:not0')
1891       apply (rule word_of_nat_less)
1892       apply (clarsimp simp: range_cover.unat_of_nat_shift[OF cover] field_simps)
1893      apply (clarsimp simp: field_simps)
1894      apply (erule range_cover_bound[OF cover])
1895     apply (rule word_plus_mono_right)
1896      apply (subst le_m1_iff_lt[THEN iffD1])
1897       apply (simp add:not0')
1898      apply (rule word_of_nat_less)
1899      apply (clarsimp simp: range_cover.unat_of_nat_shift[OF cover] field_simps)
1900     apply (clarsimp simp: field_simps)
1901      apply (erule range_cover_bound[OF cover])
1902     done
1903   show "{ptr..ptr + (of_nat n * 2 ^ us - 1)} \<subseteq> {ptr..+n * 2 ^ us}"
1904     using not0 not0'
1905     apply (clarsimp simp:intvl_def)
1906     apply (rule_tac x = "unat (x - ptr)" in exI)
1907      apply simp
1908      apply (simp add:field_simps)
1909      apply (rule unat_less_helper)
1910      apply (subst le_m1_iff_lt[THEN iffD1,symmetric])
1911      apply (simp add:field_simps not0 range_cover_not_zero_shift[unfolded shiftl_t2n,OF _ _ le_refl])
1912     apply (rule word_diff_ls')
1913      apply (simp add:field_simps)
1914     apply simp
1915    done
1916  qed
1917
1918lemma aligned_new_cap_addrs_eq_base:
1919  "is_aligned p bits \<Longrightarrow> is_aligned ptr bits
1920    \<Longrightarrow> n = 2 ^ (bits - objBitsKO ko)
1921    \<Longrightarrow> objBitsKO ko = shft
1922    \<Longrightarrow> y < of_nat n
1923    \<Longrightarrow> (p + (y << shft) \<in> set (new_cap_addrs n ptr ko)) = (p = ptr)"
1924  apply (erule is_aligned_get_word_bits)
1925   apply (rule iffI)
1926    apply (clarsimp simp: new_cap_addrs_def)
1927    apply (rule ccontr, drule(2) aligned_neq_into_no_overlap)
1928    apply (simp only: field_simps upto_intvl_eq[symmetric])
1929    apply (drule equals0D, erule notE, rule_tac c="p + (y << shft)" in IntI)
1930     apply (simp(no_asm) add: offs_in_intvl_iff)
1931     apply (rule unat_less_helper, simp, rule shiftl_less_t2n; simp)
1932    apply (simp add: offs_in_intvl_iff)
1933    apply (rule unat_less_helper, simp, rule shiftl_less_t2n; simp add: word_of_nat_less)
1934   apply (simp add: new_cap_addrs_def)
1935   apply (rule_tac x="unat y" in image_eqI; simp add: unat_less_helper)
1936  apply (erule is_aligned_get_word_bits; simp)
1937  apply (simp add: new_cap_addrs_def)
1938  apply (rule_tac x="unat y" in image_eqI; simp add: unat_less_helper)
1939  done
1940
1941lemma cmap_relation_array_add_array[OF refl]:
1942  "ptrf = Ptr \<Longrightarrow> carray_map_relation n ahp chp ptrf
1943    \<Longrightarrow> is_aligned p n
1944    \<Longrightarrow> ahp' = (\<lambda>x. if x \<in> set (new_cap_addrs sz p ko) then Some v else ahp x)
1945    \<Longrightarrow> (\<forall>x. chp x \<longrightarrow> is_aligned (ptr_val x) n \<Longrightarrow> \<forall>y. chp' y = (y = ptrf p | chp y))
1946    \<Longrightarrow> sz = 2 ^ (n - objBits v)
1947    \<Longrightarrow> objBitsKO ko = objBitsKO (injectKOS v)
1948    \<Longrightarrow> objBits v \<le> n \<Longrightarrow> n < word_bits
1949    \<Longrightarrow> carray_map_relation n ahp' chp' ptrf"
1950  apply (clarsimp simp: carray_map_relation_def objBits_koTypeOf
1951                        objBitsT_koTypeOf[symmetric]
1952                        koTypeOf_injectKO
1953              simp del: objBitsT_koTypeOf)
1954  apply (drule meta_mp)
1955   apply auto[1]
1956  apply (case_tac "pa = p"; clarsimp)
1957   apply (subst if_P; simp add: new_cap_addrs_def)
1958   apply (rule_tac x="unat ((p' && mask n) >> objBitsKO ko)" in image_eqI)
1959    apply (simp add: shiftr_shiftl1 is_aligned_andI1 add.commute
1960                     word_plus_and_or_coroll2)
1961   apply (simp, rule unat_less_helper, simp, rule shiftr_less_t2n)
1962   apply (simp add: and_mask_less_size word_size word_bits_def)
1963  apply (case_tac "chp (ptrf pa)", simp_all)
1964   apply (drule spec, drule(1) iffD2)
1965   apply (auto split: if_split)[1]
1966  apply (drule_tac x=pa in spec, clarsimp)
1967  apply (drule_tac x=p' in spec, clarsimp split: if_split_asm)
1968  apply (clarsimp simp: new_cap_addrs_def)
1969  apply (subst(asm) is_aligned_add_helper, simp_all)
1970  apply (rule shiftl_less_t2n, rule word_of_nat_less, simp_all add: word_bits_def)
1971  done
1972
1973lemma createObjects_ccorres_pte:
1974  defines "ko \<equiv> (KOArch (KOPTE (makeObject :: pte)))"
1975  shows "\<forall>\<sigma> x. (\<sigma>, x) \<in> rf_sr \<and> ptr \<noteq> 0
1976  \<and> pspace_aligned' \<sigma> \<and> pspace_distinct' \<sigma>
1977  \<and> pspace_no_overlap' ptr sz \<sigma>
1978  \<and> ret_zero ptr (2 ^ ptBits) \<sigma>
1979  \<and> region_is_zero_bytes ptr (2 ^ ptBits) x
1980  \<and> range_cover ptr sz ptBits 1
1981  \<and> valid_global_refs' s
1982  \<and> kernel_data_refs \<inter> {ptr..+ 2 ^ ptBits} = {} \<longrightarrow>
1983  (\<sigma>\<lparr>ksPSpace := foldr (\<lambda>addr. data_map_insert addr ko) (new_cap_addrs 512 ptr ko) (ksPSpace \<sigma>)\<rparr>,
1984   x\<lparr>globals := globals x
1985                 \<lparr>t_hrs_' := hrs_htd_update (ptr_retyps_gen 1 (pt_Ptr ptr) False)
1986                       (t_hrs_' (globals x))\<rparr>\<rparr>) \<in> rf_sr"
1987  (is "\<forall>\<sigma> x. ?P \<sigma> x \<longrightarrow>
1988    (\<sigma>\<lparr>ksPSpace := ?ks \<sigma>\<rparr>, x\<lparr>globals := globals x\<lparr>t_hrs_' := ?ks' x\<rparr>\<rparr>) \<in> rf_sr")
1989proof (intro impI allI)
1990  fix \<sigma> x
1991  let ?thesis = "(\<sigma>\<lparr>ksPSpace := ?ks \<sigma>\<rparr>, x\<lparr>globals := globals x\<lparr>t_hrs_' := ?ks' x\<rparr>\<rparr>) \<in> rf_sr"
1992  let ?ks = "?ks \<sigma>"
1993  let ?ks' = "?ks' x"
1994  let ?ptr = "Ptr ptr :: (pte_C[512]) ptr"
1995  assume "?P \<sigma> x"
1996  hence rf: "(\<sigma>, x) \<in> rf_sr"
1997    and cover: "range_cover ptr sz ptBits 1"
1998    and al: "is_aligned ptr ptBits"
1999    and ptr0: "ptr \<noteq> 0"
2000    and sz: "ptBits \<le> sz"
2001    and szb: "sz < word_bits"
2002    and pal: "pspace_aligned' \<sigma>"
2003    and pdst: "pspace_distinct' \<sigma>"
2004    and pno: "pspace_no_overlap' ptr sz \<sigma>"
2005    and rzo: "ret_zero ptr (2 ^ ptBits) \<sigma>"
2006    and empty: "region_is_bytes ptr (2 ^ ptBits) x"
2007    and zero: "heap_list_is_zero (hrs_mem (t_hrs_' (globals x))) ptr (2 ^ ptBits)"
2008    and kernel_data_refs_disj : "kernel_data_refs \<inter> {ptr..+ 2 ^ ptBits} = {}"
2009    by (clarsimp simp:range_cover_def[where 'a=32, folded word_bits_def])+
2010
2011    note blah[simp del] =  atLeastAtMost_iff atLeastatMost_subset_iff atLeastLessThan_iff
2012          Int_atLeastAtMost atLeastatMost_empty_iff split_paired_Ex
2013
2014  (* obj specific *)
2015  have mko: "\<And>dev. makeObjectKO dev (Inr ARM_HYP_H.PageTableObject) = Some ko" by (simp add: ko_def makeObjectKO_def)
2016
2017  have relrl:
2018    "cpte_relation makeObject (from_bytes (replicate (size_of TYPE(pte_C)) 0))"
2019    unfolding cpte_relation_def
2020    apply (simp add: Let_def makeObject_pte size_of_def pte_lift_def)
2021    apply (simp add: from_bytes_def)
2022    apply (simp add: typ_info_simps pte_C_tag_def pte_lift_def pte_get_tag_def
2023      size_td_lt_final_pad size_td_lt_ti_typ_pad_combine Let_def size_of_def)
2024    apply (simp add: final_pad_def Let_def size_td_lt_ti_typ_pad_combine Let_def
2025      size_of_def padup_def align_td_array' size_td_array update_ti_adjust_ti
2026      ti_typ_pad_combine_def Let_def ti_typ_combine_def empty_typ_info_def)
2027    apply (simp add: typ_info_array array_tag_def eval_nat_numeral)
2028    apply (simp add: array_tag_n.simps)
2029    apply (simp add: final_pad_def Let_def size_td_lt_ti_typ_pad_combine Let_def
2030      size_of_def padup_def align_td_array' size_td_array update_ti_adjust_ti
2031      ti_typ_pad_combine_def Let_def ti_typ_combine_def empty_typ_info_def)
2032    apply (simp add: update_ti_t_word32_0s pte_tag_defs)
2033    done
2034
2035  (* /obj specific *)
2036
2037  (* s/obj/obj'/ *)
2038  have szo: "size_of TYPE(pte_C[512]) = 2 ^ ptBits"
2039    by (simp add: size_of_def size_td_array table_bits_defs)
2040  have szo2: "512 * size_of TYPE(pte_C) = 2 ^ ptBits"
2041    by (simp add: szo[symmetric])
2042  have szo': "size_of TYPE(pte_C) = 2 ^ objBitsKO ko"
2043    by (simp add: objBits_simps ko_def archObjSize_def table_bits_defs)
2044
2045  note rl' = cslift_ptr_retyp_other_inst[where n=1,
2046    simplified, OF empty cover[simplified] szo[symmetric] szo]
2047
2048  have sz_weaken: "objBitsKO ko \<le> ptBits"
2049    by (simp add: objBits_simps ko_def archObjSize_def table_bits_defs)
2050  have cover': "range_cover ptr sz (objBitsKO ko) 512"
2051    apply (rule range_cover_rel[OF cover sz_weaken])
2052    apply (simp add: ptBits_def objBits_simps ko_def archObjSize_def table_bits_defs)
2053    done
2054  from sz sz_weaken have sz': "objBitsKO ko \<le> sz" by simp
2055  note al' = is_aligned_weaken[OF al sz_weaken]
2056
2057  have koT: "koTypeOf ko = ArchT PTET"
2058    by (simp add: ko_def)
2059
2060  (* rest used to be generic, but PT arrays are complicating everything *)
2061
2062  note rl = projectKO_opt_retyp_other [OF cover' pal pno ko_def]
2063  note cterl = retype_ctes_helper [OF pal pdst pno al' sz' szb mko cover']
2064
2065  have guard: "c_guard ?ptr"
2066    using al[simplified table_bits_defs]
2067    apply -
2068    apply (rule is_aligned_c_guard[where n=ptBits and m=2])
2069    apply (simp_all add:  align_td_array align_of_def table_bits_defs ptr0)
2070    done
2071
2072  have guard': "\<forall>n < 512. c_guard (pte_Ptr ptr +\<^sub>p int n)"
2073    using al[simplified table_bits_defs]
2074    apply -
2075    apply (rule retype_guard_helper [OF cover' ptr0 szo', where m=2])
2076     apply (simp_all add: objBits_simps ko_def archObjSize_def align_of_def table_bits_defs)
2077    done
2078
2079  note ptr_retyps.simps[simp del]
2080
2081  from rf have pterl: "cmap_relation (map_to_ptes (ksPSpace \<sigma>)) (cslift x) Ptr cpte_relation"
2082    unfolding rf_sr_def cstate_relation_def by (simp add: Let_def cpspace_relation_def)
2083
2084  note ht_rl = clift_eq_h_t_valid_eq[OF rl', OF tag_disj_via_td_name, simplified]
2085    uinfo_array_tag_n_m_not_le_typ_name
2086
2087  have pte_arr: "cpspace_pte_array_relation (ksPSpace \<sigma>) (t_hrs_' (globals x))
2088    \<Longrightarrow> cpspace_pte_array_relation ?ks ?ks'"
2089   apply (erule cmap_relation_array_add_array[OF _ al])
2090        apply (simp add: foldr_upd_app_if[folded data_map_insert_def])
2091        apply (rule projectKO_opt_retyp_same, simp add: ko_def projectKOs)
2092       apply (simp add: h_t_valid_clift_Some_iff dom_def split: if_split)
2093       apply (subst clift_ptr_retyps_gen_prev_memset_same[where n=1, simplified, OF guard],
2094         simp_all only: szo refl empty, simp_all add: zero)[1]
2095        apply (simp add: table_bits_defs word_bits_def)
2096       apply (auto split: if_split)[1]
2097      apply (simp_all add: objBits_simps archObjSize_def table_bits_defs
2098                           ko_def word_bits_def)
2099   done
2100
2101  from rf have "cpspace_relation (ksPSpace \<sigma>) (underlying_memory (ksMachineState \<sigma>)) (t_hrs_' (globals x))"
2102    unfolding rf_sr_def cstate_relation_def by (simp add: Let_def)
2103  hence "cpspace_relation ?ks (underlying_memory (ksMachineState \<sigma>))  ?ks'"
2104    unfolding cpspace_relation_def
2105  using pte_arr
2106  apply (clarsimp simp: rl' cterl cte_C_size tag_disj_via_td_name
2107                        foldr_upd_app_if [folded data_map_insert_def])
2108  apply (simp add: ht_rl)
2109  apply (simp add: ptr_retyp_to_array[simplified])
2110  apply (subst clift_ptr_retyps_gen_prev_memset_same[OF guard'], simp_all only: szo2 empty)
2111     apply simp
2112    apply (simp(no_asm) add: table_bits_defs word_bits_def)
2113   apply (simp add: zero)
2114  apply (simp add: rl projectKOs del: pte_C_size)
2115  apply (simp add: rl projectKO_opt_retyp_same ko_def projectKOs Let_def
2116                   ptr_add_to_new_cap_addrs [OF szo']
2117              cong: if_cong del: pte_C_size)
2118  apply (erule cmap_relation_retype)
2119  apply (insert relrl, auto)
2120  done
2121
2122  moreover
2123  from rf szb al
2124  have "ptr_span (pd_Ptr (symbol_table ''armUSGlobalPD'')) \<inter> {ptr ..+ 2 ^ ptBits} = {}"
2125    apply (clarsimp simp: valid_global_refs'_def  Let_def
2126                          valid_refs'_def ran_def rf_sr_def cstate_relation_def)
2127    apply (erule disjoint_subset)
2128    apply (simp add:kernel_data_refs_disj)
2129    done
2130
2131  ultimately
2132  show ?thesis using rf empty kernel_data_refs_disj rzo
2133    apply (simp add: rf_sr_def cstate_relation_def Let_def rl' tag_disj_via_td_name)
2134    apply (simp add: carch_state_relation_def cmachine_state_relation_def)
2135    apply (clarsimp simp add: rl' cterl tag_disj_via_td_name
2136      hrs_htd_update ht_rl foldr_upd_app_if [folded data_map_insert_def] rl projectKOs
2137      cvariable_array_ptr_retyps[OF szo]
2138      zero_ranges_ptr_retyps[where p="pt_Ptr ptr", simplified szo])
2139    apply (subst h_t_valid_ptr_retyps_gen_disjoint, assumption)
2140     apply (simp add:szo cte_C_size cte_level_bits_def)
2141     apply (erule disjoint_subset)
2142     apply (simp add: table_bits_defs del: replicate_numeral)
2143    apply (subst h_t_valid_ptr_retyps_gen_disjoint, assumption)
2144     apply (simp add:szo cte_C_size cte_level_bits_def)
2145     apply (erule disjoint_subset)
2146     apply (simp add: table_bits_defs del: replicate_numeral)
2147    by (simp add:szo ptr_retyps_htd_safe_neg hrs_htd_def
2148      kernel_data_refs_domain_eq_rotate table_bits_defs
2149      Int_ac del: replicate_numeral)
2150qed
2151
2152lemma createObjects_ccorres_pde:
2153  defines "ko \<equiv> (KOArch (KOPDE (makeObject :: pde)))"
2154  shows "\<forall>\<sigma> x. (\<sigma>, x) \<in> rf_sr \<and> ptr \<noteq> 0
2155  \<and> pspace_aligned' \<sigma> \<and> pspace_distinct' \<sigma>
2156  \<and> pspace_no_overlap' ptr sz \<sigma>
2157  \<and> ret_zero ptr (2 ^ pdBits) \<sigma>
2158  \<and> region_is_zero_bytes ptr (2 ^ pdBits) x
2159  \<and> range_cover ptr sz pdBits 1
2160  \<and> valid_global_refs' s
2161  \<and> kernel_data_refs \<inter> {ptr..+ 2 ^ pdBits} = {} \<longrightarrow>
2162  (\<sigma>\<lparr>ksPSpace := foldr (\<lambda>addr. data_map_insert addr ko) (new_cap_addrs 2048 ptr ko) (ksPSpace \<sigma>)\<rparr>,
2163   x\<lparr>globals := globals x
2164                 \<lparr>t_hrs_' := hrs_htd_update (ptr_retyps_gen 1 (Ptr ptr :: (pde_C[2048]) ptr) False)
2165                       (t_hrs_' (globals x))\<rparr>\<rparr>) \<in> rf_sr"
2166  (is "\<forall>\<sigma> x. ?P \<sigma> x \<longrightarrow>
2167    (\<sigma>\<lparr>ksPSpace := ?ks \<sigma>\<rparr>, x\<lparr>globals := globals x\<lparr>t_hrs_' := ?ks' x\<rparr>\<rparr>) \<in> rf_sr")
2168proof (intro impI allI)
2169  fix \<sigma> x
2170  let ?thesis = "(\<sigma>\<lparr>ksPSpace := ?ks \<sigma>\<rparr>, x\<lparr>globals := globals x\<lparr>t_hrs_' := ?ks' x\<rparr>\<rparr>) \<in> rf_sr"
2171  let ?ks = "?ks \<sigma>"
2172  let ?ks' = "?ks' x"
2173  let ?ptr = "Ptr ptr :: (pde_C[2048]) ptr"
2174
2175  assume "?P \<sigma> x"
2176  hence rf: "(\<sigma>, x) \<in> rf_sr" and al: "is_aligned ptr pdBits" and ptr0: "ptr \<noteq> 0"
2177    and cover: "range_cover ptr sz pdBits 1"
2178    and sz: "pdBits \<le> sz"
2179    and szb: "sz < word_bits"
2180    and pal: "pspace_aligned' \<sigma>" and pdst: "pspace_distinct' \<sigma>"
2181    and pno: "pspace_no_overlap' ptr sz \<sigma>"
2182    and rzo: "ret_zero ptr (2 ^ pdBits) \<sigma>"
2183    and empty: "region_is_bytes ptr (2 ^ pdBits) x"
2184    and zero: "heap_list_is_zero (hrs_mem (t_hrs_' (globals x))) ptr (2 ^ pdBits)"
2185    and kernel_data_refs_disj : "kernel_data_refs \<inter> {ptr..+ 2 ^ pdBits} = {}"
2186    by (clarsimp simp:range_cover_def[where 'a=32, folded word_bits_def])+
2187
2188  (* obj specific *)
2189  have mko: "\<And>dev. makeObjectKO dev (Inr ARM_HYP_H.PageDirectoryObject) = Some ko"
2190    by (simp add: ko_def makeObjectKO_def)
2191
2192  note blah[simp del] =  atLeastAtMost_iff atLeastatMost_subset_iff atLeastLessThan_iff
2193          Int_atLeastAtMost atLeastatMost_empty_iff split_paired_Ex
2194
2195  let ?blankpde =
2196    "pde_C.words_C_update (\<lambda>_.
2197       Arrays.update (Arrays.update (pde_C.words_C undefined) 0 0) 1 0
2198       ) undefined"
2199  have relrl':
2200    "from_bytes (replicate (size_of TYPE(pde_C)) 0) = ?blankpde"
2201    apply (simp add: from_bytes_def)
2202    apply (simp add: typ_info_simps pde_C_tag_def pde_lift_def pde_get_tag_def
2203      size_td_lt_final_pad size_td_lt_ti_typ_pad_combine Let_def size_of_def)
2204    apply (simp add: final_pad_def Let_def size_td_lt_ti_typ_pad_combine
2205      size_of_def padup_def align_td_array' size_td_array update_ti_adjust_ti
2206      ti_typ_pad_combine_def ti_typ_combine_def empty_typ_info_def)
2207    apply (simp add: typ_info_array array_tag_def eval_nat_numeral)
2208    apply (simp add: array_tag_n.simps)
2209    apply (simp add: final_pad_def Let_def size_td_lt_ti_typ_pad_combine Let_def
2210      size_of_def padup_def align_td_array' size_td_array update_ti_adjust_ti
2211      ti_typ_pad_combine_def Let_def ti_typ_combine_def empty_typ_info_def)
2212    apply (simp add: update_ti_t_word32_0s pde_tag_defs)
2213    done
2214
2215  have relrl:
2216    "cpde_relation makeObject (from_bytes (replicate (size_of TYPE(pde_C)) 0))"
2217    unfolding cpde_relation_def
2218    apply (simp only: relrl')
2219    apply (simp add: Let_def makeObject_pde pde_lift_def)
2220    apply (simp add: pde_lift_def pde_get_tag_def pde_pde_invalid_def)
2221    done
2222
2223  have stored_asid: "pde_stored_asid (from_bytes (replicate (size_of TYPE(pde_C)) 0))
2224                            = None"
2225    apply (simp only: relrl')
2226    apply (simp add: pde_stored_asid_def pde_lift_def pde_pde_invalid_lift_def Let_def
2227                     pde_get_tag_def pde_pde_invalid_def)
2228    done
2229
2230  (* /obj specific *)
2231
2232  (* s/obj/obj'/ *)
2233  have szo: "size_of TYPE(pde_C[2048]) = 2 ^ pdBits"
2234    by (simp add: size_of_def size_td_array table_bits_defs)
2235  have szo2: "2048 * size_of TYPE(pde_C) = 2 ^ pdBits"
2236    by (simp add: szo[symmetric])
2237  have szo': "size_of TYPE(pde_C) = 2 ^ objBitsKO ko"
2238    by (simp add: objBits_simps ko_def archObjSize_def table_bits_defs)
2239
2240  note rl' = cslift_ptr_retyp_other_inst[where n=1,
2241    simplified, OF empty cover[simplified] szo[symmetric] szo]
2242
2243  have sz_weaken: "objBitsKO ko \<le> pdBits"
2244    by (simp add: objBits_simps ko_def archObjSize_def pdBits_def pageBits_def)
2245  have cover': "range_cover ptr sz (objBitsKO ko) 2048"
2246    apply (rule range_cover_rel[OF cover sz_weaken])
2247    apply (simp add: pdBits_def objBits_simps ko_def archObjSize_def pageBits_def)
2248    done
2249  from sz sz_weaken have sz': "objBitsKO ko \<le> sz" by simp
2250  note al' = is_aligned_weaken[OF al sz_weaken]
2251
2252  have koT: "koTypeOf ko = ArchT PDET"
2253    by (simp add: ko_def)
2254
2255  (* rest used to be generic, but PD arrays are complicating everything *)
2256
2257  note rl = projectKO_opt_retyp_other [OF cover' pal pno ko_def]
2258  note cterl = retype_ctes_helper [OF pal pdst pno al' sz' szb mko cover']
2259
2260  have guard: "c_guard ?ptr"
2261    apply (rule is_aligned_c_guard[where n=pdBits and m=2])
2262        apply (simp_all add: al ptr0 align_of_def align_td_array)
2263     apply (simp_all add: table_bits_defs)
2264    done
2265
2266  have guard': "\<forall>n < 2048. c_guard (pde_Ptr ptr +\<^sub>p int n)"
2267    apply (rule retype_guard_helper [OF cover' ptr0 szo', where m=2])
2268     apply (simp_all add: objBits_simps ko_def archObjSize_def align_of_def table_bits_defs)
2269    done
2270
2271  note rl' = cslift_ptr_retyp_other_inst[OF _ cover refl szo,
2272    simplified szo, simplified, OF empty]
2273
2274  from rf have pderl: "cmap_relation (map_to_pdes (ksPSpace \<sigma>)) (cslift x) Ptr cpde_relation"
2275    unfolding rf_sr_def cstate_relation_def by (simp add: Let_def cpspace_relation_def)
2276
2277  note ht_rl = clift_eq_h_t_valid_eq[OF rl', OF tag_disj_via_td_name, simplified]
2278    uinfo_array_tag_n_m_not_le_typ_name
2279
2280  have pde_arr: "cpspace_pde_array_relation (ksPSpace \<sigma>) (t_hrs_' (globals x))
2281    \<Longrightarrow> cpspace_pde_array_relation ?ks ?ks'"
2282   apply (erule cmap_relation_array_add_array[OF _ al])
2283        apply (simp add: foldr_upd_app_if[folded data_map_insert_def])
2284        apply (rule projectKO_opt_retyp_same, simp add: ko_def projectKOs)
2285       apply (simp add: h_t_valid_clift_Some_iff dom_def split: if_split)
2286       apply (subst clift_ptr_retyps_gen_prev_memset_same[where n=1, simplified, OF guard],
2287         simp_all only: szo empty, simp_all add: zero)[1]
2288        apply (simp add: table_bits_defs word_bits_def)
2289       apply (auto split: if_split)[1]
2290      apply (simp_all add: objBits_simps archObjSize_def table_bits_defs
2291                           ko_def word_bits_def)
2292   done
2293
2294  from rf have cpsp: "cpspace_relation (ksPSpace \<sigma>) (underlying_memory (ksMachineState \<sigma>)) (t_hrs_' (globals x))"
2295    unfolding rf_sr_def cstate_relation_def by (simp add: Let_def)
2296  hence "cpspace_relation ?ks (underlying_memory (ksMachineState \<sigma>))  ?ks'"
2297    unfolding cpspace_relation_def
2298  using pde_arr
2299  apply (clarsimp simp: rl' cterl cte_C_size tag_disj_via_td_name
2300                        foldr_upd_app_if [folded data_map_insert_def])
2301  apply (simp add: ht_rl)
2302  apply (simp add: ptr_retyp_to_array[simplified])
2303  apply (subst clift_ptr_retyps_gen_prev_memset_same[OF guard'], simp_all only: szo2 empty)
2304     apply simp
2305    apply (simp(no_asm) add: table_bits_defs word_bits_def)
2306   apply (simp add: zero)
2307  apply (simp add: rl projectKOs)
2308  apply (simp add: rl projectKO_opt_retyp_same ko_def projectKOs Let_def
2309                   ptr_add_to_new_cap_addrs [OF szo']
2310              cong: if_cong)
2311  apply (erule cmap_relation_retype)
2312  apply (insert relrl, auto)
2313  done
2314
2315  moreover
2316  from rf szb al
2317  have "ptr_span (pd_Ptr (symbol_table ''armUSGlobalPD'')) \<inter> {ptr ..+ 2 ^ pdBits} = {}"
2318    apply (clarsimp simp: valid_global_refs'_def  Let_def
2319                          valid_refs'_def ran_def rf_sr_def cstate_relation_def)
2320    apply (erule disjoint_subset)
2321    apply (simp add:kernel_data_refs_disj)
2322    done
2323
2324  moreover from rf have stored_asids: "(pde_stored_asid \<circ>\<^sub>m clift ?ks')
2325                         = (pde_stored_asid \<circ>\<^sub>m cslift x)"
2326    unfolding rf_sr_def
2327    using cpsp empty
2328    apply (clarsimp simp: rl' cterl cte_C_size tag_disj_via_td_name foldr_upd_app_if [folded data_map_insert_def])
2329    apply (simp add: ptr_retyp_to_array[simplified])
2330    apply (subst clift_ptr_retyps_gen_prev_memset_same[OF guard'], simp_all only: szo2 empty)
2331       apply simp
2332      apply (simp add: table_bits_defs word_bits_def)
2333     apply (simp add: zero)
2334    apply (rule ext)
2335    apply (simp add: map_comp_def stored_asid[simplified] split: option.split if_split)
2336    apply (simp only: o_def CTypesDefs.ptr_add_def' Abs_fnat_hom_mult)
2337    apply (clarsimp simp only:)
2338    apply (drule h_t_valid_intvl_htd_contains_uinfo_t [OF h_t_valid_clift])
2339     apply (rule intvl_self, simp)
2340    apply clarsimp
2341    apply (subst (asm) empty[unfolded region_is_bytes'_def])
2342      apply (simp add: objBits_simps archObjSize_def ko_def table_bits_defs
2343                       offs_in_intvl_iff unat_word_ariths unat_of_nat)
2344     apply clarsimp
2345    apply clarsimp
2346    done
2347
2348  ultimately
2349  show ?thesis using rf empty kernel_data_refs_disj rzo
2350    apply (simp add: rf_sr_def cstate_relation_def Let_def rl'  tag_disj_via_td_name)
2351    apply (simp add: carch_state_relation_def cmachine_state_relation_def)
2352    apply (clarsimp simp add: rl' cte_C_size cterl tag_disj_via_td_name
2353                              hrs_htd_update ht_rl foldr_upd_app_if [folded data_map_insert_def]
2354                              projectKOs rl cvariable_array_ptr_retyps[OF szo]
2355                              zero_ranges_ptr_retyps[where p="pd_Ptr ptr", simplified szo])
2356    apply (subst h_t_valid_ptr_retyps_gen_disjoint)
2357      apply assumption
2358     apply (simp add:szo cte_C_size cte_level_bits_def)
2359     apply (erule disjoint_subset)
2360     apply (simp add: ko_def projectKOs objBits_simps archObjSize_def
2361                      table_bits_defs del: replicate_numeral)
2362    apply (subst h_t_valid_ptr_retyps_gen_disjoint)
2363      apply assumption
2364     apply (simp add:szo cte_C_size cte_level_bits_def)
2365     apply (erule disjoint_subset)
2366     apply (simp add: ko_def projectKOs objBits_simps archObjSize_def
2367                      table_bits_defs del: replicate_numeral)
2368    apply (simp add:szo ptr_retyps_htd_safe_neg hrs_htd_def
2369      kernel_data_refs_domain_eq_rotate
2370      ko_def projectKOs objBits_simps archObjSize_def Int_ac
2371      table_bits_defs
2372      del: replicate_numeral)
2373    done
2374qed
2375
2376definition
2377  object_type_from_H :: "object_type \<Rightarrow> word32"
2378  where
2379  "object_type_from_H tp \<equiv> case tp of
2380                              APIObjectType x \<Rightarrow>
2381                                     (case x of ArchTypes_H.apiobject_type.Untyped \<Rightarrow> scast seL4_UntypedObject
2382                                              | ArchTypes_H.apiobject_type.TCBObject \<Rightarrow> scast seL4_TCBObject
2383                                              | ArchTypes_H.apiobject_type.EndpointObject \<Rightarrow> scast seL4_EndpointObject
2384                                              | ArchTypes_H.apiobject_type.NotificationObject \<Rightarrow> scast seL4_NotificationObject
2385                                              | ArchTypes_H.apiobject_type.CapTableObject \<Rightarrow> scast seL4_CapTableObject)
2386                            | ARM_HYP_H.SmallPageObject \<Rightarrow> scast seL4_ARM_SmallPageObject
2387                            | ARM_HYP_H.LargePageObject \<Rightarrow> scast seL4_ARM_LargePageObject
2388                            | ARM_HYP_H.SectionObject \<Rightarrow> scast seL4_ARM_SectionObject
2389                            | ARM_HYP_H.SuperSectionObject \<Rightarrow> scast seL4_ARM_SuperSectionObject
2390                            | ARM_HYP_H.PageTableObject \<Rightarrow> scast seL4_ARM_PageTableObject
2391                            | ARM_HYP_H.PageDirectoryObject \<Rightarrow> scast seL4_ARM_PageDirectoryObject
2392                            | ARM_HYP_H.VCPUObject \<Rightarrow> scast seL4_ARM_VCPUObject"
2393
2394lemmas nAPIObjects_def = seL4_NonArchObjectTypeCount_def
2395
2396lemma nAPIOBjects_object_type_from_H:
2397  "(scast nAPIObjects <=s object_type_from_H tp) = (toAPIType tp = None)"
2398  by (simp add: toAPIType_def nAPIObjects_def
2399    object_type_from_H_def word_sle_def api_object_defs "StrictC'_object_defs"
2400    split: ARM_HYP_H.object_type.splits ArchTypes_H.apiobject_type.splits)
2401
2402definition
2403  object_type_to_H :: "word32 \<Rightarrow> object_type"
2404  where
2405  "object_type_to_H x \<equiv>
2406     (if (x = scast seL4_UntypedObject) then APIObjectType ArchTypes_H.apiobject_type.Untyped else (
2407      if (x = scast seL4_TCBObject) then APIObjectType ArchTypes_H.apiobject_type.TCBObject else (
2408       if (x = scast seL4_EndpointObject) then APIObjectType ArchTypes_H.apiobject_type.EndpointObject else (
2409        if (x = scast seL4_NotificationObject) then APIObjectType ArchTypes_H.apiobject_type.NotificationObject else (
2410         if (x = scast seL4_CapTableObject) then APIObjectType ArchTypes_H.apiobject_type.CapTableObject else (
2411          if (x = scast seL4_ARM_SmallPageObject) then ARM_HYP_H.SmallPageObject else (
2412           if (x = scast seL4_ARM_LargePageObject) then ARM_HYP_H.LargePageObject else (
2413            if (x = scast seL4_ARM_SectionObject) then ARM_HYP_H.SectionObject else (
2414             if (x = scast seL4_ARM_SuperSectionObject) then ARM_HYP_H.SuperSectionObject else (
2415              if (x = scast seL4_ARM_PageTableObject) then ARM_HYP_H.PageTableObject else (
2416               if (x = scast seL4_ARM_PageDirectoryObject) then ARM_HYP_H.PageDirectoryObject else (
2417                 if (x = scast seL4_ARM_VCPUObject) then ARM_HYP_H.VCPUObject else
2418                undefined))))))))))))"
2419
2420lemmas Kernel_C_defs =
2421  seL4_UntypedObject_def
2422  seL4_TCBObject_def
2423  seL4_EndpointObject_def
2424  seL4_NotificationObject_def
2425  seL4_CapTableObject_def
2426  seL4_ARM_SmallPageObject_def
2427  seL4_ARM_LargePageObject_def
2428  seL4_ARM_SectionObject_def
2429  seL4_ARM_SuperSectionObject_def
2430  seL4_ARM_PageTableObject_def
2431  seL4_ARM_PageDirectoryObject_def
2432  seL4_ARM_VCPUObject_def
2433  Kernel_C.asidLowBits_def
2434  Kernel_C.asidHighBits_def
2435
2436abbreviation(input)
2437  "Basic_htd_update f ==
2438     (Basic (globals_update (t_hrs_'_update (hrs_htd_update f))))"
2439
2440lemma object_type_to_from_H [simp]: "object_type_to_H (object_type_from_H x) = x"
2441  apply (clarsimp simp: object_type_from_H_def object_type_to_H_def Kernel_C_defs)
2442  by (clarsimp split: object_type.splits apiobject_type.splits simp: Kernel_C_defs)
2443
2444declare ptr_retyps_one[simp]
2445
2446(* FIXME: move *)
2447lemma ccorres_return_C_Seq:
2448  "ccorres_underlying sr \<Gamma> r rvxf arrel xf P P' hs X (return_C xfu v) \<Longrightarrow>
2449      ccorres_underlying sr \<Gamma> r rvxf arrel xf P P' hs X (return_C xfu v ;; Z)"
2450  apply (clarsimp simp: return_C_def)
2451  apply (erule ccorres_semantic_equiv0[rotated])
2452  apply (rule semantic_equivI)
2453  apply (clarsimp simp: exec_assoc[symmetric])
2454  apply (rule exec_Seq_cong, simp)
2455  apply (clarsimp simp: exec_assoc[symmetric])
2456  apply (rule exec_Seq_cong, simp)
2457  apply (rule iffI)
2458   apply (auto elim!:exec_Normal_elim_cases intro: exec.Throw exec.Seq)[1]
2459  apply (auto elim!:exec_Normal_elim_cases intro: exec.Throw)
2460 done
2461
2462(* FIXME: move *)
2463lemma ccorres_rewrite_while_guard:
2464  assumes rl: "\<And>s. s \<in> R \<Longrightarrow> (s \<in> P) = (s \<in> P')"
2465  and     cc: "ccorres r xf G G' hs a (While P' b)"
2466  shows   "ccorres r xf G (G' \<inter> R) hs a (While P' b)"
2467proof (rule iffD1 [OF ccorres_semantic_equiv])
2468  show "ccorres r xf G (G' \<inter> R) hs a (While P' b)"
2469    by (rule ccorres_guard_imp2 [OF cc]) simp
2470next
2471  fix s s'
2472  assume "s \<in> G' \<inter> R"
2473  hence sin: "(s \<in> P) = (s \<in> P')" using rl by simp
2474
2475  show "semantic_equiv \<Gamma> s s' (While P' b) (While P' b)"
2476    apply (rule semantic_equivI)
2477    apply (simp add: sin)
2478    done
2479qed
2480
2481(* FIXME: move *)
2482lemma ccorres_to_vcg_nf:
2483  "\<lbrakk>ccorres rrel xf P P' [] a c; no_fail Q a; \<And>s. P s \<Longrightarrow> Q s\<rbrakk>
2484   \<Longrightarrow> \<Gamma>\<turnstile> {s. P \<sigma> \<and> s \<in> P' \<and> (\<sigma>, s) \<in> rf_sr} c
2485          {s. \<exists>(rv, \<sigma>')\<in>fst (a \<sigma>). (\<sigma>', s) \<in> rf_sr \<and> rrel rv (xf s)}"
2486  apply (rule HoarePartial.conseq_exploit_pre)
2487  apply clarsimp
2488  apply (rule conseqPre)
2489  apply (drule ccorres_to_vcg')
2490    prefer 2
2491    apply simp
2492   apply (simp add: no_fail_def)
2493  apply clarsimp
2494  done
2495
2496lemma mdb_node_get_mdbNext_heap_ccorres:
2497  "ccorres (=) ret__unsigned_' \<top> UNIV hs
2498  (liftM (mdbNext \<circ> cteMDBNode) (getCTE parent))
2499  (\<acute>ret__unsigned :== CALL mdb_node_get_mdbNext(h_val
2500                           (hrs_mem \<acute>t_hrs)
2501                           (Ptr &((Ptr parent :: cte_C ptr) \<rightarrow>[''cteMDBNode_C'']))))"
2502  apply (simp add: ccorres_liftM_simp)
2503  apply (rule ccorres_add_return2)
2504  apply (rule ccorres_guard_imp2)
2505  apply (rule ccorres_getCTE)
2506   apply (rule_tac  P = "\<lambda>s. ctes_of s parent = Some x" in ccorres_from_vcg [where P' = UNIV])
2507   apply (rule allI, rule conseqPre)
2508    apply vcg
2509   apply (clarsimp simp: return_def)
2510   apply (drule cmap_relation_cte)
2511   apply (erule (1) cmap_relationE1)
2512   apply (simp add: typ_heap_simps)
2513   apply (drule ccte_relation_cmdbnode_relation)
2514   apply (erule mdbNext_CL_mdb_node_lift_eq_mdbNext [symmetric])
2515   apply simp
2516   done
2517
2518lemma getCTE_pre_cte_at:
2519  "\<lbrace>\<lambda>s. \<not> cte_at' p s \<rbrace> getCTE p \<lbrace> \<lambda>_ _. False \<rbrace>"
2520  apply (wp getCTE_wp)
2521  apply clarsimp
2522  done
2523
2524lemmas ccorres_getCTE_cte_at = ccorres_guard_from_wp [OF getCTE_pre_cte_at empty_fail_getCTE]
2525  ccorres_guard_from_wp_bind [OF getCTE_pre_cte_at empty_fail_getCTE]
2526
2527lemmas ccorres_guard_from_wp_liftM = ccorres_guard_from_wp [OF liftM_pre iffD2 [OF empty_fail_liftM]]
2528lemmas ccorres_guard_from_wp_bind_liftM = ccorres_guard_from_wp_bind [OF liftM_pre iffD2 [OF empty_fail_liftM]]
2529
2530lemmas ccorres_liftM_getCTE_cte_at = ccorres_guard_from_wp_liftM [OF getCTE_pre_cte_at empty_fail_getCTE]
2531  ccorres_guard_from_wp_bind_liftM [OF getCTE_pre_cte_at empty_fail_getCTE]
2532
2533lemma insertNewCap_ccorres_helper:
2534  notes option.case_cong_weak [cong]
2535  shows "ccap_relation cap rv'b
2536       \<Longrightarrow> ccorres dc xfdc (cte_at' slot and K (is_aligned next 3 \<and> is_aligned parent 3))
2537           UNIV hs (setCTE slot (CTE cap (MDB next parent True True)))
2538           (Basic (\<lambda>s. globals_update (t_hrs_'_update (hrs_mem_update (heap_update
2539                                    (Ptr &(Ptr slot :: cte_C ptr\<rightarrow>[''cap_C'']) :: cap_C ptr) rv'b))) s);;
2540            \<acute>ret__struct_mdb_node_C :== CALL mdb_node_new(ptr_val (Ptr next),scast true,scast true,ptr_val (Ptr parent));;
2541            Guard C_Guard \<lbrace>hrs_htd \<acute>t_hrs \<Turnstile>\<^sub>t (Ptr slot :: cte_C ptr)\<rbrace>
2542             (Basic (\<lambda>s. globals_update (t_hrs_'_update (hrs_mem_update (heap_update
2543                                                                  (Ptr &(Ptr slot :: cte_C ptr\<rightarrow>[''cteMDBNode_C'']) :: mdb_node_C ptr)
2544                                                                  (ret__struct_mdb_node_C_' s)))) s)))"
2545  apply simp
2546  apply (rule ccorres_from_vcg)
2547  apply (rule allI, rule conseqPre)
2548   apply vcg
2549  apply (clarsimp simp: Collect_const_mem cte_wp_at_ctes_of)
2550  apply (frule (1) rf_sr_ctes_of_clift)
2551  apply (clarsimp simp: typ_heap_simps)
2552  apply (rule fst_setCTE [OF ctes_of_cte_at], assumption)
2553   apply (erule bexI [rotated])
2554   apply (clarsimp simp: cte_wp_at_ctes_of)
2555   apply (clarsimp simp add: rf_sr_def cstate_relation_def typ_heap_simps
2556     Let_def cpspace_relation_def)
2557   apply (rule conjI)
2558    apply (erule (2) cmap_relation_updI)
2559    apply (simp add: ccap_relation_def ccte_relation_def cte_lift_def)
2560    subgoal by (simp add: cte_to_H_def map_option_Some_eq2 mdb_node_to_H_def to_bool_mask_to_bool_bf is_aligned_neg_mask
2561      c_valid_cte_def true_def
2562      split: option.splits)
2563   subgoal by simp
2564   apply (erule_tac t = s' in ssubst)
2565   apply (simp cong: lifth_update)
2566   apply (rule conjI)
2567    apply (erule (1) setCTE_tcb_case)
2568   by (simp add: carch_state_relation_def cmachine_state_relation_def
2569                    typ_heap_simps
2570                    cvariable_array_map_const_add_map_option[where f="tcb_no_ctes_proj"])
2571
2572definition
2573   byte_regions_unmodified :: "heap_raw_state \<Rightarrow> heap_raw_state \<Rightarrow> bool"
2574where
2575  "byte_regions_unmodified hrs hrs' \<equiv> \<forall>x. (\<forall>n td b. snd (hrs_htd hrs x) n = Some (td, b)
2576        \<longrightarrow> td = typ_uinfo_t TYPE (word8))
2577    \<longrightarrow> snd (hrs_htd hrs x) 0 \<noteq> None
2578    \<longrightarrow> hrs_mem hrs' x = hrs_mem hrs x"
2579
2580abbreviation
2581  byte_regions_unmodified' :: "globals myvars \<Rightarrow> globals myvars \<Rightarrow> bool"
2582where
2583  "byte_regions_unmodified' s t \<equiv> byte_regions_unmodified (t_hrs_' (globals s))
2584    (t_hrs_' (globals t))"
2585
2586lemma byte_regions_unmodified_refl[iff]:
2587  "byte_regions_unmodified hrs hrs"
2588  by (simp add: byte_regions_unmodified_def)
2589
2590lemma byte_regions_unmodified_trans:
2591  "byte_regions_unmodified hrs hrs'
2592    \<Longrightarrow> byte_regions_unmodified hrs' hrs''
2593    \<Longrightarrow> hrs_htd hrs' = hrs_htd hrs
2594    \<Longrightarrow> byte_regions_unmodified hrs hrs''"
2595  by (simp add: byte_regions_unmodified_def)
2596
2597lemma byte_regions_unmodified_hrs_mem_update1:
2598  "byte_regions_unmodified hrs hrs'
2599    \<Longrightarrow> hrs_htd hrs \<Turnstile>\<^sub>t (p :: ('a :: wf_type) ptr)
2600    \<Longrightarrow> hrs_htd hrs' = hrs_htd hrs
2601    \<Longrightarrow> typ_uinfo_t TYPE ('a) \<noteq> typ_uinfo_t TYPE (word8)
2602    \<Longrightarrow> byte_regions_unmodified hrs
2603      (hrs_mem_update (heap_update p v) hrs')"
2604  apply (erule byte_regions_unmodified_trans, simp_all)
2605  apply (clarsimp simp: byte_regions_unmodified_def hrs_mem_update
2606                        heap_update_def h_t_valid_def
2607                        valid_footprint_def Let_def)
2608  apply (rule heap_update_nmem_same)
2609  apply (clarsimp simp: size_of_def intvl_def)
2610  apply (drule spec, drule(1) mp, clarsimp)
2611  apply (cut_tac s="(typ_uinfo_t TYPE('a))" and n=k in ladder_set_self)
2612  apply (clarsimp dest!: in_set_list_map)
2613  apply (drule(1) map_le_trans)
2614  apply (simp add: map_le_def)
2615  apply metis
2616  done
2617
2618lemma byte_regions_unmodified_hrs_mem_update2:
2619  "byte_regions_unmodified hrs hrs'
2620    \<Longrightarrow> hrs_htd hrs \<Turnstile>\<^sub>t (p :: ('a :: wf_type) ptr)
2621    \<Longrightarrow> typ_uinfo_t TYPE ('a) \<noteq> typ_uinfo_t TYPE (word8)
2622    \<Longrightarrow> byte_regions_unmodified (hrs_mem_update (heap_update p v) hrs) hrs'"
2623  apply (erule byte_regions_unmodified_trans[rotated], simp_all)
2624  apply (clarsimp simp: byte_regions_unmodified_def hrs_mem_update
2625                        heap_update_def h_t_valid_def
2626                        valid_footprint_def Let_def)
2627  apply (rule sym, rule heap_update_nmem_same)
2628  apply (clarsimp simp: size_of_def intvl_def)
2629  apply (drule spec, drule(1) mp, clarsimp)
2630  apply (cut_tac s="(typ_uinfo_t TYPE('a))" and n=k in ladder_set_self)
2631  apply (clarsimp dest!: in_set_list_map)
2632  apply (drule(1) map_le_trans)
2633  apply (simp add: map_le_def)
2634  apply metis
2635  done
2636
2637lemmas byte_regions_unmodified_hrs_mem_update
2638  = byte_regions_unmodified_hrs_mem_update1
2639    byte_regions_unmodified_hrs_mem_update2
2640
2641lemma byte_regions_unmodified_hrs_htd_update[iff]:
2642  "byte_regions_unmodified
2643      (hrs_htd_update h hrs) hrs"
2644  by (clarsimp simp: byte_regions_unmodified_def)
2645
2646lemma byte_regions_unmodified_flip:
2647  "byte_regions_unmodified (hrs_htd_update (\<lambda>_. hrs_htd hrs) hrs') hrs
2648    \<Longrightarrow> byte_regions_unmodified hrs hrs'"
2649  by (simp add: byte_regions_unmodified_def hrs_htd_update)
2650
2651lemma mdb_node_ptr_set_mdbPrev_preserves_bytes:
2652  "\<forall>s. \<Gamma>\<turnstile>\<^bsub>/UNIV\<^esub> {s} Call mdb_node_ptr_set_mdbPrev_'proc
2653      {t. hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s))
2654         \<and> byte_regions_unmodified' s t}"
2655  apply (hoare_rule HoarePartial.ProcNoRec1)
2656  apply (rule allI, rule conseqPre, vcg)
2657  apply (clarsimp simp: )
2658  apply (intro byte_regions_unmodified_hrs_mem_update byte_regions_unmodified_refl,
2659    simp_all add: typ_heap_simps)
2660  done
2661
2662lemma mdb_node_ptr_set_mdbNext_preserves_bytes:
2663  "\<forall>s. \<Gamma>\<turnstile>\<^bsub>/UNIV\<^esub> {s} Call mdb_node_ptr_set_mdbNext_'proc
2664      {t. hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s))
2665         \<and> byte_regions_unmodified' s t}"
2666  apply (hoare_rule HoarePartial.ProcNoRec1)
2667  apply (rule allI, rule conseqPre, vcg)
2668  apply (clarsimp simp: )
2669  apply (intro byte_regions_unmodified_hrs_mem_update byte_regions_unmodified_refl,
2670    simp_all add: typ_heap_simps)
2671  done
2672
2673lemma updateNewFreeIndex_noop_ccorres:
2674  "ccorres dc xfdc (valid_objs' and cte_wp_at' (\<lambda>cte. cteCap cte = cap) slot)
2675      {s. (case untypedZeroRange cap of None \<Rightarrow> True
2676          | Some (a, b) \<Rightarrow> region_actually_is_zero_bytes a (unat ((b + 1) - a)) s)} hs
2677      (updateNewFreeIndex slot) Skip"
2678  (is "ccorres _ _ ?P ?P' hs _ _")
2679  apply (simp add: updateNewFreeIndex_def getSlotCap_def)
2680  apply (rule ccorres_guard_imp)
2681    apply (rule ccorres_pre_getCTE[where P="\<lambda>rv. cte_wp_at' ((=) rv) slot and ?P"
2682        and P'="K ?P'"])
2683    apply (case_tac "cteCap cte", simp_all add: ccorres_guard_imp[OF ccorres_return_Skip])[1]
2684    defer
2685    apply (clarsimp simp: cte_wp_at_ctes_of)
2686   apply simp
2687  apply (simp add: updateTrackedFreeIndex_def getSlotCap_def)
2688  apply (rule ccorres_guard_imp)
2689    apply (rule_tac P="\<lambda>rv. cte_wp_at' ((=) rv) slot and K (rv = cte) and ?P"
2690        in ccorres_pre_getCTE[where P'="K ?P'"])
2691    defer
2692    apply (clarsimp simp: cte_wp_at_ctes_of)
2693   apply simp
2694  apply (rule ccorres_from_vcg)
2695  apply (rule allI, rule conseqPre, vcg)
2696  apply (clarsimp simp: bind_def simpler_modify_def)
2697  apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def)
2698  apply (clarsimp simp: zero_ranges_are_zero_def
2699                        cte_wp_at_ctes_of
2700                 split: option.split)
2701  done
2702
2703lemma byte_regions_unmodified_region_is_bytes:
2704  "byte_regions_unmodified hrs hrs'
2705    \<Longrightarrow> region_actually_is_bytes' y n (hrs_htd hrs)
2706    \<Longrightarrow> x \<in> {y ..+ n}
2707    \<Longrightarrow> hrs_mem hrs' x = hrs_mem hrs x"
2708  apply (clarsimp simp: byte_regions_unmodified_def imp_conjL[symmetric])
2709  apply (drule spec, erule mp)
2710  apply (clarsimp simp: region_actually_is_bytes'_def)
2711  apply (drule(1) bspec, simp split: if_split_asm)
2712  done
2713
2714lemma insertNewCap_ccorres1:
2715  "ccorres dc xfdc (pspace_aligned' and valid_mdb' and valid_objs' and valid_cap' cap)
2716     ({s. (case untypedZeroRange cap of None \<Rightarrow> True
2717          | Some (a, b) \<Rightarrow> region_actually_is_zero_bytes a (unat ((b + 1) - a)) s)}
2718       \<inter> {s. ccap_relation cap (cap_' s)} \<inter> {s. parent_' s = Ptr parent}
2719       \<inter> {s. slot_' s = Ptr slot}) []
2720     (insertNewCap parent slot cap)
2721     (Call insertNewCap_'proc)"
2722  apply (cinit (no_ignore_call) lift: cap_' parent_' slot_')
2723  apply (rule ccorres_liftM_getCTE_cte_at)
2724   apply (rule ccorres_move_c_guard_cte)
2725   apply (simp only: )
2726   apply (rule ccorres_split_nothrow [OF mdb_node_get_mdbNext_heap_ccorres])
2727      apply ceqv
2728     apply (erule_tac s = "next" in subst)
2729     apply csymbr
2730     apply (ctac (c_lines 3) pre: ccorres_pre_getCTE ccorres_assert add: insertNewCap_ccorres_helper)
2731       apply (simp only: Ptr_not_null_pointer_not_zero)
2732       apply (ctac add: updateMDB_set_mdbPrev)
2733         apply (rule ccorres_seq_skip'[THEN iffD1])
2734         apply (ctac add: updateMDB_set_mdbNext)
2735           apply (rule updateNewFreeIndex_noop_ccorres[where cap=cap])
2736          apply (wp updateMDB_weak_cte_wp_at)
2737         apply simp
2738         apply (vcg exspec=mdb_node_ptr_set_mdbNext_preserves_bytes)
2739        apply (wp updateMDB_weak_cte_wp_at)
2740       apply clarsimp
2741       apply (vcg exspec=mdb_node_ptr_set_mdbPrev_preserves_bytes)
2742      apply (wp setCTE_weak_cte_wp_at)
2743     apply (clarsimp simp: hrs_mem_update Collect_const_mem
2744                 simp del: imp_disjL)
2745     apply vcg
2746    apply simp
2747    apply (wp getCTE_wp')
2748   apply (clarsimp simp: hrs_mem_update)
2749   apply vcg
2750  apply (rule conjI)
2751   apply (clarsimp simp: cte_wp_at_ctes_of is_aligned_3_next)
2752  apply (clarsimp split: option.split)
2753  apply (intro allI conjI impI; simp; clarsimp simp: region_actually_is_bytes)
2754   apply (erule trans[OF heap_list_h_eq2, rotated])
2755   apply (rule byte_regions_unmodified_region_is_bytes)
2756      apply (erule byte_regions_unmodified_trans[rotated]
2757         | simp
2758         | rule byte_regions_unmodified_hrs_mem_update
2759         | simp add: typ_heap_simps')+
2760  apply (erule trans[OF heap_list_h_eq2, rotated])
2761  apply (rule byte_regions_unmodified_region_is_bytes)
2762     apply (erule byte_regions_unmodified_trans[rotated]
2763        | simp
2764        | rule byte_regions_unmodified_hrs_mem_update
2765        | simp add: typ_heap_simps')+
2766  done
2767
2768 lemma insertNewCap_pre_cte_at:
2769  "\<lbrace>\<lambda>s. \<not> (cte_at' p s \<and> cte_at' p' s) \<rbrace> insertNewCap p p' cap \<lbrace> \<lambda>_ _. False \<rbrace>"
2770  unfolding insertNewCap_def
2771  apply simp
2772  apply (wp getCTE_wp)
2773  apply (clarsimp simp: cte_wp_at_ctes_of)
2774  done
2775
2776
2777end
2778
2779locale insertNewCap_i_locale = kernel
2780begin
2781
2782lemma mdb_node_get_mdbNext_spec:
2783  "\<forall>s. \<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> {s} Call mdb_node_get_mdbNext_'proc {t. i_' t = i_' s}"
2784  apply (rule allI)
2785  apply (hoare_rule HoarePartial.ProcNoRec1)
2786  apply vcg
2787  apply simp
2788  done
2789
2790lemma mdb_node_new_spec:
2791  "\<forall>s. \<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> {s} Call mdb_node_new_'proc {t. i_' t = i_' s}"
2792  apply (rule allI)
2793  apply (hoare_rule HoarePartial.ProcNoRec1)
2794  apply vcg
2795  apply simp
2796  done
2797
2798lemma mdb_node_ptr_set_mdbPrev_spec:
2799  "\<forall>s. \<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> {s} Call mdb_node_ptr_set_mdbPrev_'proc {t. i_' t = i_' s}"
2800  apply (rule allI)
2801  apply (hoare_rule HoarePartial.ProcNoRec1)
2802  apply vcg
2803  apply simp
2804  done
2805
2806lemma mdb_node_ptr_set_mdbNext_spec:
2807  "\<forall>s. \<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> {s} Call mdb_node_ptr_set_mdbNext_'proc {t. i_' t = i_' s}"
2808  apply (rule allI)
2809  apply (hoare_rule HoarePartial.ProcNoRec1)
2810  apply vcg
2811  apply simp
2812  done
2813
2814lemma insertNewCap_spec:
2815  "\<forall>s. \<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> {s} Call insertNewCap_'proc {t. i_' t = i_' s}"
2816  apply vcg
2817  apply clarsimp
2818  done
2819end
2820
2821context kernel_m
2822begin
2823
2824lemma insertNewCap_spec:
2825  "\<forall>s. \<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> {s} Call insertNewCap_'proc {t. i_' t = i_' s}"
2826  apply (rule insertNewCap_i_locale.insertNewCap_spec)
2827  apply (intro_locales)
2828  done
2829
2830lemma ccorres_fail:
2831  "ccorres r xf \<top> UNIV hs fail c"
2832  apply (rule ccorresI')
2833  apply (simp add: fail_def)
2834  done
2835
2836lemma hoarep_Cond_UNIV:
2837  "\<Gamma>\<turnstile>\<^bsub>/UNIV\<^esub> P c P', A \<Longrightarrow>
2838  \<Gamma>\<turnstile>\<^bsub>/UNIV\<^esub> P (Cond UNIV c d)  P', A"
2839  apply (rule HoarePartial.Cond [where P\<^sub>1 = P and P\<^sub>2 = "{}"])
2840    apply simp
2841   apply assumption
2842  apply (rule HoarePartial.conseq_exploit_pre)
2843  apply simp
2844  done
2845
2846lemma object_type_from_H_toAPIType_simps:
2847  "(object_type_from_H tp = scast seL4_UntypedObject) = (toAPIType tp = Some ArchTypes_H.apiobject_type.Untyped)"
2848  "(object_type_from_H tp = scast seL4_TCBObject) = (toAPIType tp = Some ArchTypes_H.apiobject_type.TCBObject)"
2849  "(object_type_from_H tp = scast seL4_EndpointObject) = (toAPIType tp = Some ArchTypes_H.apiobject_type.EndpointObject)"
2850  "(object_type_from_H tp = scast seL4_NotificationObject) = (toAPIType tp = Some ArchTypes_H.apiobject_type.NotificationObject)"
2851  "(object_type_from_H tp = scast seL4_CapTableObject) = (toAPIType tp = Some ArchTypes_H.apiobject_type.CapTableObject)"
2852  "(object_type_from_H tp = scast seL4_ARM_SmallPageObject) = (tp = ARM_HYP_H.SmallPageObject)"
2853  "(object_type_from_H tp = scast seL4_ARM_LargePageObject) = (tp = ARM_HYP_H.LargePageObject)"
2854  "(object_type_from_H tp = scast seL4_ARM_SectionObject) = (tp = ARM_HYP_H.SectionObject)"
2855  "(object_type_from_H tp = scast seL4_ARM_SuperSectionObject) = (tp = ARM_HYP_H.SuperSectionObject)"
2856  "(object_type_from_H tp = scast seL4_ARM_PageTableObject) = (tp = ARM_HYP_H.PageTableObject)"
2857  "(object_type_from_H tp = scast seL4_ARM_PageDirectoryObject) = (tp = ARM_HYP_H.PageDirectoryObject)"
2858  "(object_type_from_H tp = scast seL4_ARM_VCPUObject) = (tp = ARM_HYP_H.VCPUObject)"
2859  by (auto simp: toAPIType_def
2860                 object_type_from_H_def "StrictC'_object_defs" api_object_defs
2861          split: object_type.splits ArchTypes_H.apiobject_type.splits)
2862
2863declare Collect_const_mem [simp]
2864
2865lemma createNewCaps_untyped_if_helper:
2866  "\<forall>s s'. (s, s') \<in> rf_sr \<and> (sz < word_bits \<and> gbits < word_bits) \<and> True  \<longrightarrow>
2867             (\<not> gbits \<le> sz) = (s' \<in> \<lbrace>of_nat sz < (of_nat gbits :: word32)\<rbrace>)"
2868  by (clarsimp simp: not_le unat_of_nat32 word_less_nat_alt lt_word_bits_lt_pow)
2869
2870lemma true_mask1 [simp]:
2871  "true && mask (Suc 0) = true"
2872  unfolding true_def
2873  by (simp add: bang_eq cong: conj_cong)
2874
2875(* Levity: added (20090419 09:44:40) *)
2876declare shiftl_mask_is_0 [simp]
2877
2878lemma to_bool_simps [simp]:
2879  "to_bool true" "\<not> to_bool false"
2880  unfolding true_def false_def to_bool_def
2881  by simp_all
2882
2883lemma heap_list_update':
2884  "\<lbrakk> n = length v; length v \<le> 2 ^ word_bits \<rbrakk> \<Longrightarrow> heap_list (heap_update_list p v h) n p = v"
2885  by (simp add: heap_list_update addr_card_wb)
2886
2887lemma heap_update_field':
2888  "\<lbrakk>field_ti TYPE('a :: packed_type) f = Some t; c_guard p;
2889  export_uinfo t = export_uinfo (typ_info_t TYPE('b :: packed_type))\<rbrakk>
2890  \<Longrightarrow> heap_update (Ptr &(p\<rightarrow>f) :: 'b ptr) v hp =
2891  heap_update p (update_ti_t t (to_bytes_p v) (h_val hp p)) hp"
2892  apply (erule field_ti_field_lookupE)
2893  apply (subst packed_heap_super_field_update [unfolded typ_uinfo_t_def])
2894     apply assumption+
2895  apply (drule export_size_of [simplified typ_uinfo_t_def])
2896  apply (simp add: update_ti_t_def)
2897  done
2898
2899lemma h_t_valid_clift_Some_iff':
2900  "td \<Turnstile>\<^sub>t p = (clift (hp, td) p = Some (h_val hp p))"
2901  by (simp add: lift_t_if split: if_split)
2902
2903lemma option_noneI: "\<lbrakk> \<And>x. a = Some x \<Longrightarrow> False \<rbrakk> \<Longrightarrow> a = None"
2904  apply (case_tac a)
2905   apply clarsimp
2906  apply atomize
2907  apply clarsimp
2908  done
2909
2910lemma projectKO_opt_retyp_other':
2911  assumes pko: "\<forall>v. (projectKO_opt ko :: 'a :: pre_storable option) \<noteq> Some v"
2912  and pno: "pspace_no_overlap' ptr (objBitsKO ko) (\<sigma> :: kernel_state)"
2913  and pal: "pspace_aligned' (\<sigma> :: kernel_state)"
2914  and al: "is_aligned ptr (objBitsKO ko)"
2915  shows "projectKO_opt \<circ>\<^sub>m ((ksPSpace \<sigma>)(ptr \<mapsto> ko))
2916  = (projectKO_opt \<circ>\<^sub>m (ksPSpace \<sigma>) :: word32 \<Rightarrow> 'a :: pre_storable option)" (is "?LHS = ?RHS")
2917proof (rule ext)
2918  fix x
2919  show "?LHS x = ?RHS x"
2920  proof (cases "x = ptr")
2921    case True
2922    hence "x \<in> {ptr..(ptr && ~~ mask (objBitsKO ko)) + 2 ^ objBitsKO ko - 1}"
2923      apply (rule ssubst)
2924      apply (insert al)
2925      apply (clarsimp simp: assms)
2926      done
2927    hence "ksPSpace \<sigma> x = None" using pno
2928      apply -
2929      apply (rule option_noneI)
2930      apply (frule pspace_no_overlap_disjoint'[rotated])
2931       apply (rule pal)
2932      apply (drule domI[where a = x])
2933      apply blast
2934      done
2935    thus ?thesis using True pko by simp
2936  next
2937    case False
2938    thus ?thesis by (simp add: map_comp_def)
2939  qed
2940qed
2941
2942lemma dom_tcb_cte_cases_iff:
2943  "(x \<in> dom tcb_cte_cases) = (\<exists>y < 5. unat x = y * 16)"
2944  unfolding tcb_cte_cases_def
2945  by (auto simp: unat_arith_simps)
2946
2947lemma cmap_relation_retype2:
2948  assumes cm: "cmap_relation mp mp' Ptr rel"
2949  and   rel: "rel (mobj :: 'a :: pre_storable) ko'"
2950  shows "cmap_relation
2951        (\<lambda>x. if x \<in> ptr_val ` addrs then Some (mobj :: 'a :: pre_storable) else mp x)
2952        (\<lambda>y. if y \<in> addrs then Some ko' else mp' y)
2953        Ptr rel"
2954  using cm rel
2955  apply -
2956  apply (rule cmap_relationI)
2957   apply (simp add: dom_if cmap_relation_def image_Un)
2958  apply (case_tac "x \<in> addrs")
2959   apply (simp add: image_image)
2960  apply (simp add: image_image)
2961  apply (clarsimp split: if_split_asm)
2962   apply (erule contrapos_np)
2963   apply (erule image_eqI [rotated])
2964   apply simp
2965  apply (erule (2) cmap_relation_relI)
2966  done
2967
2968lemma ti_typ_pad_combine_empty_ti:
2969  fixes tp :: "'b :: c_type itself"
2970  shows "ti_typ_pad_combine tp lu upd fld (empty_typ_info n) =
2971  TypDesc (TypAggregate [DTPair (adjust_ti (typ_info_t TYPE('b)) lu upd) fld]) n"
2972  by (simp add: ti_typ_pad_combine_def ti_typ_combine_def empty_typ_info_def Let_def)
2973
2974lemma ti_typ_combine_empty_ti:
2975  fixes tp :: "'b :: c_type itself"
2976  shows "ti_typ_combine tp lu upd fld (empty_typ_info n) =
2977  TypDesc (TypAggregate [DTPair (adjust_ti (typ_info_t TYPE('b)) lu upd) fld]) n"
2978  by (simp add: ti_typ_combine_def empty_typ_info_def Let_def)
2979
2980lemma ti_typ_pad_combine_td:
2981  fixes tp :: "'b :: c_type itself"
2982  shows "padup (align_of TYPE('b)) (size_td_struct st) = 0 \<Longrightarrow>
2983  ti_typ_pad_combine tp lu upd fld (TypDesc st n) =
2984  TypDesc (extend_ti_struct st (adjust_ti (typ_info_t TYPE('b)) lu upd) fld) n"
2985  by (simp add: ti_typ_pad_combine_def ti_typ_combine_def Let_def)
2986
2987lemma ti_typ_combine_td:
2988  fixes tp :: "'b :: c_type itself"
2989  shows "padup (align_of TYPE('b)) (size_td_struct st) = 0 \<Longrightarrow>
2990  ti_typ_combine tp lu upd fld (TypDesc st n) =
2991  TypDesc (extend_ti_struct st (adjust_ti (typ_info_t TYPE('b)) lu upd) fld) n"
2992  by (simp add: ti_typ_combine_def Let_def)
2993
2994lemma update_ti_t_pad_combine:
2995  assumes std: "size_td td' mod 2 ^ align_td (typ_info_t TYPE('a :: c_type)) = 0"
2996  shows "update_ti_t (ti_typ_pad_combine TYPE('a :: c_type) lu upd fld td') bs v =
2997  update_ti_t (ti_typ_combine TYPE('a :: c_type) lu upd fld td') bs v"
2998  using std
2999  by (simp add: ti_typ_pad_combine_def size_td_simps Let_def)
3000
3001
3002lemma update_ti_t_ptr_0s:
3003  "update_ti_t (typ_info_t TYPE('a :: c_type ptr)) [0,0,0,0] X = NULL"
3004  apply (simp add: typ_info_ptr word_rcat_def bin_rcat_def)
3005  done
3006
3007lemma size_td_map_list:
3008  "size_td_list (map (\<lambda>n. DTPair
3009                                 (adjust_ti (typ_info_t TYPE('a :: c_type))
3010                                   (\<lambda>x. index x n)
3011                                   (\<lambda>x f. Arrays.update f n x))
3012                                 (replicate n CHR ''1''))
3013                        [0..<n]) = (size_td (typ_info_t TYPE('a :: c_type)) * n)"
3014  apply (induct n)
3015   apply simp
3016  apply simp
3017  done
3018
3019lemma update_ti_t_array_tag_n_rep:
3020  fixes x :: "'a :: c_type ['b :: finite]"
3021  shows "\<lbrakk> bs = replicate (n * size_td (typ_info_t TYPE('a))) v; n \<le> card (UNIV  :: 'b set) \<rbrakk> \<Longrightarrow>
3022  update_ti_t (array_tag_n n) bs x =
3023  foldr (\<lambda>n arr. Arrays.update arr n
3024        (update_ti_t (typ_info_t TYPE('a)) (replicate (size_td (typ_info_t TYPE('a))) v) (index arr n)))
3025        [0..<n] x"
3026  apply (induct n arbitrary: bs x)
3027   apply (simp add: array_tag_n_eq)
3028  apply (simp add: array_tag_n_eq size_td_map_list iffD2 [OF linorder_min_same1] field_simps
3029    cong: if_cong )
3030  apply (simp add: update_ti_adjust_ti)
3031  done
3032
3033lemma update_ti_t_array_rep:
3034  "bs = replicate ((card (UNIV :: 'b :: finite set)) * size_td (typ_info_t TYPE('a))) v \<Longrightarrow>
3035  update_ti_t (typ_info_t TYPE('a :: c_type['b :: finite])) bs x =
3036  foldr (\<lambda>n arr. Arrays.update arr n
3037        (update_ti_t (typ_info_t TYPE('a)) (replicate (size_td (typ_info_t TYPE('a))) v) (index arr n)))
3038        [0..<(card (UNIV :: 'b :: finite set))] x"
3039  unfolding typ_info_array array_tag_def
3040  apply (rule update_ti_t_array_tag_n_rep)
3041    apply simp
3042   apply simp
3043   done
3044
3045lemma update_ti_t_array_rep_word0:
3046  "bs = replicate ((card (UNIV :: 'b :: finite set)) * 4) 0 \<Longrightarrow>
3047  update_ti_t (typ_info_t TYPE(word32['b :: finite])) bs x =
3048  foldr (\<lambda>n arr. Arrays.update arr n 0)
3049        [0..<(card (UNIV :: 'b :: finite set))] x"
3050  apply (subst update_ti_t_array_rep)
3051   apply simp
3052  apply (simp add: update_ti_t_word32_0s)
3053  done
3054
3055lemma newContext_def2:
3056  "newContext \<equiv> (\<lambda>x. if x = register.CPSR then 0x150 else 0)"
3057proof -
3058  have "newContext = (\<lambda>x. if x = register.CPSR then 0x150 else 0)"
3059    apply (simp add: newContext_def initContext_def)
3060    apply (auto intro: ext)
3061    done
3062  thus "newContext \<equiv> (\<lambda>x. if x = register.CPSR then 0x150 else 0)" by simp
3063qed
3064
3065lemma tcb_queue_update_other:
3066  "\<lbrakk> ctcb_ptr_to_tcb_ptr p \<notin> set tcbs \<rbrakk> \<Longrightarrow>
3067  tcb_queue_relation next prev (mp(p \<mapsto> v)) tcbs qe qh =
3068  tcb_queue_relation next prev mp tcbs qe qh"
3069  apply (induct tcbs arbitrary: qh qe)
3070   apply simp
3071  apply (rename_tac a tcbs qh qe)
3072  apply simp
3073  apply (subgoal_tac "p \<noteq> tcb_ptr_to_ctcb_ptr a")
3074   apply (simp cong: conj_cong)
3075  apply clarsimp
3076  done
3077
3078lemma cmap_relation_cong':
3079  "\<lbrakk>am = am'; cm = cm';
3080   \<And>p a a' b b'.
3081      \<lbrakk>am p = Some a; am' p = Some a'; cm (f p) = Some b; cm' (f p) = Some b'\<rbrakk>
3082      \<Longrightarrow> rel a b = rel' a' b'\<rbrakk>
3083    \<Longrightarrow> cmap_relation am cm f rel = cmap_relation am' cm' f rel'"
3084  by (rule cmap_relation_cong, simp_all)
3085
3086lemma tcb_queue_update_other':
3087  "\<lbrakk> ctcb_ptr_to_tcb_ptr p \<notin> set tcbs \<rbrakk> \<Longrightarrow>
3088  tcb_queue_relation' next prev (mp(p \<mapsto> v)) tcbs qe qh =
3089  tcb_queue_relation' next prev mp tcbs qe qh"
3090  unfolding tcb_queue_relation'_def
3091  by (simp add: tcb_queue_update_other)
3092
3093lemma map_to_ko_atI2:
3094  "\<lbrakk>(projectKO_opt \<circ>\<^sub>m (ksPSpace s)) x = Some v; pspace_aligned' s; pspace_distinct' s\<rbrakk> \<Longrightarrow> ko_at' v x s"
3095  apply (clarsimp simp: map_comp_Some_iff)
3096  apply (erule (2) aligned_distinct_obj_atI')
3097  apply (simp add: project_inject)
3098  done
3099
3100lemma c_guard_tcb:
3101  assumes al: "is_aligned (ctcb_ptr_to_tcb_ptr p) tcbBlockSizeBits"
3102  and   ptr0: "ctcb_ptr_to_tcb_ptr p \<noteq> 0"
3103  shows "c_guard p"
3104  unfolding c_guard_def
3105proof (rule conjI)
3106  show "ptr_aligned p" using al
3107    apply -
3108    apply (rule is_aligned_ptr_aligned [where n = word_size_bits])
3109     apply (rule is_aligned_weaken)
3110      apply (erule ctcb_ptr_to_tcb_ptr_aligned)
3111     by (auto simp: align_of_def word_size_bits_def ctcb_size_bits_def)
3112
3113  show "c_null_guard p" using ptr0 al
3114    unfolding c_null_guard_def
3115    apply -
3116    apply (rule intvl_nowrap [where x = 0, simplified])
3117     apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_defs is_aligned_def objBits_defs)
3118    apply (drule ctcb_ptr_to_tcb_ptr_aligned)
3119    apply (erule is_aligned_no_wrap_le)
3120     by (auto simp add: word_bits_conv ctcb_size_bits_def)
3121qed
3122
3123lemma tcb_ptr_orth_cte_ptrs:
3124  "{ptr_val p..+size_of TYPE(tcb_C)} \<inter> {ctcb_ptr_to_tcb_ptr p..+5 * size_of TYPE(cte_C)} = {}"
3125  apply (rule disjointI)
3126  apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def intvl_def field_simps size_of_def ctcb_offset_defs)
3127  apply unat_arith
3128  apply (simp add: unat_of_nat32 word_bits_conv)
3129  apply (simp add: unat_of_nat32 word_bits_conv)
3130  done
3131
3132lemma tcb_ptr_orth_cte_ptrs':
3133  "ptr_span (tcb_Ptr (regionBase + 0x100)) \<inter> ptr_span (Ptr regionBase :: (cte_C[5]) ptr) = {}"
3134  apply (rule disjointI)
3135  apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def size_td_array
3136                        intvl_def field_simps size_of_def ctcb_offset_def)
3137  apply (simp add: unat_arith_simps unat_of_nat)
3138  done
3139
3140lemma region_is_typeless_weaken:
3141  "\<lbrakk> region_is_typeless a b s'; (t_hrs_' (globals s)) = (t_hrs_' (globals s')); a \<le> x; unat x + y \<le> unat a + b \<rbrakk> \<Longrightarrow> region_is_typeless x y s"
3142  by (clarsimp simp: region_is_typeless_def subsetD[OF intvl_both_le])
3143
3144lemmas ptr_retyp_htd_safe_neg
3145    = ptr_retyps_htd_safe_neg[where n="Suc 0" and arr=False,
3146    unfolded ptr_retyps_gen_def, simplified]
3147
3148lemma cnc_tcb_helper:
3149  fixes p :: "tcb_C ptr"
3150  defines "kotcb \<equiv> (KOTCB (makeObject :: tcb))"
3151  assumes rfsr: "(\<sigma>\<lparr>ksPSpace := ks\<rparr>, x) \<in> rf_sr"
3152  and      al: "is_aligned (ctcb_ptr_to_tcb_ptr p) (objBitsKO kotcb)"
3153  and ptr0: "ctcb_ptr_to_tcb_ptr p \<noteq> 0"
3154  and ptrlb: "2^ctcb_size_bits \<le> ptr_val p"
3155  and vq:  "valid_queues \<sigma>"
3156  and pal: "pspace_aligned' (\<sigma>\<lparr>ksPSpace := ks\<rparr>)"
3157  and pno: "pspace_no_overlap' (ctcb_ptr_to_tcb_ptr p) (objBitsKO kotcb) (\<sigma>\<lparr>ksPSpace := ks\<rparr>)"
3158  and pds: "pspace_distinct' (\<sigma>\<lparr>ksPSpace := ks\<rparr>)"
3159  and symref: "sym_refs (state_refs_of' (\<sigma>\<lparr>ksPSpace := ks\<rparr>))"
3160  and kssub: "dom (ksPSpace \<sigma>) \<subseteq> dom ks"
3161  and rzo: "ret_zero (ctcb_ptr_to_tcb_ptr p) (2 ^ objBitsKO kotcb) \<sigma>"
3162  and empty: "region_is_bytes (ctcb_ptr_to_tcb_ptr p) (2 ^ tcbBlockSizeBits) x"
3163  and rep0:  "heap_list (fst (t_hrs_' (globals x))) (2 ^ tcbBlockSizeBits) (ctcb_ptr_to_tcb_ptr p) = replicate (2 ^ tcbBlockSizeBits) 0"
3164  and kdr: "{ctcb_ptr_to_tcb_ptr p..+2 ^ tcbBlockSizeBits} \<inter> kernel_data_refs = {}"
3165  shows "(\<sigma>\<lparr>ksPSpace := ks(ctcb_ptr_to_tcb_ptr p \<mapsto> kotcb)\<rparr>,
3166     globals_update
3167      (t_hrs_'_update
3168        (\<lambda>a. hrs_mem_update (heap_update (Ptr &(p\<rightarrow>[''tcbTimeSlice_C'']) :: machine_word ptr) (5 :: machine_word))
3169              (hrs_mem_update
3170                (heap_update ((Ptr &((Ptr &((Ptr &(p\<rightarrow>[''tcbArch_C'']) :: arch_tcb_C ptr)\<rightarrow>[''tcbContext_C''])
3171                     :: user_context_C ptr)\<rightarrow>[''registers_C''])) :: (word32[20]) ptr)
3172                  (Arrays.update (h_val (hrs_mem a) ((Ptr &((Ptr &((Ptr &(p\<rightarrow>[''tcbArch_C'']) :: arch_tcb_C ptr)\<rightarrow>[''tcbContext_C''])
3173                       :: user_context_C ptr)\<rightarrow>[''registers_C''])) :: (word32[20]) ptr)) (unat Kernel_C.CPSR) (0x150 :: word32)))
3174                   (hrs_htd_update (\<lambda>xa. ptr_retyps_gen 1 (Ptr (ctcb_ptr_to_tcb_ptr p) :: (cte_C[5]) ptr) False
3175                       (ptr_retyps_gen 1 p False xa)) a)))) x)
3176             \<in> rf_sr"
3177  (is "(\<sigma>\<lparr>ksPSpace := ?ks\<rparr>, globals_update ?gs' x) \<in> rf_sr")
3178
3179proof -
3180  define ko where "ko \<equiv> (KOCTE (makeObject :: cte))"
3181  let ?ptr = "cte_Ptr (ctcb_ptr_to_tcb_ptr p)"
3182  let ?arr_ptr = "Ptr (ctcb_ptr_to_tcb_ptr p) :: (cte_C[5]) ptr"
3183  let ?sp = "\<sigma>\<lparr>ksPSpace := ks\<rparr>"
3184  let ?s = "\<sigma>\<lparr>ksPSpace := ?ks\<rparr>"
3185  let ?gs = "?gs' (globals x)"
3186  let ?hp = "(fst (t_hrs_' ?gs), (ptr_retyps_gen 1 p False (snd (t_hrs_' (globals x)))))"
3187
3188  note tcb_C_size[simp del]
3189
3190  from al have cover: "range_cover (ctcb_ptr_to_tcb_ptr p) (objBitsKO kotcb)
3191        (objBitsKO kotcb) (Suc 0)"
3192    by (rule range_cover_full, simp_all add: al)
3193
3194  have "\<forall>n<2 ^ (objBitsKO kotcb - objBitsKO ko). c_guard (CTypesDefs.ptr_add ?ptr (of_nat n))"
3195    apply (rule retype_guard_helper [where m = 2])
3196        apply (rule range_cover_rel[OF cover, rotated])
3197         apply simp
3198        apply (simp add: ko_def objBits_simps' kotcb_def)
3199       apply (rule ptr0)
3200      apply (simp add: ko_def objBits_simps' size_of_def)
3201     apply (simp add: ko_def objBits_simps')
3202    apply (simp add: ko_def objBits_simps align_of_def)
3203    done
3204  hence guard: "\<forall>n<5. c_guard (CTypesDefs.ptr_add ?ptr (of_nat n))"
3205    by (simp add: ko_def kotcb_def objBits_simps' align_of_def)
3206
3207  have arr_guard: "c_guard ?arr_ptr"
3208    apply (rule is_aligned_c_guard[where m=2], simp, rule al)
3209       apply (simp add: ptr0)
3210      apply (simp add: align_of_def align_td_array)
3211     apply (simp add: cte_C_size objBits_simps' kotcb_def)
3212    apply (simp add: kotcb_def objBits_simps')
3213    done
3214
3215  have heap_update_to_hrs_mem_update:
3216    "\<And>p x hp ht. (heap_update p x hp, ht) = hrs_mem_update (heap_update p x) (hp, ht)"
3217    by (simp add: hrs_mem_update_def split_def)
3218
3219  have empty_smaller:
3220    "region_is_bytes (ptr_val p) (size_of TYPE(tcb_C)) x"
3221    "region_is_bytes' (ctcb_ptr_to_tcb_ptr p) (5 * size_of TYPE(cte_C))
3222        (ptr_retyps_gen 1 p False (hrs_htd (t_hrs_' (globals x))))"
3223     using al region_is_bytes_subset[OF empty] tcb_ptr_to_ctcb_ptr_in_range'
3224     apply (simp add: objBits_simps kotcb_def)
3225    apply (clarsimp simp: region_is_bytes'_def)
3226    apply (subst(asm) ptr_retyps_gen_out)
3227     apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_defs intvl_def)
3228     apply (simp add: unat_arith_simps unat_of_nat cte_C_size tcb_C_size
3229               split: if_split_asm)
3230    apply (subst(asm) empty[unfolded region_is_bytes'_def], simp_all)
3231    apply (erule subsetD[rotated], rule intvl_start_le)
3232    apply (simp add: cte_C_size objBits_defs)
3233    done
3234
3235  note htd[simp] = hrs_htd_update_htd_update[unfolded o_def,
3236        where d="ptr_retyps_gen n p a" and d'="ptr_retyps_gen n' p' a'"
3237        for n p a n' p' a', symmetric]
3238
3239  have cgp: "c_guard p" using al
3240    apply -
3241    apply (rule c_guard_tcb [OF _ ptr0])
3242    apply (simp add: kotcb_def objBits_simps)
3243    done
3244
3245  from pal rfsr have "\<forall>x\<in>dom (cslift x :: cte_C typ_heap). is_aligned (ptr_val x) (objBitsKO ko)"
3246    apply (rule pspace_aligned_to_C_cte [OF _ cmap_relation_cte])
3247    apply (simp add: projectKOs ko_def)
3248    done
3249
3250  have "ptr_val p = ctcb_ptr_to_tcb_ptr p + ctcb_offset"
3251    by (simp add: ctcb_ptr_to_tcb_ptr_def)
3252
3253  have cte_tcb_disjoint: "\<And>y. y \<in> (CTypesDefs.ptr_add (cte_Ptr (ctcb_ptr_to_tcb_ptr p)) \<circ> of_nat) ` {k. k < 5}
3254    \<Longrightarrow> {ptr_val p..+size_of TYPE(tcb_C)} \<inter> {ptr_val y..+size_of TYPE(cte_C)} = {}"
3255    apply (rule disjoint_subset2 [OF _ tcb_ptr_orth_cte_ptrs])
3256    apply (clarsimp simp: intvl_def size_of_def)
3257    apply (rule_tac x = "x * 16 + k" in exI)
3258    apply simp
3259    done
3260
3261  have cl_cte: "(cslift (x\<lparr>globals := ?gs\<rparr>) :: cte_C typ_heap) =
3262    (\<lambda>y. if y \<in> (CTypesDefs.ptr_add (cte_Ptr (ctcb_ptr_to_tcb_ptr p)) \<circ>
3263                 of_nat) `
3264                {k. k < 5}
3265         then Some (from_bytes (replicate (size_of TYPE(cte_C)) 0)) else cslift x y)"
3266    using cgp
3267    apply (simp add: ptr_retyp_to_array[simplified] hrs_comm[symmetric])
3268    apply (subst clift_ptr_retyps_gen_prev_memset_same[OF guard],
3269           simp_all add: hrs_htd_update empty_smaller[simplified])
3270      apply (simp add: cte_C_size word_bits_def)
3271     apply (simp add: hrs_mem_update typ_heap_simps
3272                      packed_heap_update_collapse)
3273     apply (simp add: heap_update_def)
3274     apply (subst heap_list_update_disjoint_same)
3275      apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_defs intvl_def
3276                            set_eq_iff)
3277      apply (simp add: unat_arith_simps unat_of_nat cte_C_size tcb_C_size)
3278     apply (subst take_heap_list_le[symmetric])
3279      prefer 2
3280      apply (simp add: hrs_mem_def, subst rep0)
3281      apply (simp only: take_replicate, simp add: cte_C_size objBits_defs)
3282     apply (simp add: cte_C_size objBits_defs)
3283    apply (simp add: fun_eq_iff
3284              split: if_split)
3285    apply (simp add: hrs_comm packed_heap_update_collapse
3286                     typ_heap_simps)
3287    apply (subst clift_heap_update_same_td_name, simp_all,
3288      simp add: hrs_htd_update ptr_retyps_gen_def ptr_retyp_h_t_valid)+
3289    apply (subst clift_ptr_retyps_gen_other,
3290      simp_all add: empty_smaller tag_disj_via_td_name)
3291    apply (simp add: tcb_C_size word_bits_def)
3292    done
3293
3294  have tcb0: "heap_list (fst (t_hrs_' (globals x))) (size_of TYPE(tcb_C)) (ptr_val p) = replicate (size_of TYPE(tcb_C)) 0"
3295  proof -
3296    have "heap_list (fst (t_hrs_' (globals x))) (size_of TYPE(tcb_C)) (ptr_val p)
3297      = take (size_of TYPE(tcb_C))
3298             (drop (unat (ptr_val p - ctcb_ptr_to_tcb_ptr p))
3299                   (heap_list (fst (t_hrs_' (globals x))) (2 ^ tcbBlockSizeBits) (ctcb_ptr_to_tcb_ptr p)))"
3300      by (simp add: drop_heap_list_le take_heap_list_le size_of_def ctcb_ptr_to_tcb_ptr_def
3301                    ctcb_offset_defs objBits_defs)
3302    also have "\<dots> = replicate (size_of TYPE(tcb_C)) 0"
3303      apply (subst rep0)
3304      apply (simp only: take_replicate drop_replicate)
3305      apply (rule arg_cong [where f = "\<lambda>x. replicate x 0"])
3306      apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_defs size_of_def objBits_defs)
3307      done
3308    finally show "heap_list (fst (t_hrs_' (globals x))) (size_of TYPE(tcb_C)) (ptr_val p) = replicate (size_of TYPE(tcb_C)) 0" .
3309  qed
3310
3311  note alrl = pspace_aligned_to_C_tcb [OF pal cmap_relation_tcb [OF rfsr]]
3312
3313  have tdisj:
3314    "\<forall>xa\<in>dom (cslift x) \<union> {p}. \<forall>y\<in>dom (cslift x). {ptr_val xa..+size_of TYPE(tcb_C)} \<inter> {ptr_val y..+size_of TYPE(tcb_C)} \<noteq> {}
3315           \<longrightarrow> xa = y"
3316    using al
3317    apply (intro ballI impI)
3318    apply (erule contrapos_np)
3319    apply (subgoal_tac "is_aligned (ptr_val xa) ctcb_size_bits")
3320     apply (subgoal_tac "is_aligned (ptr_val y) ctcb_size_bits")
3321      apply (subgoal_tac "ctcb_size_bits < word_bits")
3322       apply (rule_tac A = "{ptr_val xa..+2 ^ ctcb_size_bits}" in disjoint_subset)
3323        apply (rule intvl_start_le)
3324        apply (simp add: size_of_def ctcb_size_bits_def)
3325       apply (rule_tac B = "{ptr_val y..+2 ^ ctcb_size_bits}" in disjoint_subset2)
3326        apply (rule intvl_start_le)
3327        apply (simp add: size_of_def ctcb_size_bits_def)
3328       apply (simp only: upto_intvl_eq)
3329       apply (rule aligned_neq_into_no_overlap [simplified field_simps])
3330         apply simp
3331        apply assumption+
3332      apply (simp add: word_bits_conv ctcb_size_bits_def)
3333     apply (erule bspec [OF alrl])
3334    apply (clarsimp)
3335    apply (erule disjE)
3336     apply (simp add: objBits_simps kotcb_def)
3337     apply (erule ctcb_ptr_to_tcb_ptr_aligned)
3338    apply (erule bspec [OF alrl])
3339    done
3340
3341  let ?new_tcb =  "(from_bytes (replicate (size_of TYPE(tcb_C)) 0)
3342                  \<lparr>tcbArch_C := tcbArch_C (from_bytes (replicate (size_of TYPE(tcb_C)) 0))
3343                    \<lparr>tcbContext_C := tcbContext_C (tcbArch_C (from_bytes (replicate (size_of TYPE(tcb_C)) 0)))
3344                     \<lparr>registers_C :=
3345                        Arrays.update (registers_C (tcbContext_C (tcbArch_C (from_bytes (replicate (size_of TYPE(tcb_C)) 0))))) (unat Kernel_C.CPSR)
3346                         0x150\<rparr>\<rparr>, tcbTimeSlice_C := 5\<rparr>)"
3347
3348  have tdisj':
3349    "\<And>y. hrs_htd (t_hrs_' (globals x)) \<Turnstile>\<^sub>t y \<Longrightarrow> ptr_span p \<inter> ptr_span y \<noteq> {} \<Longrightarrow> y = p"
3350    using tdisj by (auto simp: h_t_valid_clift_Some_iff)
3351
3352  have "ptr_retyp p (snd (t_hrs_' (globals x))) \<Turnstile>\<^sub>t p" using cgp
3353    by (rule ptr_retyp_h_t_valid)
3354  hence "clift (hrs_mem (t_hrs_' (globals x)), ptr_retyp p (snd (t_hrs_' (globals x)))) p
3355    = Some (from_bytes (replicate (size_of TYPE(tcb_C)) 0))"
3356    by (simp add: lift_t_if h_val_def tcb0 hrs_mem_def)
3357  hence cl_tcb: "(cslift (x\<lparr>globals := ?gs\<rparr>) :: tcb_C typ_heap) = (cslift x)(p \<mapsto> ?new_tcb)"
3358    using cgp
3359    apply (clarsimp simp add: typ_heap_simps
3360                              hrs_mem_update packed_heap_update_collapse_hrs)
3361    apply (simp add: hrs_comm[symmetric])
3362    apply (subst clift_ptr_retyps_gen_other, simp_all add: hrs_htd_update
3363      empty_smaller[simplified] tag_disj_via_td_name)
3364     apply (simp add: cte_C_size word_bits_def)
3365    apply (simp add: hrs_comm typ_heap_simps ptr_retyps_gen_def
3366                     hrs_htd_update ptr_retyp_h_t_valid
3367                     h_val_heap_update)
3368    apply (simp add: h_val_field_from_bytes)
3369    apply (simp add: h_val_def tcb0[folded hrs_mem_def])
3370    apply (rule ext, rename_tac p')
3371    apply (case_tac "p' = p", simp_all)
3372    apply (cut_tac clift_ptr_retyps_gen_prev_memset_same[where n=1 and arr=False, simplified,
3373      OF _ empty_smaller(1) _ refl], simp_all add: tcb0[folded hrs_mem_def])
3374     apply (simp add: ptr_retyps_gen_def)
3375    apply (simp add: tcb_C_size word_bits_def)
3376    done
3377
3378  have cl_rest:
3379    "\<lbrakk>typ_uinfo_t TYPE(tcb_C) \<bottom>\<^sub>t typ_uinfo_t TYPE('a :: mem_type);
3380      typ_uinfo_t TYPE(cte_C[5]) \<bottom>\<^sub>t typ_uinfo_t TYPE('a :: mem_type);
3381      typ_uinfo_t TYPE('a) \<noteq> typ_uinfo_t TYPE(word8) \<rbrakk> \<Longrightarrow>
3382    cslift (x\<lparr>globals := ?gs\<rparr>) = (cslift x :: 'a :: mem_type typ_heap)"
3383    using cgp
3384    apply (clarsimp simp: hrs_comm[symmetric])
3385    apply (subst clift_ptr_retyps_gen_other,
3386      simp_all add: hrs_htd_update empty_smaller[simplified],
3387      simp_all add: cte_C_size tcb_C_size word_bits_def)
3388    apply (simp add: hrs_comm ptr_retyps_gen_def)
3389    apply (simp add: clift_heap_update_same hrs_htd_update ptr_retyp_h_t_valid typ_heap_simps)
3390    apply (rule trans[OF _ clift_ptr_retyps_gen_other[where nptrs=1 and arr=False,
3391        simplified, OF empty_smaller(1)]], simp_all)
3392     apply (simp add: ptr_retyps_gen_def)
3393    apply (simp add: tcb_C_size word_bits_def)
3394    done
3395
3396  have rl:
3397    "(\<forall>v :: 'a :: pre_storable. projectKO_opt kotcb \<noteq> Some v) \<Longrightarrow>
3398    (projectKO_opt \<circ>\<^sub>m (ks(ctcb_ptr_to_tcb_ptr p \<mapsto> KOTCB makeObject)) :: word32 \<Rightarrow> 'a option)
3399    = projectKO_opt \<circ>\<^sub>m ks" using pno al
3400    apply -
3401    apply (drule(2) projectKO_opt_retyp_other'[OF _ _ pal])
3402    apply (simp add: kotcb_def)
3403    done
3404
3405  have rl_tcb: "(projectKO_opt \<circ>\<^sub>m (ks(ctcb_ptr_to_tcb_ptr p \<mapsto> KOTCB makeObject)) :: word32 \<Rightarrow> tcb option)
3406    = (projectKO_opt \<circ>\<^sub>m ks)(ctcb_ptr_to_tcb_ptr p \<mapsto> makeObject)"
3407    apply (rule ext)
3408    apply (clarsimp simp: projectKOs map_comp_def split: if_split)
3409    done
3410
3411  have mko: "\<And>dev. makeObjectKO dev (Inr (APIObjectType ArchTypes_H.apiobject_type.TCBObject)) = Some kotcb"
3412    by (simp add: makeObjectKO_def kotcb_def)
3413  note hacky_cte = retype_ctes_helper [where sz = "objBitsKO kotcb" and ko = kotcb and ptr = "ctcb_ptr_to_tcb_ptr p",
3414    OF pal pds pno al _ _ mko, simplified new_cap_addrs_def, simplified]
3415
3416  \<comment> \<open>Ugh\<close>
3417  moreover have
3418    "\<And>y. y \<in> ptr_val ` (CTypesDefs.ptr_add (cte_Ptr (ctcb_ptr_to_tcb_ptr p)) \<circ> of_nat) ` {k. k < 5}
3419      = (y && ~~ mask tcbBlockSizeBits = ctcb_ptr_to_tcb_ptr p \<and> y && mask tcbBlockSizeBits \<in> dom tcb_cte_cases)"
3420    (is "\<And>y. ?LHS y = ?RHS y")
3421  proof -
3422    fix y
3423
3424    have al_rl: "\<And>k. k < 5 \<Longrightarrow>
3425      ctcb_ptr_to_tcb_ptr p + of_nat k * of_nat (size_of TYPE(cte_C)) && mask tcbBlockSizeBits = of_nat k * of_nat (size_of TYPE(cte_C))
3426      \<and> ctcb_ptr_to_tcb_ptr p + of_nat k * of_nat (size_of TYPE(cte_C)) && ~~ mask tcbBlockSizeBits = ctcb_ptr_to_tcb_ptr p" using al
3427      apply -
3428      apply (rule is_aligned_add_helper)
3429      apply (simp add: objBits_simps kotcb_def)
3430       apply (subst Abs_fnat_hom_mult)
3431       apply (subst word_less_nat_alt)
3432       apply (subst unat_of_nat32)
3433       apply (simp add: size_of_def word_bits_conv objBits_defs)+
3434      done
3435
3436    have al_rl2: "\<And>k. k < 5 \<Longrightarrow> unat (of_nat k * of_nat (size_of TYPE(cte_C)) :: word32) = k * 2^cteSizeBits"
3437       apply (subst Abs_fnat_hom_mult)
3438       apply (subst unat_of_nat32)
3439       apply (simp add: size_of_def word_bits_conv objBits_defs)+
3440       done
3441
3442    show "?LHS y = ?RHS y" using al
3443      apply (simp add: image_image kotcb_def objBits_simps)
3444      apply rule
3445       apply (clarsimp simp: dom_tcb_cte_cases_iff al_rl al_rl2)
3446       apply (simp add: objBits_defs)
3447      apply (clarsimp simp: dom_tcb_cte_cases_iff al_rl al_rl2)
3448      apply (rule_tac x = ya in image_eqI)
3449       apply (rule mask_eqI [where n = tcbBlockSizeBits])
3450        apply (subst unat_arith_simps(3))
3451      apply (simp add: al_rl al_rl2, simp add: objBits_defs)+
3452      done
3453  qed
3454
3455  ultimately have rl_cte: "(map_to_ctes (ks(ctcb_ptr_to_tcb_ptr p \<mapsto> KOTCB makeObject)) :: word32 \<Rightarrow> cte option)
3456    = (\<lambda>x. if x \<in> ptr_val ` (CTypesDefs.ptr_add (cte_Ptr (ctcb_ptr_to_tcb_ptr p)) \<circ> of_nat) ` {k. k < 5}
3457         then Some (CTE NullCap nullMDBNode)
3458         else map_to_ctes ks x)"
3459    apply simp
3460    apply (drule_tac x = "Suc 0" in meta_spec)
3461    apply clarsimp
3462    apply (erule impE[OF impI])
3463     apply (rule range_cover_full[OF al])
3464     apply (simp add: objBits_simps' word_bits_conv machine_bits_defs archObjSize_def
3465       split:kernel_object.splits arch_kernel_object.splits)
3466    apply (simp add: fun_upd_def kotcb_def cong: if_cong)
3467    done
3468
3469  let ?tcb = "undefined
3470    \<lparr>tcbArch_C := tcbArch_C undefined
3471      \<lparr>tcbContext_C := tcbContext_C (tcbArch_C undefined)
3472         \<lparr>registers_C :=
3473           foldr (\<lambda>n arr. Arrays.update arr n 0) [0..<20]
3474             (registers_C (tcbContext_C (tcbArch_C undefined)))
3475         \<rparr>,
3476       tcbVCPU_C := vcpu_Ptr 0
3477     \<rparr>,
3478       tcbState_C :=
3479         thread_state_C.words_C_update
3480          (\<lambda>_. foldr (\<lambda>n arr. Arrays.update arr n 0) [0..<3]
3481                (thread_state_C.words_C (tcbState_C undefined)))
3482          (tcbState_C undefined),
3483       tcbFault_C :=
3484         seL4_Fault_C.words_C_update
3485          (\<lambda>_. foldr (\<lambda>n arr. Arrays.update arr n 0) [0..<2]
3486                (seL4_Fault_C.words_C (tcbFault_C undefined)))
3487          (tcbFault_C undefined),
3488       tcbLookupFailure_C :=
3489         lookup_fault_C.words_C_update
3490          (\<lambda>_. foldr (\<lambda>n arr. Arrays.update arr n 0) [0..<2]
3491                (lookup_fault_C.words_C (tcbLookupFailure_C undefined)))
3492          (tcbLookupFailure_C undefined),
3493       tcbPriority_C := 0, tcbMCP_C := 0, tcbDomain_C := 0, tcbTimeSlice_C := 0,
3494       tcbFaultHandler_C := 0, tcbIPCBuffer_C := 0,
3495       tcbSchedNext_C := tcb_Ptr 0, tcbSchedPrev_C := tcb_Ptr 0,
3496       tcbEPNext_C := tcb_Ptr 0, tcbEPPrev_C := tcb_Ptr 0,
3497       tcbBoundNotification_C := ntfn_Ptr 0\<rparr>"
3498  have fbtcb: "from_bytes (replicate (size_of TYPE(tcb_C)) 0) = ?tcb"
3499    apply (simp add: from_bytes_def)
3500    apply (simp add: typ_info_simps tcb_C_tag_def)
3501    apply (simp add: ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td align_of_def padup_def
3502      final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def)(* takes ages *)
3503    apply (simp add: update_ti_adjust_ti update_ti_t_word32_0s
3504      typ_info_simps
3505      user_context_C_tag_def thread_state_C_tag_def seL4_Fault_C_tag_def
3506      lookup_fault_C_tag_def update_ti_t_ptr_0s arch_tcb_C_tag_def
3507      ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td
3508      ti_typ_combine_empty_ti ti_typ_combine_td
3509      align_of_def padup_def
3510      final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def
3511      align_td_array' size_td_array)
3512    apply (simp add: update_ti_t_array_rep_word0)
3513    done
3514
3515  have tcb_rel:
3516    "ctcb_relation makeObject ?new_tcb"
3517    unfolding ctcb_relation_def makeObject_tcb
3518    apply (simp add: fbtcb minBound_word)
3519    apply (intro conjI)
3520        apply (simp add: cthread_state_relation_def thread_state_lift_def
3521                         eval_nat_numeral ThreadState_Inactive_def)
3522       apply (clarsimp simp: ccontext_relation_def newContext_def2 carch_tcb_relation_def
3523                             newArchTCB_def)
3524       apply (case_tac r,
3525              simp_all add: "StrictC'_register_defs" eval_nat_numeral
3526                            atcbContext_def newArchTCB_def newContext_def
3527                            initContext_def)[1] \<comment> \<open>takes ages\<close>
3528                         apply (simp add: thread_state_lift_def eval_nat_numeral atcbContextGet_def)+
3529     apply (simp add: timeSlice_def)
3530    apply (simp add: cfault_rel_def seL4_Fault_lift_def seL4_Fault_get_tag_def Let_def
3531        lookup_fault_lift_def lookup_fault_get_tag_def lookup_fault_invalid_root_def
3532        eval_nat_numeral seL4_Fault_NullFault_def option_to_ptr_def option_to_0_def
3533        split: if_split)+
3534    done
3535
3536  have pks: "ks (ctcb_ptr_to_tcb_ptr p) = None"
3537    by (rule pspace_no_overlap_base' [OF pal pno al, simplified])
3538
3539  have ep1 [simplified]: "\<And>p' list. map_to_eps (ksPSpace ?sp) p' = Some (Structures_H.endpoint.RecvEP list)
3540       \<Longrightarrow> ctcb_ptr_to_tcb_ptr p \<notin> set list"
3541    using symref pks pal pds
3542    apply -
3543    apply (frule map_to_ko_atI2)
3544      apply simp
3545     apply simp
3546    apply (drule (1) sym_refs_ko_atD')
3547    apply clarsimp
3548    apply (drule (1) bspec)
3549    apply (simp add: ko_wp_at'_def)
3550    done
3551
3552  have ep2 [simplified]: "\<And>p' list. map_to_eps (ksPSpace ?sp) p' = Some (Structures_H.endpoint.SendEP list)
3553       \<Longrightarrow> ctcb_ptr_to_tcb_ptr p \<notin> set list"
3554    using symref pks pal pds
3555    apply -
3556    apply (frule map_to_ko_atI2)
3557      apply simp
3558     apply simp
3559    apply (drule (1) sym_refs_ko_atD')
3560    apply clarsimp
3561    apply (drule (1) bspec)
3562    apply (simp add: ko_wp_at'_def)
3563    done
3564
3565  have ep3 [simplified]: "\<And>p' list boundTCB. map_to_ntfns (ksPSpace ?sp) p' = Some (Structures_H.notification.NTFN (Structures_H.ntfn.WaitingNtfn list) boundTCB)
3566       \<Longrightarrow> ctcb_ptr_to_tcb_ptr p \<notin> set list"
3567    using symref pks pal pds
3568    apply -
3569    apply (frule map_to_ko_atI2)
3570      apply simp
3571     apply simp
3572    apply (drule (1) sym_refs_ko_atD')
3573    apply clarsimp
3574    apply (drule_tac x="(ctcb_ptr_to_tcb_ptr p, NTFNSignal)" in bspec, simp)
3575    apply (simp add: ko_wp_at'_def)
3576    done
3577
3578  have pks': "ksPSpace \<sigma> (ctcb_ptr_to_tcb_ptr p) = None" using pks kssub
3579    apply -
3580    apply (erule contrapos_pp)
3581    apply (fastforce simp: dom_def)
3582    done
3583
3584  hence kstcb: "\<And>qdom prio. ctcb_ptr_to_tcb_ptr p \<notin> set (ksReadyQueues \<sigma> (qdom, prio))" using vq
3585    apply (clarsimp simp add: valid_queues_def valid_queues_no_bitmap_def)
3586    apply (drule_tac x = qdom in spec)
3587    apply (drule_tac x = prio in spec)
3588    apply clarsimp
3589    apply (drule (1) bspec)
3590    apply (simp add: obj_at'_def)
3591    done
3592
3593  have ball_subsetE:
3594    "\<And>P S R. \<lbrakk> \<forall>x \<in> S. P x; R \<subseteq> S \<rbrakk> \<Longrightarrow> \<forall>x \<in> R. P x"
3595    by blast
3596
3597  have htd_safe:
3598    "htd_safe (- kernel_data_refs) (hrs_htd (t_hrs_' (globals x)))
3599        \<Longrightarrow> htd_safe (- kernel_data_refs) (hrs_htd (t_hrs_' ?gs))"
3600    using kdr
3601    apply (simp add: hrs_htd_update)
3602    apply (intro ptr_retyp_htd_safe_neg ptr_retyps_htd_safe_neg, simp_all)
3603     apply (erule disjoint_subset[rotated])
3604     apply (simp add: ctcb_ptr_to_tcb_ptr_def size_of_def)
3605     apply (rule intvl_sub_offset[where k="ptr_val p - ctcb_offset" and x="ctcb_offset", simplified])
3606     apply (simp add: ctcb_offset_defs objBits_defs)
3607    apply (erule disjoint_subset[rotated])
3608    apply (rule intvl_start_le)
3609    apply (simp add: size_of_def objBits_defs)
3610    done
3611
3612  have zro:
3613    "zero_ranges_are_zero (gsUntypedZeroRanges \<sigma>) (t_hrs_' (globals x))"
3614    using rfsr
3615    by (clarsimp simp: rf_sr_def cstate_relation_def Let_def)
3616
3617  have h_t_valid_p:
3618    "h_t_valid (hrs_htd (t_hrs_' ?gs)) c_guard p"
3619    using fun_cong[OF cl_tcb, where x=p]
3620    by (clarsimp dest!: h_t_valid_clift)
3621
3622  have zro':
3623    "zero_ranges_are_zero (gsUntypedZeroRanges \<sigma>) (t_hrs_' ?gs)"
3624    using zro h_t_valid_p rzo al
3625    apply clarsimp
3626    apply (simp add: hrs_htd_update typ_heap_simps')
3627    apply (intro zero_ranges_ptr_retyps, simp_all)
3628     apply (erule caps_overlap_reserved'_subseteq)
3629     apply (rule order_trans, rule tcb_ptr_to_ctcb_ptr_in_range')
3630      apply (simp add: objBits_simps kotcb_def)
3631     apply (simp add: objBits_simps kotcb_def)
3632    apply (erule caps_overlap_reserved'_subseteq)
3633    apply (rule intvl_start_le)
3634    apply (simp add: cte_C_size kotcb_def objBits_simps')
3635    done
3636
3637  note ht_rest = clift_eq_h_t_valid_eq[OF cl_rest, simplified]
3638
3639  note irq = h_t_valid_eq_array_valid[where 'a=cte_C and p="ptr_coerce x" for x]
3640    h_t_array_valid_ptr_retyps_gen[where n=1, simplified, OF refl empty_smaller(1)]
3641    h_t_array_valid_ptr_retyps_gen[where p="Ptr x" for x, simplified, OF refl empty_smaller(2)]
3642
3643  from rfsr have "cpspace_relation ks (underlying_memory (ksMachineState \<sigma>)) (t_hrs_' (globals x))"
3644    unfolding rf_sr_def cstate_relation_def by (simp add: Let_def)
3645  hence "cpspace_relation ?ks (underlying_memory (ksMachineState \<sigma>))  (t_hrs_' ?gs)"
3646    unfolding cpspace_relation_def
3647    apply -
3648    apply (simp add: cl_cte [simplified] cl_tcb [simplified] cl_rest [simplified] tag_disj_via_td_name
3649                     ht_rest)
3650    apply (simp add: rl kotcb_def projectKOs rl_tcb rl_cte)
3651    apply (elim conjE)
3652    apply (intro conjI)
3653     \<comment> \<open>cte\<close>
3654     apply (erule cmap_relation_retype2)
3655     apply (simp add:ccte_relation_nullCap nullMDBNode_def nullPointer_def)
3656    \<comment> \<open>tcb\<close>
3657     apply (erule cmap_relation_updI2 [where dest = "ctcb_ptr_to_tcb_ptr p" and f = "tcb_ptr_to_ctcb_ptr", simplified])
3658     apply (rule map_comp_simps)
3659     apply (rule pks)
3660     apply (rule tcb_rel)
3661    \<comment> \<open>ep\<close>
3662     apply (erule iffD2 [OF cmap_relation_cong, OF refl refl, rotated -1])
3663     apply (simp add: cendpoint_relation_def Let_def)
3664     apply (subst endpoint.case_cong)
3665       apply (rule refl)
3666      apply (simp add: tcb_queue_update_other' ep1)
3667     apply (simp add: tcb_queue_update_other' del: tcb_queue_relation'_empty)
3668    apply (simp add: tcb_queue_update_other' ep2)
3669   apply clarsimp
3670  \<comment> \<open>ntfn\<close>
3671   apply (erule iffD2 [OF cmap_relation_cong, OF refl refl, rotated -1])
3672   apply (simp add: cnotification_relation_def Let_def)
3673     apply (subst ntfn.case_cong)
3674      apply (rule refl)
3675     apply (simp add: tcb_queue_update_other' del: tcb_queue_relation'_empty)
3676    apply (simp add: tcb_queue_update_other' del: tcb_queue_relation'_empty)
3677   apply (case_tac a, simp add: tcb_queue_update_other' ep3)
3678  apply (clarsimp simp: typ_heap_simps)
3679  done
3680
3681  moreover have "cte_array_relation \<sigma> ?gs
3682    \<and> tcb_cte_array_relation ?s ?gs"
3683    using rfsr
3684    apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
3685                          hrs_htd_update map_comp_update
3686                          kotcb_def projectKO_opt_tcb)
3687    apply (intro cvariable_array_ptr_upd conjI
3688                 cvariable_array_ptr_retyps[OF refl, where n=1, simplified],
3689           simp_all add: empty_smaller[simplified])
3690    apply (simp add: ptr_retyps_gen_def)
3691    apply (rule ptr_retyp_h_t_valid[where g=c_guard, OF arr_guard,
3692        THEN h_t_array_valid, simplified])
3693    done
3694
3695  ultimately show ?thesis
3696    using rfsr zro'
3697    apply (simp add: rf_sr_def cstate_relation_def Let_def h_t_valid_clift_Some_iff
3698      tag_disj_via_td_name carch_state_relation_def cmachine_state_relation_def irq)
3699    apply (simp add: cl_cte [simplified] cl_tcb [simplified] cl_rest [simplified] tag_disj_via_td_name)
3700    apply (clarsimp simp add: cready_queues_relation_def Let_def
3701                              htd_safe[simplified] kernel_data_refs_domain_eq_rotate)
3702    apply (simp add: kstcb tcb_queue_update_other' hrs_htd_update
3703                     ptr_retyp_to_array[simplified] irq[simplified])
3704    done
3705qed
3706
3707
3708lemma cnc_foldl_foldr:
3709  defines "ko \<equiv> (KOTCB makeObject)"
3710  shows "foldl (\<lambda>v addr. v(addr \<mapsto> ko)) mp
3711  (map (\<lambda>n. ptr + (of_nat n << tcbBlockSizeBits)) [0..< n]) =
3712  foldr (\<lambda>addr. data_map_insert addr ko) (new_cap_addrs n ptr ko) mp"
3713  by (simp add: foldr_upd_app_if foldl_conv_foldr
3714                new_cap_addrs_def objBits_simps ko_def power_minus_is_div
3715          cong: foldr_cong)
3716
3717lemma objBitsKO_gt_0:
3718  "0 < objBitsKO ko"
3719  by (simp add: objBits_simps' archObjSize_def machine_bits_defs
3720         split: kernel_object.splits arch_kernel_object.splits)
3721
3722lemma objBitsKO_gt_1:
3723  "(1 :: word32) < 2 ^ objBitsKO ko"
3724  by (simp add: objBits_simps' archObjSize_def machine_bits_defs
3725         split: kernel_object.splits arch_kernel_object.splits)
3726
3727lemma ps_clear_subset:
3728  assumes pd: "ps_clear x (objBitsKO ko) (s' \<lparr>ksPSpace := (\<lambda>x. if x \<in> as then Some (f x) else ksPSpace s' x) \<rparr>)"
3729  and    sub: "as' \<subseteq> as"
3730  and     al: "is_aligned x (objBitsKO ko)"
3731  shows  "ps_clear x (objBitsKO ko) (s' \<lparr>ksPSpace := (\<lambda>x. if x \<in> as' then Some (f x) else ksPSpace s' x) \<rparr>)"
3732  using al pd sub
3733  apply -
3734  apply (simp add: ps_clear_def3 [OF al  objBitsKO_gt_0] dom_if_Some)
3735  apply (erule disjoint_subset2 [rotated])
3736  apply fastforce
3737  done
3738
3739lemma pspace_distinct_subset:
3740  assumes pd: "pspace_distinct' (s' \<lparr>ksPSpace := (\<lambda>x. if x \<in> as then Some (f x) else ksPSpace s' x) \<rparr>)"
3741  and   pal: "pspace_aligned' (s' \<lparr>ksPSpace := (\<lambda>x. if x \<in> as then Some (f x) else ksPSpace s' x) \<rparr>)"
3742  and    sub: "as' \<subseteq> as"
3743  and  doms: "as \<inter> dom (ksPSpace s') = {}"
3744  shows  "pspace_distinct' (s' \<lparr>ksPSpace := (\<lambda>x. if x \<in> as' then Some (f x) else ksPSpace s' x) \<rparr>)"
3745  using pd sub doms pal
3746  unfolding pspace_distinct'_def pspace_aligned'_def
3747  apply -
3748  apply (rule ballI)
3749  apply (simp add: pspace_distinct'_def dom_if_Some)
3750  apply (drule_tac x = x in bspec)
3751   apply fastforce
3752  apply (drule_tac x = x in bspec)
3753   apply fastforce
3754  apply (erule disjE)
3755   apply (frule (1) subsetD)
3756   apply simp
3757   apply (erule (2) ps_clear_subset)
3758  apply (subgoal_tac "x \<notin> as")
3759   apply (frule (1) contra_subsetD)
3760   apply simp
3761   apply (erule (2) ps_clear_subset)
3762  apply fastforce
3763  done
3764
3765lemma pspace_aligned_subset:
3766  assumes pal: "pspace_aligned' (s' \<lparr>ksPSpace := (\<lambda>x. if x \<in> as then Some (f x) else ksPSpace s' x) \<rparr>)"
3767  and     sub: "as' \<subseteq> as"
3768  and    doms: "as \<inter> dom (ksPSpace s') = {}"
3769  shows  "pspace_aligned' (s' \<lparr>ksPSpace := (\<lambda>x. if x \<in> as' then Some (f x) else ksPSpace s' x) \<rparr>)"
3770  using pal sub doms unfolding pspace_aligned'_def
3771  apply -
3772  apply (rule ballI)
3773  apply (simp add: dom_if_Some)
3774  apply (drule_tac x = x in bspec)
3775   apply fastforce
3776  apply (erule disjE)
3777   apply simp
3778   apply (frule (1) subsetD)
3779   apply simp
3780  apply (subgoal_tac "x \<notin> as")
3781   apply (frule (1) contra_subsetD)
3782   apply simp
3783  apply fastforce
3784  done
3785
3786
3787lemma cslift_empty_mem_update:
3788  fixes x :: cstate and sz and ptr
3789  defines "x' \<equiv> x\<lparr>globals := globals x
3790                       \<lparr>t_hrs_' := hrs_mem_update (heap_update_list ptr (replicate sz 0)) (t_hrs_' (globals x))\<rparr>\<rparr>"
3791  assumes empty: "region_is_typeless ptr sz x"
3792  shows "cslift x' = clift (fst (t_hrs_' (globals x)), snd (t_hrs_' (globals x)))"
3793  using empty
3794  apply -
3795  apply (unfold region_is_typeless_def)
3796  apply (rule ext)
3797  apply (simp only: lift_t_if hrs_mem_update_def split_def x'_def)
3798  apply (simp add: lift_t_if hrs_mem_update_def split_def)
3799  apply (clarsimp simp: h_val_def split: if_split)
3800  apply (subst heap_list_update_disjoint_same)
3801   apply simp
3802   apply (rule disjointI)
3803   apply clarsimp
3804   apply (drule (1) bspec)
3805   apply (frule (1) h_t_valid_not_empty)
3806   apply simp
3807  apply simp
3808  done
3809
3810lemma cslift_bytes_mem_update:
3811  fixes x :: cstate and sz and ptr
3812  defines "x' \<equiv> x\<lparr>globals := globals x
3813                       \<lparr>t_hrs_' := hrs_mem_update (heap_update_list ptr (replicate sz 0)) (t_hrs_' (globals x))\<rparr>\<rparr>"
3814  assumes bytes: "region_is_bytes ptr sz x"
3815  assumes not_byte: "typ_uinfo_t TYPE ('a) \<noteq> typ_uinfo_t TYPE (word8)"
3816  shows "(cslift x' :: ('a :: mem_type) ptr \<Rightarrow> _)
3817     = clift (fst (t_hrs_' (globals x)), snd (t_hrs_' (globals x)))"
3818  using bytes
3819  apply (unfold region_is_bytes'_def)
3820  apply (rule ext)
3821  apply (simp only: lift_t_if hrs_mem_update_def split_def x'_def)
3822  apply (simp add: lift_t_if hrs_mem_update_def split_def)
3823  apply (clarsimp simp: h_val_def split: if_split)
3824  apply (subst heap_list_update_disjoint_same)
3825   apply simp
3826   apply (rule disjointI)
3827   apply clarsimp
3828   apply (drule (1) bspec)
3829   apply (frule (1) h_t_valid_intvl_htd_contains_uinfo_t)
3830   apply (clarsimp simp: hrs_htd_def not_byte)
3831  apply simp
3832  done
3833
3834lemma heap_list_eq_replicate_eq_eq:
3835  "(heap_list hp n ptr = replicate n v)
3836    = (\<forall>p \<in> {ptr ..+ n}. hp p = v)"
3837  by (induct n arbitrary: ptr, simp_all add: intvl_Suc_right)
3838
3839lemma heap_update_list_replicate_eq:
3840  "(heap_update_list x (replicate n v) hp y)
3841    = (if y \<in> {x ..+ n} then v else hp y)"
3842  apply (induct n arbitrary: x hp, simp_all add: intvl_Suc_right)
3843  apply (simp split: if_split)
3844  done
3845
3846lemma zero_ranges_are_zero_update_zero[simp]:
3847  "zero_ranges_are_zero rs hrs
3848    \<Longrightarrow> zero_ranges_are_zero rs (hrs_mem_update (heap_update_list ptr (replicate n 0)) hrs)"
3849  apply (clarsimp simp: zero_ranges_are_zero_def hrs_mem_update)
3850  apply (drule(1) bspec)
3851  apply (clarsimp simp: heap_list_eq_replicate_eq_eq heap_update_list_replicate_eq)
3852  done
3853
3854lemma rf_sr_rep0:
3855  assumes sr: "(\<sigma>, x) \<in> rf_sr"
3856  assumes empty: "region_is_bytes ptr sz x"
3857  shows "(\<sigma>, globals_update (t_hrs_'_update (hrs_mem_update (heap_update_list ptr (replicate sz 0)))) x) \<in> rf_sr"
3858  using sr
3859  by (clarsimp simp add: rf_sr_def cstate_relation_def Let_def cpspace_relation_def
3860        carch_state_relation_def cmachine_state_relation_def
3861        cslift_bytes_mem_update[OF empty, simplified] cte_C_size)
3862
3863
3864(* FIXME: generalise *)
3865lemma ccorres_already_have_rrel:
3866  "\<lbrakk> ccorres dc xfdc P P' hs a c; \<forall>s. \<Gamma> \<turnstile>\<^bsub>/UNIV\<^esub> {s} c {t. xf t = xf s} \<rbrakk>
3867  \<Longrightarrow>
3868  ccorres r xf P (P' \<inter> {s. r v (xf s)}) hs (a >>= (\<lambda>_.  return v)) c"
3869  apply (rule ccorres_return_into_rel)
3870  apply (rule ccorresI')
3871  apply (erule (2) ccorresE)
3872     apply simp
3873    apply assumption+
3874  apply (clarsimp elim!: rev_bexI)
3875  apply (simp add: unif_rrel_def)
3876  apply (drule_tac x = s' in spec)
3877  apply (drule (1) exec_handlers_use_hoare_nothrow)
3878   apply simp
3879  apply fastforce
3880  done
3881
3882lemma mapM_x_storeWord:
3883  assumes al: "is_aligned ptr 2"
3884  shows "mapM_x (\<lambda>x. storeWord (ptr + of_nat x * 4) 0) [0..<n]
3885  = modify (underlying_memory_update (\<lambda>m x. if x \<in> {ptr..+ n * 4} then 0 else m x))"
3886proof (induct n)
3887  case 0
3888  thus ?case
3889    apply (rule ext)
3890    apply (simp add: mapM_x_mapM mapM_def sequence_def
3891      modify_def get_def put_def bind_def return_def)
3892    done
3893next
3894  case (Suc n')
3895
3896  have funs_eq:
3897    "\<And>m x. (if x \<in> {ptr..+4 + n' * 4} then 0 else (m x :: word8)) =
3898           ((\<lambda>xa. if xa \<in> {ptr..+n' * 4} then 0 else m xa)
3899           (ptr + of_nat n' * 4 := word_rsplit (0 :: word32) ! 3,
3900            ptr + of_nat n' * 4 + 1 := word_rsplit (0 :: word32) ! 2,
3901            ptr + of_nat n' * 4 + 2 := word_rsplit (0 :: word32) ! Suc 0,
3902            ptr + of_nat n' * 4 + 3 := word_rsplit (0 :: word32) ! 0)) x"
3903  proof -
3904    fix m x
3905
3906    have xin': "\<And>x. (x < 4 + n' * 4) = (x < n' * 4 \<or> x = n' * 4
3907                     \<or> x = (n' * 4) + 1 \<or> x = (n' * 4) + 2 \<or> x = (n' * 4) + 3)"
3908      by (safe, simp_all)
3909
3910    have xin: "x \<in> {ptr..+4 + n' * 4} = (x \<in> {ptr..+n' * 4} \<or> x = ptr + of_nat n' * 4 \<or>
3911      x = ptr + of_nat n' * 4 + 1 \<or> x = ptr + of_nat n' * 4 + 2 \<or> x = ptr + of_nat n' * 4 + 3)"
3912      by (simp add: intvl_def xin' conj_disj_distribL
3913                    ex_disj_distrib field_simps)
3914
3915    show "?thesis m x"
3916      apply (simp add: xin word_rsplit_0 cong: if_cong)
3917      apply (simp split: if_split)
3918      done
3919  qed
3920
3921  from al have "is_aligned (ptr + of_nat n' * 4) 2"
3922    apply (rule aligned_add_aligned)
3923    apply (rule is_aligned_mult_triv2 [where n = 2, simplified])
3924    apply (simp add: word_bits_conv)+
3925    done
3926
3927  thus ?case
3928    apply (simp add: mapM_x_append bind_assoc Suc.hyps mapM_x_singleton)
3929    apply (simp add: storeWord_def assert_def is_aligned_mask modify_modify comp_def)
3930    apply (simp only: funs_eq)
3931    done
3932qed
3933
3934lemma mapM_x_storeWord_step:
3935  assumes al: "is_aligned ptr sz"
3936  and    sz2: "2 \<le> sz"
3937  and     sz: "sz < word_bits"
3938  shows "mapM_x (\<lambda>p. storeWord p 0) [ptr , ptr + 4 .e. ptr + 2 ^ sz - 1] =
3939  modify (underlying_memory_update (\<lambda>m x. if x \<in> {ptr..+2 ^ (sz - 2) * 4} then 0 else m x))"
3940  using al sz
3941  apply (simp only: upto_enum_step_def field_simps cong: if_cong)
3942  apply (subst if_not_P)
3943   apply (subst not_less)
3944   apply (erule is_aligned_no_overflow)
3945   apply (simp add: mapM_x_map comp_def upto_enum_word del: upt.simps)
3946   apply (subst div_power_helper_32 [OF sz2, simplified])
3947    apply assumption
3948   apply (simp add: word_bits_def unat_minus_one del: upt.simps)
3949   apply (subst mapM_x_storeWord)
3950   apply (erule is_aligned_weaken [OF _ sz2])
3951   apply (simp add: field_simps)
3952   done
3953
3954
3955lemma pspace_aligned_to_C_user_data:
3956  fixes v :: "user_data"
3957  assumes pal: "pspace_aligned' s"
3958  and    cmap: "cpspace_user_data_relation (ksPSpace s) (underlying_memory (ksMachineState s)) (t_hrs_' (globals x))"
3959  shows  "\<forall>x\<in>dom (cslift x :: user_data_C typ_heap). is_aligned (ptr_val x) (objBitsKO KOUserData)"
3960  (is "\<forall>x\<in>dom ?CS. is_aligned (ptr_val x) (objBitsKO KOUserData)")
3961proof
3962  fix z
3963  assume "z \<in> dom ?CS"
3964  hence "z \<in> Ptr ` dom (map_to_user_data (ksPSpace s))" using cmap
3965    by (simp add: cmap_relation_def dom_heap_to_user_data)
3966  hence pvz: "ptr_val z \<in> dom (map_to_user_data (ksPSpace s))"
3967    by clarsimp
3968  hence "projectKO_opt (the (ksPSpace s (ptr_val z))) = Some UserData"
3969    apply -
3970    apply (frule map_comp_subset_domD)
3971    apply (clarsimp simp: dom_def)+
3972    done
3973  moreover have pvz: "ptr_val z \<in> dom (ksPSpace s)" using pvz
3974    by (rule map_comp_subset_domD)
3975  ultimately show "is_aligned (ptr_val z) (objBitsKO KOUserData)" using pal
3976    unfolding pspace_aligned'_def
3977    apply -
3978    apply (drule (1) bspec)
3979    apply (simp add: projectKOs)
3980    done
3981qed
3982
3983lemma range_cover_bound_weak:
3984  "\<lbrakk>n \<noteq> 0;range_cover ptr sz us n\<rbrakk> \<Longrightarrow>
3985  ptr + (of_nat n * 2 ^ us - 1) \<le> (ptr && ~~ mask sz) + 2 ^ sz - 1"
3986 apply (frule range_cover_cell_subset[where x = "of_nat (n - 1)"])
3987  apply (simp add:range_cover_not_zero)
3988 apply (frule range_cover_subset_not_empty[rotated,where x = "of_nat (n - 1)"])
3989  apply (simp add:range_cover_not_zero)
3990 apply (clarsimp simp: field_simps)
3991 done
3992
3993lemma pspace_no_overlap_underlying_zero:
3994  "pspace_no_overlap' ptr sz \<sigma>
3995    \<Longrightarrow> valid_machine_state' \<sigma>
3996    \<Longrightarrow> x \<in> {ptr .. (ptr && ~~ mask sz) + 2 ^ sz - 1}
3997    \<Longrightarrow> underlying_memory (ksMachineState \<sigma>) x = 0"
3998  using mask_in_range[where ptr'=x and bits=pageBits and ptr="x && ~~ mask pageBits"]
3999  apply (clarsimp simp: valid_machine_state'_def)
4000  apply (drule_tac x=x in spec, clarsimp simp: pointerInUserData_def)
4001  apply (clarsimp simp: typ_at'_def ko_wp_at'_def koTypeOf_eq_UserDataT)
4002  apply (case_tac "pointerInDeviceData x \<sigma>")
4003   apply (clarsimp simp: pointerInDeviceData_def
4004                         ko_wp_at'_def obj_at'_def projectKOs
4005                  dest!: device_data_at_ko)
4006   apply (drule(1) pspace_no_overlapD')
4007   apply (drule_tac x=x in eqset_imp_iff)
4008   apply (simp add: objBits_simps)
4009  apply clarsimp
4010  apply (drule(1) pspace_no_overlapD')
4011  apply (drule_tac x=x in eqset_imp_iff, simp)
4012  apply (simp add: objBits_simps)
4013  done
4014
4015lemma range_cover_nca_neg: "\<And>x p (off :: 10 word).
4016  \<lbrakk>(x::word32) < 4; {p..+2 ^pageBits } \<inter> {ptr..ptr + (of_nat n * 2 ^ (gbits + pageBits) - 1)} = {};
4017    range_cover ptr sz (gbits + pageBits) n\<rbrakk>
4018  \<Longrightarrow> p + ucast off * 4 + x \<notin> {ptr..+n * 2 ^ (gbits + pageBits)}"
4019  apply (case_tac "n = 0")
4020   apply simp
4021  apply (subst range_cover_intvl,simp)
4022   apply simp
4023  apply (subgoal_tac "p + ucast off * 4 + x \<in>  {p..+2 ^ pageBits}")
4024   apply blast
4025  apply (clarsimp simp: intvl_def)
4026  apply (rule_tac x = "unat off * 4 + unat x" in exI)
4027  apply (simp add: ucast_nat_def)
4028  apply (rule nat_add_offset_less [where n = 2, simplified])
4029    apply (simp add: word_less_nat_alt)
4030   apply (rule unat_lt2p)
4031  apply (simp add: pageBits_def objBits_simps)
4032  done
4033
4034lemma heap_to_device_data_disj_mdf:
4035  assumes rc: "range_cover ptr sz (gbits + pageBits) n"
4036  and ko_at: "ksPSpace \<sigma> a = Some obj"
4037  and obj_size: "objBitsKO obj = pageBits"
4038  and pal: "pspace_aligned' \<sigma>" and pdst: "pspace_distinct' \<sigma>"
4039  and pno: "pspace_no_overlap' ptr sz \<sigma>"
4040  and sz: "gbits + pageBits \<le> sz"
4041  and szb: "sz < word_bits"
4042  shows "(heap_to_device_data (ksPSpace \<sigma>)
4043          (\<lambda>x. if x \<in> {ptr..+n * 2 ^ (gbits + pageBits)} then 0 else underlying_memory (ksMachineState \<sigma>) x) a)
4044          = (heap_to_device_data (ksPSpace \<sigma>) (underlying_memory (ksMachineState \<sigma>)) a)"
4045  proof -
4046  from sz have "2 \<le> sz" by (simp add: objBits_simps pageBits_def)
4047
4048  hence sz2: "2 ^ (sz - 2) * 4 = (2 :: nat) ^ sz"
4049    apply (subgoal_tac "(4 :: nat) = 2 ^ 2")
4050    apply (erule ssubst)
4051    apply (subst power_add [symmetric])
4052    apply (rule arg_cong [where f = "\<lambda>n. 2 ^ n"])
4053    apply simp
4054    apply simp
4055    done
4056  have p2dist: "n * (2::nat) ^ (gbits + pageBits) = n * 2 ^ gbits * 2 ^ pageBits" (is "?lhs = ?rhs")
4057    by (simp add: monoid_mult_class.power_add)
4058  show ?thesis
4059    apply (simp add: heap_to_device_data_def)
4060    apply (case_tac "n = 0")
4061     apply simp
4062    apply (subst map_option_byte_to_word_heap)
4063     apply (erule range_cover_nca_neg[OF _ _ rc])
4064     using range_cover_intvl[OF rc]
4065     apply (clarsimp simp add: heap_to_user_data_def Let_def
4066       byte_to_word_heap_def[abs_def] map_comp_Some_iff projectKOs)
4067     apply (cut_tac pspace_no_overlapD' [OF ko_at pno])
4068     apply (subst (asm) upto_intvl_eq [symmetric])
4069      apply (rule pspace_alignedD' [OF ko_at pal])
4070     apply (simp add: obj_size p2dist)
4071     apply (drule_tac B' = "{ptr..ptr + (of_nat n * 2 ^ (gbits + pageBits) - 1)}" in disjoint_subset2[rotated])
4072      apply (clarsimp simp: p2dist )
4073      apply (rule range_cover_bound_weak)
4074       apply simp
4075      apply (rule rc)
4076     apply simp
4077    apply simp
4078   done
4079qed
4080
4081lemma pageBitsForSize_mess_multi:
4082  "4 * (2::nat) ^ (pageBitsForSize sz - 2) = 2^(pageBitsForSize sz)"
4083  apply (subgoal_tac "(4 :: nat) = 2 ^ 2")
4084  apply (erule ssubst)
4085  apply (subst power_add [symmetric])
4086  apply (rule arg_cong [where f = "\<lambda>n. 2 ^ n"])
4087  apply (case_tac sz,simp+)
4088  done
4089
4090
4091lemma createObjects_ccorres_user_data:
4092  defines "ko \<equiv> KOUserData"
4093  shows "\<forall>\<sigma> x. (\<sigma>, x) \<in> rf_sr \<and> range_cover ptr sz (gbits + pageBits) n
4094  \<and> ptr \<noteq> 0
4095  \<and> pspace_aligned' \<sigma> \<and> pspace_distinct' \<sigma>
4096  \<and> valid_machine_state' \<sigma>
4097  \<and> ret_zero ptr (n * 2 ^ (gbits + pageBits)) \<sigma>
4098  \<and> pspace_no_overlap' ptr sz \<sigma>
4099  \<and> region_is_zero_bytes ptr (n * 2 ^ (gbits + pageBits)) x
4100  \<and> {ptr ..+ n * (2 ^ (gbits + pageBits))} \<inter> kernel_data_refs = {}
4101  \<longrightarrow>
4102  (\<sigma>\<lparr>ksPSpace :=
4103               foldr (\<lambda>addr. data_map_insert addr KOUserData)
4104                  (new_cap_addrs (n * 2^gbits) ptr KOUserData) (ksPSpace \<sigma>)\<rparr>,
4105           x\<lparr>globals := globals x\<lparr>t_hrs_' :=
4106                      hrs_htd_update
4107                       (ptr_retyps_gen (n * 2 ^ gbits) (Ptr ptr :: user_data_C ptr) arr)
4108                       ((t_hrs_' (globals x)))\<rparr> \<rparr>) \<in> rf_sr"
4109  (is "\<forall>\<sigma> x. ?P \<sigma> x \<longrightarrow>
4110    (\<sigma>\<lparr>ksPSpace := ?ks \<sigma>\<rparr>, x\<lparr>globals := globals x\<lparr>t_hrs_' := ?ks' x\<rparr>\<rparr>) \<in> rf_sr")
4111proof (intro impI allI)
4112  fix \<sigma> x
4113  let ?thesis = "(\<sigma>\<lparr>ksPSpace := ?ks \<sigma>\<rparr>, x\<lparr>globals := globals x\<lparr>t_hrs_' := ?ks' x\<rparr>\<rparr>) \<in> rf_sr"
4114  let ?ks = "?ks \<sigma>"
4115  let ?ks' = "?ks' x"
4116  let ?ptr = "Ptr ptr :: user_data_C ptr"
4117
4118  note Kernel_C.user_data_C_size [simp del]
4119
4120  assume "?P \<sigma> x"
4121  hence rf: "(\<sigma>, x) \<in> rf_sr" and al: "is_aligned ptr (gbits + pageBits)"
4122    and ptr0: "ptr \<noteq> 0"
4123    and sz: "gbits + pageBits \<le> sz"
4124    and szb: "sz < word_bits"
4125    and pal: "pspace_aligned' \<sigma>" and pdst: "pspace_distinct' \<sigma>"
4126    and pno: "pspace_no_overlap' ptr sz \<sigma>"
4127    and vms: "valid_machine_state' \<sigma>"
4128    and rzo: "ret_zero ptr (n * 2 ^ (gbits + pageBits)) \<sigma>"
4129    and empty: "region_is_bytes ptr (n * 2 ^ (gbits + pageBits)) x"
4130    and zero: "heap_list_is_zero (hrs_mem (t_hrs_' (globals x))) ptr (n * 2 ^ (gbits + pageBits))"
4131    and rc: "range_cover ptr sz (gbits + pageBits) n"
4132    and rc': "range_cover ptr sz (objBitsKO ko) (n * 2^ gbits)"
4133    and kdr: "{ptr..+n * 2 ^ (gbits + pageBits)} \<inter> kernel_data_refs = {}"
4134    by (auto simp: range_cover.aligned objBits_simps  ko_def
4135                   range_cover_rel[where sbit' = pageBits]
4136                   range_cover.sz[where 'a=32, folded word_bits_def])
4137
4138  hence al': "is_aligned ptr (objBitsKO ko)"
4139    by (clarsimp dest!: is_aligned_weaken range_cover.aligned)
4140
4141  (* This is a hack *)
4142  have mko: "\<And>dev. makeObjectKO False (Inr object_type.SmallPageObject) = Some ko"
4143    by (simp add: makeObjectKO_def ko_def)
4144
4145  from sz have "2 \<le> sz" by (simp add: objBits_simps pageBits_def ko_def)
4146
4147  hence sz2: "2 ^ (sz - 2) * 4 = (2 :: nat) ^ sz"
4148    apply (subgoal_tac "(4 :: nat) = 2 ^ 2")
4149    apply (erule ssubst)
4150    apply (subst power_add [symmetric])
4151    apply (rule arg_cong [where f = "\<lambda>n. 2 ^ n"])
4152    apply simp
4153    apply simp
4154    done
4155
4156  define big_0s where "big_0s \<equiv> (replicate (2^pageBits) 0) :: word8 list"
4157
4158  have "length big_0s = 4096" unfolding big_0s_def
4159    by simp (simp add: table_bits_defs)
4160
4161  hence i1: "\<And>off :: 10 word. index (user_data_C.words_C (from_bytes big_0s)) (unat off) = 0"
4162    apply (simp add: from_bytes_def)
4163    apply (simp add: typ_info_simps user_data_C_tag_def)
4164    apply (simp add: ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td align_of_def padup_def
4165      final_pad_def size_td_lt_ti_typ_pad_combine Let_def align_td_array' size_td_array size_of_def
4166      cong: if_cong)
4167    apply (simp add: update_ti_adjust_ti update_ti_t_word32_0s
4168      typ_info_simps update_ti_t_ptr_0s
4169      ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td
4170      ti_typ_combine_empty_ti ti_typ_combine_td
4171      align_of_def padup_def
4172      final_pad_def size_td_lt_ti_typ_pad_combine Let_def
4173      align_td_array' size_td_array cong: if_cong)
4174    apply (subst update_ti_t_array_rep_word0)
4175     apply (unfold big_0s_def)[1]
4176     apply (rule arg_cong [where f = "\<lambda>x. replicate x 0"])
4177     apply (simp (no_asm) add: size_of_def pageBits_def)
4178    apply (subst index_foldr_update)
4179      apply (rule order_less_le_trans [OF unat_lt2p])
4180      apply simp
4181     apply simp
4182    apply simp
4183    done
4184
4185  have p2dist: "n * (2::nat) ^ (gbits + pageBits) = n * 2 ^ gbits * 2 ^ pageBits" (is "?lhs = ?rhs")
4186    by (simp add:monoid_mult_class.power_add)
4187
4188  have nca: "\<And>x p (off :: 10 word). \<lbrakk> p \<in> set (new_cap_addrs (n*2^gbits) ptr KOUserData); x < 4 \<rbrakk>
4189    \<Longrightarrow> p + ucast off * 4 + x \<in> {ptr..+ n * 2 ^ (gbits + pageBits) }"
4190    using sz
4191    apply (clarsimp simp: new_cap_addrs_def objBits_simps shiftl_t2n intvl_def)
4192    apply (rule_tac x = "2 ^ pageBits * pa + unat off * 4 + unat x" in exI)
4193    apply (simp add: ucast_nat_def power_add)
4194    apply (subst mult.commute, subst add.assoc)
4195    apply (rule_tac y = "(pa + 1) * 2 ^ pageBits " in less_le_trans)
4196     apply (simp add:word_less_nat_alt)
4197    apply (rule_tac y="unat off * 4 + 4" in less_le_trans)
4198      apply simp
4199     apply (simp add:pageBits_def)
4200     apply (cut_tac x = off in unat_lt2p)
4201     apply simp
4202    apply (subst mult.assoc[symmetric])
4203    apply (rule mult_right_mono)
4204     apply simp+
4205    done
4206
4207  have nca_neg: "\<And>x p (off :: 10 word).
4208    \<lbrakk>x < 4; {p..+2 ^ objBitsKO KOUserData } \<inter> {ptr..ptr + (of_nat n * 2 ^ (gbits + pageBits) - 1)} = {}\<rbrakk>
4209     \<Longrightarrow> p + ucast off * 4 + x \<notin> {ptr..+n * 2 ^ (gbits + pageBits)}"
4210    apply (case_tac "n = 0")
4211     apply simp
4212    apply (subst range_cover_intvl[OF rc])
4213     apply simp
4214    apply (subgoal_tac " p + ucast off * 4 + x \<in>  {p..+2 ^ objBitsKO KOUserData}")
4215     apply blast
4216    apply (clarsimp simp:intvl_def)
4217    apply (rule_tac x = "unat off * 4 + unat x" in exI)
4218    apply (simp add: ucast_nat_def)
4219    apply (rule nat_add_offset_less [where n = 2, simplified])
4220      apply (simp add: word_less_nat_alt)
4221     apply (rule unat_lt2p)
4222    apply (simp add: pageBits_def objBits_simps)
4223    done
4224
4225  have zero_app: "\<And>x. x \<in> {ptr..+ n * 2 ^ (gbits + pageBits) }
4226    \<Longrightarrow> underlying_memory (ksMachineState \<sigma>) x = 0"
4227    apply (cases "n = 0")
4228     apply simp
4229    apply (rule pspace_no_overlap_underlying_zero[OF pno vms])
4230    apply (erule subsetD[rotated])
4231    apply (cases "n = 0")
4232     apply simp
4233    apply (subst range_cover_intvl[OF rc], simp)
4234    apply (rule order_trans[rotated], erule range_cover_subset'[OF rc])
4235    apply (simp add: field_simps)
4236    done
4237
4238  have cud: "\<And>p. p \<in> set (new_cap_addrs (n * 2^ gbits) ptr KOUserData) \<Longrightarrow>
4239              cuser_user_data_relation
4240                (byte_to_word_heap
4241                  (underlying_memory (ksMachineState \<sigma>)) p)
4242                (from_bytes big_0s)"
4243    unfolding cuser_user_data_relation_def
4244    apply -
4245    apply (rule allI)
4246    apply (subst i1)
4247    apply (simp add: byte_to_word_heap_def Let_def
4248                     zero_app nca nca [where x3 = 0, simplified])
4249    apply (simp add: word_rcat_bl)
4250    done
4251
4252  note blah[simp del] =  atLeastAtMost_iff atLeastatMost_subset_iff atLeastLessThan_iff
4253      Int_atLeastAtMost atLeastatMost_empty_iff split_paired_Ex
4254
4255  have cud2: "\<And>xa v y.
4256              \<lbrakk> heap_to_user_data
4257                     (\<lambda>x. if x \<in> set (new_cap_addrs (n*2^gbits) ptr KOUserData)
4258                           then Some KOUserData else ksPSpace \<sigma> x)
4259                     (underlying_memory (ksMachineState \<sigma>)) xa =
4260              Some v; xa \<notin> set (new_cap_addrs (n*2^gbits) ptr KOUserData);
4261              heap_to_user_data (ksPSpace \<sigma>) (underlying_memory (ksMachineState \<sigma>)) xa = Some y \<rbrakk> \<Longrightarrow> y = v"
4262    using range_cover_intvl[OF rc]
4263    by (clarsimp simp add: heap_to_user_data_def Let_def sz2
4264      byte_to_word_heap_def[abs_def] map_comp_Some_iff projectKOs)
4265
4266  have relrl: "cmap_relation (heap_to_user_data (ksPSpace \<sigma>) (underlying_memory (ksMachineState \<sigma>)))
4267                             (cslift x) Ptr cuser_user_data_relation
4268    \<Longrightarrow> cmap_relation
4269        (heap_to_user_data
4270          (\<lambda>x. if x \<in> set (new_cap_addrs (n * 2 ^ gbits) ptr KOUserData)
4271               then Some KOUserData else ksPSpace \<sigma> x)
4272          (underlying_memory (ksMachineState \<sigma>)))
4273        (\<lambda>y. if y \<in> Ptr ` set (new_cap_addrs (n*2^gbits) ptr KOUserData)
4274             then Some
4275                   (from_bytes (replicate (2 ^ pageBits) 0))
4276             else cslift x y)
4277        Ptr cuser_user_data_relation"
4278    apply (rule cmap_relationI)
4279    apply (clarsimp simp: dom_heap_to_user_data cmap_relation_def dom_if image_Un
4280      projectKO_opt_retyp_same projectKOs)
4281    apply (case_tac "xa \<in> set (new_cap_addrs (n*2^gbits) ptr KOUserData)")
4282    apply (clarsimp simp: heap_to_user_data_def sz2)
4283    apply (erule cud [unfolded big_0s_def])
4284    apply (subgoal_tac "(Ptr xa :: user_data_C ptr) \<notin> Ptr ` set (new_cap_addrs (n*2^gbits) ptr KOUserData)")
4285    apply simp
4286    apply (erule (1) cmap_relationE2)
4287    apply (drule (1) cud2)
4288    apply simp
4289   apply simp
4290   apply clarsimp
4291   done
4292
4293  (* /obj specific *)
4294
4295  (* s/obj/obj'/ *)
4296
4297  have szo: "size_of TYPE(user_data_C) = 2 ^ objBitsKO ko" by (simp add: size_of_def objBits_simps archObjSize_def ko_def pageBits_def)
4298  have szo': "n * 2 ^ (gbits + pageBits) = n * 2 ^ gbits * size_of TYPE(user_data_C)" using sz
4299    apply (subst szo)
4300    apply (clarsimp simp: power_add[symmetric] objBits_simps ko_def)
4301    done
4302
4303  have rb': "region_is_bytes ptr (n * 2 ^ gbits * 2 ^ objBitsKO ko) x"
4304    using empty
4305    by (simp add: mult.commute mult.left_commute power_add objBits_simps ko_def)
4306
4307  note rl' = cslift_ptr_retyp_other_inst[OF rb' rc' szo' szo, simplified]
4308
4309  (* rest is generic *)
4310
4311  note rl = projectKO_opt_retyp_other [OF rc' pal pno,unfolded ko_def]
4312  note cterl = retype_ctes_helper[OF pal pdst pno al' range_cover.sz(2)[OF rc'] range_cover.sz(1)[OF rc', folded word_bits_def] mko rc']
4313  note ht_rl = clift_eq_h_t_valid_eq[OF rl', OF tag_disj_via_td_name, simplified]
4314
4315  have guard:
4316    "\<forall>t<n * 2 ^ gbits. c_guard (CTypesDefs.ptr_add ?ptr (of_nat t))"
4317    apply (rule retype_guard_helper[OF rc' ptr0 szo,where m = 2])
4318    apply (clarsimp simp:align_of_def objBits_simps ko_def pageBits_def)+
4319    done
4320
4321  from rf have "cpspace_relation (ksPSpace \<sigma>) (underlying_memory (ksMachineState \<sigma>)) (t_hrs_' (globals x))"
4322    unfolding rf_sr_def cstate_relation_def by (simp add: Let_def)
4323
4324  hence "cpspace_relation ?ks (underlying_memory (ksMachineState \<sigma>)) ?ks'"
4325    unfolding cpspace_relation_def
4326    using empty rc' szo
4327    apply -
4328    apply (clarsimp simp: rl' tag_disj_via_td_name cte_C_size ht_rl
4329                          foldr_upd_app_if [folded data_map_insert_def])
4330    apply (simp add: rl ko_def projectKOs p2dist
4331                     cterl[unfolded ko_def])
4332    apply (subst clift_ptr_retyps_gen_prev_memset_same[OF guard])
4333         apply (simp add: pageBits_def objBits_simps)
4334        apply simp
4335       apply (simp add: pageBits_def objBits_simps)
4336      apply (cut_tac range_cover.strong_times_32[OF rc], simp_all)[1]
4337      apply (simp add: p2dist objBits_simps)
4338     apply (cut_tac zero)
4339     apply (simp add: pageBits_def power_add field_simps)
4340    apply (simp add: objBits_simps ptr_add_to_new_cap_addrs[OF szo] ko_def
4341               cong: if_cong)
4342    apply (simp add: p2dist[symmetric])
4343    apply (erule relrl[simplified])
4344    done
4345
4346  thus  ?thesis using rf empty kdr rzo
4347    apply (simp add: rf_sr_def cstate_relation_def Let_def rl' tag_disj_via_td_name )
4348    apply (simp add: carch_state_relation_def cmachine_state_relation_def)
4349    apply (simp add: tag_disj_via_td_name rl' tcb_C_size h_t_valid_clift_Some_iff)
4350    apply (clarsimp simp: hrs_htd_update szo'[symmetric])
4351    apply (simp add:szo hrs_htd_def p2dist objBits_simps ko_def ptr_retyps_htd_safe_neg
4352                    kernel_data_refs_domain_eq_rotate
4353                    rl foldr_upd_app_if [folded data_map_insert_def]
4354                    projectKOs cvariable_array_ptr_retyps
4355                    zero_ranges_ptr_retyps)
4356    done
4357qed
4358
4359lemma t_hrs_update_hrs_htd_id:
4360  "t_hrs_'_update id = id"
4361  "hrs_htd_update id = id"
4362  by (simp_all add: fun_eq_iff hrs_htd_update_def)
4363
4364lemma valid_pde_mappings_ko_atD':
4365  "\<lbrakk> ko_at' ko p s; valid_pde_mappings' s \<rbrakk>
4366       \<Longrightarrow> ko_at' ko p s \<and> valid_pde_mapping' (p && mask pdBits) ko"
4367  by (simp add: valid_pde_mappings'_def)
4368
4369lemmas clift_array_assertionE
4370    = clift_array_assertion_imp[where p="Ptr q" and p'="Ptr q" for q,
4371        OF _ refl _ exI[where x=0], simplified]
4372
4373lemma copyGlobalMappings_ccorres:
4374  "ccorres dc xfdc \<top> (UNIV \<inter> {s. newPD_' s = Ptr pd}) []
4375    (copyGlobalMappings pd) (Call copyGlobalMappings_'proc)"
4376  apply (cinit lift: newPD_' simp: ARMSectionBits_def)
4377   apply (rule ccorres_return_Skip)
4378  apply simp
4379  done
4380
4381
4382lemma getObjectSize_symb:
4383  "\<forall>s. \<Gamma> \<turnstile> {s. t_' s = object_type_from_H newType \<and> userObjSize_' s = sz} Call getObjectSize_'proc
4384  {s'. ret__unsigned_long_' s' = of_nat (getObjectSize newType (unat sz))}"
4385  apply (rule allI, rule conseqPre, vcg)
4386  apply (clarsimp simp: nAPIObjects_def Kernel_C_defs)
4387  apply (case_tac newType)
4388   apply (simp_all add:object_type_from_H_def Kernel_C_defs
4389     ARMSmallPageBits_def ARMLargePageBits_def ARMSectionBits_def ARMSuperSectionBits_def
4390     APIType_capBits_def objBits_simps)
4391   apply (rename_tac apiobject_type)
4392   apply (case_tac apiobject_type)
4393   apply (simp_all add:object_type_from_H_def Kernel_C_defs
4394     ARMSmallPageBits_def ARMLargePageBits_def ARMSectionBits_def ARMSuperSectionBits_def
4395     APIType_capBits_def objBits_simps' machine_bits_defs)
4396  apply unat_arith
4397  done
4398
4399(* If we only change local variables on the C side, nothing need be done on the abstract side. *)
4400lemma ccorres_only_change_locals:
4401  "\<lbrakk> \<And>s. \<Gamma> \<turnstile> {s} C {t. globals s = globals t} \<rbrakk> \<Longrightarrow> ccorresG rf_sr \<Gamma> dc xfdc \<top> UNIV hs (return x) C"
4402  apply (rule ccorres_from_vcg)
4403  apply (clarsimp simp: return_def)
4404  apply (clarsimp simp: rf_sr_def)
4405  apply (rule hoare_complete)
4406  apply (clarsimp simp: HoarePartialDef.valid_def)
4407  apply (erule_tac x=x in meta_allE)
4408  apply (drule hoare_sound)
4409  apply (clarsimp simp: cvalid_def HoarePartialDef.valid_def)
4410  apply auto
4411  done
4412
4413
4414lemma getObjectSize_max_size:
4415  "\<lbrakk> newType =  APIObjectType apiobject_type.Untyped \<longrightarrow> x < 32;
4416         newType =  APIObjectType apiobject_type.CapTableObject \<longrightarrow> x < 28 \<rbrakk> \<Longrightarrow> getObjectSize newType x < word_bits"
4417  apply (clarsimp simp only: getObjectSize_def apiGetObjectSize_def word_bits_def
4418                  split: ARM_HYP_H.object_type.splits apiobject_type.splits)
4419  apply (clarsimp simp: tcbBlockSizeBits_def epSizeBits_def ntfnSizeBits_def cteSizeBits_def
4420                        machine_bits_defs)
4421  done
4422
4423lemma getObjectSize_min_size:
4424  "\<lbrakk> newType =  APIObjectType apiobject_type.Untyped \<longrightarrow> minUntypedSizeBits \<le> x;
4425     newType =  APIObjectType apiobject_type.CapTableObject \<longrightarrow> 2 \<le> x \<rbrakk> \<Longrightarrow>
4426    4 \<le> getObjectSize newType x"
4427  apply (clarsimp simp only: getObjectSize_def apiGetObjectSize_def word_bits_def
4428                  split: ARM_HYP_H.object_type.splits apiobject_type.splits)
4429  apply (clarsimp simp: tcbBlockSizeBits_def epSizeBits_def ntfnSizeBits_def cteSizeBits_def
4430                        machine_bits_defs untypedBits_defs)
4431  done
4432
4433(*
4434 * Assuming "placeNewObject" doesn't fail, it is equivalent
4435 * to placing a number of objects into the PSpace.
4436 *)
4437lemma placeNewObject_eq:
4438  notes option.case_cong_weak [cong]
4439  shows
4440  "\<lbrakk> groupSizeBits < word_bits; is_aligned ptr (groupSizeBits + objBitsKO (injectKOS object));
4441    no_fail ((=) s) (placeNewObject ptr object groupSizeBits) \<rbrakk> \<Longrightarrow>
4442  ((), (s\<lparr>ksPSpace := foldr (\<lambda>addr. data_map_insert addr (injectKOS object)) (new_cap_addrs (2 ^ groupSizeBits) ptr (injectKOS object)) (ksPSpace s)\<rparr>))
4443                \<in> fst (placeNewObject ptr object groupSizeBits s)"
4444  apply (clarsimp simp: placeNewObject_def placeNewObject'_def)
4445  apply (clarsimp simp: split_def field_simps split del: if_split)
4446  apply (clarsimp simp: no_fail_def)
4447  apply (subst lookupAround2_pspace_no)
4448   apply assumption
4449  apply (subst (asm) lookupAround2_pspace_no)
4450   apply assumption
4451  apply (clarsimp simp add: in_monad' split_def bind_assoc field_simps
4452    snd_bind ball_to_all unless_def  split: option.splits if_split_asm)
4453  apply (clarsimp simp: data_map_insert_def new_cap_addrs_def)
4454  apply (subst upto_enum_red2)
4455   apply (fold word_bits_def, assumption)
4456  apply (clarsimp simp: field_simps shiftl_t2n power_add mult.commute mult.left_commute
4457           cong: foldr_cong map_cong)
4458  done
4459
4460lemma globals_list_distinct_rf_sr:
4461  "\<lbrakk> (s, s') \<in> rf_sr; S \<inter> kernel_data_refs = {} \<rbrakk>
4462    \<Longrightarrow> globals_list_distinct S symbol_table globals_list"
4463  apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def)
4464  apply (erule globals_list_distinct_subset)
4465  apply blast
4466  done
4467
4468lemma rf_sr_htd_safe:
4469  "(s, s') \<in> rf_sr \<Longrightarrow> htd_safe domain (hrs_htd (t_hrs_' (globals s')))"
4470  by (simp add: rf_sr_def cstate_relation_def Let_def)
4471
4472lemma region_actually_is_bytes_dom_s:
4473  "region_actually_is_bytes' ptr len htd
4474    \<Longrightarrow> S \<subseteq> {ptr ..+ len}
4475    \<Longrightarrow> S \<times> {SIndexVal, SIndexTyp 0} \<subseteq> dom_s htd"
4476  apply (clarsimp simp: region_actually_is_bytes'_def dom_s_def)
4477  apply fastforce
4478  done
4479
4480lemma typ_region_bytes_actually_is_bytes:
4481  "htd = typ_region_bytes ptr bits htd'
4482    \<Longrightarrow> region_actually_is_bytes' ptr (2 ^ bits) htd"
4483  by (clarsimp simp: region_actually_is_bytes'_def typ_region_bytes_def)
4484
4485(* FIXME: need a way to avoid overruling the parser on this, it's ugly *)
4486lemma memzero_modifies:
4487  "\<forall>\<sigma>. \<Gamma>\<turnstile>\<^bsub>/UNIV\<^esub> {\<sigma>} Call memzero_'proc {t. t may_only_modify_globals \<sigma> in [t_hrs]}"
4488  apply (rule allI, rule conseqPre)
4489  apply (hoare_rule HoarePartial.ProcNoRec1)
4490   apply (tactic {* HoarePackage.vcg_tac "_modifies" "false" [] @{context} 1 *})
4491  apply (clarsimp simp: mex_def meq_def simp del: split_paired_Ex)
4492  apply (intro exI globals.equality, simp_all)
4493  done
4494
4495lemma ghost_assertion_size_logic_no_unat:
4496  "sz \<le> gsMaxObjectSize s
4497    \<Longrightarrow> (s, \<sigma>) \<in> rf_sr
4498    \<Longrightarrow> gs_get_assn cap_get_capSizeBits_'proc (ghost'state_' (globals \<sigma>)) = 0 \<or>
4499            of_nat sz \<le> gs_get_assn cap_get_capSizeBits_'proc (ghost'state_' (globals \<sigma>))"
4500  apply (rule ghost_assertion_size_logic'[rotated])
4501   apply (simp add: rf_sr_def)
4502  apply (simp add: unat_of_nat)
4503  done
4504
4505lemma ccorres_placeNewObject_endpoint:
4506  "ko = (makeObject :: endpoint)
4507   \<Longrightarrow> ccorresG rf_sr \<Gamma> dc xfdc
4508   (pspace_aligned' and pspace_distinct'
4509      and pspace_no_overlap' regionBase (objBits ko)
4510      and ret_zero regionBase (2 ^ objBits ko)
4511      and (\<lambda>s. 2 ^ (objBits ko) \<le> gsMaxObjectSize s)
4512      and K (regionBase \<noteq> 0 \<and> range_cover regionBase (objBits ko) (objBits ko) 1
4513      \<and> {regionBase..+ 2 ^ (objBits ko)} \<inter> kernel_data_refs = {}))
4514   ({s. region_actually_is_zero_bytes regionBase (2 ^ objBits ko) s})
4515    hs
4516    (placeNewObject regionBase ko 0)
4517    (global_htd_update (\<lambda>_. (ptr_retyp (ep_Ptr regionBase))))"
4518  apply (rule ccorres_from_vcg_nofail)
4519  apply clarsimp
4520  apply (rule conseqPre)
4521  apply vcg
4522  apply (clarsimp simp: rf_sr_htd_safe)
4523  apply (intro conjI allI impI)
4524   apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
4525                         kernel_data_refs_domain_eq_rotate
4526                         objBits_simps'
4527                  elim!: ptr_retyp_htd_safe_neg)
4528  apply (rule bexI [OF _ placeNewObject_eq])
4529     apply (clarsimp simp: split_def)
4530     apply (clarsimp simp: new_cap_addrs_def)
4531     apply (cut_tac createObjects_ccorres_ep [where ptr=regionBase and n="1" and sz="objBitsKO (KOEndpoint makeObject)"])
4532     apply (erule_tac x=\<sigma> in allE, erule_tac x=x in allE)
4533     apply (clarsimp elim!:is_aligned_weaken simp: objBitsKO_def word_bits_def)+
4534     apply (clarsimp simp: split_def Let_def
4535         Fun.comp_def rf_sr_def new_cap_addrs_def
4536         region_actually_is_bytes ptr_retyps_gen_def
4537         objBits_simps
4538         elim!: rsubst[where P="cstate_relation s'" for s'])
4539    apply (clarsimp simp: word_bits_conv)
4540   apply (clarsimp simp: range_cover.aligned objBits_simps)
4541  apply (clarsimp simp: no_fail_def)
4542  done
4543
4544lemma ccorres_placeNewObject_notification:
4545  "ccorresG rf_sr \<Gamma> dc xfdc
4546   (pspace_aligned' and pspace_distinct' and pspace_no_overlap' regionBase 4
4547      and (\<lambda>s. 2^ntfnSizeBits \<le> gsMaxObjectSize s)
4548      and ret_zero regionBase (2^ntfnSizeBits)
4549      and K (regionBase \<noteq> 0
4550      \<and> {regionBase..+2^ntfnSizeBits} \<inter> kernel_data_refs = {}
4551      \<and> range_cover regionBase ntfnSizeBits ntfnSizeBits 1))
4552   ({s. region_actually_is_zero_bytes regionBase (2^ntfnSizeBits) s})
4553    hs
4554    (placeNewObject regionBase (makeObject :: Structures_H.notification) 0)
4555    (global_htd_update (\<lambda>_. (ptr_retyp (ntfn_Ptr regionBase))))"
4556  apply (rule ccorres_from_vcg_nofail)
4557  apply clarsimp
4558  apply (rule conseqPre)
4559  apply vcg
4560  apply (clarsimp simp: rf_sr_htd_safe)
4561  apply (intro conjI allI impI)
4562   apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
4563                         kernel_data_refs_domain_eq_rotate objBits_defs
4564                  elim!: ptr_retyp_htd_safe_neg)
4565  apply (rule bexI [OF _ placeNewObject_eq])
4566     apply (clarsimp simp: split_def new_cap_addrs_def)
4567     apply (cut_tac createObjects_ccorres_ntfn [where ptr=regionBase and n="1" and sz="objBitsKO (KONotification makeObject)"])
4568     apply (erule_tac x=\<sigma> in allE, erule_tac x=x in allE)
4569     apply (clarsimp elim!: is_aligned_weaken simp: objBitsKO_def word_bits_def)+
4570     apply (clarsimp simp: split_def objBitsKO_def Let_def
4571                           Fun.comp_def rf_sr_def new_cap_addrs_def)
4572     apply (clarsimp simp: cstate_relation_def carch_state_relation_def split_def
4573                           Let_def cmachine_state_relation_def cpspace_relation_def
4574                           region_actually_is_bytes ptr_retyps_gen_def objBits_defs)
4575    apply (clarsimp simp: word_bits_conv)
4576   apply (clarsimp simp: objBits_simps range_cover.aligned)
4577  apply (clarsimp simp: no_fail_def)
4578  done
4579
4580lemma htd_update_list_dom_better [rule_format]:
4581  "(\<forall>p d. dom_s (htd_update_list p xs d) =
4582          (dom_s d) \<union> dom_tll p xs)"
4583apply(induct_tac xs)
4584 apply simp
4585apply clarsimp
4586apply(auto split: if_split_asm)
4587 apply(erule notE)
4588 apply(clarsimp simp: dom_s_def)
4589apply(case_tac y)
4590 apply clarsimp+
4591apply(clarsimp simp: dom_s_def)
4592done
4593
4594lemma ptr_array_retyps_htd_safe_neg:
4595  "\<lbrakk> htd_safe (- D) htd;
4596    {ptr_val ptr ..+ n * size_of TYPE('a :: mem_type)}
4597        \<inter> D = {} \<rbrakk>
4598   \<Longrightarrow> htd_safe (- D) (ptr_arr_retyps n (ptr :: 'a ptr) htd)"
4599  apply (simp add: htd_safe_def ptr_arr_retyps_def htd_update_list_dom_better)
4600  apply (auto simp: dom_tll_def intvl_def)
4601  done
4602
4603lemma ccorres_placeNewObject_captable:
4604  "ccorresG rf_sr \<Gamma> dc xfdc
4605   (pspace_aligned' and pspace_distinct' and pspace_no_overlap' regionBase (unat userSize + cteSizeBits)
4606      and (\<lambda>s. 2 ^ (unat userSize + cteSizeBits) \<le> gsMaxObjectSize s)
4607      and ret_zero regionBase (2 ^ (unat userSize + cteSizeBits))
4608      and K (regionBase \<noteq> 0 \<and> range_cover regionBase (unat userSize + cteSizeBits) (unat userSize + cteSizeBits) 1
4609      \<and> ({regionBase..+2 ^ (unat userSize + cteSizeBits)} \<inter> kernel_data_refs = {})))
4610    ({s. region_actually_is_zero_bytes regionBase (2 ^ (unat userSize + cteSizeBits)) s})
4611    hs
4612    (placeNewObject regionBase (makeObject :: cte) (unat (userSize::word32)))
4613    (global_htd_update (\<lambda>_. (ptr_arr_retyps (2 ^ (unat userSize)) (cte_Ptr regionBase))))"
4614  apply (rule ccorres_from_vcg_nofail)
4615  apply clarsimp
4616  apply (rule conseqPre)
4617  apply vcg
4618  apply (clarsimp simp: rf_sr_htd_safe)
4619  apply (intro conjI allI impI)
4620   apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
4621                         kernel_data_refs_domain_eq_rotate
4622                  elim!: ptr_array_retyps_htd_safe_neg)
4623   apply (simp add: size_of_def power_add objBits_defs)
4624  apply (frule range_cover_rel[where sbit' = cteSizeBits])
4625    apply simp
4626   apply simp
4627  apply (frule range_cover.unat_of_nat_shift[where gbits = cteSizeBits, OF _ le_refl le_refl])
4628  apply (subgoal_tac "region_is_bytes regionBase (2 ^ (unat userSize + cteSizeBits)) x")
4629   apply (rule bexI [OF _ placeNewObject_eq])
4630      apply (clarsimp simp: split_def new_cap_addrs_def)
4631      apply (cut_tac createObjects_ccorres_cte
4632                       [where ptr=regionBase and n="2 ^ unat userSize"
4633                          and sz="unat userSize + objBitsKO (KOCTE makeObject)"])
4634      apply (erule_tac x=\<sigma> in allE, erule_tac x=x in allE)
4635      apply (clarsimp elim!:is_aligned_weaken simp: objBitsKO_def word_bits_def)+
4636      apply (clarsimp simp: split_def objBitsKO_def
4637                            Fun.comp_def rf_sr_def Let_def
4638                            new_cap_addrs_def field_simps power_add ptr_retyps_gen_def
4639                   elim!: rsubst[where P="cstate_relation s'" for s'])
4640     apply (clarsimp simp: word_bits_conv range_cover_def)
4641    apply (clarsimp simp: objBitsKO_def range_cover.aligned)
4642   apply (clarsimp simp: no_fail_def)
4643  apply (simp add: region_actually_is_bytes)
4644 done
4645
4646lemma rf_sr_helper:
4647  "\<And>a b P X. ((a, globals_update P (b\<lparr>tcb_' := X\<rparr>)) \<in> rf_sr) = ((a, globals_update P b) \<in> rf_sr)"
4648  apply (clarsimp simp: rf_sr_def)
4649  done
4650
4651lemma rf_sr_domain_eq:
4652  "(\<sigma>, s) \<in> rf_sr \<Longrightarrow> htd_safe domain = htd_safe (- kernel_data_refs)"
4653  by (simp add: rf_sr_def cstate_relation_def Let_def
4654                kernel_data_refs_domain_eq_rotate)
4655
4656declare replicate_numeral [simp del]
4657
4658lemma ccorres_placeNewObject_tcb:
4659  "ccorresG rf_sr \<Gamma> dc xfdc
4660   (pspace_aligned' and pspace_distinct' and pspace_no_overlap' regionBase tcbBlockSizeBits
4661      and valid_queues and (\<lambda>s. sym_refs (state_refs_of' s))
4662      and (\<lambda>s. 2 ^ tcbBlockSizeBits \<le> gsMaxObjectSize s)
4663      and ret_zero regionBase (2 ^ tcbBlockSizeBits)
4664      and K (regionBase \<noteq> 0 \<and> range_cover regionBase tcbBlockSizeBits tcbBlockSizeBits 1
4665      \<and>  {regionBase..+2^tcbBlockSizeBits} \<inter> kernel_data_refs = {}))
4666   ({s. region_actually_is_zero_bytes regionBase (2^tcbBlockSizeBits) s})
4667    hs
4668   (placeNewObject regionBase (makeObject :: tcb) 0)
4669   (\<acute>tcb :== tcb_Ptr (regionBase + 0x100);;
4670        (global_htd_update (\<lambda>s. ptr_retyp (Ptr (ptr_val (tcb_' s) - ctcb_offset) :: (cte_C[5]) ptr)
4671            \<circ> ptr_retyp (tcb_' s)));;
4672        (Guard C_Guard \<lbrace>hrs_htd \<acute>t_hrs \<Turnstile>\<^sub>t \<acute>tcb\<rbrace>
4673           (call (\<lambda>s. s\<lparr>context_' := Ptr &((Ptr &(tcb_' s\<rightarrow>[''tcbArch_C'']) :: arch_tcb_C ptr)\<rightarrow>[''tcbContext_C''])\<rparr>) Arch_initContext_'proc (\<lambda>s t. s\<lparr>globals := globals t\<rparr>) (\<lambda>s' s''. Basic (\<lambda>s. s))));;
4674        (Guard C_Guard \<lbrace>hrs_htd \<acute>t_hrs \<Turnstile>\<^sub>t \<acute>tcb\<rbrace>
4675           (Basic (\<lambda>s. globals_update (t_hrs_'_update (hrs_mem_update (heap_update (Ptr &((tcb_' s)\<rightarrow>[''tcbTimeSlice_C''])) (5::word32)))) s))))"
4676  apply -
4677  apply (rule ccorres_from_vcg_nofail)
4678  apply clarsimp
4679  apply (rule conseqPre)
4680   apply vcg
4681  apply (clarsimp simp: rf_sr_htd_safe ctcb_offset_defs)
4682  apply (subgoal_tac "c_guard (tcb_Ptr (regionBase + 0x100))")
4683   apply (subgoal_tac "hrs_htd (hrs_htd_update (ptr_retyp (Ptr regionBase :: (cte_C[5]) ptr)
4684                                 \<circ> ptr_retyp (tcb_Ptr (regionBase + 0x100)))
4685                  (t_hrs_' (globals x))) \<Turnstile>\<^sub>t tcb_Ptr (regionBase + 0x100)")
4686    prefer 2
4687    apply (clarsimp simp: hrs_htd_update)
4688    apply (rule h_t_valid_ptr_retyps_gen_disjoint[where n=1 and arr=False,
4689                unfolded ptr_retyps_gen_def, simplified])
4690     apply (rule ptr_retyp_h_t_valid)
4691     apply simp
4692    apply (rule tcb_ptr_orth_cte_ptrs')
4693   apply (intro conjI allI impI)
4694         apply (simp only: rf_sr_domain_eq)
4695         apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
4696                               kernel_data_refs_domain_eq_rotate)
4697         apply (intro ptr_retyps_htd_safe_neg ptr_retyp_htd_safe_neg, simp_all add: size_of_def)[1]
4698          apply (erule disjoint_subset[rotated])
4699          apply (rule intvl_sub_offset, simp add: objBits_defs)
4700         apply (erule disjoint_subset[rotated],
4701                simp add: intvl_start_le size_td_array cte_C_size objBits_defs)
4702        apply (clarsimp simp: hrs_htd_update)
4703       apply (clarsimp simp: CPSR_def word_sle_def)+
4704     apply (clarsimp simp: hrs_htd_update)
4705     apply (rule h_t_valid_field[rotated], simp+)+
4706    apply (clarsimp simp: hrs_htd_update)
4707   apply (clarsimp simp: hrs_htd_update)
4708   apply (rule bexI [OF _ placeNewObject_eq])
4709      apply (clarsimp simp: split_def new_cap_addrs_def)
4710      apply (cut_tac \<sigma>=\<sigma> and x=x
4711                   and ks="ksPSpace \<sigma>" and p="tcb_Ptr (regionBase + 0x100)" in cnc_tcb_helper)
4712                    apply clarsimp
4713                   apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_defs
4714                                         objBitsKO_def range_cover.aligned)
4715                  apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_defs objBitsKO_def)
4716                 apply (simp add:olen_add_eqv[symmetric] ctcb_size_bits_def)
4717                 apply (erule is_aligned_no_wrap'[OF range_cover.aligned])
4718                 apply (simp add: objBits_defs)
4719                apply simp
4720               apply clarsimp
4721              apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_defs objBitsKO_def)
4722             apply (clarsimp)
4723            apply simp
4724           apply clarsimp
4725          apply (clarsimp simp: objBits_simps ctcb_ptr_to_tcb_ptr_def ctcb_offset_defs)
4726         apply (frule region_actually_is_bytes)
4727         apply (clarsimp simp: region_is_bytes'_def ctcb_ptr_to_tcb_ptr_def ctcb_offset_defs split_def
4728                               hrs_mem_update_def hrs_htd_def)
4729        apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_defs hrs_mem_update_def split_def)
4730        apply (simp add: hrs_mem_def)
4731       apply (simp add: ctcb_ptr_to_tcb_ptr_def ctcb_offset_defs)
4732      apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_defs hrs_mem_update_def split_def)
4733      apply (clarsimp simp: rf_sr_def ptr_retyps_gen_def cong: Kernel_C.globals.unfold_congs
4734                            StateSpace.state.unfold_congs kernel_state.unfold_congs)
4735     apply (clarsimp simp: word_bits_def)
4736    apply (clarsimp simp: objBitsKO_def range_cover.aligned)
4737   apply (clarsimp simp: no_fail_def)
4738  apply (rule c_guard_tcb)
4739   apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_defs range_cover.aligned)
4740  apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_defs)
4741  done
4742
4743(* FIXME this is harcoded madness: 12 = ptBits, 9 = ptBits - pteBits, 512 = 2 ^ 9 *)
4744lemma placeNewObject_pte:
4745  "ccorresG rf_sr \<Gamma> dc xfdc
4746   ( valid_global_refs' and pspace_aligned' and pspace_distinct' and pspace_no_overlap' regionBase 12
4747      and (\<lambda>s. 2 ^ 12 \<le> gsMaxObjectSize s)
4748      and ret_zero regionBase (2 ^ 12)
4749      and K (regionBase \<noteq> 0 \<and> range_cover regionBase 12 12 1
4750      \<and> ({regionBase..+2 ^ 12} \<inter> kernel_data_refs = {})
4751      ))
4752    ({s. region_actually_is_zero_bytes regionBase (2 ^ 12) s})
4753    hs
4754    (placeNewObject regionBase (makeObject :: pte) 9)
4755    (global_htd_update (\<lambda>_. (ptr_retyp (Ptr regionBase :: (pte_C[512]) ptr))))"
4756  apply (rule ccorres_from_vcg_nofail)
4757  apply clarsimp
4758  apply (rule conseqPre)
4759  apply vcg
4760  apply (clarsimp simp: rf_sr_htd_safe)
4761  apply (intro conjI allI impI)
4762   apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
4763                         kernel_data_refs_domain_eq_rotate
4764                  elim!: ptr_retyp_htd_safe_neg)
4765  apply (frule range_cover_rel[where sbit' = 3])
4766    apply simp+
4767  apply (frule range_cover.unat_of_nat_shift[where gbits = 3 ])
4768     apply simp+
4769    apply (rule le_refl)
4770  apply (subgoal_tac "region_is_bytes regionBase 4096 x")
4771   apply (rule bexI [OF _ placeNewObject_eq])
4772      apply (clarsimp simp: split_def new_cap_addrs_def)
4773      apply (cut_tac s=\<sigma> in createObjects_ccorres_pte [where ptr=regionBase and sz=12])
4774      apply (erule_tac x=\<sigma> in allE, erule_tac x=x in allE)
4775      apply (clarsimp elim!:is_aligned_weaken simp: objBitsKO_def word_bits_def)+
4776      apply (clarsimp simp: split_def objBitsKO_def archObjSize_def
4777          Fun.comp_def rf_sr_def split_def Let_def ptr_retyps_gen_def
4778          new_cap_addrs_def field_simps power_add
4779          cong: globals.unfold_congs)
4780      apply (simp add: Int_ac table_bits_defs)
4781     apply (clarsimp simp: word_bits_conv range_cover_def archObjSize_def)
4782    apply (clarsimp simp: objBitsKO_def range_cover.aligned archObjSize_def table_bits_defs)
4783   apply (clarsimp simp: no_fail_def)
4784  apply (simp add: region_actually_is_bytes)
4785 done
4786
4787lemma placeNewObject_pde:
4788  "ccorresG rf_sr \<Gamma> dc xfdc
4789   (valid_global_refs' and pspace_aligned' and pspace_distinct' and pspace_no_overlap' regionBase 14
4790      and (\<lambda>s. 2 ^ 14 \<le> gsMaxObjectSize s)
4791      and ret_zero regionBase (2 ^ 14)
4792      and K (regionBase \<noteq> 0 \<and> range_cover regionBase 14 14 1
4793      \<and> ({regionBase..+2 ^ 14}
4794          \<inter> kernel_data_refs = {})
4795      ))
4796    ({s. region_actually_is_zero_bytes regionBase (2 ^ 14) s})
4797    hs
4798    (placeNewObject regionBase (makeObject :: pde) 11)
4799    (global_htd_update (\<lambda>_. (ptr_retyp (pd_Ptr regionBase))))"
4800  apply (rule ccorres_from_vcg_nofail)
4801  apply clarsimp
4802  apply (rule conseqPre)
4803  apply vcg
4804  apply (clarsimp simp: rf_sr_htd_safe)
4805  apply (intro conjI allI impI)
4806   apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
4807                         kernel_data_refs_domain_eq_rotate
4808                  elim!: ptr_retyp_htd_safe_neg)
4809  apply (frule range_cover_rel[where sbit' = 3])
4810    apply simp+
4811  apply (frule range_cover.unat_of_nat_shift[where gbits = 3 ])
4812   apply simp+
4813   apply (rule le_refl)
4814
4815   apply (subgoal_tac "region_is_bytes regionBase 16384 x")
4816   apply (rule bexI [OF _ placeNewObject_eq])
4817      apply (clarsimp simp: split_def new_cap_addrs_def)
4818      apply (cut_tac s=\<sigma> in createObjects_ccorres_pde [where ptr=regionBase and sz=14])
4819      apply (erule_tac x=\<sigma> in allE, erule_tac x=x in allE)
4820      apply (clarsimp elim!:is_aligned_weaken simp: objBitsKO_def word_bits_def table_bits_defs)+
4821      apply (clarsimp simp: split_def objBitsKO_def archObjSize_def
4822          Fun.comp_def rf_sr_def Let_def ptr_retyps_gen_def
4823          new_cap_addrs_def field_simps power_add
4824          cong: globals.unfold_congs)
4825      apply (simp add: Int_ac table_bits_defs)
4826     apply (clarsimp simp: word_bits_conv range_cover_def archObjSize_def)
4827    apply (clarsimp simp: objBitsKO_def range_cover.aligned archObjSize_def table_bits_defs)
4828   apply (clarsimp simp: no_fail_def)
4829  apply (simp add: region_actually_is_bytes)
4830 done
4831
4832end
4833
4834context begin interpretation Arch . (*FIXME: arch_split*)
4835end
4836
4837lemma dom_disj_union:
4838  "dom (\<lambda>x. if P x \<or> Q x then Some (G x) else None) = dom (\<lambda>x. if P x then Some (G x) else None)
4839  \<union> dom (\<lambda>x. if Q x then Some (G x) else None)"
4840  by (auto split:if_splits)
4841context kernel_m begin
4842
4843lemma createObjects_ccorres_user_data_device:
4844  defines "ko \<equiv> KOUserDataDevice"
4845  shows "\<forall>\<sigma> x. (\<sigma>, x) \<in> rf_sr \<and> range_cover ptr sz (gbits + pageBits) n
4846  \<and> ptr \<noteq> 0
4847  \<and> pspace_aligned' \<sigma> \<and> pspace_distinct' \<sigma>
4848  \<and> pspace_no_overlap' ptr sz \<sigma>
4849  \<and> ret_zero ptr (n * 2 ^ (gbits + pageBits)) \<sigma>
4850  \<and> region_is_bytes ptr (n * 2 ^ (gbits + pageBits)) x
4851  \<and> {ptr ..+ n * (2 ^ (gbits + pageBits))} \<inter> kernel_data_refs = {}
4852  \<longrightarrow>
4853  (\<sigma>\<lparr>ksPSpace :=
4854               foldr (\<lambda>addr. data_map_insert addr KOUserDataDevice) (new_cap_addrs (n * 2^gbits) ptr KOUserDataDevice) (ksPSpace \<sigma>)\<rparr>,
4855           x\<lparr>globals := globals x\<lparr>t_hrs_' :=
4856                      hrs_htd_update
4857                       (ptr_retyps_gen (n * 2 ^ gbits) (Ptr ptr :: user_data_device_C ptr) arr)
4858                       (t_hrs_' (globals x))\<rparr> \<rparr>) \<in> rf_sr"
4859  (is "\<forall>\<sigma> x. ?P \<sigma> x \<longrightarrow>
4860    (\<sigma>\<lparr>ksPSpace := ?ks \<sigma>\<rparr>, x\<lparr>globals := globals x\<lparr>t_hrs_' := ?ks' x\<rparr>\<rparr>) \<in> rf_sr")
4861proof (intro impI allI)
4862  fix \<sigma> x
4863  let ?thesis = "(\<sigma>\<lparr>ksPSpace := ?ks \<sigma>\<rparr>, x\<lparr>globals := globals x\<lparr>t_hrs_' := ?ks' x\<rparr>\<rparr>) \<in> rf_sr"
4864  let ?ks = "?ks \<sigma>"
4865  let ?ks' = "?ks' x"
4866  let ?ptr = "Ptr ptr :: user_data_device_C ptr"
4867
4868  note Kernel_C.user_data_C_size [simp del]
4869
4870  assume "?P \<sigma> x"
4871  hence rf: "(\<sigma>, x) \<in> rf_sr" and al: "is_aligned ptr (gbits + pageBits)"
4872    and ptr0: "ptr \<noteq> 0"
4873    and sz: "gbits + pageBits \<le> sz"
4874    and szb: "sz < word_bits"
4875    and pal: "pspace_aligned' \<sigma>" and pdst: "pspace_distinct' \<sigma>"
4876    and pno: "pspace_no_overlap' ptr sz \<sigma>"
4877    and rzo: "ret_zero ptr (n * 2 ^ (gbits + pageBits)) \<sigma>"
4878    and empty: "region_is_bytes ptr (n * 2 ^ (gbits + pageBits)) x"
4879    and rc: "range_cover ptr sz (gbits + pageBits) n"
4880    and rc': "range_cover ptr sz (objBitsKO ko) (n * 2^ gbits)"
4881    and kdr: "{ptr..+n * 2 ^ (gbits + pageBits)} \<inter> kernel_data_refs = {}"
4882    by (auto simp: range_cover.aligned objBits_simps  ko_def
4883                   range_cover_rel[where sbit' = pageBits]
4884                   range_cover.sz[where 'a=32, folded word_bits_def])
4885
4886
4887  hence al': "is_aligned ptr (objBitsKO ko)"
4888    by (clarsimp dest!:is_aligned_weaken range_cover.aligned)
4889
4890  note range_cover.no_overflow_n[OF rc']
4891  hence sz_word_bits:
4892    "n * 2 ^ gbits * size_of TYPE(user_data_device_C)  < 2 ^ word_bits"
4893      by (simp add:word_bits_def objBits_simps ko_def pageBits_def)
4894
4895  (* This is a hack *)
4896  have mko: "\<And>dev. makeObjectKO True (Inr object_type.SmallPageObject) = Some ko"
4897    by (simp add: makeObjectKO_def ko_def)
4898
4899  from sz have "2 \<le> sz" by (simp add: objBits_simps pageBits_def ko_def)
4900
4901  hence sz2: "2 ^ (sz - 2) * 4 = (2 :: nat) ^ sz"
4902    apply (subgoal_tac "(4 :: nat) = 2 ^ 2")
4903    apply (erule ssubst)
4904    apply (subst power_add [symmetric])
4905    apply (rule arg_cong [where f = "\<lambda>n. 2 ^ n"])
4906    apply simp
4907    apply simp
4908    done
4909
4910  have p2dist: "n * (2::nat) ^ (gbits + pageBits) = n * 2 ^ gbits * 2 ^ pageBits" (is "?lhs = ?rhs")
4911    by (simp add:monoid_mult_class.power_add)
4912
4913  note blah[simp del] =  atLeastAtMost_iff atLeastatMost_subset_iff atLeastLessThan_iff
4914      Int_atLeastAtMost atLeastatMost_empty_iff split_paired_Ex
4915
4916  (* /obj specific *)
4917
4918  (* s/obj/obj'/ *)
4919
4920  have szo: "size_of TYPE(user_data_device_C) = 2 ^ objBitsKO ko"
4921    by (simp add: size_of_def objBits_simps archObjSize_def ko_def pageBits_def)
4922  have szo': "n * 2 ^ (gbits + pageBits) = n * 2 ^ gbits * size_of TYPE(user_data_device_C)" using sz
4923    apply (subst szo)
4924    apply (clarsimp simp: power_add[symmetric] objBits_simps ko_def)
4925    done
4926
4927  have rb': "region_is_bytes ptr (n * 2 ^ gbits * 2 ^ objBitsKO ko) x"
4928    using empty
4929    by (simp add: mult.commute mult.left_commute power_add objBits_simps ko_def)
4930
4931  from rb' have rbu: "region_is_bytes ptr (n * 2 ^ gbits * size_of TYPE(user_data_device_C)) x"
4932    by (simp add:szo[symmetric])
4933
4934  note rl' = clift_ptr_retyps_gen_other[where p = "Ptr ptr",simplified, OF rbu  sz_word_bits]
4935
4936  (* rest is generic *)
4937
4938  note rl = projectKO_opt_retyp_other [OF rc' pal pno,unfolded ko_def]
4939  note cterl = retype_ctes_helper[OF pal pdst pno al' range_cover.sz(2)[OF rc'] range_cover.sz(1)[OF rc', folded word_bits_def] mko rc']
4940  note ht_rl = clift_eq_h_t_valid_eq[OF rl', OF tag_disj_via_td_name, simplified]
4941
4942  have guard:
4943    "\<forall>t<n * 2 ^ gbits. c_guard (CTypesDefs.ptr_add ?ptr (of_nat t))"
4944    apply (rule retype_guard_helper[OF rc' ptr0 szo,where m = 2])
4945    apply (clarsimp simp: align_of_def objBits_simps ko_def pageBits_def)+
4946    done
4947
4948  have cud2: "\<And>xa v y.
4949              \<lbrakk> heap_to_device_data
4950                     (\<lambda>x. if x \<in> set (new_cap_addrs (n*2^gbits) ptr KOUserDataDevice)
4951                           then Some KOUserData else ksPSpace \<sigma> x)
4952                     (underlying_memory (ksMachineState \<sigma>)) xa =
4953              Some v; xa \<notin> set (new_cap_addrs (n*2^gbits) ptr KOUserDataDevice);
4954              heap_to_device_data (ksPSpace \<sigma>) (underlying_memory (ksMachineState \<sigma>)) xa = Some y \<rbrakk> \<Longrightarrow> y = v"
4955    using range_cover_intvl[OF rc]
4956    by (clarsimp simp add: heap_to_device_data_def Let_def sz2
4957      byte_to_word_heap_def[abs_def] map_comp_Some_iff projectKOs)
4958
4959  note ptr_retyps_valid = h_t_valid_ptr_retyps_gen_same[OF guard rbu,unfolded addr_card_wb,OF _ sz_word_bits,simplified]
4960
4961  from rf have "cpspace_relation (ksPSpace \<sigma>) (underlying_memory (ksMachineState \<sigma>)) (t_hrs_' (globals x))"
4962    unfolding rf_sr_def cstate_relation_def by (simp add: Let_def)
4963
4964  hence "cpspace_relation ?ks (underlying_memory (ksMachineState \<sigma>)) ?ks'"
4965    unfolding cpspace_relation_def
4966    using empty rc' szo
4967    apply -
4968    apply (clarsimp simp: rl' tag_disj_via_td_name cte_C_size ht_rl
4969                          clift_ptr_retyps_gen_other
4970                          foldr_upd_app_if [folded data_map_insert_def])
4971    apply (simp add: rl ko_def projectKOs p2dist
4972                     cterl[unfolded ko_def])
4973    apply (rule cmap_relationI)
4974     apply (clarsimp simp: dom_heap_to_device_data cmap_relation_def dom_if image_Un
4975                           projectKO_opt_retyp_same projectKOs liftt_if[folded hrs_mem_def hrs_htd_def]
4976                           hrs_htd_update hrs_mem_update ptr_retyps_valid dom_disj_union ptr_add_to_new_cap_addrs)
4977    apply (simp add: heap_to_device_data_def cuser_user_data_device_relation_def)
4978    done (* dont need to track all the device memory *)
4979
4980  thus  ?thesis using rf empty kdr rzo
4981    apply (simp add: rf_sr_def cstate_relation_def Let_def rl' tag_disj_via_td_name )
4982    apply (simp add: carch_state_relation_def cmachine_state_relation_def)
4983    apply (simp add: tag_disj_via_td_name rl' tcb_C_size h_t_valid_clift_Some_iff)
4984    apply (clarsimp simp: hrs_htd_update szo'[symmetric] cvariable_array_ptr_retyps[OF szo] rb')
4985    apply (subst zero_ranges_ptr_retyps, simp_all only: szo'[symmetric] power_add,
4986      simp)
4987    apply (simp add:szo  p2dist objBits_simps ko_def ptr_retyps_htd_safe_neg
4988                    kernel_data_refs_domain_eq_rotate
4989                    rl foldr_upd_app_if [folded data_map_insert_def]
4990                    projectKOs cvariable_array_ptr_retyps)
4991    apply (subst cvariable_array_ptr_retyps[OF szo])
4992    apply (simp add: rb'  ptr_retyps_htd_safe_neg)+
4993    apply (erule ptr_retyps_htd_safe_neg)
4994    apply (simp add:pageBits_def field_simps)
4995    done
4996qed
4997
4998lemma placeNewObject_user_data:
4999  "ccorresG rf_sr \<Gamma> dc xfdc
5000  (pspace_aligned' and pspace_distinct' and pspace_no_overlap' regionBase (pageBits+us)
5001  and valid_queues and valid_machine_state'
5002  and ret_zero regionBase (2 ^ (pageBits+us))
5003  and (\<lambda>s. sym_refs (state_refs_of' s))
5004  and (\<lambda>s. 2^(pageBits +  us) \<le> gsMaxObjectSize s)
5005  and K (regionBase \<noteq> 0 \<and> range_cover regionBase (pageBits + us) (pageBits+us) (Suc 0)
5006  \<and> us < word_bits
5007  \<and>  {regionBase..+2^(pageBits +  us)} \<inter> kernel_data_refs = {}))
5008  ({s. region_actually_is_zero_bytes regionBase (2^(pageBits+us)) s})
5009  hs
5010  (placeNewObject regionBase UserData us)
5011  (global_htd_update (\<lambda>s. (ptr_retyps (2^us) (Ptr regionBase :: user_data_C ptr))))"
5012  apply (rule ccorres_from_vcg_nofail)
5013  apply (clarsimp simp:)
5014  apply (rule conseqPre)
5015  apply vcg
5016  apply (clarsimp simp: rf_sr_htd_safe)
5017  apply (intro conjI allI impI)
5018   apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
5019                         kernel_data_refs_domain_eq_rotate
5020                  elim!: ptr_retyps_htd_safe_neg[where arr=False,
5021                        unfolded ptr_retyps_gen_def, simplified])
5022   apply (simp add: size_of_def pageBits_def power_add mult.commute mult.left_commute)
5023  apply (frule range_cover.unat_of_nat_shift[where gbits = "pageBits + us"])
5024    apply simp
5025   apply (clarsimp simp:size_of_def power_add pageBits_def
5026     rf_sr_def cstate_relation_def Let_def field_simps)
5027   apply blast
5028  apply (frule range_cover.aligned)
5029  apply (rule bexI [OF _ placeNewObject_eq], simp_all)
5030    apply (cut_tac ptr=regionBase and sz="pageBits + us" and gbits=us and arr=False
5031                in createObjects_ccorres_user_data[rule_format])
5032     apply (rule conjI, assumption, clarsimp)
5033     apply (fastforce simp: pageBits_def field_simps region_actually_is_bytes)
5034    apply (clarsimp elim!: rsubst[where P="\<lambda>x. (\<sigma>, x) \<in> rf_sr" for \<sigma>]
5035                     simp: field_simps objBitsKO_def ptr_retyps_gen_def)
5036   apply (simp add: objBitsKO_def field_simps)
5037  apply (rule no_fail_pre, rule no_fail_placeNewObject)
5038  apply (clarsimp simp: objBitsKO_def)
5039  done
5040
5041
5042definition
5043  createObject_hs_preconds :: "word32 \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> kernel_state \<Rightarrow> bool"
5044where
5045  "createObject_hs_preconds regionBase newType userSize d \<equiv>
5046     (invs' and pspace_no_overlap' regionBase (getObjectSize newType userSize)
5047           and caps_overlap_reserved' {regionBase ..+ 2 ^ (getObjectSize newType userSize)}
5048           and (\<lambda>s. 2 ^ (getObjectSize newType userSize) \<le> gsMaxObjectSize s)
5049           and K(regionBase \<noteq> 0
5050                   \<and> ({regionBase..+2 ^ (getObjectSize newType userSize)} \<inter> kernel_data_refs = {})
5051                   \<and> range_cover regionBase (getObjectSize newType userSize) (getObjectSize newType userSize) (Suc 0)
5052                   \<and> (newType = APIObjectType apiobject_type.Untyped \<longrightarrow> userSize \<le> maxUntypedSizeBits)
5053                   \<and> (newType = APIObjectType apiobject_type.CapTableObject \<longrightarrow> userSize < 28)
5054                   \<and> (newType = APIObjectType apiobject_type.Untyped \<longrightarrow> minUntypedSizeBits \<le> userSize)
5055                   \<and> (newType = APIObjectType apiobject_type.CapTableObject \<longrightarrow> 0 < userSize)
5056                   \<and> (d \<longrightarrow> newType = APIObjectType apiobject_type.Untyped \<or> isFrameType newType)
5057           ))"
5058
5059abbreviation
5060  "region_actually_is_dev_bytes ptr len devMem s
5061    \<equiv> region_actually_is_bytes ptr len s
5062        \<and> (\<not> devMem \<longrightarrow> heap_list_is_zero (hrs_mem (t_hrs_' (globals s))) ptr len)"
5063
5064(* these preconds actually used throughout the proof *)
5065abbreviation(input)
5066  createObject_c_preconds1 :: "word32 \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
5067where
5068  "createObject_c_preconds1 regionBase newType userSize deviceMemory \<equiv>
5069    {s. region_actually_is_dev_bytes regionBase (2 ^ getObjectSize newType userSize) deviceMemory s}"
5070
5071(* these preconds used at start of proof *)
5072definition
5073  createObject_c_preconds :: "word32 \<Rightarrow> ArchTypes_H.object_type \<Rightarrow> nat \<Rightarrow> bool \<Rightarrow> (globals myvars) set"
5074where
5075  "createObject_c_preconds regionBase newType userSize deviceMemory \<equiv>
5076  (createObject_c_preconds1 regionBase newType userSize deviceMemory
5077           \<inter> {s. object_type_from_H newType = t_' s}
5078           \<inter> {s. Ptr regionBase = regionBase_' s}
5079           \<inter> {s. unat (scast (userSize_' s) :: word32) = userSize}
5080           \<inter> {s. to_bool (deviceMemory_' s) = deviceMemory}
5081     )"
5082
5083lemma ccorres_apiType_split:
5084  "\<lbrakk> apiType = apiobject_type.Untyped \<Longrightarrow> ccorres rr xf P1 P1' hs X Y;
5085     apiType = apiobject_type.TCBObject \<Longrightarrow> ccorres rr xf P2 P2' hs X Y;
5086     apiType = apiobject_type.EndpointObject \<Longrightarrow> ccorres rr xf P3 P3' hs X Y;
5087     apiType = apiobject_type.NotificationObject \<Longrightarrow> ccorres rr xf P4 P4' hs X Y;
5088     apiType = apiobject_type.CapTableObject \<Longrightarrow> ccorres rr xf P5 P5' hs X Y
5089   \<rbrakk> \<Longrightarrow> ccorres rr xf
5090         ((\<lambda>s. apiType = apiobject_type.Untyped \<longrightarrow> P1 s)
5091         and (\<lambda>s. apiType = apiobject_type.TCBObject \<longrightarrow> P2 s)
5092         and (\<lambda>s. apiType = apiobject_type.EndpointObject \<longrightarrow> P3 s)
5093         and (\<lambda>s. apiType = apiobject_type.NotificationObject \<longrightarrow> P4 s)
5094         and (\<lambda>s. apiType = apiobject_type.CapTableObject \<longrightarrow> P5 s))
5095         ({s. apiType = apiobject_type.Untyped \<longrightarrow> s \<in> P1'}
5096         \<inter> {s. apiType = apiobject_type.TCBObject \<longrightarrow> s \<in> P2'}
5097         \<inter> {s. apiType = apiobject_type.EndpointObject \<longrightarrow> s \<in> P3'}
5098         \<inter> {s. apiType = apiobject_type.NotificationObject \<longrightarrow> s \<in> P4'}
5099         \<inter> {s. apiType = apiobject_type.CapTableObject \<longrightarrow> s \<in> P5'})
5100         hs X Y"
5101  apply (case_tac apiType, simp_all)
5102  done
5103
5104
5105lemma range_cover_simpleI:
5106  "\<lbrakk> is_aligned (ptr :: 'a :: len word) a; a < len_of TYPE('a); c = Suc 0 \<rbrakk>
5107  \<Longrightarrow> range_cover ptr a a c"
5108  apply (clarsimp simp: range_cover_def)
5109  apply (metis shiftr_0 is_aligned_mask unat_0)
5110  done
5111
5112lemma range_coverI:
5113  "\<lbrakk>is_aligned (ptr :: 'a :: len word) a; b \<le> a; a < len_of TYPE('a);
5114    c \<le> 2 ^ (a - b)\<rbrakk>
5115  \<Longrightarrow> range_cover ptr a b c"
5116  apply (clarsimp simp: range_cover_def field_simps)
5117  apply (rule conjI)
5118   apply (erule(1) is_aligned_weaken)
5119  apply (subst mask_zero, simp)
5120  apply simp
5121  done
5122
5123(* FIXME: with the current state of affairs, we could simplify gs_new_frames *)
5124lemma gsUserPages_update_ccorres:
5125  "ccorresG rf_sr G dc xf (\<lambda>_. sz = pageBitsForSize pgsz) UNIV hs
5126     (modify (gsUserPages_update (\<lambda>m a. if a = ptr then Some pgsz else m a)))
5127     (Basic (globals_update (ghost'state_'_update
5128                  (gs_new_frames pgsz ptr sz))))"
5129  apply (rule ccorres_from_vcg)
5130  apply vcg_step
5131  apply (clarsimp simp: split_def simpler_modify_def gs_new_frames_def)
5132  apply (case_tac "ghost'state_' (globals x)")
5133  apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def fun_upd_def
5134                        carch_state_relation_def cmachine_state_relation_def
5135                        ghost_size_rel_def ghost_assertion_data_get_def
5136                  cong: if_cong)
5137  done
5138
5139lemma placeNewObject_user_data_device:
5140  "ccorresG rf_sr \<Gamma> dc xfdc
5141  (pspace_aligned' and pspace_distinct'
5142    and ret_zero regionBase (2 ^ (pageBits + us))
5143    and pspace_no_overlap' regionBase (pageBits+us) and valid_queues
5144    and (\<lambda>s. sym_refs (state_refs_of' s))
5145    and (\<lambda>s. 2^(pageBits +  us) \<le> gsMaxObjectSize s)
5146    and K (regionBase \<noteq> 0 \<and> range_cover regionBase (pageBits + us) (pageBits+us) (Suc 0)
5147    \<and>  {regionBase..+2^(pageBits +  us)} \<inter> kernel_data_refs = {}))
5148  ({s. region_actually_is_bytes regionBase (2^(pageBits+us)) s})
5149  hs
5150  (placeNewObject regionBase UserDataDevice us )
5151  (global_htd_update (\<lambda>s. (ptr_retyps (2^us) (Ptr regionBase :: user_data_device_C ptr))))"
5152  apply (rule ccorres_from_vcg_nofail)
5153  apply (clarsimp simp:)
5154  apply (rule conseqPre)
5155  apply vcg
5156  apply (clarsimp simp: rf_sr_htd_safe)
5157  apply (intro conjI allI impI)
5158   apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
5159                         kernel_data_refs_domain_eq_rotate
5160                  elim!: ptr_retyps_htd_safe_neg[where arr=False,
5161                        unfolded ptr_retyps_gen_def, simplified])
5162   apply (simp add: size_of_def pageBits_def power_add mult.commute mult.left_commute)
5163  apply (frule range_cover.unat_of_nat_shift[where gbits = "pageBits + us"])
5164    apply simp
5165   apply (clarsimp simp:size_of_def power_add pageBits_def
5166     rf_sr_def cstate_relation_def Let_def field_simps)
5167   apply blast
5168  apply (frule range_cover.aligned)
5169  apply (frule range_cover.sz(1), fold word_bits_def)
5170  apply (rule bexI [OF _ placeNewObject_eq], simp_all)
5171    apply (cut_tac ptr=regionBase and sz="pageBits + us" and gbits=us and arr=False
5172                in createObjects_ccorres_user_data_device[rule_format])
5173     apply (rule conjI, assumption, clarsimp)
5174     apply (fastforce simp: pageBits_def field_simps region_actually_is_bytes)
5175    apply (clarsimp elim!: rsubst[where P="\<lambda>x. (\<sigma>, x) \<in> rf_sr" for \<sigma>]
5176                     simp: field_simps objBitsKO_def ptr_retyps_gen_def)
5177   apply (simp add: objBitsKO_def field_simps)
5178  apply (rule no_fail_pre, rule no_fail_placeNewObject)
5179  apply (clarsimp simp: objBitsKO_def)
5180  done
5181
5182lemma createObjects'_page_directory_at_global:
5183  "\<lbrace> \<lambda>s. n \<noteq> 0 \<and> range_cover ptr sz (objBitsKO val + gbits) n
5184      \<and> pspace_aligned' s \<and> pspace_distinct' s \<and> pspace_no_overlap' ptr sz s
5185      \<and> page_directory_at' (armKSGlobalPD (ksArchState s)) s \<rbrace>
5186    createObjects' ptr n val gbits
5187  \<lbrace> \<lambda>rv s. page_directory_at' (armKSGlobalPD (ksArchState s)) s \<rbrace>"
5188  apply (simp add: page_directory_at'_def)
5189  apply (rule hoare_pre, wp hoare_vcg_all_lift hoare_vcg_const_imp_lift)
5190   apply (wps createObjects'_ksArch)
5191   apply (wp createObjects'_typ_at[where sz=sz])
5192  apply simp
5193  done
5194
5195lemma gsUserPages_update:
5196    "\<And>f. (\<lambda>s. s\<lparr>gsUserPages := f(gsUserPages s)\<rparr>) = gsUserPages_update f"
5197    by (rule ext) simp
5198
5199lemma modify_gsUserPages_update:
5200  "modify (\<lambda>s. s\<lparr>gsUserPages := f(gsUserPages s)\<rparr>) = modify (gsUserPages_update f)"
5201  by (simp only: gsUserPages_update)
5202
5203method arch_create_data_obj_corres_helper =
5204  (match conclusion in "ccorres ?rel ?var ?P ?P' ?hs
5205            (ARM_HYP_H.createObject object_type.SmallPageObject ?regionBase sz ?deviceMemory)
5206            (Call Arch_createObject_'proc)
5207             " for sz \<Rightarrow> \<open>(simp add: toAPIType_def ARM_HYP_H.createObject_def
5208                placeNewDataObject_def bind_assoc
5209               ARMLargePageBits_def),subst gsUserPages_update,((rule ccorres_gen_asm)+) \<close>)
5210
5211lemma placeNewDataObject_ccorres:
5212  "ccorresG rf_sr \<Gamma> dc xfdc
5213  (createObject_hs_preconds regionBase newType us devMem
5214      and K (APIType_capBits newType us = pageBits + us))
5215  ({s. region_actually_is_bytes regionBase (2 ^ (pageBits + us)) s
5216      \<and> (\<not> devMem \<longrightarrow> heap_list_is_zero (hrs_mem (t_hrs_' (globals s))) regionBase
5217          (2 ^ (pageBits + us)))})
5218  hs
5219  (placeNewDataObject regionBase us devMem)
5220  (Cond {s. devMem}
5221    (global_htd_update (\<lambda>s. (ptr_retyps (2^us) (Ptr regionBase :: user_data_device_C ptr))))
5222    (global_htd_update (\<lambda>s. (ptr_retyps (2^us) (Ptr regionBase :: user_data_C ptr))))
5223  )"
5224  apply (cases devMem)
5225   apply (simp add: placeNewDataObject_def ccorres_cond_univ_iff)
5226   apply (rule ccorres_guard_imp, rule placeNewObject_user_data_device, simp_all)
5227   apply (clarsimp simp: createObject_hs_preconds_def invs'_def
5228                         valid_state'_def valid_pspace'_def)
5229  apply (simp add: placeNewDataObject_def ccorres_cond_empty_iff)
5230  apply (rule ccorres_guard_imp, rule placeNewObject_user_data, simp_all)
5231  apply (clarsimp simp: createObject_hs_preconds_def invs'_def
5232                        valid_state'_def valid_pspace'_def)
5233  apply (frule range_cover.sz(1), simp add: word_bits_def)
5234  done
5235
5236lemma cond_second_eq_seq_ccorres:
5237  "ccorres_underlying sr Gamm r xf arrel axf G G' hs m
5238        (Cond P (a ;; c) (b ;; c) ;; d)
5239    = ccorres_underlying sr Gamm r xf arrel axf G G' hs m
5240        (Cond P a b ;; c ;; d)"
5241  apply (rule ccorres_semantic_equiv)
5242  apply (rule semantic_equivI)
5243  apply (auto elim!: exec_Normal_elim_cases intro: exec.Seq exec.CondTrue exec.CondFalse)
5244  done
5245
5246
5247lemma cvariable_array_ptr_kill:
5248  "cvariable_array_map_relation m ns ptrfun htd
5249    \<Longrightarrow> cvariable_array_map_relation (m(x := None))
5250          ns (ptrfun :: _ \<Rightarrow> ('b :: mem_type) ptr) htd"
5251  by (clarsimp simp: cvariable_array_map_relation_def
5252              split: if_split)
5253
5254(* FIXME move, depends on setObject_modify which lives in kernel_m *)
5255(* FIXME would be interesting to generalise this kind of lemma *)
5256lemma monadic_rewrite_setObject_vcpu_modify:
5257  fixes vcpu vcpupre :: vcpu
5258  shows "monadic_rewrite True False (vcpu_at' v)
5259           (setObject v vcpu)
5260           (modify (ksPSpace_update (\<lambda>ps. ps(v \<mapsto> injectKOS vcpu))))"
5261  by (clarsimp simp: monadic_rewrite_def setObject_modify objBits_simps archObjSize_def
5262                     vcpuBits_def machine_bits_defs typ_at_to_obj_at_arches)
5263
5264(* FIXME move, depends on setObject_modify which lives in kernel_m *)
5265lemma monadic_rewrite_modify_setObject_vcpu:
5266  fixes vcpu vcpupre :: vcpu
5267  shows "monadic_rewrite True False (vcpu_at' v)
5268           (modify (ksPSpace_update (\<lambda>ps. ps(v \<mapsto> injectKOS vcpu))))
5269           (setObject v vcpu)"
5270  by (clarsimp simp: monadic_rewrite_def setObject_modify objBits_simps archObjSize_def
5271                     vcpuBits_def machine_bits_defs typ_at_to_obj_at_arches)
5272
5273lemma monadic_rewrite_placeNewObject_vcpu_decompose:
5274  fixes vcpu vcpupre :: vcpu
5275  shows "monadic_rewrite True False \<top>
5276           (placeNewObject v vcpu 0)
5277           (do placeNewObject v vcpupre 0;
5278               setObject v vcpu
5279            od)"
5280  apply (rule monadic_rewrite_imp)
5281   apply (rule monadic_rewrite_trans[rotated])
5282    apply (rule monadic_rewrite_bind_tail)
5283     apply clarsimp
5284     apply (rule monadic_rewrite_modify_setObject_vcpu)
5285    apply (rule hoare_post_imp[OF _ placeNewObject_creates_object_vcpu])
5286    apply (fastforce simp: ko_at_vcpu_at'D)
5287   apply (clarsimp simp: placeNewObject_def placeNewObject'_def bind_assoc split_def)
5288   apply (clarsimp simp: objBits_simps' archObjSize_def)
5289   apply (rule monadic_rewrite_bind_tail)+
5290      apply (rule monadic_rewrite_is_refl)
5291      apply (rule ext)
5292      apply (clarsimp simp: exec_modify)
5293      apply (fastforce simp: simpler_modify_def comp_def)
5294     apply (wpsimp wp: hoare_vcg_imp_lift' hoare_vcg_ex_lift)+
5295  done
5296
5297(* FIXME would be interesting to generalise this *)
5298lemma monadic_rewrite_setObject_vcpu_twice:
5299  fixes vcpu vcpupre :: vcpu
5300  shows "monadic_rewrite True False (vcpu_at' v)
5301           (setObject v vcpu)
5302           (do setObject v vcpupre;
5303               setObject v vcpu
5304            od)"
5305  supply fun_upd_apply[simp del]
5306  apply (rule monadic_rewrite_imp)
5307   apply (rule monadic_rewrite_trans[rotated])
5308    apply (rule monadic_rewrite_bind_head)
5309    apply (rule monadic_rewrite_modify_setObject_vcpu)
5310   apply (rule monadic_rewrite_trans[rotated])
5311    apply (rule monadic_rewrite_bind_tail)
5312     apply clarsimp
5313     apply (rule monadic_rewrite_modify_setObject_vcpu)
5314    apply wp
5315   apply (rule monadic_rewrite_trans)
5316    apply (rule monadic_rewrite_setObject_vcpu_modify)
5317   apply (rule monadic_rewrite_is_refl)
5318   apply (rule ext)
5319   apply (clarsimp simp: exec_modify)
5320   apply (fastforce simp: simpler_modify_def comp_def)
5321  apply clarsimp
5322  apply (clarsimp simp: vcpu_at_ko'_eq[symmetric] obj_at'_def projectKOs fun_upd_apply
5323                        objBits_simps archObjSize_def vcpuBits_def machine_bits_defs ps_clear_def
5324                  split: if_splits)
5325  done
5326
5327(* The usual way we deal with this is in terms of default object construction, but when you
5328   put in an arbitrary object that's not cte/tcb related, we don't need to care about putting
5329   it in memory for cte_wp_at'
5330   If you do have a cte or tcb *)
5331lemma cte_wp_at_retype'_not_cte_tcb:
5332  fixes obj :: kernel_object
5333  fixes s :: kernel_state
5334  fixes addrs :: "addr list"
5335  defines "s' \<equiv> s\<lparr>ksPSpace := (\<lambda>x. if x \<in> set addrs then Some obj else ksPSpace s x) \<rparr>"
5336  assumes pv: "pspace_aligned' s" "pspace_distinct' s"
5337  and pv': "pspace_aligned' s'" "pspace_distinct' s'"
5338  and al: "\<forall>x \<in> set addrs. is_aligned x (objBitsKO obj)"
5339  and pn: "\<forall>x \<in> set addrs. ksPSpace s x = None"
5340  and irrelko: "koTypeOf obj \<noteq> CTET" "koTypeOf obj \<noteq> TCBT"
5341  shows "cte_wp_at' P p s' = cte_wp_at' P p s"
5342proof -
5343  have not_tcb_cte: "\<And>ko. obj \<noteq> KOTCB ko" "\<And>ko. obj \<noteq> KOCTE ko"
5344    using irrelko by (clarsimp simp: irrelko)+
5345
5346  show ?thesis
5347    apply (subgoal_tac "\<forall>p \<in> set addrs. \<forall>(P :: cte \<Rightarrow> bool). \<not> obj_at' P p s")
5348     prefer 2
5349     apply (fastforce elim!: obj_atE' simp: pn)
5350    apply (subgoal_tac "\<forall>p \<in> set addrs. \<forall>(P :: tcb \<Rightarrow> bool). \<not> obj_at' P p s")
5351     prefer 2
5352     apply (fastforce elim!: obj_atE' simp: pn)
5353    apply (simp only: cte_wp_at_obj_cases_mask)
5354    apply (clarsimp simp: s'_def)
5355    apply (clarsimp simp: foldr_update_obj_at'[OF pv pv'[simplified s'_def]] al)
5356    apply (clarsimp cong: if_cong split: if_split)
5357    apply (clarsimp simp: projectKOs dom_def)
5358    apply (case_tac obj; clarsimp simp: not_tcb_cte)
5359    done
5360qed
5361
5362lemma ctes_of_retype_not_cte_tcb:
5363  fixes obj :: kernel_object
5364  fixes s :: kernel_state
5365  fixes addrs :: "addr list"
5366  defines "s' \<equiv> s\<lparr>ksPSpace := (\<lambda>x. if x \<in> set addrs then Some obj else ksPSpace s x) \<rparr>"
5367  assumes pv: "pspace_aligned' s" "pspace_distinct' s"
5368  and pv': "pspace_aligned' s'" "pspace_distinct' s'"
5369  and al: "\<forall>x \<in> set addrs. is_aligned x (objBitsKO obj)"
5370  and pn: "\<forall>x \<in> set addrs. ksPSpace s x = None"
5371  and irrelko: "koTypeOf obj \<noteq> CTET" "koTypeOf obj \<noteq> TCBT"
5372  shows "map_to_ctes (\<lambda>x. if x \<in> set addrs then Some obj else ksPSpace s x)
5373          = (\<lambda>x. map_to_ctes (ksPSpace s) x)"
5374  (is "map_to_ctes ?ps' = ?map'")
5375  using cte_wp_at_retype'_not_cte_tcb[where P="(=) cte" for cte, OF pv _ _ al pn] pv' irrelko
5376        arg_cong [where f=Not, OF cte_wp_at_retype'_not_cte_tcb [OF pv _ _ al pn, where P="\<top>"]]
5377  apply simp
5378  apply (simp add: s'_def)
5379  apply (simp(no_asm_use) add: cte_wp_at_ctes_of cong: if_cong)
5380  apply (rule ext)
5381  apply (case_tac "map_to_ctes ?ps' x")
5382   apply (simp(no_asm_simp))
5383   apply (drule_tac x=x in meta_spec)+
5384   apply fastforce
5385  apply (simp split: if_splits)
5386  done
5387
5388(* for when we do not have direct access to s *)
5389lemma magnitudeCheck_assert2':
5390  "\<lbrakk> is_aligned x n; (1 :: word32) < 2 ^ n; ksPSpace s x = Some v ;
5391     ksPSpace (s::kernel_state) = ps \<rbrakk> \<Longrightarrow>
5392   magnitudeCheck x (snd (lookupAround2 x ps)) n
5393     = assert (ps_clear x n s)"
5394  by (clarsimp simp: magnitudeCheck_assert2)
5395
5396(* VCPU when recast from memset 0 *)
5397abbreviation (input)
5398  fromzeroVCPU :: vcpu
5399where
5400  "fromzeroVCPU \<equiv> vcpu.VCPUObj None (VGICInterface 0 0 0 (\<lambda>_. 0)) (const 0)"
5401
5402lemma monadic_rewrite_setObject_vcpu_as_init:
5403  defines "vcpu0 \<equiv> fromzeroVCPU"
5404  defines "vcpu1 \<equiv> (vcpuRegs_update (\<lambda>_. (vcpuRegs vcpu0)(VCPURegSCTLR := sctlrDefault)) vcpu0)"
5405  defines "vcpu2 \<equiv> (vcpuRegs_update (\<lambda>_. (vcpuRegs vcpu1)(VCPURegACTLR := actlrDefault)) vcpu1)"
5406  shows
5407  "monadic_rewrite True False (K (v \<noteq> 0) and ko_at' fromzeroVCPU v)
5408     (setObject v makeVCPUObject)
5409     (do vcpuWriteReg v VCPURegSCTLR sctlrDefault;
5410         vcpuWriteReg v VCPURegACTLR actlrDefault;
5411         vgicUpdate v (vgicHCR_update (\<lambda>_. vgicHCREN))
5412      od)
5413     "
5414  supply fun_upd_apply[simp del]
5415  apply (simp add: K_def)
5416  apply (rule monadic_rewrite_gen_asm)
5417  apply (rule monadic_rewrite_imp)
5418   apply (simp add: vcpuWriteReg_def vgicUpdate_def bind_assoc)
5419   apply (rule monadic_rewrite_trans[rotated])
5420    apply (clarsimp simp: vcpuUpdate_def bind_assoc)
5421    apply (rule monadic_rewrite_symb_exec_r)
5422       apply wp+
5423     apply (rename_tac vcpu)
5424     apply (rule_tac P="vcpu = vcpu0" in monadic_rewrite_gen_asm, simp)
5425     apply (rule monadic_rewrite_bind_tail)
5426      apply (rule monadic_rewrite_symb_exec_r)
5427         apply wp+
5428       apply (rename_tac vcpu')
5429       apply (rule_tac P="vcpu' = vcpu1" in monadic_rewrite_gen_asm)
5430       apply simp
5431       apply (rule monadic_rewrite_bind_tail)
5432        apply (rule monadic_rewrite_symb_exec_r)
5433           apply wp+
5434         apply (rename_tac vcpu'')
5435         apply (rule_tac P="vcpu'' = vcpu2" in monadic_rewrite_gen_asm, simp)
5436         apply (rule monadic_rewrite_refl)
5437        apply (wpsimp wp: getObject_vcpu_prop simp: vcpu1_def vcpu2_def vcpu0_def)+
5438       apply (wp setObject_sets_object_vcpu)
5439      apply (wpsimp wp: getObject_vcpu_prop)+
5440     apply (wpsimp wp: getObject_vcpu_prop simp: vcpu1_def vcpu2_def vcpu0_def)+
5441     apply (wp setObject_sets_object_vcpu)
5442    apply (wpsimp wp: getObject_vcpu_prop)+
5443   (* now we have three setObjects in a row *)
5444   apply (rule monadic_rewrite_trans[rotated])
5445    apply (rule monadic_rewrite_bind_tail)
5446     apply (rule monadic_rewrite_setObject_vcpu_twice[simplified])
5447    apply wp
5448   apply (rule monadic_rewrite_trans[rotated])
5449    apply (rule monadic_rewrite_setObject_vcpu_twice[simplified])
5450   apply (rule monadic_rewrite_is_refl)
5451   apply (fastforce simp: vcpu2_def vcpu1_def vcpu0_def makeVCPUObject_def)
5452  apply (fastforce simp: vcpu0_def ko_at_vcpu_at'D)
5453  done
5454
5455lemma ptr_retyp_fromzeroVCPU:
5456  fixes p' :: "vcpu_C ptr"
5457  defines "vcpu0 \<equiv> fromzeroVCPU"
5458  defines "ko_vcpu \<equiv> KOArch (KOVCPU vcpu0)"
5459  assumes "valid_global_refs' \<sigma>"
5460  assumes pal: "pspace_aligned' \<sigma>"
5461  assumes pdst: "pspace_distinct' \<sigma>"
5462  assumes pno: "pspace_no_overlap' p vcpu_bits \<sigma>"
5463  assumes "2 ^ vcpu_bits \<le> gsMaxObjectSize \<sigma>"
5464  assumes vo: "valid_objs' \<sigma>"
5465  assumes urz: "untyped_ranges_zero' \<sigma>"
5466  assumes cor: "caps_overlap_reserved' {p ..+ 2 ^ vcpu_bits} \<sigma>"
5467  assumes ptr0: "p \<noteq> 0"
5468  assumes kdr: "{p ..+ 2 ^ vcpu_bits} \<inter> kernel_data_refs = {}"
5469  assumes subr: "{p ..+ 428} \<subseteq> {p ..+ 2 ^ vcpu_bits}"
5470  assumes act_bytes: "region_actually_is_bytes p (2 ^ vcpu_bits) \<sigma>'"
5471  assumes rep0: "heap_list (hrs_mem (t_hrs_' (globals \<sigma>'))) (2 ^ vcpu_bits) p = replicate (2 ^ vcpu_bits) 0"
5472  assumes "\<not> snd (placeNewObject p vcpu0 0 \<sigma>)"
5473  assumes cover: "range_cover p vcpu_bits vcpu_bits 1"
5474  assumes al: "is_aligned p vcpu_bits"
5475  assumes sr: "(\<sigma>, \<sigma>') \<in> rf_sr"
5476  shows "(\<sigma>\<lparr>ksPSpace := ksPSpace \<sigma>(p \<mapsto> ko_vcpu)\<rparr>,
5477           globals_update (t_hrs_'_update (hrs_htd_update (ptr_retyp (vcpu_Ptr p)))) \<sigma>')
5478         \<in> rf_sr"
5479         (is "(\<sigma>\<lparr>ksPSpace := ?ks\<rparr>, globals_update ?gs' \<sigma>') \<in> rf_sr")
5480proof -
5481  let ?gs = "?gs' (globals \<sigma>')"
5482  let ?s' = "\<sigma>\<lparr>ksPSpace := ?ks\<rparr>"
5483  let ?htdret = "(hrs_htd_update (ptr_retyp (vcpu_Ptr p)) (t_hrs_' (globals \<sigma>')))"
5484  let ?zeros = "from_bytes (replicate (size_of TYPE(vcpu_C)) 0) :: vcpu_C"
5485
5486  have "size_of TYPE(vcpu_C) = 428" (is "_ = ?vcpusz")
5487    by simp
5488
5489  have ptr_al:
5490    "ptr_aligned (vcpu_Ptr p)" using al
5491    by (auto simp: align_of_def vcpu_bits_def pageBits_def
5492           intro!: is_aligned_ptr_aligned[of _ 2]
5493            elim!: is_aligned_weaken)
5494
5495  have "c_null_guard (vcpu_Ptr p)" using ptr0 al
5496    by (auto simp: c_null_guard_def vcpu_bits_def pageBits_def vcpu_C_size
5497            elim!: intvl_nowrap[where x = 0, simplified] is_aligned_no_wrap_le)
5498
5499  have cguard: "c_guard (vcpu_Ptr p)"
5500    using \<open>ptr_aligned (vcpu_Ptr p)\<close> \<open>c_null_guard (vcpu_Ptr p)\<close>
5501    by (simp add: c_guard_def)
5502
5503  have zro:
5504    "zero_ranges_are_zero (gsUntypedZeroRanges \<sigma>) (t_hrs_' (globals \<sigma>'))"
5505    using sr
5506    by (clarsimp simp: rf_sr_def cstate_relation_def Let_def)
5507
5508  have retyp_p': "ptr_retyp (vcpu_Ptr p) (snd (t_hrs_' (globals \<sigma>'))) \<Turnstile>\<^sub>t (vcpu_Ptr p)"
5509    using cguard by (rule ptr_retyp_h_t_valid)
5510
5511  have clift_retyp_p':
5512    "clift (hrs_mem (t_hrs_' (globals \<sigma>')), ptr_retyp (vcpu_Ptr p) (snd (t_hrs_' (globals \<sigma>')))) (vcpu_Ptr p)
5513          = Some (from_bytes (replicate (size_of TYPE(vcpu_C)) 0))"
5514    (is "?hl = ?rep0")
5515    using retyp_p' rep0
5516    apply (simp add: lift_t_if h_val_def)
5517    apply (subst take_heap_list_le[where k="?vcpusz" and n="2^vcpu_bits", symmetric])
5518     apply (simp add: vcpu_bits_def')+
5519    done
5520
5521  have htd_safe:
5522      "htd_safe (- kernel_data_refs) (hrs_htd (t_hrs_' (globals \<sigma>')))
5523          \<Longrightarrow> htd_safe (- kernel_data_refs) (hrs_htd (t_hrs_' ?gs))"
5524    using kdr
5525    apply (simp add: hrs_htd_update)
5526    apply (intro ptr_retyp_htd_safe_neg ptr_retyps_htd_safe_neg, simp_all)
5527    apply (erule disjoint_subset[rotated])
5528    apply (clarsimp; rule intvl_mem_weaken[where n="2 ^ vcpu_bits - size_of (TYPE (vcpu_C))"])
5529    by (clarsimp simp: vcpu_bits_def')
5530
5531  have pks: "map_to_vcpus (ksPSpace \<sigma>) p = None"
5532    apply (rule map_comp_simps)
5533    apply (rule pspace_no_overlap_base'[OF pal pno al, simplified])
5534    done
5535
5536  note cmap_vcpus = cmap_relation_updI2[where am="map_to_vcpus (ksPSpace \<sigma>)"
5537                                        and cm="cslift \<sigma>'"
5538                                        and dest=p and f=vcpu_Ptr,
5539                                        OF _ pks, simplified]
5540
5541  have map_vcpus:
5542    "cmap_relation (map_to_vcpus (ksPSpace \<sigma>)) (cslift \<sigma>') vcpu_Ptr cvcpu_relation
5543     \<Longrightarrow> cmap_relation (map_to_vcpus (ksPSpace \<sigma>)(p \<mapsto> vcpu0))
5544                      (cslift \<sigma>'(vcpu_Ptr p \<mapsto> ?zeros)) vcpu_Ptr cvcpu_relation"
5545    apply (erule cmap_vcpus)
5546    apply (simp add: vcpu0_def from_bytes_def)
5547    apply (simp add: typ_info_simps vcpu_C_tag_def)
5548    apply (simp add: ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td align_of_def padup_def
5549                     final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def)
5550    apply (simp add: update_ti_adjust_ti update_ti_t_word32_0s
5551                     typ_info_simps gicVCpuIface_C_tag_def virq_C_tag_def
5552                     update_ti_t_ptr_0s ti_typ_pad_combine_empty_ti ti_typ_pad_combine_td
5553                     ti_typ_combine_empty_ti ti_typ_combine_td
5554                     align_of_def padup_def replicate_def update_ti_t_array_rep
5555                     final_pad_def size_td_lt_ti_typ_pad_combine Let_def size_of_def
5556                     align_td_array' size_td_array)
5557    apply (clarsimp simp: cvcpu_relation_def cvcpu_regs_relation_def option_to_ctcb_ptr_def
5558                          cvgic_relation_def virq_to_H_def)
5559    apply (rule conjI; clarsimp)
5560     (* regs_C array initialisation *)
5561     using le_imp_less_Suc[OF maxBound_is_bound[where 'a=vcpureg,
5562                                                simplified fromEnum_maxBound_vcpureg_def]]
5563     apply (case_tac r; clarsimp simp: index_foldr_update)
5564    (* vgic_C array initialisation *)
5565    apply (subst index_fold_update ; clarsimp)
5566    done
5567
5568  have is_bytes: "region_is_bytes p (size_of TYPE(vcpu_C)) \<sigma>'"
5569    using region_actually_is_bytes[OF act_bytes] region_is_bytes_subset[OF _ subr]
5570    by (simp add: vcpu_C_size)
5571
5572  have zero_bytes [simplified]:
5573    "heap_list (hrs_mem (t_hrs_' (globals \<sigma>'))) (size_of TYPE(vcpu_C)) p
5574      = replicate (size_of TYPE(vcpu_C)) 0"
5575    apply (subst take_heap_list_le[where n="2^vcpu_bits", symmetric])
5576     prefer 2
5577     apply (subst rep0)
5578     apply (simp add: vcpu_bits_def')+
5579    done
5580
5581  have clift_retyp_p':
5582    "clift (?htdret) (vcpu_Ptr p) = Some (from_bytes (replicate (size_of TYPE(vcpu_C)) 0))"
5583    using cguard zero_bytes
5584    by (case_tac "t_hrs_' (globals \<sigma>')")
5585      (auto simp: lift_t_if hrs_htd_update_def h_val_def hrs_mem_def ptr_retyp_h_t_valid)
5586
5587  have cl_vcpu:
5588    "(clift (?htdret) :: vcpu_C typ_heap) = (cslift \<sigma>')(vcpu_Ptr p \<mapsto> ?zeros)"
5589    using cguard clift_retyp_p'
5590    apply clarsimp
5591    apply (rule ext, rename_tac p')
5592    apply (case_tac "p' = vcpu_Ptr p", fastforce)
5593    using clift_ptr_retyps_gen_prev_memset_same[
5594             where 'a=vcpu_C and arr=False and n=1 and hrs="(t_hrs_' (globals \<sigma>'))",
5595             simplified, OF cguard _ refl] is_bytes[simplified]
5596    apply (simp add: word_bits_def zero_bytes ptr_retyps_gen_def)
5597    done
5598
5599  have zro':
5600    "zero_ranges_are_zero (gsUntypedZeroRanges \<sigma>) (t_hrs_' ?gs)"
5601    using zro vo urz cor al
5602    apply (clarsimp simp: hrs_htd_update typ_heap_simps')
5603    apply (intro zero_ranges_ptr_retyps[where n=1 and arr=False,
5604                                          simplified ptr_retyps_gen_def, simplified])
5605       apply simp_all
5606    apply (erule caps_overlap_reserved'_subseteq; clarsimp)
5607    apply (rule intvl_mem_weaken[where n="2 ^ vcpu_bits - size_of (TYPE (vcpu_C))"])
5608    apply (clarsimp simp: vcpu_bits_def')
5609    done
5610
5611  have cl_rest:
5612    "\<lbrakk> typ_uinfo_t TYPE(vcpu_C) \<bottom>\<^sub>t typ_uinfo_t TYPE('z); typ_uinfo_t TYPE('z) \<noteq> typ_uinfo_t TYPE(word8) \<rbrakk>
5613     \<Longrightarrow> (clift ?htdret) =  (cslift \<sigma>' :: 'z::mem_type typ_heap)"
5614    using cguard
5615    apply (clarsimp simp: clift_heap_update_same hrs_htd_update
5616                            ptr_retyp_h_t_valid typ_heap_simps)
5617    apply (subst ptr_retyp_gen_one[symmetric])
5618    apply (rule clift_ptr_retyps_gen_other)
5619    using region_actually_is_bytes[OF act_bytes] region_is_bytes_subset[OF _ subr]
5620    apply (auto simp: word_bits_def vcpu_bits_def')
5621    done
5622
5623  have arel: "cte_array_relation \<sigma> ?gs \<and> tcb_cte_array_relation ?s' ?gs"
5624    using sr
5625    apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
5626                           hrs_htd_update ko_vcpu_def map_comp_update)
5627    apply (intro conjI cvariable_array_ptr_kill
5628                  cvariable_array_ptr_retyps[OF refl, where n=1 and arr=False,
5629                    simplified ptr_retyps_gen_def, simplified])
5630    using  region_actually_is_bytes[OF act_bytes] region_is_bytes_subset[OF _ subr]
5631    by auto
5632
5633  note ht_rest = clift_eq_h_t_valid_eq[OF cl_rest, simplified ko_vcpu_def, simplified]
5634
5635  have objBitsKO_vcpu:
5636    "\<And>v. objBitsKO (KOArch (KOVCPU v)) = vcpu_bits"
5637    by (simp add: objBitsKO_def archObjSize_def vcpu_bits_def' vcpuBits_def')
5638
5639  have rl_vcpu:
5640    "(projectKO_opt \<circ>\<^sub>m (ksPSpace \<sigma>(p \<mapsto> KOArch (KOVCPU vcpu0))) :: word32 \<Rightarrow> vcpu option)
5641      = (projectKO_opt \<circ>\<^sub>m ksPSpace \<sigma>)(p \<mapsto> vcpu0)"
5642    by (rule ext)
5643       (clarsimp simp: projectKOs map_comp_def vcpu0_def split: if_split)
5644
5645  have ctes:
5646    "map_to_ctes (ksPSpace \<sigma>(p \<mapsto> KOArch (KOVCPU vcpu0))) = ctes_of \<sigma>"
5647    using pal pdst al pno
5648    apply (clarsimp simp: fun_upd_def)
5649    apply (frule (2) pspace_no_overlap_base')
5650    apply (rule ctes_of_retype_not_cte_tcb[where addrs="[p]", simplified]
5651            ; simp?)
5652      apply (simp add: pspace_aligned'_def)
5653      apply (clarsimp split: if_splits simp: objBitsKO_vcpu)
5654      apply (erule_tac x=x in ballE; (simp add: dom_def)?)
5655     apply (simp add: pspace_distinct'_def)
5656     apply (clarsimp split: if_splits simp: objBitsKO_vcpu)
5657      apply (subst ps_clear_ksPSpace_upd_same[simplified fun_upd_def])
5658      apply (rule ps_clear_entire_slotI)
5659      apply (drule (1) pspace_no_overlap_disjoint'[where n=vcpu_bits])
5660      apply fastforce
5661     apply (clarsimp simp: ps_clear_def)
5662     apply (subst dom_fun_upd[simplified fun_upd_def])
5663     apply (simp only: option.distinct if_False)
5664     apply (subgoal_tac "({x..x + 2 ^ objBitsKO y - 1} - {x}) \<inter> {p} = {}")
5665      apply fastforce
5666     apply (drule (2) pspace_no_overlapD3')
5667     apply (simp only: obj_range'_def)
5668     apply (erule disjoint_subset_both[rotated 2])
5669      apply blast
5670     apply fastforce
5671    apply (clarsimp simp: objBitsKO_def archObjSize_def machine_bits_defs)
5672    done
5673
5674  have csrel: "cpspace_relation (ksPSpace \<sigma>) (underlying_memory (ksMachineState \<sigma>)) (t_hrs_' (globals \<sigma>'))"
5675    using sr unfolding rf_sr_def cstate_relation_def by (simp add: Let_def)
5676
5677  have csrel': "cpspace_relation ?ks (underlying_memory (ksMachineState \<sigma>))  (t_hrs_' ?gs)"
5678    using csrel map_vcpus pno pal al unfolding cpspace_relation_def
5679    apply (clarsimp simp: cl_rest ht_rest tag_disj_via_td_name)
5680    apply (simp add: ctes rl ko_vcpu_def projectKOs rl_vcpu objBitsKO_vcpu)
5681    apply (simp add: projectKO_opt_retyp_other' objBitsKO_vcpu cl_vcpu)
5682    done
5683
5684  show ?thesis
5685    using assms zro' csrel' arel csrel map_vcpus
5686    apply (clarsimp simp: ko_vcpu_def vcpu0_def)
5687    apply (clarsimp simp: rf_sr_def cstate_relation_def carch_state_relation_def
5688                          cmachine_state_relation_def Let_def h_t_valid_clift_Some_iff)
5689    apply (subgoal_tac "region_is_bytes p 428 \<sigma>'")
5690     prefer 2
5691     apply (fastforce simp: region_actually_is_bytes[OF act_bytes]
5692                            region_is_bytes_subset[OF _ subr])
5693    apply (simp add: projectKO_opt_retyp_other' objBitsKO_vcpu cl_vcpu
5694                     htd_safe[simplified] kernel_data_refs_domain_eq_rotate)
5695    apply (subst ptr_retyp_gen_one[symmetric])+
5696    apply clarsimp
5697    apply (subst clift_ptr_retyps_gen_other[where 'a=vcpu_C and arr=False and nptrs=1,
5698                                            simplified word_bits_def, simplified]
5699           ; simp add: tag_disj_via_td_name)+
5700    done
5701qed
5702
5703(* retyping to a vcpu from after a memset zero *)
5704lemma placeNewObject_vcpu_fromzero_ccorres:
5705  "ccorres dc xfdc
5706    (valid_global_refs' and pspace_aligned' and pspace_distinct'
5707      and pspace_no_overlap' regionBase vcpu_bits
5708      and (\<lambda>s. 2 ^ vcpu_bits \<le> gsMaxObjectSize s)
5709      and ret_zero regionBase (2 ^ vcpu_bits)
5710      and K (regionBase \<noteq> 0 \<and> range_cover regionBase vcpu_bits vcpu_bits 1
5711             \<and> ({regionBase..+2 ^ vcpu_bits} \<inter> kernel_data_refs = {})))
5712    ({s. region_actually_is_zero_bytes regionBase (2 ^ vcpu_bits) s})
5713    hs
5714    (placeNewObject regionBase fromzeroVCPU 0)
5715    (global_htd_update (\<lambda>_. (ptr_retyp (vcpu_Ptr regionBase))))"
5716  apply (rule ccorres_from_vcg_nofail, clarsimp)
5717  apply (rule conseqPre, vcg)
5718  apply (clarsimp simp: rf_sr_htd_safe)
5719  apply (subgoal_tac "{regionBase..+428} \<subseteq> {regionBase..+2^vcpu_bits}")
5720   prefer 2
5721   apply clarsimp
5722   apply (drule intvlD, clarsimp)
5723   apply (rule intvlI, simp add: vcpu_bits_def pageBits_def)
5724  apply (intro conjI allI impI)
5725    apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
5726                          kernel_data_refs_domain_eq_rotate
5727                   elim!: ptr_retyp_htd_safe_neg)
5728   apply blast
5729  apply (rule bexI [OF _ placeNewObject_eq])
5730     apply (clarsimp simp: split_def new_cap_addrs_def)
5731     apply (rule ptr_retyp_fromzeroVCPU[simplified] ; simp?)
5732     apply (clarsimp simp: range_cover_def)
5733    apply (clarsimp simp: word_bits_conv)
5734   apply (clarsimp simp: objBits_simps range_cover.aligned archObjSize_def vcpuBits_def vcpu_bits_def)
5735  apply (clarsimp simp: no_fail_def)
5736  done
5737
5738lemma vcpu_init_ccorres:
5739  "ccorres dc xfdc
5740       (ko_at' fromzeroVCPU vcpuptr and K (vcpuptr \<noteq> 0))
5741       (UNIV \<inter> \<lbrace> \<acute>vcpu = vcpu_Ptr vcpuptr \<rbrace>) hs
5742     (setObject vcpuptr (makeObject :: vcpu))
5743     (Call vcpu_init_'proc)"
5744  supply dc_simp[simp del]
5745  apply (cinit' lift: vcpu_' simp: makeObject_vcpu)
5746   apply clarsimp
5747   apply (rule monadic_rewrite_ccorres_assemble[OF _ monadic_rewrite_setObject_vcpu_as_init])
5748   apply (ctac (no_vcg) add: vcpu_write_reg_ccorres)
5749    apply clarsimp
5750    apply (ctac (no_vcg) add: vcpu_write_reg_ccorres)
5751     apply (clarsimp simp: vgicHCREN_def)
5752     apply (rule ccorres_move_c_guard_vcpu)
5753     apply (ctac add: vgicUpdate_HCR_ccorres)
5754    apply wpsimp+
5755  apply (fastforce simp: fromEnum_def enum_vcpureg seL4_VCPUReg_defs dc_def const_def
5756                         sctlrDefault_def actlrDefault_def ko_at_vcpu_at'D)
5757  done
5758
5759lemma placeNewObject_vcpu_ccorres:
5760  "ccorres dc xfdc
5761    (valid_global_refs' and pspace_aligned' and pspace_distinct'
5762      and pspace_no_overlap' regionBase vcpu_bits
5763      and (\<lambda>s. 2 ^ vcpu_bits \<le> gsMaxObjectSize s)
5764      and ret_zero regionBase (2 ^ vcpu_bits)
5765      and K (regionBase \<noteq> 0 \<and> range_cover regionBase vcpu_bits vcpu_bits 1
5766             \<and> ({regionBase..+2 ^ vcpu_bits} \<inter> kernel_data_refs = {})))
5767    (UNIV \<inter> {s. region_actually_is_zero_bytes regionBase (2 ^ vcpu_bits) s})
5768    hs
5769    (placeNewObject regionBase (makeObject :: vcpu) 0)
5770    (global_htd_update (\<lambda>_. (ptr_retyp (vcpu_Ptr regionBase)));; CALL vcpu_init(vcpu_Ptr regionBase))"
5771  supply dc_simp[simp del]
5772  apply (rule ccorres_guard_imp)
5773    apply (rule monadic_rewrite_ccorres_assemble[OF _
5774                  monadic_rewrite_placeNewObject_vcpu_decompose[where vcpupre=fromzeroVCPU]])
5775    apply (rule ccorres_split_nothrow)
5776        apply (rule placeNewObject_vcpu_fromzero_ccorres)
5777       apply ceqv
5778      apply clarsimp
5779      apply (ctac (no_vcg) add: vcpu_init_ccorres)
5780     apply (wp placeNewObject_creates_object_vcpu)
5781    apply clarsimp
5782    apply vcg
5783   apply clarsimp
5784  apply fastforce
5785  done
5786
5787lemma Arch_createObject_ccorres:
5788  assumes t: "toAPIType newType = None"
5789  notes is_aligned_neg_mask_eq[simp del]
5790    and is_aligned_neg_mask_weaken[simp del]
5791  shows "ccorres (\<lambda>a b. ccap_relation (ArchObjectCap a) b) ret__struct_cap_C_'
5792     (createObject_hs_preconds regionBase newType userSize deviceMemory)
5793     (createObject_c_preconds regionBase newType userSize deviceMemory)
5794     []
5795     (Arch.createObject newType regionBase userSize deviceMemory)
5796     (Call Arch_createObject_'proc)"
5797proof -
5798  note if_cong[cong]
5799
5800  show ?thesis
5801    apply (clarsimp simp: createObject_c_preconds_def
5802                          createObject_hs_preconds_def)
5803    apply (rule ccorres_gen_asm)
5804    apply clarsimp
5805    apply (frule range_cover.aligned)
5806    apply (cut_tac t)
5807    apply (case_tac newType,
5808           simp_all add: toAPIType_def
5809               bind_assoc
5810               ARMLargePageBits_def)
5811         apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
5812          apply (simp add: object_type_from_H_def Kernel_C_defs)
5813          apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff
5814                      ARMLargePageBits_def ARMSmallPageBits_def
5815                      ARMSectionBits_def ARMSuperSectionBits_def asidInvalid_def
5816                      sle_positive APIType_capBits_def shiftL_nat objBits_simps
5817                      ptBits_def archObjSize_def pageBits_def word_sle_def word_sless_def
5818                      fold_eq_0_to_bool)
5819          apply (ccorres_remove_UNIV_guard)
5820          apply (clarsimp simp: hrs_htd_update ptBits_def objBits_simps archObjSize_def
5821            ARM_HYP_H.createObject_def pageBits_def
5822            cond_second_eq_seq_ccorres modify_gsUserPages_update
5823            intro!: ccorres_rhs_assoc)
5824          apply ((rule ccorres_return_C | simp | wp | vcg
5825            | (rule match_ccorres, ctac add:
5826                    placeNewDataObject_ccorres[where us=0 and newType=newType, simplified]
5827                    gsUserPages_update_ccorres[folded modify_gsUserPages_update])
5828            | (rule match_ccorres, csymbr))+)[1]
5829         apply (intro conjI)
5830          apply (clarsimp simp: createObject_hs_preconds_def
5831                                APIType_capBits_def pageBits_def)
5832         apply (clarsimp simp: pageBits_def ccap_relation_def APIType_capBits_def
5833                    framesize_to_H_def cap_to_H_simps cap_small_frame_cap_lift
5834                    vmrights_to_H_def vm_rights_defs is_aligned_neg_mask_eq, simp add: mask_def)
5835
5836        \<comment> \<open>Page objects: could possibly fix the duplication here\<close>
5837        apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
5838         apply (simp add: object_type_from_H_def Kernel_C_defs)
5839         apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff
5840                     ARMLargePageBits_def ARMSmallPageBits_def
5841                     ARMSectionBits_def ARMSuperSectionBits_def asidInvalid_def
5842                     sle_positive APIType_capBits_def shiftL_nat objBits_simps
5843                     ptBits_def archObjSize_def pageBits_def word_sle_def word_sless_def
5844                     fold_eq_0_to_bool)
5845         apply (ccorres_remove_UNIV_guard)
5846         apply (clarsimp simp: hrs_htd_update ptBits_def objBits_simps archObjSize_def
5847           ARM_HYP_H.createObject_def pageBits_def
5848           cond_second_eq_seq_ccorres modify_gsUserPages_update
5849           intro!: ccorres_rhs_assoc)
5850         apply ((rule ccorres_return_C | simp | wp | vcg
5851           | (rule match_ccorres, ctac add:
5852                   placeNewDataObject_ccorres[where us=4 and newType=newType, simplified]
5853                   gsUserPages_update_ccorres[folded modify_gsUserPages_update])
5854           | (rule match_ccorres, csymbr))+)[1]
5855        apply (intro conjI)
5856         apply (clarsimp simp: createObject_hs_preconds_def
5857                               APIType_capBits_def pageBits_def)
5858        apply (clarsimp simp: pageBits_def ccap_relation_def APIType_capBits_def
5859                   framesize_to_H_def cap_to_H_simps cap_frame_cap_lift
5860                   vmrights_to_H_def mask_def vm_rights_defs vm_page_size_defs
5861                   cl_valid_cap_def c_valid_cap_def
5862                   is_aligned_neg_mask_eq_concrete[THEN sym])
5863
5864       apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
5865        apply (simp add: object_type_from_H_def Kernel_C_defs)
5866        apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff
5867                    ARMLargePageBits_def ARMSmallPageBits_def
5868                    ARMSectionBits_def ARMSuperSectionBits_def asidInvalid_def
5869                    sle_positive APIType_capBits_def shiftL_nat objBits_simps
5870                    ptBits_def archObjSize_def pageBits_def word_sle_def word_sless_def
5871                    fold_eq_0_to_bool)
5872        apply (ccorres_remove_UNIV_guard)
5873        apply (clarsimp simp: hrs_htd_update ptBits_def objBits_simps archObjSize_def
5874          ARM_HYP_H.createObject_def pageBits_def
5875          cond_second_eq_seq_ccorres modify_gsUserPages_update
5876          intro!: ccorres_rhs_assoc)
5877        apply ((rule ccorres_return_C | simp | wp | vcg
5878          | (rule match_ccorres, ctac add:
5879                  placeNewDataObject_ccorres[where us=9 and newType=newType, simplified]
5880                  gsUserPages_update_ccorres[folded modify_gsUserPages_update])
5881          | (rule match_ccorres, csymbr))+)[1]
5882       apply (intro conjI)
5883        apply (clarsimp simp: createObject_hs_preconds_def
5884                              APIType_capBits_def pageBits_def)
5885       apply (clarsimp simp: pageBits_def ccap_relation_def APIType_capBits_def
5886                  framesize_to_H_def cap_to_H_simps cap_frame_cap_lift
5887                  vmrights_to_H_def mask_def vm_rights_defs vm_page_size_defs
5888                  cl_valid_cap_def c_valid_cap_def
5889                  is_aligned_neg_mask_eq_concrete[THEN sym])
5890
5891      apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
5892       apply (simp add: object_type_from_H_def Kernel_C_defs)
5893       apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff
5894                   ARMLargePageBits_def ARMSmallPageBits_def
5895                   ARMSectionBits_def ARMSuperSectionBits_def asidInvalid_def
5896                   sle_positive APIType_capBits_def shiftL_nat objBits_simps
5897                   ptBits_def archObjSize_def pageBits_def word_sle_def word_sless_def
5898                   fold_eq_0_to_bool)
5899       apply (ccorres_remove_UNIV_guard)
5900       apply (clarsimp simp: hrs_htd_update ptBits_def objBits_simps archObjSize_def
5901         ARM_HYP_H.createObject_def pageBits_def
5902         cond_second_eq_seq_ccorres modify_gsUserPages_update
5903         intro!: ccorres_rhs_assoc)
5904       apply ((rule ccorres_return_C | simp | wp | vcg
5905         | (rule match_ccorres, ctac add:
5906                 placeNewDataObject_ccorres[where us=13 and newType=newType, simplified]
5907                 gsUserPages_update_ccorres[folded modify_gsUserPages_update])
5908         | (rule match_ccorres, csymbr))+)[1]
5909      apply (intro conjI)
5910       apply (clarsimp simp: createObject_hs_preconds_def
5911                             APIType_capBits_def pageBits_def)
5912      apply (clarsimp simp: pageBits_def ccap_relation_def APIType_capBits_def
5913                 framesize_to_H_def cap_to_H_simps cap_frame_cap_lift
5914                 vmrights_to_H_def mask_def vm_rights_defs vm_page_size_defs
5915                 cl_valid_cap_def c_valid_cap_def
5916                 is_aligned_neg_mask_eq_concrete[THEN sym])
5917
5918     \<comment> \<open>PageTableObject\<close>
5919     apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
5920      apply (simp add: object_type_from_H_def Kernel_C_defs)
5921      apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff
5922                  ARMLargePageBits_def ARMSmallPageBits_def
5923                  ARMSectionBits_def ARMSuperSectionBits_def asidInvalid_def
5924                  sle_positive APIType_capBits_def shiftL_nat objBits_simps
5925                  ptBits_def archObjSize_def pageBits_def word_sle_def word_sless_def)
5926      apply (ccorres_remove_UNIV_guard)
5927      apply (rule ccorres_rhs_assoc)+
5928      apply (clarsimp simp: hrs_htd_update ptBits_def objBits_simps archObjSize_def
5929        ARM_HYP_H.createObject_def pageBits_def)
5930      apply (ctac pre only: add: placeNewObject_pte[simplified])
5931        apply csymbr
5932        apply (rule ccorres_return_C)
5933          apply simp
5934         apply simp
5935        apply simp
5936       apply wp
5937      apply vcg
5938     apply clarify
5939     apply (intro conjI)
5940      apply (clarsimp simp: invs_pspace_aligned' invs_pspace_distinct' invs_valid_global'
5941                            APIType_capBits_def invs_queues invs_valid_objs'
5942                            invs_urz)
5943     apply clarsimp
5944     apply (clarsimp simp: pageBits_def ccap_relation_def APIType_capBits_def
5945                framesize_to_H_def cap_to_H_simps cap_page_table_cap_lift
5946                is_aligned_neg_mask_eq vmrights_to_H_def
5947                Kernel_C.VMReadWrite_def Kernel_C.VMNoAccess_def
5948                Kernel_C.VMKernelOnly_def Kernel_C.VMReadOnly_def)
5949     apply (clarsimp simp: to_bool_def false_def isFrameType_def)
5950     apply (rule sym)
5951     apply (simp add: is_aligned_neg_mask_eq'[symmetric] is_aligned_weaken)
5952
5953    \<comment> \<open>PageDirectoryObject\<close>
5954    apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
5955     apply (simp add: object_type_from_H_def Kernel_C_defs)
5956     apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff
5957                asidInvalid_def sle_positive APIType_capBits_def shiftL_nat
5958                objBits_simps archObjSize_def
5959                ptBits_def pageBits_def pdBits_def word_sle_def word_sless_def)
5960     apply (ccorres_remove_UNIV_guard)
5961     apply (rule ccorres_rhs_assoc)+
5962     apply (clarsimp simp: hrs_htd_update ptBits_def objBits_simps archObjSize_def
5963        ARM_HYP_H.createObject_def pageBits_def pdBits_def)
5964     apply (ctac pre only: add: placeNewObject_pde[simplified])
5965       apply (ctac add: copyGlobalMappings_ccorres)
5966         apply csymbr
5967         apply (ctac add: cleanCacheRange_PoU_ccorres)
5968           apply csymbr
5969           apply (rule ccorres_return_C)
5970             apply simp
5971            apply simp
5972           apply simp
5973          apply wp
5974         apply (clarsimp simp: false_def)
5975         apply vcg
5976        apply wp
5977       apply (clarsimp simp: pageBits_def ccap_relation_def APIType_capBits_def
5978                  framesize_to_H_def cap_to_H_simps cap_page_directory_cap_lift
5979                  is_aligned_neg_mask_eq vmrights_to_H_def
5980                  Kernel_C.VMReadWrite_def Kernel_C.VMNoAccess_def
5981                  Kernel_C.VMKernelOnly_def Kernel_C.VMReadOnly_def)
5982       apply (vcg exspec=copyGlobalMappings_modifies)
5983      apply (clarsimp simp:placeNewObject_def2)
5984      apply (wp createObjects'_pde_mappings' createObjects'_page_directory_at_global[where sz=pdBits]
5985                createObjects'_page_directory_at'[where n=0, simplified])
5986     apply clarsimp
5987     apply vcg
5988    apply (clarsimp simp: invs_pspace_aligned' invs_pspace_distinct'
5989               archObjSize_def invs_valid_global' makeObject_pde pdBits_def
5990               pageBits_def range_cover.aligned projectKOs APIType_capBits_def
5991               object_type_from_H_def objBits_simps
5992               invs_valid_objs' isFrameType_def)
5993    apply (frule invs_arch_state')
5994    apply (frule range_cover.aligned)
5995    apply (frule is_aligned_addrFromPPtr_n, simp)
5996    apply (intro conjI, simp_all add: table_bits_defs)[1]
5997        apply fastforce
5998       apply ((clarsimp simp: is_aligned_no_overflow'[where n=14, simplified]
5999                             field_simps is_aligned_mask[symmetric] mask_AND_less_0)+)[3]
6000    \<comment> \<open>VCPU\<close>
6001    apply (cinit' lift: t_' regionBase_' userSize_' deviceMemory_')
6002     apply (simp add: object_type_from_H_def Kernel_C_defs)
6003     apply ccorres_rewrite
6004     apply (simp add: ccorres_cond_univ_iff ccorres_cond_empty_iff
6005                asidInvalid_def sle_positive APIType_capBits_def shiftL_nat
6006                objBits_simps archObjSize_def word_sle_def word_sless_def)
6007     apply (clarsimp simp: hrs_htd_update ptBits_def objBits_simps archObjSize_def
6008                           ARM_HYP_H.createObject_def pageBits_def pdBits_def)
6009     apply (rule ccorres_rhs_assoc)+
6010     apply (rule ccorres_rhs_assoc2)
6011     apply (ctac add: placeNewObject_vcpu_ccorres)
6012       apply csymbr
6013       apply (rule ccorres_return_C; simp)
6014      apply wp
6015     apply (vcg exspec=vcpu_init_modifies)
6016    apply (clarsimp simp: invs_pspace_aligned' invs_pspace_distinct'
6017               invs_valid_global' range_cover.aligned APIType_capBits_def
6018               invs_valid_objs' isFrameType_def invs_urz)
6019    apply (frule range_cover.aligned)
6020    apply (clarsimp simp: ccap_relation_def cap_vcpu_cap_lift cap_to_H_def)
6021    apply (rule sym, rule is_aligned_neg_mask_eq)
6022    apply (simp add: vcpu_bits_def table_bits_defs is_aligned_weaken)
6023    done
6024qed
6025
6026
6027(* FIXME: with the current state of affairs, we could simplify gs_new_cnodes *)
6028lemma gsCNodes_update_ccorres:
6029  "ccorresG rf_sr G dc xf (\<lambda>_. bits = sz + 4)
6030        \<lbrace> h_t_array_valid (hrs_htd \<acute>t_hrs) (cte_Ptr ptr) (2 ^ sz) \<rbrace> hs
6031     (modify (gsCNodes_update (\<lambda>m a. if a = ptr then Some sz else m a)))
6032     (Basic (globals_update (ghost'state_'_update
6033                  (gs_new_cnodes sz ptr bits))))"
6034  apply (rule ccorres_from_vcg)
6035  apply vcg_step
6036  apply (clarsimp simp: split_def simpler_modify_def gs_new_cnodes_def)
6037  apply (case_tac "ghost'state_' (globals x)")
6038  apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def fun_upd_def
6039                        carch_state_relation_def cmachine_state_relation_def
6040                        ghost_size_rel_def ghost_assertion_data_get_def
6041                 cong: if_cong)
6042  apply (rule cvariable_array_ptr_upd[unfolded fun_upd_def], simp_all)
6043  done
6044
6045(* FIXME: move *)
6046lemma map_to_tcbs_upd:
6047  "map_to_tcbs (ksPSpace s(t \<mapsto> KOTCB tcb')) = map_to_tcbs (ksPSpace s)(t \<mapsto> tcb')"
6048  apply (rule ext)
6049  apply (clarsimp simp: map_comp_def projectKOs split: option.splits if_splits)
6050  done
6051
6052(* FIXME: move *)
6053lemma cmap_relation_updI:
6054  "\<lbrakk>cmap_relation am cm f rel; am dest = Some ov; rel nv nv'; inj f\<rbrakk> \<Longrightarrow> cmap_relation (am(dest \<mapsto> nv)) (cm(f dest \<mapsto> nv')) f rel"
6055  apply (clarsimp simp: cmap_relation_def)
6056  apply (rule conjI)
6057   apply (drule_tac t="dom cm" in sym)
6058   apply fastforce
6059  apply clarsimp
6060  apply (case_tac "x = dest")
6061   apply simp
6062  apply clarsimp
6063  apply (subgoal_tac "f x \<noteq> f dest")
6064   apply simp
6065   apply force
6066  apply clarsimp
6067  apply (drule (1) injD)
6068  apply simp
6069  done
6070
6071lemma cep_relations_drop_fun_upd:
6072  "\<lbrakk> f x = Some v; tcbEPNext_C v' = tcbEPNext_C v; tcbEPPrev_C v' = tcbEPPrev_C v \<rbrakk>
6073      \<Longrightarrow> cendpoint_relation (f (x \<mapsto> v')) = cendpoint_relation f"
6074  "\<lbrakk> f x = Some v; tcbEPNext_C v' = tcbEPNext_C v; tcbEPPrev_C v' = tcbEPPrev_C v \<rbrakk>
6075      \<Longrightarrow> cnotification_relation (f (x \<mapsto> v')) = cnotification_relation f"
6076  by (intro ext cendpoint_relation_upd_tcb_no_queues[where thread=x]
6077                cnotification_relation_upd_tcb_no_queues[where thread=x]
6078          | simp split: if_split)+
6079
6080lemma threadSet_domain_ccorres [corres]:
6081  "ccorres dc xfdc (tcb_at' thread) {s. thread' s = tcb_ptr_to_ctcb_ptr thread \<and> d' s = ucast d} hs
6082           (threadSet (tcbDomain_update (\<lambda>_. d)) thread)
6083           (Basic (\<lambda>s. globals_update (t_hrs_'_update (hrs_mem_update (heap_update (Ptr &(thread' s\<rightarrow>[''tcbDomain_C''])::word32 ptr) (d' s)))) s))"
6084  apply (rule ccorres_guard_imp2)
6085   apply (rule threadSet_ccorres_lemma4 [where P=\<top> and P'=\<top>])
6086    apply vcg
6087   prefer 2
6088   apply (rule conjI, simp)
6089   apply assumption
6090  apply clarsimp
6091  apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def)
6092  apply (clarsimp simp: cmachine_state_relation_def carch_state_relation_def cpspace_relation_def)
6093  apply (clarsimp simp: update_tcb_map_tos typ_heap_simps')
6094  apply (simp add: map_to_ctes_upd_tcb_no_ctes map_to_tcbs_upd tcb_cte_cases_def)
6095  apply (simp add: cep_relations_drop_fun_upd
6096                   cvariable_relation_upd_const ko_at_projectKO_opt)
6097  apply (rule conjI)
6098   defer
6099   apply (erule cready_queues_relation_not_queue_ptrs)
6100    apply (rule ext, simp split: if_split)
6101   apply (rule ext, simp split: if_split)
6102  apply (drule ko_at_projectKO_opt)
6103  apply (erule (2) cmap_relation_upd_relI)
6104    subgoal by (simp add: ctcb_relation_def)
6105   apply assumption
6106  by simp
6107
6108lemma createObject_ccorres:
6109  notes APITypecapBits_simps[simp] =
6110          APIType_capBits_def[split_simps
6111          object_type.split apiobject_type.split]
6112  shows
6113    "ccorres ccap_relation ret__struct_cap_C_'
6114     (createObject_hs_preconds regionBase newType userSize isdev)
6115     (createObject_c_preconds regionBase newType userSize isdev)
6116     []
6117     (createObject newType regionBase userSize isdev)
6118     (Call createObject_'proc)"
6119proof -
6120  note if_cong[cong]
6121
6122  have gsCNodes_update:
6123    "\<And>f. (\<lambda>ks. ks \<lparr>gsCNodes := f (gsCNodes ks)\<rparr>) = gsCNodes_update f"
6124    by (rule ext) simp
6125
6126  show ?thesis
6127  apply (clarsimp simp: createObject_c_preconds_def
6128                        createObject_hs_preconds_def)
6129  apply (rule ccorres_gen_asm_state)
6130  apply (cinit lift: t_' regionBase_' userSize_' deviceMemory_')
6131   apply (rule ccorres_cond_seq)
6132   (* Architecture specific objects. *)
6133   apply (rule_tac
6134           Q="createObject_hs_preconds regionBase newType userSize isdev" and
6135           S="createObject_c_preconds1 regionBase newType userSize isdev" and
6136           R="createObject_hs_preconds regionBase newType userSize isdev" and
6137           T="createObject_c_preconds1 regionBase newType userSize isdev"
6138           in ccorres_Cond_rhs)
6139    apply (subgoal_tac "toAPIType newType = None")
6140     apply clarsimp
6141     apply (rule ccorres_rhs_assoc)+
6142     apply (rule ccorres_guard_imp)
6143       apply (ctac (no_vcg) add: Arch_createObject_ccorres)
6144        apply (rule ccorres_return_C_Seq)
6145        apply (rule ccorres_return_C)
6146          apply clarsimp
6147         apply clarsimp
6148        apply clarsimp
6149       apply wp[1]
6150      apply clarsimp
6151     apply (clarsimp simp: createObject_c_preconds_def
6152                           region_actually_is_bytes
6153                           region_actually_is_bytes_def)
6154    apply (clarsimp simp: object_type_from_H_def
6155      ARM_HYP_H.toAPIType_def Kernel_C_defs toAPIType_def
6156      nAPIObjects_def word_sle_def createObject_c_preconds_def
6157      word_le_nat_alt split:
6158      apiobject_type.splits object_type.splits)
6159   apply (subgoal_tac "\<exists>apiType. newType = APIObjectType apiType")
6160    apply clarsimp
6161    apply (rule ccorres_guard_imp)
6162      apply (rule_tac apiType=apiType in ccorres_apiType_split)
6163
6164          (* Untyped *)
6165          apply (clarsimp simp: Kernel_C_defs object_type_from_H_def
6166            toAPIType_def ARM_HYP_H.toAPIType_def nAPIObjects_def
6167            word_sle_def intro!: Corres_UL_C.ccorres_cond_empty
6168            Corres_UL_C.ccorres_cond_univ ccorres_rhs_assoc)
6169          apply (rule_tac
6170             A ="createObject_hs_preconds regionBase
6171                   (APIObjectType apiobject_type.Untyped)
6172                    (unat (userSizea :: word32)) isdev" and
6173             A'=UNIV in
6174             ccorres_guard_imp)
6175            apply (rule ccorres_symb_exec_r)
6176              apply (rule ccorres_return_C, simp, simp, simp)
6177             apply vcg
6178            apply (rule conseqPre, vcg, clarsimp)
6179           apply simp
6180          apply (clarsimp simp: ccap_relation_def cap_to_H_def ARM_HYP_H.getObjectSize_def
6181                     apiGetObjectSize_def cap_untyped_cap_lift to_bool_eq_0 true_def
6182                     aligned_add_aligned
6183                   split: option.splits)
6184          apply (subst is_aligned_neg_mask_eq [OF is_aligned_weaken])
6185            apply (erule range_cover.aligned)
6186           apply (clarsimp simp:APIType_capBits_def untypedBits_defs)
6187          apply (clarsimp simp: cap_untyped_cap_lift_def)
6188          apply (subst word_le_mask_eq, clarsimp simp: mask_def, unat_arith,
6189                 auto simp: to_bool_eq_0 word_bits_conv untypedBits_defs split:if_splits)[1]
6190
6191         (* TCB *)
6192         apply (clarsimp simp: Kernel_C_defs object_type_from_H_def
6193           toAPIType_def ARM_HYP_H.toAPIType_def nAPIObjects_def
6194           word_sle_def intro!: Corres_UL_C.ccorres_cond_empty
6195           Corres_UL_C.ccorres_cond_univ ccorres_rhs_assoc)
6196         apply (rule_tac
6197           A ="createObject_hs_preconds regionBase
6198                 (APIObjectType apiobject_type.TCBObject) (unat userSizea) isdev" and
6199           A'="createObject_c_preconds1 regionBase
6200                 (APIObjectType apiobject_type.TCBObject) (unat userSizea) isdev" in
6201            ccorres_guard_imp2)
6202          apply (rule ccorres_symb_exec_r)
6203            apply (ccorres_remove_UNIV_guard)
6204            apply (simp add: hrs_htd_update)
6205            apply (ctac (c_lines 4) add: ccorres_placeNewObject_tcb[simplified])
6206              apply simp
6207              apply (rule ccorres_pre_curDomain)
6208              apply ctac
6209                apply (rule ccorres_symb_exec_r)
6210                  apply (rule ccorres_return_C, simp, simp, simp)
6211                 apply vcg
6212                apply (rule conseqPre, vcg, clarsimp)
6213               apply wp
6214              apply vcg
6215             apply (simp add: obj_at'_real_def)
6216             apply (wp placeNewObject_ko_wp_at')
6217            apply vcg
6218           apply (clarsimp simp: dc_def)
6219           apply vcg
6220          apply (clarsimp simp: CPSR_def)
6221          apply (rule conseqPre, vcg, clarsimp)
6222         apply (clarsimp simp: createObject_hs_preconds_def
6223                               createObject_c_preconds_def)
6224         apply (frule invs_pspace_aligned')
6225         apply (frule invs_pspace_distinct')
6226         apply (frule invs_queues)
6227         apply (frule invs_sym')
6228         apply (simp add: getObjectSize_def objBits_simps word_bits_conv
6229                          ARM_HYP_H.getObjectSize_def apiGetObjectSize_def
6230                          tcbBlockSizeBits_def new_cap_addrs_def projectKO_opt_tcb)
6231         apply (clarsimp simp: range_cover.aligned
6232                               region_actually_is_bytes_def APIType_capBits_def)
6233         apply (frule(1) ghost_assertion_size_logic_no_unat)
6234         apply (clarsimp simp: ccap_relation_def cap_to_H_def
6235                    getObjectSize_def ARM_HYP_H.getObjectSize_def
6236                    apiGetObjectSize_def Collect_const_mem
6237                    cap_thread_cap_lift to_bool_def true_def
6238                    aligned_add_aligned
6239                  split: option.splits)
6240         apply (clarsimp simp: ctcb_ptr_to_tcb_ptr_def ctcb_offset_defs
6241                               tcb_ptr_to_ctcb_ptr_def
6242                               invs_valid_objs' invs_urz isFrameType_def)
6243         apply (subst is_aligned_neg_mask_weaken)
6244           apply (rule is_aligned_add[where n=8])
6245             apply (clarsimp elim!: is_aligned_weaken
6246                             dest!: range_cover.aligned)
6247            apply (clarsimp simp: is_aligned_def)
6248           apply (clarsimp simp: word_bits_def)
6249          apply simp
6250         apply clarsimp
6251
6252        (* Endpoint *)
6253        apply (clarsimp simp: Kernel_C_defs object_type_from_H_def
6254          toAPIType_def ARM_HYP_H.toAPIType_def nAPIObjects_def
6255          word_sle_def intro!: ccorres_cond_empty ccorres_cond_univ
6256          ccorres_rhs_assoc)
6257        apply (rule_tac
6258           A ="createObject_hs_preconds regionBase
6259                 (APIObjectType apiobject_type.EndpointObject)
6260                 (unat (userSizea :: word32)) isdev" and
6261           A'="createObject_c_preconds1 regionBase
6262                 (APIObjectType apiobject_type.EndpointObject)
6263                 (unat userSizea) isdev" in
6264           ccorres_guard_imp2)
6265         apply (ccorres_remove_UNIV_guard)
6266         apply (simp add: hrs_htd_update)
6267         apply (ctac (no_vcg) pre only: add: ccorres_placeNewObject_endpoint)
6268           apply (rule ccorres_symb_exec_r)
6269             apply (rule ccorres_return_C, simp, simp, simp)
6270            apply vcg
6271           apply (rule conseqPre, vcg, clarsimp)
6272          apply wp
6273         apply (clarsimp simp: ccap_relation_def cap_to_H_def
6274                    getObjectSize_def ARM_HYP_H.getObjectSize_def
6275                    objBits_simps apiGetObjectSize_def epSizeBits_def
6276                    Collect_const_mem cap_endpoint_cap_lift
6277                    to_bool_def true_def
6278                  split: option.splits   dest!: range_cover.aligned)
6279        apply (clarsimp simp: createObject_hs_preconds_def isFrameType_def)
6280        apply (frule invs_pspace_aligned')
6281        apply (frule invs_pspace_distinct')
6282        apply (frule invs_queues)
6283        apply (frule invs_sym')
6284        apply (auto simp: getObjectSize_def objBits_simps
6285                    ARM_HYP_H.getObjectSize_def apiGetObjectSize_def
6286                    epSizeBits_def word_bits_conv
6287                  elim!: is_aligned_no_wrap'   intro!: range_cover_simpleI)[1]
6288
6289       (* Notification *)
6290       apply (clarsimp simp: createObject_c_preconds_def)
6291       apply (clarsimp simp: getObjectSize_def objBits_simps
6292                  ARM_HYP_H.getObjectSize_def apiGetObjectSize_def
6293                  epSizeBits_def word_bits_conv word_sle_def word_sless_def)
6294       apply (clarsimp simp: Kernel_C_defs object_type_from_H_def
6295         toAPIType_def ARM_HYP_H.toAPIType_def nAPIObjects_def
6296         word_sle_def intro!: ccorres_cond_empty ccorres_cond_univ
6297         ccorres_rhs_assoc)
6298       apply (rule_tac
6299         A ="createObject_hs_preconds regionBase
6300               (APIObjectType apiobject_type.NotificationObject)
6301               (unat (userSizea :: word32)) isdev" and
6302         A'="createObject_c_preconds1 regionBase
6303               (APIObjectType apiobject_type.NotificationObject)
6304               (unat userSizea) isdev" in
6305         ccorres_guard_imp2)
6306        apply (ccorres_remove_UNIV_guard)
6307        apply (simp add: hrs_htd_update)
6308        apply (ctac (no_vcg) pre only: add: ccorres_placeNewObject_notification)
6309          apply (rule ccorres_symb_exec_r)
6310            apply (rule ccorres_return_C, simp, simp, simp)
6311           apply vcg
6312          apply (rule conseqPre, vcg, clarsimp)
6313         apply wp
6314        apply (clarsimp simp: ccap_relation_def cap_to_H_def
6315            getObjectSize_def ARM_HYP_H.getObjectSize_def
6316            apiGetObjectSize_def ntfnSizeBits_def objBits_simps
6317            Collect_const_mem cap_notification_cap_lift to_bool_def true_def
6318            dest!: range_cover.aligned split: option.splits)
6319       apply (clarsimp simp: createObject_hs_preconds_def isFrameType_def)
6320       apply (frule invs_pspace_aligned')
6321       apply (frule invs_pspace_distinct')
6322       apply (frule invs_queues)
6323       apply (frule invs_sym')
6324       apply (auto simp: getObjectSize_def objBits_simps
6325                   ARM_HYP_H.getObjectSize_def apiGetObjectSize_def
6326                   ntfnSizeBits_def word_bits_conv
6327                elim!: is_aligned_no_wrap'  intro!: range_cover_simpleI)[1]
6328
6329      (* CapTable *)
6330      apply (clarsimp simp: createObject_c_preconds_def)
6331      apply (clarsimp simp: getObjectSize_def objBits_simps
6332                  ARM_HYP_H.getObjectSize_def apiGetObjectSize_def
6333                  ntfnSizeBits_def word_bits_conv)
6334      apply (clarsimp simp: Kernel_C_defs object_type_from_H_def
6335                 toAPIType_def ARM_HYP_H.toAPIType_def nAPIObjects_def
6336                 word_sle_def word_sless_def zero_le_sint_32
6337               intro!: ccorres_cond_empty ccorres_cond_univ ccorres_rhs_assoc
6338                       ccorres_move_c_guards ccorres_Guard_Seq)
6339      apply (rule_tac
6340         A ="createObject_hs_preconds regionBase
6341               (APIObjectType apiobject_type.CapTableObject)
6342               (unat (userSizea :: word32)) isdev" and
6343         A'="createObject_c_preconds1 regionBase
6344               (APIObjectType apiobject_type.CapTableObject)
6345               (unat userSizea) isdev" in
6346         ccorres_guard_imp2)
6347       apply (simp add:field_simps hrs_htd_update)
6348       apply (ccorres_remove_UNIV_guard)
6349       apply (ctac pre only: add: ccorres_placeNewObject_captable)
6350         apply (subst gsCNodes_update)
6351         apply (ctac add: gsCNodes_update_ccorres)
6352           apply (rule ccorres_symb_exec_r)
6353             apply (rule ccorres_return_C, simp, simp, simp)
6354            apply vcg
6355           apply (rule conseqPre, vcg, clarsimp)
6356          apply (rule hoare_triv[of \<top>], simp add:hoare_TrueI)
6357         apply vcg
6358        apply wp
6359       apply vcg
6360      apply (rule conjI)
6361       apply (clarsimp simp: createObject_hs_preconds_def isFrameType_def)
6362       apply (frule invs_pspace_aligned')
6363       apply (frule invs_pspace_distinct')
6364       apply (frule invs_queues)
6365       apply (frule invs_sym')
6366       apply (frule(1) ghost_assertion_size_logic_no_unat)
6367       apply (clarsimp simp: getObjectSize_def objBits_simps
6368                  ARM_HYP_H.getObjectSize_def apiGetObjectSize_def
6369                  cteSizeBits_def word_bits_conv add.commute createObject_c_preconds_def
6370                  region_actually_is_bytes_def
6371                  invs_valid_objs' invs_urz
6372                 elim!: is_aligned_no_wrap'
6373                dest: word_of_nat_le  intro!: range_coverI)[1]
6374      apply (clarsimp simp: createObject_hs_preconds_def hrs_htd_update isFrameType_def)
6375      apply (frule range_cover.strong_times_32[folded addr_card_wb], simp+)
6376      apply (subst h_t_array_valid_retyp, simp+)
6377       apply (simp add: power_add cte_C_size objBits_defs)
6378      apply (frule range_cover.aligned)
6379      apply (clarsimp simp: ccap_relation_def cap_to_H_def cap_cnode_cap_lift to_bool_def true_def
6380                            getObjectSize_def apiGetObjectSize_def objBits_simps' field_simps
6381                            is_aligned_power2 addr_card_wb is_aligned_weaken[where y=word_size_bits]
6382                            is_aligned_neg_mask
6383                     split: option.splits)
6384      apply (subst word_le_mask_eq[symmetric, THEN eqTrueI])
6385        apply (clarsimp simp: mask_def)
6386        apply unat_arith
6387       apply (clarsimp simp: word_bits_conv)
6388      apply simp
6389     apply unat_arith
6390     apply auto[1]
6391    apply (clarsimp simp: createObject_c_preconds_def)
6392    apply (clarsimp simp:nAPIOBjects_object_type_from_H)?
6393    apply (intro impI conjI, simp_all)[1]
6394   apply (clarsimp simp: nAPIObjects_def object_type_from_H_def Kernel_C_defs
6395                  split: object_type.splits)
6396  apply (clarsimp simp: createObject_c_preconds_def
6397                        createObject_hs_preconds_def)
6398  done
6399qed
6400
6401lemma ccorres_guard_impR:
6402  "\<lbrakk>ccorres_underlying sr \<Gamma> r xf arrel axf W Q' hs f g; (\<And>s s'. \<lbrakk>(s, s') \<in> sr; s' \<in> A'\<rbrakk> \<Longrightarrow> s' \<in> Q')\<rbrakk>
6403  \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf W A' hs f g"
6404  by (rule ccorres_guard_imp2,simp+)
6405
6406lemma typ_clear_region_dom:
6407 "dom (clift (hrs_htd_update (typ_clear_region ptr bits) hp) :: 'b :: mem_type typ_heap)
6408  \<subseteq>  dom ((clift hp) :: 'b :: mem_type typ_heap)"
6409   apply (clarsimp simp:lift_t_def lift_typ_heap_def Fun.comp_def)
6410   apply (clarsimp simp:lift_state_def)
6411   apply (case_tac hp)
6412   apply (clarsimp simp:)
6413   apply (case_tac x)
6414   apply (clarsimp simp:s_valid_def h_t_valid_def)
6415    apply (clarsimp simp:valid_footprint_def Let_def)
6416    apply (drule spec)
6417    apply (erule(1) impE)
6418   apply clarsimp
6419   apply (rule conjI)
6420    apply (clarsimp simp add:map_le_def)
6421    apply (drule_tac x = aa in bspec)
6422     apply simp
6423    apply (drule sym)
6424     apply simp
6425    apply (clarsimp simp:proj_d_def)
6426    apply (clarsimp simp:hrs_htd_update_def typ_clear_region_def
6427     split:if_splits option.splits)
6428   apply (clarsimp simp:proj_d_def)
6429   apply (clarsimp simp:hrs_htd_update_def typ_clear_region_def
6430     split:if_splits option.splits)
6431  done
6432
6433lemma tcb_range_subseteq:
6434  "is_aligned x (objBitsKO (KOTCB ko))
6435   \<Longrightarrow> {ptr_val (tcb_ptr_to_ctcb_ptr x)..+size_of TYPE(tcb_C)} \<subseteq> {x..x + 2 ^ objBitsKO (KOTCB ko) - 1}"
6436  apply (simp add: tcb_ptr_to_ctcb_ptr_def)
6437  apply (rule subset_trans)
6438   apply (rule intvl_sub_offset[where z = "2^objBitsKO (KOTCB ko)"])
6439   apply (simp add: ctcb_offset_defs size_of_def objBits_simps')
6440  apply (subst intvl_range_conv)
6441    apply simp
6442   apply (simp add: objBits_simps' word_bits_conv)
6443  apply simp
6444  done
6445
6446lemma pspace_no_overlap_induce_tcb:
6447  "\<lbrakk>cpspace_relation (ksPSpace (s::kernel_state))
6448      (underlying_memory (ksMachineState s)) hp;
6449    pspace_aligned' s; clift hp xa = Some (v::tcb_C);
6450    is_aligned ptr bits; bits < word_bits;
6451    pspace_no_overlap' ptr bits s\<rbrakk>
6452   \<Longrightarrow> {ptr_val xa..+size_of TYPE(tcb_C)} \<inter> {ptr..+2 ^ bits} = {}"
6453  apply (clarsimp simp:cpspace_relation_def)
6454  apply (clarsimp simp:cmap_relation_def)
6455  apply (subgoal_tac "xa\<in>tcb_ptr_to_ctcb_ptr ` dom (map_to_tcbs (ksPSpace s))")
6456    prefer 2
6457    apply (simp add:domI)
6458  apply (thin_tac "S = dom K" for S K)+
6459  apply (thin_tac "\<forall>x\<in> S. K x" for S K)+
6460  apply (clarsimp simp: image_def projectKO_opt_tcb map_comp_def
6461                 split: option.splits kernel_object.split_asm)
6462  apply (frule(1) pspace_no_overlapD')
6463  apply (rule disjoint_subset[OF tcb_range_subseteq[simplified]])
6464   apply (erule(1) pspace_alignedD')
6465  apply (subst intvl_range_conv)
6466   apply (simp add: word_bits_def)+
6467  done
6468
6469lemma pspace_no_overlap_induce_endpoint:
6470  "\<lbrakk>cpspace_relation (ksPSpace (s::kernel_state))
6471      (underlying_memory (ksMachineState s)) hp;
6472    pspace_aligned' s; clift hp xa = Some (v::endpoint_C);
6473    is_aligned ptr bits; bits < word_bits;
6474    pspace_no_overlap' ptr bits s\<rbrakk>
6475   \<Longrightarrow> {ptr_val xa..+size_of TYPE(endpoint_C)} \<inter> {ptr..+2 ^ bits} = {}"
6476  supply
6477    is_aligned_neg_mask_eq[simp del]
6478    is_aligned_neg_mask_weaken[simp del]
6479  apply (clarsimp simp: cpspace_relation_def)
6480  apply (clarsimp simp: cmap_relation_def)
6481  apply (subgoal_tac "xa\<in>ep_Ptr ` dom (map_to_eps (ksPSpace s))")
6482   prefer 2
6483   subgoal by (simp add: domI)
6484  apply (thin_tac "S = dom K" for S K)+
6485  apply (thin_tac "\<forall>x\<in> S. K x" for S K)+
6486  apply (clarsimp simp: image_def projectKO_opt_ep map_comp_def
6487                 split: option.splits kernel_object.split_asm)
6488  apply (frule(1) pspace_no_overlapD')
6489  apply (subst intvl_range_conv)
6490    apply simp
6491   apply (simp add: word_bits_def)
6492  apply (simp add: size_of_def)
6493  apply (subst intvl_range_conv[where bits = epSizeBits, simplified epSizeBits_def, simplified])
6494    apply (drule(1) pspace_alignedD')
6495    apply (simp add: objBits_simps' archObjSize_def
6496              split: arch_kernel_object.split_asm)
6497   apply (simp add: word_bits_conv)
6498  apply (simp add: objBits_simps' archObjSize_def is_aligned_neg_mask_eq
6499            split: arch_kernel_object.split_asm, safe)
6500  done
6501
6502lemma pspace_no_overlap_induce_notification:
6503  "\<lbrakk>cpspace_relation (ksPSpace (s::kernel_state))
6504      (underlying_memory (ksMachineState s)) hp;
6505    pspace_aligned' s; clift hp xa = Some (v::notification_C);
6506    is_aligned ptr bits; bits < word_bits;
6507    pspace_no_overlap' ptr bits s\<rbrakk>
6508   \<Longrightarrow> {ptr_val xa..+size_of TYPE(notification_C)} \<inter> {ptr..+2 ^ bits} = {}"
6509  apply (clarsimp simp: cpspace_relation_def)
6510  apply (clarsimp simp: cmap_relation_def size_of_def)
6511  apply (subgoal_tac "xa\<in>ntfn_Ptr ` dom (map_to_ntfns (ksPSpace s))")
6512   prefer 2
6513   apply (simp add: domI)
6514  apply (thin_tac "S = dom K" for S K)+
6515  apply (thin_tac "\<forall>x\<in> S. K x" for S K)+
6516  apply (clarsimp simp: image_def projectKO_opt_ntfn map_comp_def
6517                 split: option.splits kernel_object.split_asm)
6518  apply (frule(1) pspace_no_overlapD')
6519  apply (subst intvl_range_conv)
6520    apply simp
6521   apply (simp add: word_bits_def)
6522  apply (subst intvl_range_conv[where bits = ntfnSizeBits, simplified ntfnSizeBits_def, simplified])
6523    apply (drule(1) pspace_alignedD')
6524    apply (simp add: objBits_simps' archObjSize_def
6525              split: arch_kernel_object.split_asm)
6526   apply (simp add: word_bits_conv)
6527  apply (simp add: objBits_simps' archObjSize_def
6528            split: arch_kernel_object.split_asm)
6529  done
6530
6531lemma vcpu_range_subseteq:
6532  "is_aligned x (objBitsKO (KOArch (KOVCPU ko)))
6533   \<Longrightarrow> {(x::machine_word)..+size_of TYPE(vcpu_C)} \<subseteq> {x..x + 2 ^ objBitsKO (KOArch (KOVCPU ko)) - 1}"
6534  apply simp
6535  apply (rule subset_trans)
6536  apply (rule intvl_start_le[where y = "2^objBitsKO ((KOArch (KOVCPU ko)))"])
6537   apply (simp add: objBits_simps archObjSize_def machine_bits_defs)
6538  apply (subst intvl_range_conv)
6539    apply simp
6540   apply (simp add:objBits_simps word_bits_conv archObjSize_def machine_bits_defs)+
6541  done
6542
6543lemma pspace_no_overlap_induce_vcpu:
6544  "\<lbrakk>cpspace_relation (ksPSpace (s::kernel_state))
6545      (underlying_memory (ksMachineState s)) hp;
6546    pspace_aligned' s; clift hp xa = Some (v::vcpu_C);
6547    is_aligned ptr bits; bits < word_bits;
6548    pspace_no_overlap' ptr bits s\<rbrakk>
6549   \<Longrightarrow> {ptr_val xa..+size_of TYPE(vcpu_C)} \<inter> {ptr..+2 ^ bits} = {}"
6550  apply (clarsimp simp:cpspace_relation_def)
6551  apply (clarsimp simp:cmap_relation_def)
6552  apply (subgoal_tac "xa\<in> vcpu_Ptr ` dom (map_to_vcpus (ksPSpace s))")
6553    prefer 2
6554    apply (simp add:domI)
6555  apply (thin_tac "S = dom K" for S K)+
6556  apply (thin_tac "\<forall>x\<in> S. K x" for S K)+
6557  apply (clarsimp simp: image_def projectKO_opt_tcb map_comp_def
6558                 split: option.splits kernel_object.split_asm)
6559  apply (subst intvl_range_conv)
6560   apply (simp add: word_bits_def)+
6561  (* prevent intersection in assumption becoming predicates *)
6562  supply Int_atLeastAtMost[simp del]
6563  apply (frule(1) pspace_no_overlapD')
6564  apply (drule(1) pspace_alignedD')
6565  apply (clarsimp simp: image_def projectKO_opt_vcpu map_comp_def
6566                 split: option.splits kernel_object.split_asm  arch_kernel_object.split_asm)
6567  apply (erule (1) disjoint_subset[OF vcpu_range_subseteq[simplified]])
6568  done
6569
6570lemma ctes_of_ko_at_strong:
6571  "\<lbrakk>ctes_of s p = Some a; is_aligned p cteSizeBits\<rbrakk> \<Longrightarrow>
6572  (\<exists>ptr ko. (ksPSpace s ptr = Some ko \<and> {p ..+ 2^cteSizeBits} \<subseteq> obj_range' ptr ko))"
6573  apply (clarsimp simp: map_to_ctes_def Let_def split: if_split_asm)
6574   apply (intro exI conjI, assumption)
6575   apply (simp add: obj_range'_def objBits_simps is_aligned_no_wrap' field_simps)
6576   apply (subst intvl_range_conv[where bits=cteSizeBits])
6577      apply simp
6578     apply (simp add: word_bits_def objBits_defs)
6579    apply (simp add: field_simps)
6580  apply (intro exI conjI, assumption)
6581  apply (clarsimp simp: objBits_simps obj_range'_def word_and_le2)
6582  apply (cut_tac intvl_range_conv[where bits=cteSizeBits and ptr=p, simplified])
6583    defer
6584    apply simp
6585   apply (simp add: word_bits_conv objBits_defs)
6586  apply (intro conjI)
6587   apply (rule order_trans[OF word_and_le2])
6588   apply clarsimp
6589  apply clarsimp
6590  apply (thin_tac "P \<or> Q" for P Q)
6591  apply (erule order_trans)
6592  apply (subst word_plus_and_or_coroll2[where x=p and w="mask tcbBlockSizeBits",symmetric])
6593  apply (clarsimp simp: tcb_cte_cases_def field_simps split: if_split_asm;
6594         simp only: p_assoc_help;
6595         rule word_plus_mono_right[OF _ is_aligned_no_wrap', OF _ Aligned.is_aligned_neg_mask[OF le_refl]];
6596         simp add: objBits_defs)
6597  done
6598
6599lemma pspace_no_overlap_induce_cte:
6600  "\<lbrakk>cpspace_relation (ksPSpace (s::kernel_state))
6601      (underlying_memory (ksMachineState s)) hp;
6602    pspace_aligned' s; clift hp xa = Some (v::cte_C);
6603    is_aligned ptr bits; bits < word_bits;
6604    pspace_no_overlap' ptr bits s\<rbrakk>
6605   \<Longrightarrow> {ptr_val xa..+size_of TYPE(cte_C)} \<inter> {ptr..+2 ^ bits} = {}"
6606  apply (clarsimp simp: cpspace_relation_def)
6607  apply (clarsimp simp: cmap_relation_def size_of_def)
6608  apply (subgoal_tac "xa\<in>cte_Ptr ` dom (ctes_of s)")
6609   prefer 2
6610   apply (simp add:domI)
6611  apply (thin_tac "S = dom K" for S K)+
6612  apply (thin_tac "\<forall>x\<in> S. K x" for S K)+
6613  apply (clarsimp simp: image_def projectKO_opt_cte map_comp_def
6614                 split: option.splits kernel_object.split_asm)
6615  apply (frule ctes_of_is_aligned)
6616  apply (simp add: objBits_simps)
6617  apply (drule ctes_of_ko_at_strong)
6618   apply simp
6619  apply (clarsimp simp: objBits_defs)
6620  apply (erule disjoint_subset)
6621  apply (frule(1) pspace_no_overlapD')
6622  apply (subst intvl_range_conv)
6623    apply simp
6624   apply (simp add: word_bits_def)
6625  apply (simp add: obj_range'_def)
6626  done
6627
6628lemma pspace_no_overlap_induce_asidpool:
6629  "\<lbrakk>cpspace_relation (ksPSpace (s::kernel_state)) (underlying_memory (ksMachineState s)) hp;
6630    pspace_aligned' s; clift hp xa = Some (v::asid_pool_C);
6631    is_aligned ptr bits; bits < word_bits;
6632    pspace_no_overlap' ptr bits s\<rbrakk>
6633     \<Longrightarrow> {ptr_val xa..+size_of TYPE(asid_pool_C)} \<inter> {ptr..+2 ^ bits} = {}"
6634  apply (clarsimp simp:cpspace_relation_def)
6635  apply (clarsimp simp:cmap_relation_def size_of_def)
6636  apply (subgoal_tac "xa\<in>ap_Ptr ` dom (map_to_asidpools (ksPSpace s))")
6637    prefer 2
6638    apply (simp add:domI)
6639  apply (thin_tac "S = dom K" for S K)+
6640  apply (thin_tac "\<forall>x\<in> S. K x" for S K)+
6641  apply (clarsimp simp:image_def projectKO_opt_asidpool
6642    map_comp_def split:option.splits kernel_object.split_asm)
6643  apply (frule(1) pspace_no_overlapD')
6644   apply (subst intvl_range_conv)
6645     apply simp
6646    apply (simp add: word_bits_def)
6647   apply (subst intvl_range_conv[where bits = 12,simplified])
6648    apply (drule(1) pspace_alignedD')
6649    apply (simp add: objBits_simps archObjSize_def pageBits_def split:arch_kernel_object.split_asm)
6650    apply (clarsimp elim!:is_aligned_weaken)
6651  apply (simp only: is_aligned_neg_mask_eq)
6652  apply (erule disjoint_subset[rotated])
6653  apply (clarsimp simp: field_simps)
6654  apply (simp add: p_assoc_help)
6655   apply (rule word_plus_mono_right)
6656   apply (clarsimp simp:objBits_simps archObjSize_def pageBits_def split:arch_kernel_object.split_asm)+
6657  done
6658
6659lemma pspace_no_overlap_induce_user_data:
6660  "\<lbrakk>cpspace_relation (ksPSpace (s::kernel_state)) (underlying_memory (ksMachineState s)) hp;
6661    pspace_aligned' s; clift hp xa = Some (v::user_data_C);
6662    is_aligned ptr bits; bits < word_bits;
6663    pspace_no_overlap' ptr bits s\<rbrakk>
6664     \<Longrightarrow> {ptr_val xa..+size_of TYPE(user_data_C)} \<inter> {ptr..+2 ^ bits} = {}"
6665  apply (clarsimp simp:cpspace_relation_def)
6666  apply (clarsimp simp:cmap_relation_def size_of_def)
6667  apply (subgoal_tac "xa\<in>Ptr ` dom (heap_to_user_data (ksPSpace s) (underlying_memory (ksMachineState s)))")
6668    prefer 2
6669    apply (simp add:domI)
6670  apply (thin_tac "S = dom K" for S K)+
6671  apply (thin_tac "\<forall>x\<in> S. K x" for S K)+
6672  apply (clarsimp simp: image_def heap_to_user_data_def projectKO_opt_user_data map_comp_def
6673                 split: option.splits kernel_object.splits)
6674  apply (frule(1) pspace_no_overlapD')
6675  apply (clarsimp simp: word_bits_def)
6676   apply (subst intvl_range_conv[where bits = 12,simplified])
6677    apply (drule(1) pspace_alignedD')
6678    apply (simp add:objBits_simps archObjSize_def pageBits_def split:arch_kernel_object.split_asm)
6679    apply (clarsimp elim!:is_aligned_weaken)
6680  apply (subst intvl_range_conv, simp, simp)
6681  apply (clarsimp simp: field_simps)
6682  apply (simp add: p_assoc_help)
6683  apply (clarsimp simp: objBits_simps archObjSize_def pageBits_def
6684                        is_aligned_no_overflow field_simps
6685                 split: arch_kernel_object.split_asm)
6686  done
6687
6688lemma pspace_no_overlap_induce_device_data:
6689  "\<lbrakk>cpspace_relation (ksPSpace (s::kernel_state)) (underlying_memory (ksMachineState s)) hp;
6690    pspace_aligned' s; clift hp xa = Some (v::user_data_device_C);
6691    is_aligned ptr bits; bits < word_bits;
6692    pspace_no_overlap' ptr bits s\<rbrakk>
6693     \<Longrightarrow> {ptr_val xa..+size_of TYPE(user_data_C)} \<inter> {ptr..+2 ^ bits} = {}"
6694  apply (clarsimp simp: cpspace_relation_def)
6695  apply (clarsimp simp: cmap_relation_def size_of_def)
6696  apply (subgoal_tac "xa\<in>Ptr ` dom (heap_to_device_data (ksPSpace s) (underlying_memory (ksMachineState s)))")
6697    prefer 2
6698    apply (simp add: domI)
6699  apply (thin_tac "S = dom K" for S K)+
6700  apply (thin_tac "\<forall>x\<in> S. K x" for S K)+
6701  apply (clarsimp simp: image_def heap_to_device_data_def projectKO_opt_user_data_device map_comp_def
6702                 split: option.splits kernel_object.splits)
6703  apply (frule(1) pspace_no_overlapD')
6704  apply (clarsimp simp: word_bits_def)
6705   apply (subst intvl_range_conv[where bits = 12,simplified])
6706    apply (drule(1) pspace_alignedD')
6707    apply (simp add: objBits_simps archObjSize_def pageBits_def split: arch_kernel_object.split_asm)
6708    apply (clarsimp elim!: is_aligned_weaken)
6709  apply (subst intvl_range_conv, simp, simp)
6710  apply (clarsimp simp: field_simps)
6711  apply (simp add: p_assoc_help)
6712  apply (fastforce simp: objBits_simps archObjSize_def pageBits_def field_simps
6713                   split:arch_kernel_object.split_asm)
6714  done
6715
6716lemma typ_region_bytes_dom:
6717 "typ_uinfo_t TYPE('b) \<noteq> typ_uinfo_t TYPE (word8)
6718    \<Longrightarrow> dom (clift (hrs_htd_update (typ_region_bytes ptr bits) hp) :: 'b :: mem_type typ_heap)
6719  \<subseteq>  dom ((clift hp) :: 'b :: mem_type typ_heap)"
6720  apply (clarsimp simp: liftt_if split: if_splits)
6721  apply (case_tac "{ptr_val x ..+ size_of TYPE('b)} \<inter> {ptr ..+ 2 ^ bits} = {}")
6722   apply (clarsimp simp: h_t_valid_def valid_footprint_def Let_def
6723                         hrs_htd_update_def split_def typ_region_bytes_def)
6724   apply (drule spec, drule(1) mp)
6725   apply (simp add: size_of_def split: if_split_asm)
6726   apply (drule subsetD[OF equalityD1], rule IntI, erule intvlI, simp)
6727   apply simp
6728  apply (clarsimp simp: set_eq_iff)
6729  apply (drule(1) h_t_valid_intvl_htd_contains_uinfo_t)
6730  apply (clarsimp simp: hrs_htd_update_def typ_region_bytes_def split_def
6731                 split: if_split_asm)
6732  done
6733
6734lemma lift_t_typ_region_bytes_none:
6735  "\<lbrakk> \<And>x (v :: 'a). lift_t g hp x = Some v
6736    \<Longrightarrow> {ptr_val x ..+ size_of TYPE('a)} \<inter> {ptr ..+ 2 ^ bits} = {};
6737     typ_uinfo_t TYPE('a) \<noteq> typ_uinfo_t TYPE(8 word) \<rbrakk> \<Longrightarrow>
6738  lift_t g (hrs_htd_update (typ_region_bytes ptr bits) hp)
6739    = (lift_t g hp :: (('a :: mem_type) ptr) \<Rightarrow> _)"
6740  apply atomize
6741  apply (subst lift_t_typ_region_bytes, simp_all)
6742   apply (clarsimp simp: liftt_if hrs_htd_def split: if_splits)
6743  apply (rule ext, simp add: restrict_map_def)
6744  apply (rule ccontr, clarsimp split: if_splits)
6745  apply (clarsimp simp: liftt_if hrs_htd_def split: if_splits)
6746  apply (clarsimp simp: set_eq_iff intvl_self)
6747  done
6748
6749lemma typ_bytes_cpspace_relation_clift_userdata:
6750assumes "cpspace_relation (ksPSpace s) (underlying_memory (ksMachineState (s::kernel_state))) hp"
6751  and "is_aligned ptr bits" "bits < word_bits"
6752  and "pspace_aligned' s"
6753  and "pspace_no_overlap' ptr bits s"
6754shows "clift (hrs_htd_update (typ_region_bytes ptr bits) hp) = ((clift hp) :: (user_data_C ptr \<rightharpoonup> user_data_C))"
6755  (is "?lhs = ?rhs")
6756  using assms
6757  apply -
6758  apply (rule lift_t_typ_region_bytes_none, simp_all)
6759  apply (rule pspace_no_overlap_induce_user_data[simplified], auto)
6760  done
6761
6762
6763lemma typ_bytes_cpspace_relation_clift_devicedata:
6764assumes "cpspace_relation (ksPSpace s) (underlying_memory (ksMachineState (s::kernel_state))) hp"
6765  and "is_aligned ptr bits" "bits < word_bits"
6766  and "pspace_aligned' s"
6767  and "pspace_no_overlap' ptr bits s"
6768 shows "clift (hrs_htd_update (typ_region_bytes ptr bits) hp) = ((clift hp) :: (user_data_device_C ptr \<rightharpoonup> user_data_device_C))"
6769  (is "?lhs = ?rhs")
6770  using assms
6771  apply -
6772  apply (rule lift_t_typ_region_bytes_none, simp_all)
6773  apply (rule pspace_no_overlap_induce_device_data[simplified], auto)
6774  done
6775
6776
6777lemma pspace_no_overlap_induce_pte:
6778  "\<lbrakk>cpspace_relation (ksPSpace (s::kernel_state)) (underlying_memory (ksMachineState s)) hp;
6779    pspace_aligned' s; clift hp xa = Some (v::pte_C);
6780    is_aligned ptr bits; bits < word_bits;
6781    pspace_no_overlap' ptr bits s\<rbrakk>
6782     \<Longrightarrow> {ptr_val xa..+size_of TYPE(pte_C)} \<inter> {ptr..+2 ^ bits} = {}"
6783  apply (clarsimp simp:cpspace_relation_def)
6784  apply (clarsimp simp:cmap_relation_def)
6785  apply (subgoal_tac "xa\<in>pte_Ptr ` dom (map_to_ptes (ksPSpace s))")
6786    prefer 2
6787    apply (simp add:domI)
6788  apply (thin_tac "S = dom K" for S K)+
6789  apply (thin_tac "\<forall>x\<in> S. K x" for S K)+
6790  apply (clarsimp simp: image_def projectKO_opt_pte map_comp_def
6791                 split: option.splits kernel_object.split_asm)
6792  apply (frule(1) pspace_no_overlapD')
6793   apply (subst intvl_range_conv)
6794     apply simp
6795    apply (simp add: word_bits_def)
6796   apply (subst intvl_range_conv[where bits = 3,simplified])
6797    apply (drule(1) pspace_alignedD')
6798    apply (simp add: objBits_simps archObjSize_def table_bits_defs split:arch_kernel_object.split_asm)
6799   apply (simp add: word_bits_conv)
6800  apply (simp add: objBits_simps archObjSize_def table_bits_defs split:arch_kernel_object.split_asm)
6801  done
6802
6803lemma pspace_no_overlap_induce_pde:
6804  "\<lbrakk>cpspace_relation (ksPSpace (s::kernel_state)) (underlying_memory (ksMachineState s)) hp;
6805    pspace_aligned' s; clift hp xa = Some (v::pde_C);
6806    is_aligned ptr bits; bits < word_bits;
6807    pspace_no_overlap' ptr bits s\<rbrakk>
6808     \<Longrightarrow> {ptr_val xa..+size_of TYPE(pde_C)} \<inter> {ptr..+2 ^ bits} = {}"
6809  apply (clarsimp simp:cpspace_relation_def)
6810  apply (clarsimp simp:cmap_relation_def)
6811  apply (subgoal_tac "xa\<in>pde_Ptr ` dom (map_to_pdes (ksPSpace s))")
6812    prefer 2
6813    subgoal by (simp add:domI)
6814  apply (thin_tac "S = dom K" for S K)+
6815  apply (thin_tac "\<forall>x\<in> S. K x" for S K)+
6816  apply (clarsimp simp:image_def projectKO_opt_pde
6817    map_comp_def split:option.splits kernel_object.split_asm)
6818  apply (frule(1) pspace_no_overlapD')
6819   apply (subst intvl_range_conv)
6820     apply simp
6821    apply (simp add: word_bits_def)
6822   apply (subst intvl_range_conv[where bits = 3,simplified])
6823    apply (drule(1) pspace_alignedD')
6824    apply (simp add: objBits_simps archObjSize_def table_bits_defs split:arch_kernel_object.split_asm)
6825   apply (simp add:word_bits_conv)
6826  by (simp add: objBits_simps archObjSize_def table_bits_defs split:arch_kernel_object.split_asm)
6827
6828
6829lemma typ_bytes_cpspace_relation_clift_tcb:
6830assumes "cpspace_relation (ksPSpace s) (underlying_memory (ksMachineState (s::kernel_state))) hp"
6831  and "is_aligned ptr bits" "bits < word_bits"
6832  and "pspace_aligned' s"
6833  and "pspace_no_overlap' ptr bits s"
6834 shows "clift (hrs_htd_update (typ_region_bytes ptr bits) hp) = ((clift hp) :: (tcb_C ptr \<rightharpoonup> tcb_C))"
6835  (is "?lhs = ?rhs")
6836  using assms
6837  apply -
6838  apply (rule lift_t_typ_region_bytes_none, simp_all)
6839  apply (erule(5) pspace_no_overlap_induce_tcb[simplified])
6840  done
6841
6842lemma typ_bytes_cpspace_relation_clift_pde:
6843assumes "cpspace_relation (ksPSpace s) (underlying_memory (ksMachineState (s::kernel_state))) hp"
6844  and "is_aligned ptr bits" "bits < word_bits"
6845  and "pspace_aligned' s"
6846  and "pspace_no_overlap' ptr bits s"
6847 shows "clift (hrs_htd_update (typ_region_bytes ptr bits) hp) = ((clift hp) :: (pde_C ptr \<rightharpoonup> pde_C))"
6848  (is "?lhs = ?rhs")
6849  using assms
6850  apply -
6851  apply (rule lift_t_typ_region_bytes_none, simp_all)
6852  apply (erule(5) pspace_no_overlap_induce_pde[unfolded size_of_def,simplified])
6853  done
6854
6855lemma typ_bytes_cpspace_relation_clift_pte:
6856assumes "cpspace_relation (ksPSpace s) (underlying_memory (ksMachineState (s::kernel_state))) hp"
6857  and "is_aligned ptr bits" "bits < word_bits"
6858  and "pspace_aligned' s"
6859  and "pspace_no_overlap' ptr bits s"
6860 shows "clift (hrs_htd_update (typ_region_bytes ptr bits) hp) = ((clift hp) :: (pte_C ptr \<rightharpoonup> pte_C))"
6861  (is "?lhs = ?rhs")
6862  using assms
6863  apply -
6864  apply (rule lift_t_typ_region_bytes_none, simp_all)
6865  apply (erule(5) pspace_no_overlap_induce_pte[unfolded size_of_def,simplified])
6866  done
6867
6868lemma typ_bytes_cpspace_relation_clift_endpoint:
6869assumes "cpspace_relation (ksPSpace s) (underlying_memory (ksMachineState (s::kernel_state))) hp"
6870  and "is_aligned ptr bits" "bits < word_bits"
6871  and "pspace_aligned' s"
6872  and "pspace_no_overlap' ptr bits s"
6873 shows "clift (hrs_htd_update (typ_region_bytes ptr bits) hp) = ((clift hp) :: (endpoint_C ptr \<rightharpoonup> endpoint_C))"
6874  (is "?lhs = ?rhs")
6875  using assms
6876  apply -
6877  apply (rule lift_t_typ_region_bytes_none, simp_all)
6878  apply (erule(5) pspace_no_overlap_induce_endpoint[simplified])
6879  done
6880
6881lemma typ_bytes_cpspace_relation_clift_notification:
6882assumes "cpspace_relation (ksPSpace s) (underlying_memory (ksMachineState (s::kernel_state))) hp"
6883  and "is_aligned ptr bits" "bits < word_bits"
6884  and "pspace_aligned' s"
6885  and "pspace_no_overlap' ptr bits s"
6886 shows "clift (hrs_htd_update (typ_region_bytes ptr bits) hp) = ((clift hp) :: (notification_C ptr \<rightharpoonup> notification_C))"
6887  (is "?lhs = ?rhs")
6888  using assms
6889  apply -
6890  apply (rule lift_t_typ_region_bytes_none, simp_all)
6891  apply (erule(5) pspace_no_overlap_induce_notification[simplified])
6892  done
6893
6894lemma typ_bytes_cpspace_relation_clift_vcpu:
6895assumes "cpspace_relation (ksPSpace s) (underlying_memory (ksMachineState (s::kernel_state))) hp"
6896  and "is_aligned ptr bits" "bits < word_bits"
6897  and "pspace_aligned' s"
6898  and "pspace_no_overlap' ptr bits s"
6899 shows "clift (hrs_htd_update (typ_region_bytes ptr bits) hp) = ((clift hp) :: (vcpu_C ptr \<rightharpoonup> vcpu_C))"
6900  (is "?lhs = ?rhs")
6901  using assms
6902  apply -
6903  apply (rule lift_t_typ_region_bytes_none, simp_all)
6904  apply (erule(5) pspace_no_overlap_induce_vcpu[simplified])
6905  done
6906
6907lemma typ_bytes_cpspace_relation_clift_asid_pool:
6908assumes "cpspace_relation (ksPSpace s) (underlying_memory (ksMachineState (s::kernel_state))) hp"
6909  and "is_aligned ptr bits" "bits < word_bits"
6910  and "pspace_aligned' s"
6911  and "pspace_no_overlap' ptr bits s"
6912 shows "clift (hrs_htd_update (typ_region_bytes ptr bits) hp) = ((clift hp) :: (asid_pool_C ptr \<rightharpoonup> asid_pool_C))"
6913  (is "?lhs = ?rhs")
6914  using assms
6915  apply -
6916  apply (rule lift_t_typ_region_bytes_none, simp_all)
6917  apply (erule(5) pspace_no_overlap_induce_asidpool[simplified])
6918  done
6919
6920lemma typ_bytes_cpspace_relation_clift_cte:
6921assumes "cpspace_relation (ksPSpace s) (underlying_memory (ksMachineState (s::kernel_state))) hp"
6922  and "is_aligned ptr bits" "bits < word_bits"
6923  and "pspace_aligned' s"
6924  and "pspace_no_overlap' ptr bits s"
6925 shows "clift (hrs_htd_update (typ_region_bytes ptr bits) hp) = ((clift hp) :: (cte_C ptr \<rightharpoonup> cte_C))"
6926  (is "?lhs = ?rhs")
6927  using assms
6928  apply -
6929  apply (rule lift_t_typ_region_bytes_none)
6930   apply (erule(5) pspace_no_overlap_induce_cte)
6931  apply (simp add: cte_C_size)
6932  done
6933
6934lemma pspace_no_overlap_obj_atD':
6935  "obj_at' P p s \<Longrightarrow> pspace_no_overlap' ptr bits s
6936    \<Longrightarrow> \<exists>ko. P ko \<and> is_aligned p (objBitsKO (injectKOS ko))
6937        \<and> {p .. p + (2 ^ objBitsKO (injectKOS ko)) - 1}
6938            \<inter> {ptr .. (ptr && ~~ mask bits) + 2 ^ bits - 1} = {}"
6939  apply (clarsimp simp: obj_at'_def)
6940  apply (drule(1) pspace_no_overlapD')
6941  apply (clarsimp simp: projectKOs project_inject)
6942  apply auto
6943  done
6944
6945lemma typ_bytes_cpspace_relation_clift_gptr:
6946assumes "cpspace_relation (ksPSpace s) (underlying_memory (ksMachineState (s::kernel_state))) hp"
6947  and "is_aligned ptr bits" "bits < word_bits"
6948  and "pspace_aligned' s"
6949  and "kernel_data_refs \<inter> {ptr ..+ 2^bits} = {}"
6950  and "ptr_span (ptr' :: 'a ptr) \<subseteq> kernel_data_refs"
6951  and "typ_uinfo_t TYPE('a :: mem_type) \<noteq> typ_uinfo_t TYPE(8 word)"
6952 shows "clift (hrs_htd_update (typ_region_bytes ptr bits) hp)
6953    ptr'
6954  = (clift hp) ptr'"
6955  (is "?lhs = ?rhs ptr'")
6956  using assms
6957  apply -
6958   apply (case_tac "ptr' \<notin> dom ?rhs")
6959   apply (frule contra_subsetD[OF typ_region_bytes_dom[where ptr = ptr and bits = bits], rotated])
6960    apply simp
6961   apply fastforce
6962  apply (clarsimp simp: liftt_if hrs_htd_update_def split_def split: if_splits)
6963  apply (simp add: h_t_valid_typ_region_bytes)
6964  apply blast
6965  done
6966
6967lemma cmap_array_typ_region_bytes_triv[OF refl]:
6968  "ptrf = (Ptr :: _ \<Rightarrow> 'b ptr)
6969    \<Longrightarrow> carray_map_relation bits' (map_comp f (ksPSpace s)) (h_t_valid htd c_guard) ptrf
6970    \<Longrightarrow> is_aligned ptr bits
6971    \<Longrightarrow> pspace_no_overlap' ptr bits s
6972    \<Longrightarrow> pspace_aligned' s
6973    \<Longrightarrow> typ_uinfo_t TYPE('b :: c_type) \<noteq> typ_uinfo_t TYPE(8 word)
6974    \<Longrightarrow> size_of TYPE('b) = 2 ^ bits'
6975    \<Longrightarrow> objBitsT (koType TYPE('a :: pspace_storable)) \<le> bits
6976    \<Longrightarrow> objBitsT (koType TYPE('a :: pspace_storable)) \<le> bits'
6977    \<Longrightarrow> bits' < word_bits
6978    \<Longrightarrow> carray_map_relation bits' (map_comp (f :: _ \<Rightarrow> 'a option) (ksPSpace s))
6979        (h_t_valid (typ_region_bytes ptr bits htd) c_guard) ptrf"
6980  apply (frule(7) cmap_array_typ_region_bytes[where ptrf=ptrf])
6981  apply (subst(asm) restrict_map_subdom, simp_all)
6982  apply (drule(1) pspace_no_overlap_disjoint')
6983  apply (simp add: upto_intvl_eq)
6984  apply (rule order_trans[OF map_comp_subset_dom])
6985  apply auto
6986  done
6987
6988lemma intvl_mult_is_union:
6989  "{p..+n * m} = (\<Union>i < m. {p + of_nat (i * n)..+ n})"
6990  apply (cases "n = 0")
6991   apply simp
6992  apply (simp add: intvl_def, safe, simp_all)
6993   apply (rule_tac x="k div n" in bexI)
6994    apply (rule_tac x="k mod n" in exI)
6995    apply (simp only: Abs_fnat_hom_mult Abs_fnat_hom_add, simp)
6996   apply (simp add: Word_Miscellaneous.td_gal_lt[symmetric] mult.commute)
6997  apply (rule_tac x="xa * n + k" in exI, simp)
6998  apply (subst add.commute, rule order_less_le_trans, erule add_less_mono1)
6999  apply (case_tac m, simp_all)
7000  done
7001
7002lemma h_t_array_first_element_at:
7003  "h_t_array_valid htd p n
7004    \<Longrightarrow> 0 < n
7005    \<Longrightarrow> gd p
7006    \<Longrightarrow> h_t_valid htd gd (p :: ('a :: wf_type) ptr)"
7007  apply (clarsimp simp: h_t_array_valid_def h_t_valid_def valid_footprint_def
7008                        Let_def CTypes.sz_nzero[unfolded size_of_def])
7009  apply(drule_tac x="y" in spec, erule impE)
7010   apply (erule order_less_le_trans, simp add: size_of_def)
7011  apply (clarsimp simp: uinfo_array_tag_n_m_def upt_conv_Cons)
7012  apply (erule map_le_trans[rotated])
7013  apply (simp add: list_map_mono split: if_split)
7014  done
7015
7016lemma aligned_intvl_disjointI:
7017  "is_aligned p sz \<Longrightarrow> is_aligned q sz'
7018    \<Longrightarrow> p \<notin> {q ..+ 2 ^ sz'}
7019    \<Longrightarrow> q \<notin> {p ..+ 2 ^ sz}
7020    \<Longrightarrow> {p..+2 ^ sz} \<inter> {q..+2 ^ sz'} = {}"
7021  apply (frule(1) aligned_ranges_subset_or_disjoint[where p=p and p'=q])
7022  apply (simp add: upto_intvl_eq[symmetric])
7023  apply (elim disjE, simp_all)
7024   apply (erule notE, erule subsetD, simp add: intvl_self)
7025  apply (erule notE, erule subsetD, simp add: intvl_self)
7026  done
7027
7028end
7029
7030definition
7031  "cnodes_retype_have_size R bits cns
7032    = (\<forall>ptr' sz'. cns ptr' = Some sz'
7033        \<longrightarrow> is_aligned ptr' (cte_level_bits + sz')
7034            \<and> ({ptr' ..+ 2 ^ (cte_level_bits + sz')} \<inter> R = {}
7035                \<or> cte_level_bits + sz' = bits))"
7036
7037lemma cnodes_retype_have_size_mono:
7038  "cnodes_retype_have_size T bits cns \<and> S \<subseteq> T
7039    \<longrightarrow> cnodes_retype_have_size S bits cns"
7040  by (auto simp add: cnodes_retype_have_size_def)
7041
7042context kernel_m begin
7043
7044lemma gsCNodes_typ_region_bytes:
7045  "cvariable_array_map_relation (gsCNodes \<sigma>) ((^) 2) cte_Ptr (hrs_htd hrs)
7046    \<Longrightarrow> cnodes_retype_have_size {ptr..+2 ^ bits} bits (gsCNodes \<sigma>)
7047    \<Longrightarrow> 0 \<notin> {ptr..+2 ^ bits} \<Longrightarrow> is_aligned ptr bits
7048    \<Longrightarrow> clift (hrs_htd_update (typ_region_bytes ptr bits) hrs)
7049        = (clift hrs :: cte_C ptr \<Rightarrow> _)
7050    \<Longrightarrow> cvariable_array_map_relation (gsCNodes \<sigma>) ((^) 2) cte_Ptr
7051        (typ_region_bytes ptr bits (hrs_htd hrs))"
7052  apply (clarsimp simp: cvariable_array_map_relation_def
7053                        h_t_array_valid_def)
7054  apply (elim allE, drule(1) mp)
7055  apply (subst valid_footprint_typ_region_bytes)
7056   apply (simp add: uinfo_array_tag_n_m_def typ_uinfo_t_def typ_info_word)
7057  apply (clarsimp simp: cnodes_retype_have_size_def field_simps)
7058  apply (elim allE, drule(1) mp)
7059  apply (subgoal_tac "size_of TYPE(cte_C) * 2 ^ v = 2 ^ (cte_level_bits + v)")
7060  prefer 2
7061   apply (simp add: cte_C_size cte_level_bits_def power_add)
7062  apply (clarsimp simp add: upto_intvl_eq[symmetric] field_simps)
7063  apply (case_tac "p \<in> {ptr ..+ 2 ^ bits}")
7064   apply (drule h_t_array_first_element_at[where p="Ptr p" and gd=c_guard for p,
7065       unfolded h_t_array_valid_def, simplified])
7066     apply simp
7067    apply (rule is_aligned_c_guard[where m=2], simp+)
7068       apply clarsimp
7069      apply (simp add: align_of_def)
7070     apply (simp add: size_of_def cte_level_bits_def power_add)
7071    apply (simp add: cte_level_bits_def)
7072   apply (drule_tac x="cte_Ptr p" in fun_cong)
7073   apply (simp add: liftt_if[folded hrs_htd_def] hrs_htd_update
7074                    h_t_valid_def valid_footprint_typ_region_bytes
7075             split: if_split_asm)
7076   apply (subgoal_tac "p \<in> {p ..+ size_of TYPE(cte_C)}")
7077    apply (simp add: cte_C_size)
7078    apply blast
7079   apply (simp add: intvl_self)
7080  apply (simp only: upto_intvl_eq mask_in_range[symmetric])
7081  apply (rule aligned_ranges_subset_or_disjoint_coroll, simp_all)
7082  done
7083
7084lemma tcb_ctes_typ_region_bytes:
7085  "cvariable_array_map_relation (map_to_tcbs (ksPSpace \<sigma>))
7086      (\<lambda>x. 5) cte_Ptr (hrs_htd hrs)
7087    \<Longrightarrow> pspace_no_overlap' ptr bits \<sigma>
7088    \<Longrightarrow> pspace_aligned' \<sigma>
7089    \<Longrightarrow> is_aligned ptr bits
7090    \<Longrightarrow> cpspace_tcb_relation (ksPSpace \<sigma>) hrs
7091    \<Longrightarrow> cvariable_array_map_relation (map_to_tcbs (ksPSpace \<sigma>)) (\<lambda>x. 5)
7092        cte_Ptr (typ_region_bytes ptr bits (hrs_htd hrs))"
7093  apply (clarsimp simp: cvariable_array_map_relation_def
7094                        h_t_array_valid_def)
7095  apply (drule spec, drule mp, erule exI)
7096  apply (subst valid_footprint_typ_region_bytes)
7097   apply (simp add: uinfo_array_tag_n_m_def typ_uinfo_t_def typ_info_word)
7098  apply (clarsimp simp only: map_comp_Some_iff projectKOs
7099                        pspace_no_overlap'_def is_aligned_neg_mask_eq simp_thms
7100                        field_simps upto_intvl_eq[symmetric])
7101  apply (elim allE, drule(1) mp)
7102  apply (drule(1) pspace_alignedD')
7103  apply (erule disjoint_subset[rotated])
7104  apply (simp add: upto_intvl_eq[symmetric])
7105  apply (rule intvl_start_le)
7106  apply (simp add: objBits_simps' cte_C_size)
7107  done
7108
7109lemma ccorres_typ_region_bytes_dummy:
7110  "ccorresG rf_sr
7111     AnyGamma dc xfdc
7112     (invs' and ct_active' and sch_act_simple and
7113      pspace_no_overlap' ptr bits and
7114      (cnodes_retype_have_size S bits o gsCNodes)
7115      and K (bits < word_bits \<and> is_aligned ptr bits \<and> 4 \<le> bits
7116         \<and> 0 \<notin> {ptr..+2 ^ bits}
7117         \<and> {ptr ..+ 2 ^ bits} \<subseteq> S
7118         \<and> kernel_data_refs \<inter> {ptr..+2 ^ bits} = {}))
7119     UNIV hs
7120     (return ())
7121     (global_htd_update (\<lambda>_. (typ_region_bytes ptr bits)))"
7122  apply (rule ccorres_from_vcg)
7123  apply (clarsimp simp: return_def)
7124  apply (simp add: rf_sr_def)
7125  apply vcg
7126  apply (clarsimp simp: cstate_relation_def Let_def)
7127  apply (frule typ_bytes_cpspace_relation_clift_tcb)
7128      apply (simp add: invs_pspace_aligned')+
7129  apply (frule typ_bytes_cpspace_relation_clift_pte)
7130      apply (simp add: invs_pspace_aligned')+
7131  apply (frule typ_bytes_cpspace_relation_clift_pde)
7132      apply (simp add: invs_pspace_aligned')+
7133  apply (frule typ_bytes_cpspace_relation_clift_endpoint)
7134      apply (simp add: invs_pspace_aligned')+
7135  apply (frule typ_bytes_cpspace_relation_clift_notification)
7136      apply (simp add: invs_pspace_aligned')+
7137  apply (frule typ_bytes_cpspace_relation_clift_asid_pool)
7138      apply (simp add: invs_pspace_aligned')+
7139  apply (frule typ_bytes_cpspace_relation_clift_cte)
7140      apply (simp add: invs_pspace_aligned')+
7141  apply (frule typ_bytes_cpspace_relation_clift_userdata)
7142      apply (simp add: invs_pspace_aligned')+
7143  apply (frule typ_bytes_cpspace_relation_clift_devicedata)
7144      apply (simp add: invs_pspace_aligned')+
7145  apply (frule typ_bytes_cpspace_relation_clift_vcpu)
7146      apply (simp add: invs_pspace_aligned')+
7147  apply (frule typ_bytes_cpspace_relation_clift_gptr[where
7148            ptr'="pd_Ptr (symbol_table ''armUSGlobalPD'')"])
7149        apply (simp add: invs_pspace_aligned')+
7150  apply (frule typ_bytes_cpspace_relation_clift_gptr[where
7151            ptr'="ptr_coerce x :: (cte_C[256]) ptr" for x])
7152        apply (simp add: invs_pspace_aligned')+
7153    apply (simp add: cte_level_bits_def cte_C_size, simp+)
7154  apply (simp add: carch_state_relation_def cmachine_state_relation_def)
7155  apply (simp add: cpspace_relation_def htd_safe_typ_region_bytes)
7156  apply (simp add: h_t_valid_clift_Some_iff)
7157  apply (simp add: hrs_htd_update gsCNodes_typ_region_bytes
7158                   cnodes_retype_have_size_mono[where T=S]
7159                   tcb_ctes_typ_region_bytes[OF _ _ invs_pspace_aligned'])
7160  apply (simp add: cmap_array_typ_region_bytes_triv
7161               invs_pspace_aligned' table_bits_defs
7162               objBitsT_simps word_bits_def
7163               zero_ranges_are_zero_typ_region_bytes)
7164  apply (rule htd_safe_typ_region_bytes, simp)
7165  apply blast
7166  done
7167
7168lemma region_is_typeless_cong:
7169  "t_hrs_' (globals t) = t_hrs_' (globals s)
7170   \<Longrightarrow> region_is_typeless ptr sz s = region_is_typeless ptr sz t"
7171  by (simp add:region_is_typeless_def)
7172
7173lemma region_is_bytes_cong:
7174  "t_hrs_' (globals t) = t_hrs_' (globals s)
7175   \<Longrightarrow> region_is_bytes ptr sz s = region_is_bytes ptr sz t"
7176  by (simp add:region_is_bytes'_def)
7177
7178lemma insertNewCap_sch_act_simple[wp]:
7179 "\<lbrace>sch_act_simple\<rbrace>insertNewCap a b c\<lbrace>\<lambda>_. sch_act_simple\<rbrace>"
7180  by (simp add:sch_act_simple_def,wp)
7181
7182lemma updateMDB_ctes_of_cap:
7183  "\<lbrace>\<lambda>s. (\<forall>x\<in>ran(ctes_of s). P (cteCap x)) \<and> no_0 (ctes_of s)\<rbrace>
7184    updateMDB srcSlot t
7185  \<lbrace>\<lambda>r s. \<forall>x\<in>ran (ctes_of s). P (cteCap x)\<rbrace>"
7186  apply (rule hoare_pre)
7187  apply wp
7188  apply (clarsimp)
7189  apply (erule ranE)
7190  apply (clarsimp simp:modify_map_def split:if_splits)
7191   apply (drule_tac x = z in bspec)
7192    apply fastforce
7193   apply simp
7194  apply (drule_tac x = x in bspec)
7195   apply fastforce
7196  apply simp
7197  done
7198
7199lemma insertNewCap_caps_no_overlap'':
7200notes blah[simp del] =  atLeastAtMost_iff atLeastatMost_subset_iff atLeastLessThan_iff
7201      Int_atLeastAtMost atLeastatMost_empty_iff split_paired_Ex
7202shows "\<lbrace>cte_wp_at' (\<lambda>_. True) cptr and valid_pspace'
7203        and caps_no_overlap'' ptr us
7204        and K  (cptr \<noteq> (0::word32)) and K (untypedRange x \<inter> {ptr..(ptr && ~~ mask us) + 2 ^ us - 1} = {})\<rbrace>
7205 insertNewCap srcSlot cptr x
7206          \<lbrace>\<lambda>rv s. caps_no_overlap'' ptr us s\<rbrace>"
7207  apply (clarsimp simp:insertNewCap_def caps_no_overlap''_def)
7208  apply (rule hoare_pre)
7209   apply (wp getCTE_wp updateMDB_ctes_of_cap)
7210  apply (clarsimp simp:cte_wp_at_ctes_of valid_pspace'_def
7211    valid_mdb'_def valid_mdb_ctes_def no_0_def split:if_splits)
7212  apply (erule ranE)
7213  apply (clarsimp split:if_splits)
7214  apply (frule_tac c=  "(cteCap xa)" and q = xb in caps_no_overlapD''[rotated])
7215   apply (clarsimp simp:cte_wp_at_ctes_of)
7216  apply clarsimp
7217  apply blast
7218  done
7219
7220lemma insertNewCap_caps_overlap_reserved':
7221notes blah[simp del] =  atLeastAtMost_iff atLeastatMost_subset_iff atLeastLessThan_iff
7222      Int_atLeastAtMost atLeastatMost_empty_iff split_paired_Ex
7223shows "\<lbrace>cte_wp_at' (\<lambda>_. True) cptr and valid_pspace' and caps_overlap_reserved' S
7224        and valid_cap' x and K  (cptr \<noteq> (0::word32)) and K (untypedRange x \<inter> S = {})\<rbrace>
7225       insertNewCap srcSlot cptr x
7226       \<lbrace>\<lambda>rv s. caps_overlap_reserved' S s\<rbrace>"
7227   apply (clarsimp simp:insertNewCap_def caps_overlap_reserved'_def)
7228   apply (rule hoare_pre)
7229   apply (wp getCTE_wp updateMDB_ctes_of_cap)
7230   apply (clarsimp simp:cte_wp_at_ctes_of valid_pspace'_def
7231    valid_mdb'_def valid_mdb_ctes_def no_0_def split:if_splits)
7232   apply (erule ranE)
7233   apply (clarsimp split:if_splits)
7234   apply (drule usableRange_subseteq[rotated])
7235     apply (simp add:valid_cap'_def)
7236    apply blast
7237   apply (drule_tac p = xaa in caps_overlap_reserved'_D)
7238     apply simp
7239    apply simp
7240   apply blast
7241  done
7242
7243lemma insertNewCap_pspace_no_overlap':
7244notes blah[simp del] =  atLeastAtMost_iff atLeastatMost_subset_iff atLeastLessThan_iff
7245      Int_atLeastAtMost atLeastatMost_empty_iff split_paired_Ex
7246shows "\<lbrace>pspace_no_overlap' ptr sz and pspace_aligned'
7247  and pspace_distinct' and cte_wp_at' (\<lambda>_. True) cptr\<rbrace>
7248  insertNewCap srcSlot cptr x
7249  \<lbrace>\<lambda>rv s. pspace_no_overlap' ptr sz s\<rbrace>"
7250   apply (clarsimp simp:insertNewCap_def)
7251   apply (rule hoare_pre)
7252   apply (wp updateMDB_pspace_no_overlap'
7253     setCTE_pspace_no_overlap' getCTE_wp)
7254   apply (clarsimp simp:cte_wp_at_ctes_of)
7255   done
7256
7257lemma insertNewCap_cte_at:
7258  "\<lbrace>cte_at' p\<rbrace> insertNewCap srcSlot q cap
7259   \<lbrace>\<lambda>rv. cte_at' p\<rbrace>"
7260  apply (clarsimp simp:insertNewCap_def)
7261  apply (wp getCTE_wp)
7262  apply (clarsimp simp:cte_wp_at_ctes_of)
7263  done
7264
7265lemma createObject_invs':
7266  "\<lbrace>\<lambda>s. invs' s \<and> ct_active' s \<and> pspace_no_overlap' ptr (APIType_capBits ty us) s
7267          \<and> caps_no_overlap'' ptr (APIType_capBits ty us) s \<and> ptr \<noteq> 0 \<and>
7268          caps_overlap_reserved' {ptr..ptr + 2 ^ APIType_capBits ty us - 1} s \<and>
7269          (ty = APIObjectType apiobject_type.CapTableObject \<longrightarrow> 0 < us) \<and>
7270          is_aligned ptr (APIType_capBits ty us) \<and> APIType_capBits ty us < word_bits \<and>
7271          {ptr..ptr + 2 ^ APIType_capBits ty us - 1} \<inter> kernel_data_refs = {} \<and>
7272          0 < gsMaxObjectSize s
7273    \<rbrace> createObject ty ptr us dev\<lbrace>\<lambda>r s. invs' s \<rbrace>"
7274  apply (simp add:createObject_def3)
7275  apply (rule hoare_pre)
7276  apply (wp createNewCaps_invs'[where sz = "APIType_capBits ty us"])
7277  apply (clarsimp simp:range_cover_full)
7278  done
7279
7280lemma createObject_sch_act_simple[wp]:
7281  "\<lbrace>\<lambda>s. sch_act_simple s
7282    \<rbrace>createObject ty ptr us dev\<lbrace>\<lambda>r s. sch_act_simple s \<rbrace>"
7283 apply (simp add:sch_act_simple_def)
7284 apply wp
7285 done
7286
7287lemma createObject_ct_active'[wp]:
7288  "\<lbrace>\<lambda>s. ct_active' s \<and> pspace_aligned' s \<and> pspace_distinct' s
7289     \<and>  pspace_no_overlap' ptr (APIType_capBits ty us) s
7290     \<and>  is_aligned ptr (APIType_capBits ty us) \<and> APIType_capBits ty us < word_bits
7291    \<rbrace>createObject ty ptr us dev\<lbrace>\<lambda>r s. ct_active' s \<rbrace>"
7292 apply (simp add:ct_in_state'_def createObject_def3)
7293 apply (rule hoare_pre)
7294 apply wp
7295 apply wps
7296 apply (wp createNewCaps_pred_tcb_at')
7297 apply (intro conjI)
7298 apply (auto simp:range_cover_full)
7299 done
7300
7301lemma createObject_notZombie[wp]:
7302  "\<lbrace>\<top>\<rbrace>createObject ty ptr us dev \<lbrace>\<lambda>r s. \<not> isZombie r\<rbrace>"
7303  apply (rule hoare_pre)
7304  apply (simp add:createObject_def)
7305   apply wpc
7306    apply (wp| clarsimp simp add:isCap_simps)+
7307   apply wpc
7308    apply (wp| clarsimp simp add:isCap_simps)+
7309  done
7310
7311lemma createObject_valid_cap':
7312  "\<lbrace>\<lambda>s. pspace_no_overlap' ptr (APIType_capBits ty us) s \<and>
7313         valid_pspace' s \<and>
7314         is_aligned ptr (APIType_capBits ty us) \<and>
7315          APIType_capBits ty us < word_bits \<and>
7316         (ty = APIObjectType apiobject_type.CapTableObject \<longrightarrow> 0 < us) \<and>
7317         (ty = APIObjectType apiobject_type.Untyped \<longrightarrow> minUntypedSizeBits \<le> us \<and> us \<le> maxUntypedSizeBits) \<and> ptr \<noteq> 0\<rbrace>
7318    createObject ty ptr us dev \<lbrace>\<lambda>r s. s \<turnstile>' r\<rbrace>"
7319  apply (simp add:createObject_def3)
7320  apply (rule hoare_pre)
7321  apply wp
7322   apply (rule_tac Q = "\<lambda>r s. r \<noteq> [] \<and> Q r s" for Q in hoare_strengthen_post)
7323   apply (rule hoare_vcg_conj_lift)
7324     apply (rule hoare_strengthen_post[OF createNewCaps_ret_len])
7325      apply clarsimp
7326     apply (rule hoare_strengthen_post[OF createNewCaps_valid_cap'[where sz = "APIType_capBits ty us"]])
7327    apply assumption
7328   apply clarsimp
7329  apply (clarsimp simp add:word_bits_conv range_cover_full)
7330  done
7331
7332lemma createObject_untypedRange:
7333  assumes split:
7334    "\<lbrace>P\<rbrace> createObject ty ptr us dev
7335     \<lbrace>\<lambda>m s. (toAPIType ty = Some apiobject_type.Untyped \<longrightarrow>
7336                            Q {ptr..ptr + 2 ^ us - 1} s) \<and>
7337            (toAPIType ty \<noteq> Some apiobject_type.Untyped \<longrightarrow> Q {} s)\<rbrace>"
7338  shows "\<lbrace>P\<rbrace> createObject ty ptr us dev\<lbrace>\<lambda>m s. Q (untypedRange m) s\<rbrace>"
7339  including no_pre
7340  using split
7341  apply (simp add: createObject_def)
7342  apply (case_tac "toAPIType ty")
7343   apply (simp add: split | wp)+
7344   apply (simp add: valid_def return_def bind_def split_def)
7345  apply (case_tac a, simp_all)
7346      apply (simp add: valid_def return_def simpler_gets_def simpler_modify_def
7347                       bind_def split_def curDomain_def)+
7348  done
7349
7350lemma createObject_capRange:
7351shows "\<lbrace>P\<rbrace>createObject ty ptr us dev \<lbrace>\<lambda>m s. capRange m = {ptr.. ptr + 2 ^ (APIType_capBits ty us) - 1}\<rbrace>"
7352  apply (simp add:createObject_def)
7353  apply (case_tac "ty")
7354    apply (simp_all add:toAPIType_def ARM_HYP_H.toAPIType_def)
7355        apply (rule hoare_pre)
7356         apply wpc
7357             apply wp
7358        apply (simp add:split untypedRange.simps objBits_simps capRange_def APIType_capBits_def | wp)+
7359       apply (wpsimp simp: ARM_HYP_H.createObject_def capRange_def APIType_capBits_def
7360                        machine_bits_defs acapClass.simps)+
7361  done
7362
7363lemma createObject_capRange_helper:
7364assumes static: "\<lbrace>P\<rbrace>createObject ty ptr us dev \<lbrace>\<lambda>m s. Q {ptr.. ptr + 2 ^ (APIType_capBits ty us) - 1} s\<rbrace>"
7365shows "\<lbrace>P\<rbrace>createObject ty ptr us dev \<lbrace>\<lambda>m s. Q (capRange m) s\<rbrace>"
7366  apply (rule hoare_pre)
7367   apply (rule hoare_strengthen_post[OF hoare_vcg_conj_lift])
7368     apply (rule static)
7369    apply (rule createObject_capRange)
7370   apply simp
7371  apply simp
7372  done
7373
7374lemma createObject_caps_overlap_reserved':
7375  "\<lbrace>\<lambda>s. caps_overlap_reserved' S s \<and>
7376         pspace_aligned' s \<and>
7377         pspace_distinct' s \<and> pspace_no_overlap' ptr (APIType_capBits ty us) s \<and>
7378         is_aligned ptr (APIType_capBits ty us) \<and> APIType_capBits ty us < word_bits
7379    \<rbrace>createObject ty ptr us dev \<lbrace>\<lambda>rv. caps_overlap_reserved' S\<rbrace>"
7380  apply (simp add:createObject_def3)
7381  apply (wp createNewCaps_caps_overlap_reserved'[where sz = "APIType_capBits ty us"])
7382  apply (clarsimp simp:range_cover_full)
7383  done
7384
7385lemma createObject_caps_overlap_reserved_ret':
7386  "\<lbrace>\<lambda>s.  caps_overlap_reserved' {ptr..ptr + 2 ^ APIType_capBits ty us - 1} s \<and>
7387         pspace_aligned' s \<and>
7388         pspace_distinct' s \<and> pspace_no_overlap' ptr (APIType_capBits ty us) s \<and>
7389         is_aligned ptr (APIType_capBits ty us) \<and> APIType_capBits ty us < word_bits
7390    \<rbrace>createObject ty ptr us dev \<lbrace>\<lambda>rv. caps_overlap_reserved' (untypedRange rv)\<rbrace>"
7391  apply (simp add:createObject_def3)
7392  apply (rule hoare_pre)
7393  apply wp
7394   apply (rule_tac Q = "\<lambda>r s. r \<noteq> [] \<and> Q r s" for Q in hoare_strengthen_post)
7395   apply (rule hoare_vcg_conj_lift)
7396     apply (rule hoare_strengthen_post[OF createNewCaps_ret_len])
7397      apply clarsimp
7398     apply (rule hoare_strengthen_post[OF createNewCaps_caps_overlap_reserved_ret'[where sz = "APIType_capBits ty us"]])
7399    apply assumption
7400   apply (case_tac r,simp)
7401   apply clarsimp
7402   apply (erule caps_overlap_reserved'_subseteq)
7403   apply (rule untypedRange_in_capRange)
7404  apply (clarsimp simp add:word_bits_conv range_cover_full)
7405  done
7406
7407lemma createObject_descendants_range':
7408  "\<lbrace>\<lambda>s.  descendants_range_in' {ptr..ptr + 2 ^ APIType_capBits ty us - 1} q (ctes_of s) \<and>
7409         pspace_aligned' s \<and>
7410         pspace_distinct' s \<and> pspace_no_overlap' ptr (APIType_capBits ty us) s \<and>
7411         is_aligned ptr (APIType_capBits ty us) \<and> APIType_capBits ty us < word_bits
7412    \<rbrace>createObject ty ptr us dev \<lbrace>\<lambda>rv s. descendants_range' rv q (ctes_of s)\<rbrace>"
7413  apply (simp add:createObject_def3)
7414  apply (rule hoare_pre)
7415  apply wp
7416   apply (rule_tac Q = "\<lambda>r s. r \<noteq> [] \<and> Q r s" for Q in hoare_strengthen_post)
7417   apply (rule hoare_vcg_conj_lift)
7418     apply (rule hoare_strengthen_post[OF createNewCaps_ret_len])
7419      apply clarsimp
7420     apply (rule hoare_strengthen_post[OF createNewCaps_descendants_range_ret'[where sz = "APIType_capBits ty us"]])
7421    apply assumption
7422   apply fastforce
7423  apply (clarsimp simp add:word_bits_conv range_cover_full)
7424  done
7425
7426lemma createObject_descendants_range_in':
7427  "\<lbrace>\<lambda>s.  descendants_range_in' S q (ctes_of s) \<and>
7428         pspace_aligned' s \<and>
7429         pspace_distinct' s \<and> pspace_no_overlap' ptr (APIType_capBits ty us) s \<and>
7430         is_aligned ptr (APIType_capBits ty us) \<and> APIType_capBits ty us < word_bits
7431    \<rbrace>createObject ty ptr us dev \<lbrace>\<lambda>rv s. descendants_range_in' S q (ctes_of s)\<rbrace>"
7432  apply (simp add:createObject_def3 descendants_range_in'_def2)
7433  apply (wp createNewCaps_null_filter')
7434  apply clarsimp
7435  apply (intro conjI)
7436   apply simp
7437  apply (simp add:range_cover_full)
7438  done
7439
7440lemma createObject_idlethread_range:
7441  "\<lbrace>\<lambda>s. is_aligned ptr (APIType_capBits ty us) \<and> APIType_capBits ty us < word_bits
7442        \<and> ksIdleThread s \<notin> {ptr..ptr + 2 ^ (APIType_capBits ty us) - 1}\<rbrace>
7443   createObject ty ptr us dev \<lbrace>\<lambda>cap s. ksIdleThread s \<notin> capRange cap\<rbrace>"
7444  apply (simp add:createObject_def3)
7445  apply (rule hoare_pre)
7446  apply wp
7447   apply (rule_tac Q = "\<lambda>r s. r \<noteq> [] \<and> Q r s" for Q in hoare_strengthen_post)
7448   apply (rule hoare_vcg_conj_lift)
7449     apply (rule hoare_strengthen_post[OF createNewCaps_ret_len])
7450      apply clarsimp
7451     apply (rule hoare_strengthen_post[OF createNewCaps_idlethread_ranges[where sz = "APIType_capBits ty us"]])
7452    apply assumption
7453   apply clarsimp
7454  apply (clarsimp simp: word_bits_conv range_cover_full)
7455  done
7456
7457lemma caps_overlap_reserved_empty'[simp]:
7458  "caps_overlap_reserved' {} s = True"
7459  by (simp add:caps_overlap_reserved'_def)
7460
7461lemma createObject_IRQHandler:
7462  "\<lbrace>\<top>\<rbrace> createObject ty ptr us dev
7463    \<lbrace>\<lambda>rv s. rv = IRQHandlerCap x \<longrightarrow> P rv s x\<rbrace>"
7464  apply (simp add:createObject_def3)
7465  apply (rule hoare_pre)
7466  apply wp
7467   apply (rule_tac Q = "\<lambda>r s. r \<noteq> [] \<and> Q r s" for Q in hoare_strengthen_post)
7468   apply (rule hoare_vcg_conj_lift)
7469     apply (rule hoare_strengthen_post[OF createNewCaps_ret_len])
7470      apply clarsimp
7471     apply (rule hoare_strengthen_post[OF createNewCaps_IRQHandler[where irq = x and P = "\<lambda>_ _. False"]])
7472    apply assumption
7473   apply (case_tac r,clarsimp+)
7474  apply (clarsimp simp:word_bits_conv)
7475  done
7476
7477lemma createObject_capClass[wp]:
7478  "\<lbrace> \<lambda>s. is_aligned ptr (APIType_capBits ty us) \<and> APIType_capBits ty us < word_bits
7479   \<rbrace> createObject ty ptr us dev
7480   \<lbrace>\<lambda>rv s. capClass rv = PhysicalClass\<rbrace>"
7481  apply (simp add:createObject_def3)
7482  apply (rule hoare_pre)
7483  apply wp
7484   apply (rule_tac Q = "\<lambda>r s. r \<noteq> [] \<and> Q r s" for Q in hoare_strengthen_post)
7485   apply (rule hoare_vcg_conj_lift)
7486     apply (rule hoare_strengthen_post[OF createNewCaps_ret_len])
7487      apply clarsimp
7488     apply (rule hoare_strengthen_post[OF createNewCaps_range_helper])
7489    apply assumption
7490   apply (case_tac r,clarsimp+)
7491  apply (clarsimp simp:word_bits_conv )
7492  apply (rule range_cover_full)
7493   apply (simp add:word_bits_conv)+
7494  done
7495
7496lemma createObject_child:
7497  "\<lbrace>\<lambda>s.
7498     is_aligned ptr (APIType_capBits ty us) \<and> APIType_capBits ty us < word_bits \<and>
7499     {ptr .. ptr + (2^APIType_capBits ty us) - 1} \<subseteq> (untypedRange cap) \<and> isUntypedCap cap
7500   \<rbrace> createObject ty ptr us dev
7501   \<lbrace>\<lambda>rv s. sameRegionAs cap rv\<rbrace>"
7502  apply (rule hoare_assume_pre)
7503  apply (simp add:createObject_def3)
7504  apply wp
7505   apply (rule hoare_chain [OF createNewCaps_range_helper[where sz = "APIType_capBits ty us"]])
7506    apply (fastforce simp:range_cover_full)
7507   apply clarsimp
7508   apply (drule_tac x = ptr in spec)
7509   apply (case_tac "(capfn ptr)")
7510              apply (simp_all add: capUntypedPtr_def sameRegionAs_def Let_def isCap_simps)+
7511        apply clarsimp+
7512    apply (rename_tac arch_capability d v0 v1 f)
7513    apply (simp add: ARM_HYP_H.capUntypedSize_def machine_bits_defs)+
7514    apply (case_tac arch_capability,
7515           auto simp: ARM_HYP_H.capUntypedSize_def machine_bits_defs
7516                      is_aligned_no_wrap' field_simps
7517               split: arch_capability.split)+
7518  done
7519
7520lemma createObject_parent_helper:
7521  "\<lbrace>\<lambda>s. cte_wp_at' (\<lambda>cte. isUntypedCap (cteCap cte)
7522         \<and> {ptr .. ptr + (2^APIType_capBits ty us) - 1} \<subseteq> untypedRange (cteCap cte)) p s \<and>
7523         pspace_aligned' s \<and>
7524         pspace_distinct' s \<and>
7525         pspace_no_overlap' ptr (APIType_capBits ty us) s \<and>
7526         is_aligned ptr (APIType_capBits ty us) \<and> APIType_capBits ty us < word_bits \<and>
7527         (ty = APIObjectType apiobject_type.CapTableObject \<longrightarrow> 0 < us)
7528    \<rbrace>
7529    createObject ty ptr us dev
7530    \<lbrace>\<lambda>rv. cte_wp_at' (\<lambda>cte. isUntypedCap (cteCap cte) \<and> (sameRegionAs (cteCap cte) rv)) p\<rbrace>"
7531  apply (rule hoare_post_imp [where Q="\<lambda>rv s. \<exists>cte. cte_wp_at' ((=) cte) p s
7532                                           \<and> isUntypedCap (cteCap cte) \<and>
7533                                sameRegionAs (cteCap cte) rv"])
7534  apply (clarsimp simp:cte_wp_at_ctes_of)
7535  apply (wp hoare_vcg_ex_lift)
7536   apply (rule hoare_vcg_conj_lift)
7537   apply (simp add:createObject_def3)
7538    apply (wp createNewCaps_cte_wp_at')
7539   apply (wp createObject_child)
7540  apply (clarsimp simp:cte_wp_at_ctes_of)
7541  apply (intro conjI)
7542   apply (erule range_cover_full)
7543    apply simp
7544  apply simp
7545  done
7546
7547lemma insertNewCap_untypedRange:
7548  "\<lbrace>\<lambda>s. cte_wp_at' (\<lambda>cte. isUntypedCap (cteCap cte) \<and> P untypedRange (cteCap cte)) srcSlot s\<rbrace>
7549    insertNewCap srcSlot destSlot x
7550   \<lbrace>\<lambda>rv s. cte_wp_at' (\<lambda>cte. isUntypedCap (cteCap cte) \<and> P untypedRange (cteCap cte)) srcSlot s\<rbrace>"
7551  apply (simp add:insertNewCap_def)
7552  apply (wp updateMDB_weak_cte_wp_at setCTE_cte_wp_at_other getCTE_wp)
7553  apply (clarsimp simp:cte_wp_at_ctes_of)
7554  done
7555
7556lemma createObject_caps_no_overlap'':
7557  " \<lbrace>\<lambda>s. caps_no_overlap'' (ptr + (1 + of_nat n << APIType_capBits newType userSize))
7558                     sz s \<and>
7559     pspace_aligned' s \<and> pspace_distinct' s \<and>
7560     pspace_no_overlap' (ptr + (of_nat n << APIType_capBits newType userSize)) (APIType_capBits newType userSize) s
7561     \<and> is_aligned ptr (APIType_capBits newType userSize)
7562     \<and> APIType_capBits newType userSize < word_bits\<rbrace>
7563   createObject newType (ptr + (of_nat n << APIType_capBits newType userSize)) userSize dev
7564   \<lbrace>\<lambda>rv s. caps_no_overlap'' (ptr + (1 + of_nat n << APIType_capBits newType userSize))
7565                     sz s \<rbrace>"
7566  apply (clarsimp simp:createObject_def3 caps_no_overlap''_def2)
7567  apply (wp createNewCaps_null_filter')
7568  apply clarsimp
7569  apply (intro conjI)
7570   apply simp
7571  apply (rule range_cover_full)
7572   apply (erule aligned_add_aligned)
7573     apply (rule is_aligned_shiftl_self)
7574    apply simp
7575   apply simp
7576  done
7577
7578lemma createObject_ex_cte_cap_wp_to:
7579  "\<lbrace>\<lambda>s. ex_cte_cap_wp_to' P p s \<and> is_aligned ptr (APIType_capBits ty us) \<and> pspace_aligned' s
7580    \<and> pspace_distinct' s \<and> (APIType_capBits ty us) < word_bits  \<and> pspace_no_overlap' ptr (APIType_capBits ty us) s \<rbrace>
7581    createObject ty ptr us dev
7582   \<lbrace>\<lambda>rv s. ex_cte_cap_wp_to' P p s \<rbrace>"
7583  apply (clarsimp simp:ex_cte_cap_wp_to'_def createObject_def3)
7584  apply (rule hoare_pre)
7585   apply (wp hoare_vcg_ex_lift)
7586   apply wps
7587   apply (wp createNewCaps_cte_wp_at')
7588  apply clarsimp
7589  apply (intro exI conjI)
7590      apply assumption
7591     apply (rule range_cover_full)
7592    apply (clarsimp simp:cte_wp_at_ctes_of)
7593   apply simp
7594  apply simp
7595  done
7596
7597
7598lemma range_cover_one:
7599  "\<lbrakk>is_aligned (ptr :: 'a :: len word) us; us\<le> sz;sz < len_of TYPE('a)\<rbrakk>
7600  \<Longrightarrow> range_cover ptr sz us (Suc 0)"
7601  apply (clarsimp simp:range_cover_def)
7602  apply (rule Suc_leI)
7603  apply (rule unat_less_power)
7604   apply simp
7605  apply (rule shiftr_less_t2n)
7606   apply simp
7607  apply (rule le_less_trans[OF word_and_le1])
7608  apply (simp add:mask_def)
7609  done
7610
7611lemma createObject_no_inter:
7612notes blah[simp del] =  atLeastAtMost_iff atLeastatMost_subset_iff atLeastLessThan_iff
7613      Int_atLeastAtMost atLeastatMost_empty_iff split_paired_Ex
7614shows
7615  "\<lbrace>\<lambda>s. range_cover ptr sz (APIType_capBits newType userSize) (n + 2) \<and> ptr \<noteq> 0\<rbrace>
7616  createObject newType (ptr + (of_nat n << APIType_capBits newType userSize)) userSize dev
7617  \<lbrace>\<lambda>rv s. untypedRange rv \<inter>
7618  {ptr + (1 + of_nat n << APIType_capBits newType userSize) ..
7619   ptrend } =
7620  {}\<rbrace>"
7621  apply (rule createObject_untypedRange)
7622  apply (clarsimp | wp)+
7623  apply (clarsimp simp: blah toAPIType_def APIType_capBits_def
7624    ARM_HYP_H.toAPIType_def split: object_type.splits)
7625  apply (clarsimp simp:shiftl_t2n field_simps)
7626  apply (drule word_eq_zeroI)
7627  apply (drule(1) range_cover_no_0[where p = "Suc n"])
7628   apply simp
7629  apply (simp add:field_simps)
7630  done
7631
7632lemma range_cover_bound'':
7633  "\<lbrakk>range_cover ptr sz us n; x < of_nat n\<rbrakk>
7634  \<Longrightarrow> ptr + x * 2 ^ us + 2 ^ us - 1 \<le> (ptr && ~~ mask sz) + 2 ^ sz - 1"
7635  apply (frule range_cover_cell_subset)
7636   apply assumption
7637  apply (drule(1) range_cover_subset_not_empty)
7638   apply (clarsimp simp: field_simps)
7639  done
7640
7641lemma caps_no_overlap''_cell:
7642  "\<lbrakk>range_cover ptr sz us n;caps_no_overlap'' ptr sz s;p < n\<rbrakk>
7643    \<Longrightarrow> caps_no_overlap'' (ptr + (of_nat p << us)) us s"
7644  apply (clarsimp simp:caps_no_overlap''_def)
7645  apply (drule(1) bspec)
7646  apply (subgoal_tac  "{ptr + (of_nat p << us)..(ptr + (of_nat p << us) && ~~ mask us) + 2 ^ us - 1}
7647                      \<subseteq>  {ptr..(ptr && ~~ mask sz) + 2 ^ sz - 1}")
7648   apply (erule impE)
7649    apply (rule ccontr)
7650    apply clarify
7651    apply (drule(1) disjoint_subset2[rotated -1])
7652    apply simp
7653   apply (erule subsetD)+
7654   apply simp
7655  apply (subst is_aligned_neg_mask_eq)
7656   apply (rule aligned_add_aligned[OF range_cover.aligned],assumption)
7657     apply (simp add:is_aligned_shiftl_self)
7658    apply (simp add:range_cover_sz')
7659   apply simp
7660  apply (frule range_cover_cell_subset[where x = "of_nat p"])
7661   apply (rule word_of_nat_less)
7662   apply (simp add:range_cover.unat_of_nat_n)
7663  apply (simp add:shiftl_t2n field_simps)
7664  done
7665
7666lemma caps_no_overlap''_le:
7667  "\<lbrakk>caps_no_overlap'' ptr sz s;us \<le> sz;sz < word_bits\<rbrakk>
7668    \<Longrightarrow> caps_no_overlap'' ptr us s"
7669  apply (clarsimp simp:caps_no_overlap''_def)
7670  apply (drule(1) bspec)
7671  apply (subgoal_tac  "{ptr..(ptr && ~~ mask us) + 2 ^ us - 1}
7672                      \<subseteq>  {ptr..(ptr && ~~ mask sz) + 2 ^ sz - 1}")
7673   apply (erule impE)
7674    apply (rule ccontr)
7675    apply clarify
7676    apply (drule(1) disjoint_subset2[rotated -1])
7677    apply simp
7678   apply (erule subsetD)+
7679   apply simp
7680  apply clarsimp
7681  apply (frule neg_mask_diff_bound[where ptr = ptr])
7682  apply (simp add:p_assoc_help)
7683   apply (rule word_plus_mcs[where x = "2 ^ us - 1 + (ptr && ~~ mask sz)"])
7684    apply (simp add:field_simps)
7685   apply (simp add:field_simps)
7686   apply (simp add:p_assoc_help)
7687   apply (rule word_plus_mono_right)
7688   apply (simp add: word_bits_def)
7689   apply (erule two_power_increasing)
7690   apply simp
7691  apply (rule is_aligned_no_overflow')
7692   apply (simp add:is_aligned_neg_mask)
7693  done
7694
7695lemma caps_no_overlap''_le2:
7696  "\<lbrakk>caps_no_overlap'' ptr sz s;ptr \<le> ptr'; ptr' && ~~ mask sz = ptr && ~~ mask sz\<rbrakk>
7697    \<Longrightarrow> caps_no_overlap'' ptr' sz s"
7698  apply (clarsimp simp:caps_no_overlap''_def)
7699  apply (drule(1) bspec)
7700  apply (subgoal_tac  "{ptr'..(ptr' && ~~ mask sz) + 2 ^ sz - 1}
7701                      \<subseteq>  {ptr..(ptr && ~~ mask sz) + 2 ^ sz - 1}")
7702   apply (erule impE)
7703    apply (rule ccontr)
7704    apply clarify
7705    apply (drule(1) disjoint_subset2[rotated -1])
7706    apply simp
7707   apply (erule subsetD)+
7708   apply simp
7709  apply clarsimp
7710  done
7711
7712lemma range_cover_head_mask:
7713  "\<lbrakk>range_cover (ptr :: word32) sz us (Suc n); ptr \<noteq> 0\<rbrakk>
7714  \<Longrightarrow> ptr + (of_nat n << us) && ~~ mask sz = ptr && ~~ mask sz"
7715  apply (case_tac n)
7716   apply clarsimp
7717  apply (clarsimp simp:range_cover_tail_mask)
7718  done
7719
7720lemma pspace_no_overlap'_strg:
7721  "pspace_no_overlap' ptr sz s \<and> sz' \<le> sz \<and> sz < word_bits \<longrightarrow> pspace_no_overlap' ptr sz' s"
7722  apply clarsimp
7723  apply (erule(2) pspace_no_overlap'_le)
7724  done
7725
7726lemma cte_wp_at_no_0:
7727  "\<lbrakk>invs' s; cte_wp_at' (\<lambda>_. True) ptr s\<rbrakk> \<Longrightarrow> ptr \<noteq> 0"
7728  by (clarsimp dest!:invs_mdb' simp:valid_mdb'_def valid_mdb_ctes_def no_0_def cte_wp_at_ctes_of)
7729
7730lemma insertNewCap_descendants_range_in':
7731  "\<lbrace>\<lambda>s. valid_pspace' s \<and> descendants_range_in' S p (ctes_of s)
7732    \<and> capRange x \<inter> S = {}
7733    \<and> cte_wp_at' (\<lambda>cte. isUntypedCap (cteCap cte) \<and> sameRegionAs (cteCap cte) x) p s
7734    \<and> cte_wp_at' (\<lambda>cte. cteCap cte = capability.NullCap) dslot s
7735    \<and> descendants_range' x p (ctes_of s) \<and> capClass x = PhysicalClass
7736   \<rbrace> insertNewCap p dslot x
7737    \<lbrace>\<lambda>rv s. descendants_range_in' S p (ctes_of s)\<rbrace>"
7738  apply (clarsimp simp:insertNewCap_def descendants_range_in'_def)
7739  apply (wp getCTE_wp)
7740  apply (clarsimp simp:cte_wp_at_ctes_of)
7741  apply (intro conjI allI)
7742   apply (clarsimp simp:valid_pspace'_def valid_mdb'_def
7743     valid_mdb_ctes_def no_0_def split:if_splits)
7744  apply (clarsimp simp: descendants_of'_mdbPrev split:if_splits)
7745  apply (cut_tac p = p and m = "ctes_of s" and parent = p and s = s
7746        and parent_cap = "cteCap cte" and parent_node = "cteMDBNode cte"
7747        and site = dslot and site_cap = capability.NullCap and site_node = "cteMDBNode ctea"
7748        and c' = x
7749    in mdb_insert_again_child.descendants)
7750   apply (case_tac cte ,case_tac ctea)
7751   apply (rule mdb_insert_again_child.intro[OF mdb_insert_again.intro])
7752      apply (simp add:mdb_ptr_def vmdb_def valid_pspace'_def valid_mdb'_def
7753            mdb_ptr_axioms_def mdb_insert_again_axioms_def )+
7754    apply (intro conjI allI impI)
7755      apply clarsimp
7756      apply (erule(1) ctes_of_valid_cap')
7757     apply (clarsimp simp:valid_mdb_ctes_def)
7758    apply clarsimp
7759   apply (rule mdb_insert_again_child_axioms.intro)
7760   apply (clarsimp simp: nullPointer_def)+
7761   apply (clarsimp simp:isMDBParentOf_def valid_pspace'_def
7762      valid_mdb'_def valid_mdb_ctes_def)
7763   apply (frule(2) ut_revocableD'[rotated 1])
7764   apply (clarsimp simp:isCap_simps)
7765  apply (clarsimp cong: if_cong)
7766  done
7767
7768lemma insertNewCap_cte_wp_at_other:
7769  "\<lbrace>cte_wp_at' (\<lambda>cte. P (cteCap cte)) p and K (slot \<noteq> p)\<rbrace> insertNewCap srcSlot slot x
7770            \<lbrace>\<lambda>rv. cte_wp_at' (\<lambda>cte. P (cteCap cte)) p \<rbrace>"
7771  apply (clarsimp simp:insertNewCap_def)
7772  apply (wp updateMDB_weak_cte_wp_at setCTE_cte_wp_at_other getCTE_wp)
7773  apply (clarsimp simp:cte_wp_at_ctes_of)
7774  done
7775
7776
7777lemma range_cover_bound3:
7778  "\<lbrakk>range_cover ptr sz us n; x < of_nat n\<rbrakk>
7779  \<Longrightarrow> ptr + x * 2 ^ us + 2 ^ us - 1 \<le> ptr + (of_nat n) * 2 ^ us - 1"
7780  apply (frule range_cover_subset[where p = "unat x"])
7781    apply (simp add:unat_less_helper)
7782   apply (rule ccontr,simp)
7783  apply (drule(1) range_cover_subset_not_empty)
7784   apply (clarsimp simp: field_simps)
7785  done
7786
7787lemma range_cover_gsMaxObjectSize:
7788  "cte_wp_at' (\<lambda>cte. cteCap cte = UntypedCap dev (ptr &&~~ mask sz) sz idx) srcSlot s
7789    \<Longrightarrow> range_cover ptr sz (APIType_capBits newType userSize) (length destSlots)
7790    \<Longrightarrow> valid_global_refs' s
7791    \<Longrightarrow> unat num = length destSlots
7792    \<Longrightarrow> unat (num << (APIType_capBits newType userSize) :: word32) \<le> gsMaxObjectSize s
7793        \<and> 2 ^ APIType_capBits newType userSize \<le> gsMaxObjectSize s"
7794  apply (clarsimp simp: cte_wp_at_ctes_of)
7795  apply (drule (1) valid_global_refsD_with_objSize)
7796  apply clarsimp
7797  apply (rule conjI)
7798   apply (frule range_cover.range_cover_compare_bound)
7799   apply (drule range_cover.unat_of_nat_n_shift, rule order_refl)
7800   apply (drule_tac s="unat num" in sym)
7801   apply simp
7802  apply (clarsimp simp: range_cover_def)
7803  apply (erule order_trans[rotated])
7804  apply simp
7805  done
7806
7807lemma APIType_capBits_min:
7808  "(tp = APIObjectType apiobject_type.Untyped \<longrightarrow> minUntypedSizeBits \<le> userSize)
7809    \<Longrightarrow> 4 \<le> APIType_capBits tp userSize"
7810  by (simp add: APIType_capBits_def objBits_simps' machine_bits_defs untypedBits_defs
7811         split: object_type.split ArchTypes_H.apiobject_type.split)
7812
7813end
7814
7815context begin interpretation Arch . (*FIXME: arch_split*)
7816
7817crunch gsCNodes[wp]: insertNewCap, Arch_createNewCaps, threadSet,
7818        "Arch.createObject" "\<lambda>s. P (gsCNodes s)"
7819  (wp: crunch_wps setObject_ksPSpace_only
7820     simp: unless_def updateObject_default_def crunch_simps
7821   ignore: getObject setObject)
7822
7823lemma createNewCaps_1_gsCNodes_p:
7824  "\<lbrace>\<lambda>s. P (gsCNodes s p) \<and> p \<noteq> ptr\<rbrace> createNewCaps newType ptr 1 n dev\<lbrace>\<lambda>rv s. P (gsCNodes s p)\<rbrace>"
7825  apply (simp add: createNewCaps_def)
7826  apply (rule hoare_pre)
7827   apply (wp mapM_x_wp' | wpc | simp add: createObjects_def)+
7828  done
7829
7830lemma createObject_gsCNodes_p:
7831  "\<lbrace>\<lambda>s. P (gsCNodes s p) \<and> p \<noteq> ptr\<rbrace> createObject t ptr sz dev\<lbrace>\<lambda>rv s. P (gsCNodes s p)\<rbrace>"
7832  apply (simp add: createObject_def)
7833  apply (rule hoare_pre)
7834   apply (wp mapM_x_wp' | wpc | simp add: createObjects_def)+
7835  done
7836
7837lemma createObject_cnodes_have_size:
7838  "\<lbrace>\<lambda>s. is_aligned ptr (APIType_capBits newType userSize)
7839      \<and> cnodes_retype_have_size R (APIType_capBits newType userSize) (gsCNodes s)\<rbrace>
7840    createObject newType ptr userSize dev
7841  \<lbrace>\<lambda>rv s. cnodes_retype_have_size R (APIType_capBits newType userSize) (gsCNodes s)\<rbrace>"
7842  apply (simp add: createObject_def)
7843  apply (rule hoare_pre)
7844   apply (wp mapM_x_wp' | wpc | simp add: createObjects_def)+
7845  apply (cases newType, simp_all add: ARM_HYP_H.toAPIType_def)
7846  apply (clarsimp simp: APIType_capBits_def objBits_simps'
7847                              cnodes_retype_have_size_def cte_level_bits_def
7848                       split: if_split_asm)
7849  done
7850
7851lemma range_cover_not_in_neqD:
7852  "\<lbrakk> x \<notin> {ptr..ptr + (of_nat n << APIType_capBits newType userSize) - 1};
7853    range_cover ptr sz (APIType_capBits newType userSize) n; n' < n \<rbrakk>
7854  \<Longrightarrow> x \<noteq> ptr + (of_nat n' << APIType_capBits newType userSize)"
7855  apply (clarsimp simp only: shiftl_t2n mult.commute)
7856  apply (erule notE, rule subsetD, erule_tac p=n' in range_cover_subset)
7857    apply simp+
7858  apply (rule is_aligned_no_overflow)
7859  apply (rule aligned_add_aligned)
7860    apply (erule range_cover.aligned)
7861   apply (simp add: is_aligned_mult_triv2)
7862  apply simp
7863  done
7864
7865crunch gsMaxObjectSize[wp]: createObject "\<lambda>s. P (gsMaxObjectSize s)"
7866  (simp: crunch_simps unless_def wp: crunch_wps)
7867
7868end
7869
7870context kernel_m begin
7871
7872lemma nat_le_induct [consumes 1, case_names base step]:
7873  assumes le: "i \<le> (k::nat)" and
7874        base: "P(k)" and
7875        step: "\<And>i. \<lbrakk>i \<le> k; P i; 0 < i\<rbrakk> \<Longrightarrow> P(i - 1)"
7876  shows "P i"
7877proof -
7878  obtain j where jk: "j \<le> k" and j_eq: "i = k - j"
7879    using le
7880    apply (drule_tac x="k - i" in meta_spec)
7881    apply simp
7882    done
7883
7884  have "j \<le> k \<Longrightarrow> P (k - j)"
7885    apply (induct j)
7886     apply (simp add: base)
7887    apply simp
7888    apply (drule step[rotated], simp+)
7889    done
7890
7891  thus "P i" using jk j_eq
7892    by simp
7893qed
7894
7895lemma ceqv_restore_as_guard:
7896  "ceqv Gamma xf' rv' t t' d (Guard C_Guard {s. xf' s = rv'} d)"
7897  apply (simp add: ceqv_def)
7898  apply (auto elim!: exec_Normal_elim_cases intro: exec.Guard)
7899  done
7900
7901lemma insertNewCap_preserves_bytes:
7902  "\<forall>s. \<Gamma>\<turnstile>\<^bsub>/UNIV\<^esub> {s} Call insertNewCap_'proc
7903      {t. hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s))
7904         \<and> byte_regions_unmodified' s t}"
7905  apply (hoare_rule HoarePartial.ProcNoRec1)
7906  apply (rule allI, rule conseqPre, vcg exspec=mdb_node_ptr_set_mdbPrev_preserves_bytes
7907    exspec=mdb_node_ptr_set_mdbNext_preserves_bytes
7908    exspec=mdb_node_get_mdbNext_modifies exspec=mdb_node_new_modifies)
7909  apply (safe intro!: byte_regions_unmodified_hrs_mem_update
7910    elim!: byte_regions_unmodified_trans byte_regions_unmodified_trans[rotated],
7911    simp_all add: h_t_valid_field)
7912  done
7913
7914lemma byte_regions_unmodified_flip_eq:
7915  "byte_regions_unmodified hrs' hrs
7916    \<Longrightarrow> hrs_htd hrs' = hrs_htd hrs
7917    \<Longrightarrow> byte_regions_unmodified hrs hrs'"
7918  by (simp add: byte_regions_unmodified_def)
7919
7920lemma insertNewCap_preserves_bytes_flip:
7921  "\<forall>s. \<Gamma>\<turnstile>\<^bsub>/UNIV\<^esub> {s} Call insertNewCap_'proc
7922      {t. hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s))
7923         \<and> byte_regions_unmodified' t s}"
7924  by (rule allI, rule conseqPost,
7925    rule insertNewCap_preserves_bytes[rule_format],
7926    auto elim: byte_regions_unmodified_flip_eq)
7927
7928lemma copyGlobalMappings_preserves_bytes:
7929  "\<forall>s. \<Gamma>\<turnstile>\<^bsub>/UNIV\<^esub> {s} Call copyGlobalMappings_'proc
7930      {t. hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s))
7931         \<and> byte_regions_unmodified' s t}"
7932  apply (hoare_rule HoarePartial.ProcNoRec1)
7933  apply vcg
7934  apply clarsimp
7935  done
7936
7937lemma cleanByVA_PoU_preserves_bytes:
7938  "\<forall>s. \<Gamma>\<turnstile>\<^bsub>/UNIV\<^esub> {s} Call cleanByVA_PoU_'proc
7939      {t. hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s))
7940         \<and> byte_regions_unmodified' s t}"
7941  apply (rule allI, rule conseqPost,
7942    rule cleanByVA_PoU_preserves_kernel_bytes[rule_format])
7943   apply simp_all
7944  apply (clarsimp simp: byte_regions_unmodified_def)
7945  done
7946
7947lemma cleanCacheRange_PoU_preserves_bytes:
7948  "\<forall>s. \<Gamma>\<turnstile>\<^bsub>/UNIV\<^esub> {s} Call cleanCacheRange_PoU_'proc
7949      {t. hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s))
7950         \<and> byte_regions_unmodified' s t}"
7951  apply (hoare_rule HoarePartial.ProcNoRec1)
7952  apply (clarsimp simp only: whileAnno_def)
7953  apply (subst whileAnno_def[symmetric, where V=undefined
7954       and I="{t. hrs_htd (t_hrs_' (globals t)) = hrs_htd (t_hrs_' (globals s))
7955         \<and> byte_regions_unmodified' s t}" for s])
7956  apply (rule conseqPre, vcg exspec=cleanByVA_PoU_preserves_bytes)
7957  apply (safe intro!: byte_regions_unmodified_hrs_mem_update
7958    elim!: byte_regions_unmodified_trans byte_regions_unmodified_trans[rotated],
7959    (simp_all add: h_t_valid_field)+)
7960  done
7961
7962lemma hrs_htd_update_canon:
7963  "hrs_htd_update (\<lambda>_. f (hrs_htd hrs)) hrs = hrs_htd_update f hrs"
7964  by (cases hrs, simp add: hrs_htd_update_def hrs_htd_def)
7965
7966lemma Arch_createObject_preserves_bytes:
7967  "\<forall>s. \<Gamma>\<turnstile>\<^bsub>/UNIV\<^esub> {s} Call Arch_createObject_'proc
7968      {t. \<forall>nt. t_' s = object_type_from_H nt
7969         \<longrightarrow> (\<forall>x \<in> - {ptr_val (regionBase_' s) ..+ 2 ^ getObjectSize nt (unat (userSize_' s))}.
7970             hrs_htd (t_hrs_' (globals t)) x = hrs_htd (t_hrs_' (globals s)) x)
7971         \<and> byte_regions_unmodified' t s}"
7972  apply (hoare_rule HoarePartial.ProcNoRec1)
7973  apply clarsimp
7974  apply (rule conseqPre, vcg exspec=cap_small_frame_cap_new_modifies
7975    exspec=cap_frame_cap_new_modifies
7976    exspec=cap_page_table_cap_new_modifies
7977    exspec=copyGlobalMappings_preserves_bytes
7978    exspec=addrFromPPtr_modifies
7979    exspec=cleanCacheRange_PoU_preserves_bytes
7980    exspec=cap_page_directory_cap_new_modifies
7981    exspec=cap_vcpu_cap_new_modifies)
7982  apply (safe intro!: byte_regions_unmodified_hrs_mem_update,
7983         (simp_all add: h_t_valid_field hrs_htd_update)+)
7984             apply (safe intro!: ptr_retyp_d ptr_retyps_out)
7985             apply (simp_all add: object_type_from_H_def Kernel_C_defs APIType_capBits_def
7986                           split: object_type.split_asm ArchTypes_H.apiobject_type.split_asm)
7987   apply (rule byte_regions_unmodified_flip, simp)
7988   apply (rule byte_regions_unmodified_trans[rotated],
7989          assumption, simp_all add: hrs_htd_update_canon hrs_htd_update)
7990  apply (drule intvlD)
7991  apply clarsimp
7992  apply (erule notE, rule intvlI)
7993  apply (simp add: vcpu_bits_def pageBits_def)
7994  done
7995
7996lemma ptr_arr_retyps_eq_outside_dom:
7997  "x \<notin> {ptr_val (p :: 'a ptr) ..+ n * size_of TYPE ('a :: wf_type)}
7998    \<Longrightarrow> ptr_arr_retyps n p htd x = htd x"
7999  by (simp add: ptr_arr_retyps_def htd_update_list_same2)
8000
8001lemma createObject_preserves_bytes:
8002  "\<forall>s. \<Gamma>\<turnstile>\<^bsub>/UNIV\<^esub> {s} Call createObject_'proc
8003      {t. \<forall>nt. t_' s = object_type_from_H nt
8004         \<longrightarrow> (\<forall>x \<in> - {ptr_val (regionBase_' s) ..+ 2 ^ getObjectSize nt (unat (userSize_' s))}.
8005             hrs_htd (t_hrs_' (globals t)) x = hrs_htd (t_hrs_' (globals s)) x)
8006         \<and> byte_regions_unmodified' t s}"
8007  apply (hoare_rule HoarePartial.ProcNoRec1)
8008  apply clarsimp
8009  apply (rule conseqPre, vcg exspec=Arch_createObject_preserves_bytes
8010    exspec=cap_thread_cap_new_modifies
8011    exspec=cap_endpoint_cap_new_modifies
8012    exspec=cap_notification_cap_new_modifies
8013    exspec=cap_cnode_cap_new_modifies
8014    exspec=cap_untyped_cap_new_modifies)
8015  apply (safe intro!: byte_regions_unmodified_hrs_mem_update,
8016    (simp_all add: h_t_valid_field hrs_htd_update)+)
8017  apply (safe intro!: ptr_retyp_d ptr_retyps_out trans[OF ptr_retyp_d ptr_retyp_d]
8018                      ptr_arr_retyps_eq_outside_dom)
8019  apply (simp_all add: object_type_from_H_def Kernel_C_defs APIType_capBits_def
8020                       objBits_simps' cte_C_size power_add ctcb_offset_defs
8021    split: object_type.split_asm ArchTypes_H.apiobject_type.split_asm)
8022   apply (erule notE, erule subsetD[rotated],
8023     rule intvl_start_le intvl_sub_offset, simp)+
8024  done
8025
8026
8027lemma offset_intvl_first_chunk_subsets:
8028  "range_cover (p :: addr) sz bits n
8029    \<Longrightarrow> i < of_nat n
8030    \<Longrightarrow> {p + (i << bits) ..+ 2 ^ bits} \<subseteq> {p + (i << bits) ..+ (n - unat i) * 2 ^ bits}
8031        \<and> {p + ((i + 1) << bits) ..+ (n - unat (i + 1)) * 2 ^ bits}
8032            \<le> {p + (i << bits) ..+ (n - unat i) * 2 ^ bits}
8033        \<and> {p + (i << bits) ..+ 2 ^ bits}
8034            \<inter> {p + ((i + 1) << bits) ..+ (n - unat (i + 1)) * 2 ^ bits}
8035            = {}"
8036  apply (strengthen intvl_start_le)
8037  apply (strengthen order_trans[OF _
8038      intvl_sub_offset[where x="2 ^ bits" and y="(n - unat (i + 1)) * 2 ^ bits"]])
8039  apply (frule range_cover_sz')
8040  apply (cut_tac n=i in unatSuc)
8041   apply unat_arith
8042  apply (simp add: word_shiftl_add_distrib field_simps TWO)
8043  apply (simp add: mult_Suc[symmetric] del: mult_Suc)
8044  apply (frule unat_less_helper)
8045  apply (cut_tac p="p + (i << bits)" and k="2 ^ bits"
8046    and z="(n - unat (i + 1)) * 2 ^ bits" in init_intvl_disj)
8047   apply (simp add: field_simps)
8048   apply (drule range_cover.strong_times_32, simp)
8049   apply (simp add: addr_card_def word_bits_def card_word)
8050   apply (erule order_le_less_trans[rotated])
8051   apply (simp add: mult_Suc[symmetric] del: mult_Suc)
8052  apply (simp add: Int_commute field_simps)
8053  apply unat_arith
8054  done
8055
8056lemma offset_intvl_first_chunk_subsets_unat:
8057  "range_cover (p :: addr) sz bits n
8058    \<Longrightarrow> unat n' = n
8059    \<Longrightarrow> i < of_nat n
8060    \<Longrightarrow> {p + (i << bits) ..+ 2 ^ bits} \<subseteq> {p + (i << bits) ..+ unat (n' - i) * 2 ^ bits}
8061        \<and> {p + ((i + 1) << bits) ..+ unat (n' - (i + 1)) * 2 ^ bits}
8062            \<le> {p + (i << bits) ..+ unat (n' - i) * 2 ^ bits}
8063        \<and> {p + (i << bits) ..+ 2 ^ bits}
8064            \<inter> {p + ((i + 1) << bits) ..+ unat (n' - (i + 1)) * 2 ^ bits}
8065            = {}"
8066  apply (subgoal_tac "unat (n' - (i + 1)) = unat n' - unat (i + 1)
8067        \<and> unat (n' - i) = unat n' - unat i")
8068   apply (frule(1) offset_intvl_first_chunk_subsets)
8069   apply simp
8070  apply (intro conjI unat_sub)
8071   apply (rule minus_one_helper2, simp)
8072   apply (simp add: word_less_nat_alt unat_of_nat)
8073  apply (simp add: word_le_nat_alt word_less_nat_alt unat_of_nat)
8074  done
8075
8076lemma retype_offs_region_actually_is_zero_bytes:
8077  "\<lbrakk> ctes_of s p = Some cte; (s, s') \<in> rf_sr; untyped_ranges_zero' s;
8078      cteCap cte = UntypedCap False (ptr &&~~ mask sz) sz idx;
8079      idx \<le> unat (ptr && mask sz);
8080      range_cover ptr sz (getObjectSize newType userSize) num_ret \<rbrakk>
8081    \<Longrightarrow> region_actually_is_zero_bytes ptr
8082            (num_ret * 2 ^ APIType_capBits newType userSize) s'"
8083  using word_unat_mask_lt[where w=ptr and m=sz]
8084  apply -
8085  apply (frule range_cover.sz(1))
8086  apply (drule(2) ctes_of_untyped_zero_rf_sr)
8087   apply (simp add: untypedZeroRange_def max_free_index_def word_size)
8088  apply clarify
8089  apply (strengthen heap_list_is_zero_mono2[mk_strg I E]
8090      region_actually_is_bytes_subset[mk_strg I E])
8091  apply (simp add: getFreeRef_def word_size)
8092  apply (rule intvl_both_le)
8093   apply (rule order_trans, rule word_plus_mono_right, erule word_of_nat_le)
8094    apply (simp add: word_plus_and_or_coroll2 add.commute word_and_le2)
8095   apply (simp add: word_plus_and_or_coroll2 add.commute)
8096  apply (subst unat_plus_simple[THEN iffD1], rule is_aligned_no_wrap',
8097    rule is_aligned_neg_mask2)
8098   apply (rule word_of_nat_less, simp)
8099  apply (simp add: unat_of_nat_eq[OF order_less_trans, OF _ power_strict_increasing[where n=sz]]
8100    unat_sub[OF word_of_nat_le])
8101  apply (subst word_plus_and_or_coroll2[where x=ptr and w="mask sz", symmetric])
8102  apply (subst unat_plus_simple[THEN iffD1],
8103    simp add: word_plus_and_or_coroll2 add.commute word_and_le2)
8104  apply simp
8105  apply (rule order_trans[rotated], erule range_cover.range_cover_compare_bound)
8106  apply simp
8107  done
8108
8109lemma createNewCaps_valid_cap_hd:
8110    "\<lbrace>\<lambda>s. pspace_no_overlap' ptr sz s \<and>
8111        valid_pspace' s \<and> n \<noteq> 0 \<and>
8112        range_cover ptr sz (APIType_capBits ty us) n \<and>
8113        (ty = APIObjectType ArchTypes_H.CapTableObject \<longrightarrow> 0 < us) \<and>
8114        (ty = APIObjectType ArchTypes_H.apiobject_type.Untyped \<longrightarrow> minUntypedSizeBits \<le> us \<and> us \<le> maxUntypedSizeBits) \<and>
8115       ptr \<noteq> 0 \<rbrace>
8116    createNewCaps ty ptr n us dev
8117  \<lbrace>\<lambda>r s. s \<turnstile>' hd r\<rbrace>"
8118  apply (cases "n = 0")
8119   apply simp
8120  apply (rule hoare_chain)
8121    apply (rule hoare_vcg_conj_lift)
8122     apply (rule createNewCaps_ret_len)
8123    apply (rule createNewCaps_valid_cap'[where sz=sz])
8124   apply (clarsimp simp: range_cover_n_wb)
8125  apply simp
8126  done
8127
8128lemma insertNewCap_ccorres:
8129  "ccorres dc xfdc (pspace_aligned' and valid_mdb' and cte_wp_at' (\<lambda>_. True) slot
8130          and valid_objs' and valid_cap' cap)
8131     ({s. cap_get_tag (cap_' s) = scast cap_untyped_cap
8132         \<longrightarrow> (case untypedZeroRange (cap_to_H (the (cap_lift (cap_' s)))) of None \<Rightarrow> True
8133          | Some (a, b) \<Rightarrow> region_actually_is_zero_bytes a (unat ((b + 1) - a)) s)}
8134       \<inter> {s. ccap_relation cap (cap_' s)} \<inter> {s. parent_' s = Ptr parent}
8135       \<inter> {s. slot_' s = Ptr slot}) []
8136     (insertNewCap parent slot cap)
8137     (Call insertNewCap_'proc)"
8138  (is "ccorres _ _ ?P ?P' _ _ _")
8139  apply (rule ccorres_guard_imp2, rule insertNewCap_ccorres1)
8140  apply (clarsimp simp: cap_get_tag_isCap)
8141  apply (clarsimp simp: ccap_relation_def map_option_Some_eq2)
8142  apply (simp add: untypedZeroRange_def Let_def)
8143  done
8144
8145lemma createObject_untyped_region_is_zero_bytes:
8146  "\<forall>\<sigma>. \<Gamma>\<turnstile>\<^bsub>/UNIV\<^esub> {s. let tp = (object_type_to_H (t_' s));
8147          sz = APIType_capBits tp (unat (userSize_' s))
8148      in (\<not> to_bool (deviceMemory_' s)
8149              \<longrightarrow> region_actually_is_zero_bytes (ptr_val (regionBase_' s)) (2 ^ sz) s)
8150          \<and> is_aligned (ptr_val (regionBase_' s)) sz
8151          \<and> sz < 32 \<and> (tp = APIObjectType ArchTypes_H.apiobject_type.Untyped \<longrightarrow> sz \<ge> minUntypedSizeBits)}
8152      Call createObject_'proc
8153   {t. cap_get_tag (ret__struct_cap_C_' t) = scast cap_untyped_cap
8154         \<longrightarrow> (case untypedZeroRange (cap_to_H (the (cap_lift (ret__struct_cap_C_' t)))) of None \<Rightarrow> True
8155          | Some (a, b) \<Rightarrow> region_actually_is_zero_bytes a (unat ((b + 1) - a)) t)}"
8156  apply (rule allI, rule conseqPre, vcg exspec=copyGlobalMappings_modifies
8157      exspec=Arch_initContext_modifies
8158      exspec=cleanCacheRange_PoU_modifies)
8159  apply (clarsimp simp: cap_tag_defs Let_def)
8160  apply (simp add: cap_lift_untyped_cap cap_tag_defs cap_to_H_simps
8161                   cap_untyped_cap_lift_def object_type_from_H_def)
8162  apply (simp add: untypedZeroRange_def split: if_split)
8163  apply (clarsimp simp: getFreeRef_def Let_def object_type_to_H_def untypedBits_defs)
8164  apply (simp add:  APIType_capBits_def
8165                   less_mask_eq word_less_nat_alt)
8166  done
8167
8168lemma createNewObjects_ccorres:
8169notes blah[simp del] =  atLeastAtMost_iff atLeastatMost_subset_iff atLeastLessThan_iff
8170      Int_atLeastAtMost atLeastatMost_empty_iff split_paired_Ex
8171and   hoare_TrueI[simp add]
8172defines "unat_eq a b \<equiv> unat a = b"
8173shows  "ccorres dc xfdc
8174     (invs' and sch_act_simple and ct_active'
8175                  and (cte_wp_at' (\<lambda>cte. cteCap cte = UntypedCap isdev (ptr &&~~ mask sz) sz idx ) srcSlot)
8176                  and (\<lambda>s. \<forall>slot\<in>set destSlots. cte_wp_at' (\<lambda>c. cteCap c = NullCap) slot s)
8177                  and (\<lambda>s. \<forall>slot\<in>set destSlots. ex_cte_cap_wp_to' (\<lambda>_. True) slot s)
8178                  and (\<lambda>s. \<exists>n. gsCNodes s cnodeptr = Some n \<and> unat start + length destSlots \<le> 2 ^ n)
8179                  and (pspace_no_overlap' ptr sz)
8180                  and caps_no_overlap'' ptr sz
8181                  and caps_overlap_reserved' {ptr .. ptr + of_nat (length destSlots) * 2^ (getObjectSize newType userSize) - 1}
8182                  and (\<lambda>s. descendants_range_in' {ptr..(ptr && ~~ mask sz) + 2 ^ sz - 1} srcSlot (ctes_of s))
8183                  and cnodes_retype_have_size {ptr .. ptr + of_nat (length destSlots) * 2^ (getObjectSize newType userSize) - 1}
8184                      (APIType_capBits newType userSize) o gsCNodes
8185                  and (K (srcSlot \<notin> set destSlots
8186                    \<and> destSlots \<noteq> []
8187                    \<and> range_cover ptr sz (getObjectSize newType userSize) (length destSlots )
8188                    \<and> ptr \<noteq> 0
8189                    \<and> {ptr .. ptr + of_nat (length destSlots) * 2^ (getObjectSize newType userSize) - 1}
8190                      \<inter> kernel_data_refs = {}
8191                    \<and> cnodeptr \<notin> {ptr .. ptr + (of_nat (length destSlots)<< APIType_capBits newType userSize) - 1}
8192                    \<and> 0 \<notin> {ptr..(ptr && ~~ mask sz) + 2 ^ sz - 1}
8193                    \<and> is_aligned ptr 4
8194                    \<and> (newType = APIObjectType apiobject_type.Untyped \<longrightarrow> userSize \<le> maxUntypedSizeBits)
8195                    \<and> (newType = APIObjectType apiobject_type.CapTableObject \<longrightarrow> userSize < 28)
8196                    \<and> (newType = APIObjectType apiobject_type.Untyped \<longrightarrow> minUntypedSizeBits \<le> userSize)
8197                    \<and> (newType = APIObjectType apiobject_type.CapTableObject \<longrightarrow> 0 < userSize)
8198                    \<and> (isdev \<longrightarrow> newType = APIObjectType ArchTypes_H.apiobject_type.Untyped \<or>
8199                                           isFrameType newType)
8200                    \<and> (unat num = length destSlots)
8201                    )))
8202    ({s. (\<not> isdev \<longrightarrow> region_actually_is_zero_bytes ptr
8203            (length destSlots * 2 ^ APIType_capBits newType userSize) s)}
8204           \<inter> {s. t_' s = object_type_from_H newType}
8205           \<inter> {s. parent_' s = cte_Ptr srcSlot}
8206           \<inter> {s. slots_' s = slot_range_C (cte_Ptr cnodeptr) start num
8207                     \<and> unat num \<noteq> 0
8208                     \<and> (\<forall>n. n < length destSlots \<longrightarrow> destSlots ! n = cnodeptr + ((start + of_nat n) * 0x10))
8209                     }
8210           \<inter> {s. regionBase_' s = Ptr ptr }
8211           \<inter> {s. unat_eq (userSize_' s) userSize}
8212           \<inter> {s. to_bool (deviceMemory_' s) = isdev}
8213     ) []
8214     (createNewObjects newType srcSlot destSlots ptr userSize isdev)
8215     (Call createNewObjects_'proc)"
8216  apply (rule ccorres_gen_asm_state)
8217  apply clarsimp
8218  apply (subgoal_tac "unat (of_nat (getObjectSize newType userSize)) = getObjectSize newType userSize")
8219   prefer 2
8220   apply (subst unat_of_nat32)
8221    apply (rule less_le_trans [OF getObjectSize_max_size], auto simp: word_bits_def untypedBits_defs)[1]
8222   apply simp
8223  apply (cinit lift: t_' parent_' slots_' regionBase_' userSize_' deviceMemory_')
8224   apply (rule ccorres_rhs_assoc2)+
8225   apply (rule ccorres_rhs_assoc)
8226   apply (rule_tac Q' = "Q'
8227     \<inter> {s. objectSize_' s = of_nat (APIType_capBits newType userSize)}
8228     \<inter> {s. nextFreeArea_' s = Ptr ptr } "
8229     and R="(\<lambda>s. unat (num << (APIType_capBits newType userSize) :: word32)
8230        \<le> gsMaxObjectSize s) and R''"
8231     for Q' R'' in ccorres_symb_exec_r)
8232     apply (rule ccorres_guard_imp[where A="X and Q"
8233         and A'=Q' and Q=Q and Q'=Q' for X Q Q', rotated]
8234         (* this moves the gsMaxObjectSize bit into the ccorres_symb_exec_r
8235            vcg proof *))
8236       apply clarsimp
8237      apply clarsimp
8238     apply (cinitlift objectSize_' nextFreeArea_')
8239     apply simp
8240     apply (clarsimp simp: whileAnno_def)
8241     apply (rule ccorres_rel_imp)
8242      apply (rule_tac Q="{s. \<not> isdev \<longrightarrow> region_actually_is_zero_bytes
8243            (ptr + (i_' s << APIType_capBits newType userSize))
8244            (unat (num - i_' s) * 2 ^ APIType_capBits newType userSize) s}"
8245            in ccorres_zipWithM_x_while_genQ[where j=1, OF _ _ _ _ _ i_xf_for_sequence, simplified])
8246          apply clarsimp
8247          apply (subst upt_enum_offset_trivial)
8248            apply (rule minus_one_helper)
8249             apply (rule word_of_nat_le)
8250             apply (drule range_cover.range_cover_n_less)
8251             apply (simp add:word_bits_def minus_one_norm)
8252            apply (erule range_cover_not_zero[rotated],simp)
8253           apply simp
8254          apply (rule ccorres_guard_impR)
8255           apply (rule_tac xf'=i_' in ccorres_abstract, ceqv)
8256           apply (rule_tac P="rv' = of_nat n" in ccorres_gen_asm2, simp)
8257           apply (rule ccorres_rhs_assoc)+
8258           apply (rule ccorres_add_return)
8259           apply (simp only: dc_def[symmetric] hrs_htd_update)
8260           apply ((rule ccorres_Guard_Seq[where S=UNIV])+)?
8261           apply (rule ccorres_split_nothrow,
8262                rule_tac S="{ptr .. ptr + of_nat (length destSlots) * 2^ (getObjectSize newType userSize) - 1}"
8263                  in ccorres_typ_region_bytes_dummy, ceqv)
8264             apply (rule ccorres_Guard_Seq)+
8265             apply (ctac add:createObject_ccorres)
8266               apply (rule ccorres_move_array_assertion_cnode_ctes
8267                           ccorres_move_c_guard_cte)+
8268               apply (rule ccorres_add_return2)
8269               apply (ctac (no_vcg) add: insertNewCap_ccorres)
8270                apply (rule ccorres_move_array_assertion_cnode_ctes
8271                            ccorres_return_Skip')+
8272               apply wp
8273              apply (clarsimp simp:createObject_def3 conj_ac)
8274              apply (wp createNewCaps_valid_pspace_extras[where sz = sz]
8275                createNewCaps_cte_wp_at[where sz = sz]
8276                createNewCaps_valid_cap_hd[where sz = sz])
8277                apply (rule range_cover_one)
8278                  apply (rule aligned_add_aligned[OF is_aligned_shiftl_self])
8279                   apply (simp add:range_cover.aligned)
8280                  apply (simp add:range_cover_def)
8281                 apply (simp add:range_cover_def)
8282                apply (simp add:range_cover_def)
8283               apply (simp add:range_cover.sz)
8284              apply (wp createNewCaps_1_gsCNodes_p[simplified]
8285                        createNewCaps_cte_wp_at'[where sz=sz])[1]
8286             apply clarsimp
8287             apply (vcg exspec=createObject_untyped_region_is_zero_bytes)
8288            apply (simp add:size_of_def)
8289            apply (rule_tac P = "\<lambda>s. cte_wp_at' (\<lambda>cte. isUntypedCap (cteCap cte) \<and>
8290              {ptr .. ptr + (of_nat (length destSlots)<< APIType_capBits newType userSize) - 1} \<subseteq> untypedRange (cteCap cte)) srcSlot s
8291              \<and> pspace_no_overlap'  ((of_nat n << APIType_capBits newType userSize) + ptr) sz s
8292              \<and> caps_no_overlap'' ((of_nat n << APIType_capBits newType userSize) + ptr) sz s
8293              \<and> caps_overlap_reserved'  {(of_nat n << APIType_capBits newType userSize) +
8294                 ptr.. ptr + of_nat (length destSlots) * 2^ (getObjectSize newType userSize) - 1 } s
8295              \<and> kernel_data_refs \<inter> {ptr .. ptr + (of_nat (length destSlots) << APIType_capBits newType userSize) - 1} = {}
8296              \<and> (\<forall>n < length destSlots. cte_at' (cnodeptr + (start * 0x10 + of_nat n * 0x10)) s
8297                    \<and> ex_cte_cap_wp_to' (\<lambda>_. True) (cnodeptr + (start * 0x10 + of_nat n * 0x10)) s)
8298              \<and> invs' s
8299              \<and> 2 ^ APIType_capBits newType userSize \<le> gsMaxObjectSize s
8300              \<and> (\<exists>cn. gsCNodes s cnodeptr = Some cn \<and> unat start + length destSlots \<le> 2 ^ cn)
8301              \<and> cnodeptr \<notin> {ptr .. ptr + (of_nat (length destSlots)<< APIType_capBits newType userSize) - 1}
8302              \<and> (\<forall>k < length destSlots - n.
8303                 cte_wp_at' (\<lambda>c. cteCap c = NullCap)
8304                 (cnodeptr + (of_nat k * 0x10 + start * 0x10 + of_nat n * 0x10)) s)
8305              \<and> descendants_range_in' {(of_nat n << APIType_capBits newType userSize) +
8306                 ptr.. (ptr && ~~ mask sz) + 2 ^ sz  - 1} srcSlot (ctes_of s)"
8307              in hoare_pre(1))
8308             apply wp
8309            apply (clarsimp simp:createObject_hs_preconds_def field_simps conj_comms
8310                   invs_valid_pspace' invs_pspace_distinct' invs_pspace_aligned'
8311                   invs_ksCurDomain_maxDomain')
8312            apply (subst intvl_range_conv)
8313              apply (rule aligned_add_aligned[OF range_cover.aligned],assumption)
8314               subgoal by (simp add:is_aligned_shiftl_self)
8315              apply (fold_subgoals (prefix))[2]
8316              subgoal premises prems using prems
8317                        by (simp_all add:range_cover_sz'[where 'a=32, folded word_bits_def]
8318                                   word_bits_def range_cover_def)+
8319            apply (simp add: range_cover_not_in_neqD)
8320            apply (intro conjI)
8321                  apply (drule_tac p = n in range_cover_no_0)
8322                    apply (simp add:shiftl_t2n field_simps)+
8323                 apply (cut_tac x=num in unat_lt2p, simp)
8324                 apply (simp add: unat_arith_simps unat_of_nat, simp split: if_split)
8325                 apply (intro impI, erule order_trans[rotated], simp)
8326                apply (erule pspace_no_overlap'_le)
8327                 apply (fold_subgoals (prefix))[2]
8328                 subgoal premises prems using prems
8329                           by (simp add:range_cover.sz[where 'a=32, folded word_bits_def])+
8330               apply (rule range_cover_one)
8331                 apply (rule aligned_add_aligned[OF range_cover.aligned],assumption)
8332                  apply (simp add:is_aligned_shiftl_self)
8333                 apply (fold_subgoals (prefix))[2]
8334                 subgoal premises prems using prems
8335                           by (simp add: range_cover_sz'[where 'a=32, folded word_bits_def]
8336                                         range_cover.sz[where 'a=32, folded word_bits_def])+
8337               apply (simp add:  word_bits_def range_cover_def)
8338              apply (rule range_cover_full)
8339               apply (rule aligned_add_aligned[OF range_cover.aligned],assumption)
8340                apply (simp add:is_aligned_shiftl_self)
8341               apply (fold_subgoals (prefix))[2]
8342               subgoal premises prems using prems
8343                         by (simp add: range_cover_sz'[where 'a=32, folded word_bits_def]
8344                                       range_cover.sz[where 'a=32, folded word_bits_def])+
8345              apply (erule caps_overlap_reserved'_subseteq)
8346              apply (frule_tac x = "of_nat n" in range_cover_bound3)
8347               apply (rule word_of_nat_less)
8348               apply (simp add:range_cover.unat_of_nat_n)
8349              apply (clarsimp simp:field_simps shiftl_t2n blah)
8350             apply (erule disjoint_subset[rotated])
8351             apply (rule_tac p1 = n in subset_trans[OF _ range_cover_subset])
8352                apply (simp add: upto_intvl_eq is_aligned_add range_cover.aligned is_aligned_shiftl)
8353                apply (simp add:field_simps shiftl_t2n)
8354               apply simp+
8355            apply (erule caps_overlap_reserved'_subseteq)
8356            apply (frule_tac x = "of_nat n" in range_cover_bound3)
8357             apply (rule word_of_nat_less)
8358             apply (simp add:range_cover.unat_of_nat_n)
8359            apply (clarsimp simp: field_simps shiftl_t2n blah)
8360           apply (clarsimp simp:createObject_c_preconds_def field_simps cong:region_is_bytes_cong)
8361           apply vcg
8362          apply (clarsimp simp: cte_C_size conj_comms untypedBits_defs)
8363          apply (simp cong: conj_cong)
8364          apply (intro conjI impI)
8365              apply (simp add: unat_eq_def)
8366             apply (drule range_cover_sz')
8367             apply (simp add: unat_eq_def word_less_nat_alt)
8368            apply (simp add: hrs_htd_update typ_region_bytes_actually_is_bytes)
8369           apply clarsimp
8370           apply (erule heap_list_is_zero_mono)
8371           apply (subgoal_tac "unat (num - of_nat n) \<noteq> 0")
8372            apply simp
8373           apply (simp only: unat_eq_0, clarsimp simp: unat_of_nat)
8374          apply (frule range_cover_sz')
8375          apply (clarsimp simp: Let_def hrs_htd_update
8376                                APIType_capBits_def[where ty="APIObjectType ArchTypes_H.apiobject_type.Untyped"])
8377          apply (subst is_aligned_add, erule range_cover.aligned)
8378           apply (simp add: is_aligned_shiftl)+
8379         apply (subst range_cover.unat_of_nat_n)
8380          apply (erule range_cover_le)
8381          subgoal by simp
8382         subgoal by (simp add:word_unat.Rep_inverse')
8383        apply clarsimp
8384        apply (rule conseqPre, vcg exspec=insertNewCap_preserves_bytes_flip
8385            exspec=createObject_preserves_bytes)
8386        apply (clarsimp simp del: imp_disjL)
8387        apply (frule(1) offset_intvl_first_chunk_subsets_unat,
8388          erule order_less_le_trans)
8389         apply (drule range_cover.weak)
8390         apply (simp add: word_le_nat_alt unat_of_nat)
8391
8392        apply (drule spec, drule mp, rule refl[where t="object_type_from_H newType"])
8393        apply clarsimp
8394        apply (rule context_conjI)
8395         apply (simp add: hrs_htd_update)
8396         apply (simp add: region_actually_is_bytes'_def, rule ballI)
8397         apply (drule bspec, erule(1) subsetD)
8398         apply (drule(1) orthD2)
8399         apply (simp add: Ball_def unat_eq_def typ_bytes_region_out)
8400        apply (erule trans[OF heap_list_h_eq2 heap_list_is_zero_mono2, rotated])
8401         apply (simp add: word_shiftl_add_distrib field_simps)
8402        apply (rule sym, rule byte_regions_unmodified_region_is_bytes)
8403          apply (erule byte_regions_unmodified_trans, simp_all)[1]
8404          apply (simp add: byte_regions_unmodified_def)
8405         apply simp
8406        apply assumption
8407
8408       apply (clarsimp simp:conj_comms field_simps
8409                       createObject_hs_preconds_def range_cover_sz')
8410       apply (subgoal_tac "is_aligned (ptr + (1 + of_nat n << APIType_capBits newType userSize))
8411         (APIType_capBits newType userSize)")
8412        prefer 2
8413        apply (rule aligned_add_aligned[OF range_cover.aligned],assumption)
8414         apply (rule is_aligned_shiftl_self)
8415        apply (simp)
8416       apply (simp add: range_cover_one[OF _  range_cover.sz(2) range_cover.sz(1)])
8417       including no_pre
8418       apply (wp insertNewCap_invs' insertNewCap_valid_pspace' insertNewCap_caps_overlap_reserved'
8419                 insertNewCap_pspace_no_overlap' insertNewCap_caps_no_overlap'' insertNewCap_descendants_range_in'
8420                 insertNewCap_untypedRange hoare_vcg_all_lift insertNewCap_cte_at static_imp_wp)
8421         apply (wp insertNewCap_cte_wp_at_other)
8422        apply (wp hoare_vcg_all_lift static_imp_wp insertNewCap_cte_at)
8423       apply (clarsimp simp:conj_comms |
8424         strengthen invs_valid_pspace' invs_pspace_aligned'
8425         invs_pspace_distinct')+
8426       apply (frule range_cover.range_cover_n_less)
8427       apply (subst upt_enum_offset_trivial)
8428         apply (rule minus_one_helper[OF word_of_nat_le])
8429          apply (fold_subgoals (prefix))[3]
8430          subgoal premises prems using prems
8431                    by (simp add:word_bits_conv minus_one_norm range_cover_not_zero[rotated])+
8432       apply (simp add: intvl_range_conv aligned_add_aligned[OF range_cover.aligned]
8433              is_aligned_shiftl_self range_cover_sz')
8434       apply (subst intvl_range_conv)
8435         apply (erule aligned_add_aligned[OF range_cover.aligned])
8436          apply (rule is_aligned_shiftl_self, rule le_refl)
8437        apply (erule range_cover_sz')
8438       apply (subst intvl_range_conv)
8439         apply (erule aligned_add_aligned[OF range_cover.aligned])
8440          apply (rule is_aligned_shiftl_self, rule le_refl)
8441        apply (erule range_cover_sz')
8442       apply (rule hoare_pre)
8443        apply (strengthen pspace_no_overlap'_strg[where sz = sz])
8444        apply (clarsimp simp:range_cover.sz conj_comms)
8445        apply (wp createObject_invs'
8446                  createObject_caps_overlap_reserved_ret' createObject_valid_cap'
8447                  createObject_descendants_range' createObject_idlethread_range
8448                  hoare_vcg_all_lift createObject_IRQHandler createObject_parent_helper
8449                  createObject_caps_overlap_reserved' createObject_caps_no_overlap''
8450                  createObject_pspace_no_overlap' createObject_cte_wp_at'
8451                  createObject_ex_cte_cap_wp_to createObject_descendants_range_in'
8452                  createObject_caps_overlap_reserved'
8453                  hoare_vcg_prop createObject_gsCNodes_p createObject_cnodes_have_size)
8454        apply (rule hoare_vcg_conj_lift[OF createObject_capRange_helper])
8455         apply (wp createObject_cte_wp_at' createObject_ex_cte_cap_wp_to
8456                   createObject_no_inter[where sz = sz] hoare_vcg_all_lift static_imp_wp)+
8457       apply (clarsimp simp:invs_pspace_aligned' invs_pspace_distinct' invs_valid_pspace'
8458         field_simps range_cover.sz conj_comms range_cover.aligned range_cover_sz'
8459         is_aligned_shiftl_self aligned_add_aligned[OF range_cover.aligned])
8460       apply (drule_tac x = n and  P = "\<lambda>x. x< length destSlots \<longrightarrow> Q x" for Q in spec)+
8461       apply clarsimp
8462       apply (simp add: range_cover_not_in_neqD)
8463       apply (intro conjI)
8464                          subgoal by (simp add: word_bits_def range_cover_def)
8465                         subgoal by (clarsimp simp: cte_wp_at_ctes_of invs'_def valid_state'_def
8466                                               valid_global_refs'_def cte_at_valid_cap_sizes_0)
8467                        apply (erule range_cover_le,simp)
8468                       apply (drule_tac p = "n" in range_cover_no_0)
8469                         apply (simp add:field_simps shiftl_t2n)+
8470                      apply (erule caps_no_overlap''_le)
8471                       apply (simp add:range_cover.sz[where 'a=32, folded word_bits_def])+
8472                     apply (erule caps_no_overlap''_le2)
8473                      apply (erule range_cover_compare_offset,simp+)
8474                     apply (simp add:range_cover_tail_mask[OF range_cover_le] range_cover_head_mask[OF range_cover_le])
8475                    apply (rule contra_subsetD)
8476                     apply (rule order_trans[rotated], erule range_cover_cell_subset,
8477                       erule of_nat_mono_maybe[rotated], simp)
8478                     apply (simp add: upto_intvl_eq shiftl_t2n mult.commute
8479                                      aligned_add_aligned[OF range_cover.aligned is_aligned_mult_triv2])
8480                    subgoal by simp
8481                   apply (simp add:cte_wp_at_no_0)
8482                  apply (rule disjoint_subset2[where B="{ptr .. foo}" for foo, rotated], simp add: Int_commute)
8483                  apply (rule order_trans[rotated], erule_tac p="Suc n" in range_cover_subset, simp+)
8484                  subgoal by (simp add: upto_intvl_eq shiftl_t2n mult.commute
8485                                   aligned_add_aligned[OF range_cover.aligned is_aligned_mult_triv2])
8486                 apply (drule_tac x = 0 in spec)
8487                 subgoal by simp
8488                apply (erule caps_overlap_reserved'_subseteq)
8489                subgoal by (clarsimp simp:range_cover_compare_offset blah)
8490               apply (erule descendants_range_in_subseteq')
8491               subgoal by (clarsimp simp:range_cover_compare_offset blah)
8492              apply (erule caps_overlap_reserved'_subseteq)
8493              apply (clarsimp simp:range_cover_compare_offset blah)
8494              apply (frule_tac x = "of_nat n" in range_cover_bound3)
8495               subgoal by (simp add:word_of_nat_less range_cover.unat_of_nat_n blah)
8496              subgoal by (simp add:field_simps shiftl_t2n blah)
8497             apply (simp add:shiftl_t2n field_simps)
8498             apply (rule contra_subsetD)
8499              apply (rule_tac x1 = 0 in subset_trans[OF _ range_cover_cell_subset,rotated ])
8500                apply (erule_tac p = n in range_cover_offset[rotated])
8501                subgoal by simp
8502               apply simp
8503               apply (rule less_diff_gt0)
8504               subgoal by (simp add:word_of_nat_less range_cover.unat_of_nat_n blah)
8505              apply (clarsimp simp: field_simps)
8506               apply (clarsimp simp: valid_idle'_def pred_tcb_at'_def
8507               dest!:invs_valid_idle' elim!: obj_atE')
8508             apply (drule(1) pspace_no_overlapD')
8509             apply (erule_tac x = "ksIdleThread s" in in_empty_interE[rotated])
8510              prefer 2
8511              apply (simp add:Int_ac)
8512             subgoal by (clarsimp simp: blah)
8513            subgoal by blast
8514           apply (erule descendants_range_in_subseteq')
8515           apply (clarsimp simp: blah)
8516           apply (rule order_trans[rotated], erule_tac x="of_nat n" in range_cover_bound'')
8517            subgoal by (simp add: word_less_nat_alt unat_of_nat)
8518           subgoal by (simp add: shiftl_t2n field_simps)
8519          apply (rule order_trans[rotated],
8520            erule_tac p="Suc n" in range_cover_subset, simp_all)[1]
8521          subgoal by (simp add: upto_intvl_eq shiftl_t2n mult.commute
8522                   aligned_add_aligned[OF range_cover.aligned is_aligned_mult_triv2])
8523         apply (erule cte_wp_at_weakenE')
8524         apply (clarsimp simp:shiftl_t2n field_simps)
8525         apply (erule subsetD)
8526         apply (erule subsetD[rotated])
8527         apply (rule_tac p1 = n in subset_trans[OF _ range_cover_subset])
8528            prefer 2
8529            apply (simp add:field_simps )
8530           apply (fold_subgoals (prefix))[2]
8531           subgoal premises prems using prems by (simp add:field_simps )+
8532        apply (clarsimp simp: word_shiftl_add_distrib)
8533        apply (clarsimp simp:blah field_simps shiftl_t2n)
8534        apply (drule word_eq_zeroI)
8535        apply (drule_tac p = "Suc n" in range_cover_no_0)
8536          apply (simp add:field_simps)+
8537       apply clarsimp
8538       apply (rule conjI)
8539        apply (drule_tac n = "x+1" and gbits = 4 in range_cover_not_zero_shift[OF _ range_cover_le,rotated])
8540           apply simp
8541          subgoal by (case_tac newType; simp add: objBits_simps
8542                       APIType_capBits_def range_cover_def split:apiobject_type.splits)
8543         subgoal by simp
8544        subgoal by (simp add:word_of_nat_plus word_shiftl_add_distrib field_simps shiftl_t2n)
8545       apply (drule_tac x = "Suc x" in spec)
8546       subgoal by (clarsimp simp: field_simps)
8547      apply clarsimp
8548      apply (subst range_cover.unat_of_nat_n)
8549       apply (erule range_cover_le)
8550       apply simp
8551      apply (simp add:word_unat.Rep_inverse')
8552      subgoal by (clarsimp simp:range_cover.range_cover_n_less[where 'a=32, simplified])
8553     subgoal by clarsimp
8554    apply vcg
8555   apply (rule conseqPre, vcg, clarsimp)
8556   apply (frule(1) ghost_assertion_size_logic)
8557   apply (drule range_cover_sz')
8558   subgoal by (intro conjI impI; simp add: o_def word_of_nat_less)
8559  apply (rule conjI)
8560   apply (frule range_cover.aligned)
8561   apply (frule range_cover_full[OF range_cover.aligned])
8562    apply (simp add:range_cover_def word_bits_def)
8563   apply (clarsimp simp: invs_valid_pspace' conj_comms intvl_range_conv
8564        createObject_hs_preconds_def range_cover.aligned range_cover_full)
8565   apply (frule(1) range_cover_gsMaxObjectSize, fastforce, assumption)
8566   apply (simp add: intvl_range_conv[OF range_cover.aligned range_cover_sz']
8567                    order_trans[OF _ APIType_capBits_min])
8568   apply (intro conjI)
8569           subgoal by (simp add: word_bits_def range_cover_def)
8570          apply (clarsimp simp:rf_sr_def cstate_relation_def Let_def)
8571          apply (erule pspace_no_overlap'_le)
8572           apply (fold_subgoals (prefix))[2]
8573           subgoal premises prems using prems
8574                     by (simp add:range_cover.sz[where 'a=32, simplified] word_bits_def)+
8575         apply (erule contra_subsetD[rotated])
8576         subgoal by (rule order_trans[rotated], rule range_cover_subset'[where n=1],
8577           erule range_cover_le, simp_all, (clarsimp simp: neq_Nil_conv)+)
8578        apply (rule disjoint_subset2[rotated])
8579         apply (simp add:Int_ac)
8580        apply (erule range_cover_subset[where p = 0,simplified])
8581         subgoal by simp
8582        subgoal by simp
8583       subgoal by (simp add: Int_commute shiftl_t2n mult.commute)
8584      apply (erule cte_wp_at_weakenE')
8585      apply (clarsimp simp:blah word_and_le2 shiftl_t2n field_simps)
8586      apply (frule range_cover_bound''[where x = "of_nat (length destSlots) - 1"])
8587       subgoal by (simp add: range_cover_not_zero[rotated])
8588      subgoal by (simp add:field_simps)
8589     subgoal by (erule range_cover_subset[where p=0, simplified]; simp)
8590    apply clarsimp
8591    apply (drule_tac x = k in spec)
8592    apply simp
8593    apply (drule(1) bspec[OF _ nth_mem])+
8594    subgoal by (clarsimp simp: field_simps)
8595   apply clarsimp
8596   apply (drule(1) bspec[OF _ nth_mem])+
8597   subgoal by (clarsimp simp:cte_wp_at_ctes_of)
8598  apply clarsimp
8599  apply (frule range_cover_sz')
8600  apply (frule(1) range_cover_gsMaxObjectSize, fastforce, assumption)
8601  apply (clarsimp simp: cte_wp_at_ctes_of)
8602  apply (drule(1) ghost_assertion_size_logic)+
8603  apply (simp add: o_def)
8604  apply (case_tac newType,simp_all add:object_type_from_H_def Kernel_C_defs
8605             nAPIObjects_def APIType_capBits_def o_def split:apiobject_type.splits)[1]
8606          subgoal by (simp add:unat_eq_def word_unat.Rep_inverse' word_less_nat_alt)
8607         subgoal by (clarsimp simp:objBits_simps',unat_arith)
8608        apply (fold_subgoals (prefix))[3]
8609        subgoal premises prems using prems
8610                  by (clarsimp simp: objBits_simps' unat_eq_def word_unat.Rep_inverse'
8611                                     word_less_nat_alt)+
8612     by (clarsimp simp: ARMSmallPageBits_def ARMLargePageBits_def
8613                        ARMSectionBits_def ARMSuperSectionBits_def machine_bits_defs)+
8614
8615end
8616
8617end
8618