1(*
2 * Copyright 2014, General Dynamics C4 Systems
3 *
4 * This software may be distributed and modified according to the terms of
5 * the GNU General Public License version 2. Note that NO WARRANTY is provided.
6 * See "LICENSE_GPLv2.txt" for details.
7 *
8 * @TAG(GD_GPL)
9 *)
10
11theory Arch_C
12imports Recycle_C
13begin
14
15context begin interpretation Arch . (*FIXME: arch_split*)
16
17crunches unmapPageTable, unmapPageDirectory, unmapPDPT
18  for ctes_of[wp]:  "\<lambda>s. P (ctes_of s)"
19  and gsMaxObjectSize[wp]: "\<lambda>s. P (gsMaxObjectSize s)"
20  (wp: crunch_wps simp: crunch_simps ignore: getObject setObject)
21
22end
23
24context kernel_m begin
25
26lemma storePTE_def':
27  "storePTE slot pte = setObject slot pte"
28  unfolding storePTE_def
29  by (simp add: tailM_def headM_def)
30
31lemma storePDE_def':
32  "storePDE slot pde = setObject slot pde"
33  unfolding storePDE_def
34  by (simp add: tailM_def headM_def)
35
36lemma storePDPTE_def':
37  "storePDPTE slot pdpte = setObject slot pdpte"
38  unfolding storePDPTE_def
39  by (simp add: tailM_def headM_def)
40
41lemma storePML4E_def':
42  "storePML4E slot pml4e = setObject slot pml4e"
43  unfolding storePML4E_def
44  by (simp add: tailM_def headM_def)
45
46lemma objBits_InvalidPTE:
47  "objBits X64_H.InvalidPTE = word_size_bits"
48  apply (simp add: objBits_simps archObjSize_def word_size_bits_def)
49  done
50
51lemma performPageTableInvocationUnmap_ccorres:
52  "ccorres (K (K \<bottom>) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
53       (invs' and cte_wp_at' (diminished' (ArchObjectCap cap) \<circ> cteCap) ctSlot
54              and (\<lambda>_. isPageTableCap cap))
55       (UNIV \<inter> \<lbrace>ccap_relation (ArchObjectCap cap) \<acute>cap\<rbrace> \<inter> \<lbrace>\<acute>ctSlot = Ptr ctSlot\<rbrace>)
56       []
57       (liftE (performPageTableInvocation (PageTableUnmap cap ctSlot)))
58       (Call performX86PageTableInvocationUnmap_'proc)"
59  apply (simp only: liftE_liftM ccorres_liftM_simp)
60  apply (rule ccorres_gen_asm)
61  apply (cinit lift: cap_' ctSlot_')
62   apply csymbr
63   apply (simp del: Collect_const)
64   apply (rule ccorres_split_nothrow_novcg_dc)
65      apply (subgoal_tac "capPTMappedAddress cap
66                           = (\<lambda>cp. if to_bool (capPTIsMapped_CL cp)
67                              then Some (capPTMappedASID_CL cp, capPTMappedAddress_CL cp)
68                              else None) (cap_page_table_cap_lift capa)")
69       apply (rule ccorres_Cond_rhs)
70        apply (simp add: to_bool_def)
71        apply (rule ccorres_rhs_assoc)+
72        apply csymbr
73        apply csymbr
74        apply csymbr
75        apply csymbr
76        apply (ctac add: unmapPageTable_ccorres)
77          apply csymbr
78          apply (simp add: storePTE_def' swp_def)
79          apply clarsimp
80          apply(simp only: dc_def[symmetric] bit_simps_corres[symmetric])
81          apply (ctac add: clearMemory_setObject_PTE_ccorres)
82         apply wp
83        apply (simp del: Collect_const)
84        apply (vcg exspec=unmapPageTable_modifies)
85       apply (simp add: to_bool_def)
86       apply (rule ccorres_return_Skip')
87      apply (simp add: cap_get_tag_isCap_ArchObject[symmetric])
88      apply (clarsimp simp: cap_lift_page_table_cap cap_to_H_def
89                            cap_page_table_cap_lift_def
90                     elim!: ccap_relationE cong: if_cong)
91     apply (simp add: liftM_def getSlotCap_def updateCap_def
92                 del: Collect_const)
93     apply (rule ccorres_move_c_guard_cte)
94     apply (rule ccorres_getCTE)+
95     apply (rule_tac P="cte_wp_at' ((=) rv) ctSlot
96                          and (\<lambda>_. rv = rva \<and> isArchCap isPageTableCap (cteCap rv))"
97                in ccorres_from_vcg_throws [where P'=UNIV])
98     apply (rule allI, rule conseqPre, vcg)
99     apply (clarsimp simp: cte_wp_at_ctes_of cap_get_tag_isCap_ArchObject)
100     apply (rule cmap_relationE1[OF cmap_relation_cte], assumption+)
101     apply (frule ccte_relation_ccap_relation)
102     apply (clarsimp simp: typ_heap_simps cap_get_tag_isCap_ArchObject)
103     apply (rule fst_setCTE [OF ctes_of_cte_at], assumption)
104     apply (erule rev_bexI)
105     apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
106                           typ_heap_simps')
107     apply (rule conjI)
108      apply (clarsimp simp: cpspace_relation_def typ_heap_simps')
109      apply (subst setCTE_tcb_case, assumption+)
110      apply (clarsimp dest!: ksPSpace_update_eq_ExD)
111      apply (erule cmap_relation_updI, assumption)
112       apply (simp add: cap_get_tag_isCap_ArchObject[symmetric])
113       apply (clarsimp simp: ccte_relation_def c_valid_cte_def
114                      elim!: ccap_relationE)
115       apply (subst cteCap_update_cte_to_H)
116         apply (clarsimp simp: map_option_Some_eq2)
117         apply (rule trans, rule sym, rule option.sel, rule sym, erule arg_cong)
118        apply (erule iffD1[OF cap_page_table_cap_lift])
119       apply (clarsimp simp: map_option_Some_eq2 cap_get_tag_isCap_ArchObject[symmetric]
120                             cap_lift_page_table_cap cap_to_H_def
121                             cap_page_table_cap_lift_def)
122      apply simp
123     apply (clarsimp simp: carch_state_relation_def cmachine_state_relation_def
124                           fpu_null_state_heap_update_tag_disj_simps
125                           global_ioport_bitmap_heap_update_tag_disj_simps
126                           cvariable_array_map_const_add_map_option[where f="tcb_no_ctes_proj"]
127                    dest!: ksPSpace_update_eq_ExD)
128    apply (simp add: cte_wp_at_ctes_of)
129    apply (wp mapM_x_wp' | wpc | simp)+
130   apply (simp add: guard_is_UNIV_def)
131  apply (clarsimp simp: cap_get_tag_isCap_ArchObject[symmetric] cte_wp_at_ctes_of)
132  apply (frule ctes_of_valid', clarsimp)
133  apply (frule_tac x=s in fun_cong[OF diminished_valid'])
134  apply (frule valid_global_refsD_with_objSize, clarsimp)
135  apply (clarsimp simp: cap_lift_page_table_cap cap_to_H_def
136                        cap_page_table_cap_lift_def isCap_simps
137                        valid_cap'_def get_capSizeBits_CL_def
138                        bit_simps capAligned_def
139                        to_bool_def mask_def page_table_at'_def
140                        capRange_def Int_commute asid_bits_def
141                 elim!: ccap_relationE cong: if_cong
142                 dest!: diminished_capMaster)
143  apply (drule spec[where x=0])
144  apply (auto simp add: word_and_le1)
145  done
146
147lemma clearMemory_setObject_PDE_ccorres:
148  "ccorres dc xfdc (page_directory_at' ptr
149                and (\<lambda>s. 2 ^ pdBits \<le> gsMaxObjectSize s)
150                and (\<lambda>_. is_aligned ptr pdBits \<and> ptr \<noteq> 0 \<and> pstart = addrFromPPtr ptr))
151            (UNIV \<inter> {s. ptr___ptr_to_void_' s = Ptr ptr} \<inter> {s. bits_' s = of_nat pdBits}) []
152       (mapM_x (\<lambda>a. setObject a X64_H.InvalidPDE)
153                       [ptr , ptr + 2 ^ objBits X64_H.InvalidPDE .e. ptr + 2 ^ pdBits - 1])
154       (Call clearMemory_'proc)"
155  apply (rule ccorres_gen_asm)+
156  apply (cinit' lift: ptr___ptr_to_void_' bits_')
157   apply (rule_tac P="page_directory_at' ptr and (\<lambda>s. 2 ^ pdBits \<le> gsMaxObjectSize s)"
158               in ccorres_from_vcg_nofail[where P'=UNIV])
159   apply (rule allI, rule conseqPre, vcg)
160   apply clarsimp
161   apply (subst ghost_assertion_size_logic[unfolded o_def])
162     apply (simp add: bit_simps)
163    apply simp
164   apply (clarsimp simp: replicateHider_def[symmetric] bit_simps)
165   apply (frule is_aligned_no_overflow', simp)
166   apply (intro conjI)
167      apply (erule is_aligned_weaken, simp)
168     apply (clarsimp simp: is_aligned_def)
169    apply (erule (1) page_directory_at_rf_sr_dom_s[unfolded pdBits_def bit_simps, simplified])
170   apply (clarsimp simp add: bit_simps
171                      cong: StateSpace.state.fold_congs globals.fold_congs)
172   apply (simp add: upto_enum_step_def objBits_simps bit_simps
173                    field_simps linorder_not_less[symmetric] archObjSize_def
174                    upto_enum_word split_def)
175  apply (erule mapM_x_store_memset_ccorres_assist
176                      [unfolded split_def, OF _ _ _ _ _ _ subset_refl],
177         simp_all add: shiftl_t2n hd_map objBits_simps archObjSize_def bit_simps)[1]
178   apply (rule cmap_relationE1, erule rf_sr_cpde_relation, erule ko_at_projectKO_opt)
179   apply (subst coerce_memset_to_heap_update_pde)
180   apply (clarsimp simp: rf_sr_def Let_def cstate_relation_def typ_heap_simps)
181   apply (rule conjI)
182    apply (simp add: cpspace_relation_def typ_heap_simps update_pde_map_tos
183                     update_pde_map_to_pdes carray_map_relation_upd_triv)
184    apply (rule cmap_relation_updI, simp_all)[1]
185    subgoal by (simp add: cpde_relation_def Let_def pde_lift_def pde_get_tag_def pde_tag_defs
186                   split: if_splits pde.split_asm)
187   apply (simp add: carch_state_relation_def cmachine_state_relation_def
188                    fpu_null_state_heap_update_tag_disj_simps
189                    global_ioport_bitmap_heap_update_tag_disj_simps
190                    update_pde_map_tos)
191  apply simp
192  done
193
194lemma performPageDirectoryInvocationUnmap_ccorres:
195  "ccorres (K (K \<bottom>) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
196       (invs' and cte_wp_at' (diminished' (ArchObjectCap cap) \<circ> cteCap) ctSlot
197              and (\<lambda>_. isPageDirectoryCap cap))
198       (UNIV \<inter> \<lbrace>ccap_relation (ArchObjectCap cap) \<acute>cap\<rbrace> \<inter> \<lbrace>\<acute>ctSlot = Ptr ctSlot\<rbrace>)
199       []
200       (liftE (performPageDirectoryInvocation (PageDirectoryUnmap cap ctSlot)))
201       (Call performX64PageDirectoryInvocationUnmap_'proc)"
202  apply (simp only: liftE_liftM ccorres_liftM_simp)
203  apply (rule ccorres_gen_asm)
204  apply (cinit lift: cap_' ctSlot_')
205   apply csymbr
206   apply (simp del: Collect_const)
207   apply (rule ccorres_split_nothrow_novcg_dc)
208      apply (subgoal_tac "capPDMappedAddress cap
209                           = (\<lambda>cp. if to_bool (capPDIsMapped_CL cp)
210                              then Some (capPDMappedASID_CL cp, capPDMappedAddress_CL cp)
211                              else None) (cap_page_directory_cap_lift capa)")
212       apply (rule ccorres_Cond_rhs)
213        apply (simp add: to_bool_def)
214        apply (rule ccorres_rhs_assoc)+
215        apply csymbr
216        apply csymbr
217        apply csymbr
218        apply csymbr
219        apply (ctac add: unmapPageDirectory_ccorres)
220          apply csymbr
221          apply (simp add: storePDE_def' swp_def)
222          apply clarsimp
223          apply(simp only: dc_def[symmetric] bit_simps_corres[symmetric])
224          apply (ctac add: clearMemory_setObject_PDE_ccorres)
225         apply wp
226        apply (simp del: Collect_const)
227        apply (vcg exspec=unmapPageDirectory_modifies)
228       apply (simp add: to_bool_def)
229       apply (rule ccorres_return_Skip')
230      apply (simp add: cap_get_tag_isCap_ArchObject[symmetric])
231      apply (clarsimp simp: cap_lift_page_directory_cap cap_to_H_def
232                            cap_page_directory_cap_lift_def
233                     elim!: ccap_relationE cong: if_cong)
234     apply (simp add: liftM_def getSlotCap_def updateCap_def
235                 del: Collect_const)
236     apply (rule ccorres_move_c_guard_cte)
237     apply (rule ccorres_getCTE)+
238     apply (rule_tac P="cte_wp_at' ((=) rv) ctSlot
239                          and (\<lambda>_. rv = rva \<and> isArchCap isPageDirectoryCap (cteCap rv))"
240                in ccorres_from_vcg_throws [where P'=UNIV])
241     apply (rule allI, rule conseqPre, vcg)
242     apply (clarsimp simp: cte_wp_at_ctes_of cap_get_tag_isCap_ArchObject)
243     apply (rule cmap_relationE1[OF cmap_relation_cte], assumption+)
244     apply (frule ccte_relation_ccap_relation)
245     apply (clarsimp simp: typ_heap_simps cap_get_tag_isCap_ArchObject)
246     apply (rule fst_setCTE [OF ctes_of_cte_at], assumption)
247     apply (erule rev_bexI)
248     apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
249                           typ_heap_simps')
250     apply (rule conjI)
251      apply (clarsimp simp: cpspace_relation_def typ_heap_simps')
252      apply (subst setCTE_tcb_case, assumption+)
253      apply (clarsimp dest!: ksPSpace_update_eq_ExD)
254      apply (erule cmap_relation_updI, assumption)
255       apply (simp add: cap_get_tag_isCap_ArchObject[symmetric])
256       apply (clarsimp simp: ccte_relation_def c_valid_cte_def
257                      elim!: ccap_relationE)
258       apply (subst cteCap_update_cte_to_H)
259         apply (clarsimp simp: map_option_Some_eq2)
260         apply (rule trans, rule sym, rule option.sel, rule sym, erule arg_cong)
261        apply (erule iffD1[OF cap_page_directory_cap_lift])
262       apply (clarsimp simp: map_option_Some_eq2 cap_get_tag_isCap_ArchObject[symmetric]
263                             cap_lift_page_directory_cap cap_to_H_def
264                             cap_page_directory_cap_lift_def)
265      apply simp
266     apply (clarsimp simp: carch_state_relation_def cmachine_state_relation_def
267                           fpu_null_state_heap_update_tag_disj_simps
268                           global_ioport_bitmap_heap_update_tag_disj_simps
269                           cvariable_array_map_const_add_map_option[where f="tcb_no_ctes_proj"]
270                    dest!: ksPSpace_update_eq_ExD)
271    apply (simp add: cte_wp_at_ctes_of)
272    apply (wp mapM_x_wp' | wpc | simp)+
273   apply (simp add: guard_is_UNIV_def)
274  apply (clarsimp simp: cap_get_tag_isCap_ArchObject[symmetric] cte_wp_at_ctes_of)
275  apply (frule ctes_of_valid', clarsimp)
276  apply (frule_tac x=s in fun_cong[OF diminished_valid'])
277  apply (frule valid_global_refsD_with_objSize, clarsimp)
278  apply (clarsimp simp: cap_lift_page_directory_cap cap_to_H_def
279                        cap_page_directory_cap_lift_def isCap_simps
280                        valid_cap'_def get_capSizeBits_CL_def
281                        bit_simps capAligned_def
282                        to_bool_def mask_def page_directory_at'_def
283                        capRange_def Int_commute asid_bits_def
284                 elim!: ccap_relationE cong: if_cong
285                 dest!: diminished_capMaster)
286  apply (drule spec[where x=0])
287  apply (auto simp add: word_and_le1)
288  done
289
290lemma clearMemory_setObject_PDPTE_ccorres:
291  "ccorres dc xfdc (pd_pointer_table_at' ptr
292                and (\<lambda>s. 2 ^ pdptBits \<le> gsMaxObjectSize s)
293                and (\<lambda>_. is_aligned ptr pdptBits \<and> ptr \<noteq> 0 \<and> pstart = addrFromPPtr ptr))
294            (UNIV \<inter> {s. ptr___ptr_to_void_' s = Ptr ptr} \<inter> {s. bits_' s = of_nat pdptBits}) []
295       (mapM_x (\<lambda>a. setObject a X64_H.InvalidPDPTE)
296                       [ptr , ptr + 2 ^ objBits X64_H.InvalidPDPTE .e. ptr + 2 ^ pdptBits - 1])
297       (Call clearMemory_'proc)"
298  apply (rule ccorres_gen_asm)+
299  apply (cinit' lift: ptr___ptr_to_void_' bits_')
300   apply (rule_tac P="pd_pointer_table_at' ptr and (\<lambda>s. 2 ^ pdptBits \<le> gsMaxObjectSize s)"
301               in ccorres_from_vcg_nofail[where P'=UNIV])
302   apply (rule allI, rule conseqPre, vcg)
303   apply clarsimp
304   apply (subst ghost_assertion_size_logic[unfolded o_def])
305     apply (simp add: bit_simps)
306    apply simp
307   apply (clarsimp simp: replicateHider_def[symmetric] bit_simps)
308   apply (frule is_aligned_no_overflow', simp)
309   apply (intro conjI)
310      apply (erule is_aligned_weaken, simp)
311     apply (clarsimp simp: is_aligned_def)
312    apply (erule (1) pd_pointer_table_at_rf_sr_dom_s[unfolded pdptBits_def bit_simps, simplified])
313   apply (clarsimp simp add: bit_simps
314                      cong: StateSpace.state.fold_congs globals.fold_congs)
315   apply (simp add: upto_enum_step_def objBits_simps bit_simps
316                    field_simps linorder_not_less[symmetric] archObjSize_def
317                    upto_enum_word split_def)
318  apply (erule mapM_x_store_memset_ccorres_assist
319                      [unfolded split_def, OF _ _ _ _ _ _ subset_refl],
320         simp_all add: shiftl_t2n hd_map objBits_simps archObjSize_def bit_simps)[1]
321   apply (rule cmap_relationE1, erule rf_sr_cpdpte_relation, erule ko_at_projectKO_opt)
322   apply (subst coerce_memset_to_heap_update_pdpte)
323   apply (clarsimp simp: rf_sr_def Let_def cstate_relation_def typ_heap_simps)
324   apply (rule conjI)
325    apply (simp add: cpspace_relation_def typ_heap_simps update_pdpte_map_tos
326                     update_pdpte_map_to_pdptes carray_map_relation_upd_triv)
327    apply (rule cmap_relation_updI, simp_all)[1]
328    subgoal by (simp add: cpdpte_relation_def Let_def pdpte_lift_def pdpte_get_tag_def pdpte_tag_defs
329                   split: if_splits pdpte.split_asm)
330   apply (simp add: carch_state_relation_def cmachine_state_relation_def
331                    fpu_null_state_heap_update_tag_disj_simps
332                    global_ioport_bitmap_heap_update_tag_disj_simps
333                    update_pdpte_map_tos)
334  apply simp
335  done
336
337lemma performPDPTInvocationUnmap_ccorres:
338  "ccorres (K (K \<bottom>) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
339       (invs' and cte_wp_at' (diminished' (ArchObjectCap cap) \<circ> cteCap) ctSlot
340              and (\<lambda>_. isPDPointerTableCap cap))
341       (UNIV \<inter> \<lbrace>ccap_relation (ArchObjectCap cap) \<acute>cap\<rbrace> \<inter> \<lbrace>\<acute>ctSlot = Ptr ctSlot\<rbrace>)
342       []
343       (liftE (performPDPTInvocation (PDPTUnmap cap ctSlot)))
344       (Call performX64PDPTInvocationUnmap_'proc)"
345  apply (simp only: liftE_liftM ccorres_liftM_simp)
346  apply (rule ccorres_gen_asm)
347  apply (cinit lift: cap_' ctSlot_')
348   apply csymbr
349   apply (simp del: Collect_const)
350   apply (rule ccorres_split_nothrow_novcg_dc)
351      apply (subgoal_tac "capPDPTMappedAddress cap
352                           = (\<lambda>cp. if to_bool (capPDPTIsMapped_CL cp)
353                              then Some (capPDPTMappedASID_CL cp, capPDPTMappedAddress_CL cp)
354                              else None) (cap_pdpt_cap_lift capa)")
355       apply (rule ccorres_Cond_rhs)
356        apply (simp add: to_bool_def)
357        apply (rule ccorres_rhs_assoc)+
358        apply csymbr
359        apply csymbr
360        apply csymbr
361        apply csymbr
362        apply (ctac add: unmapPDPointerTable_ccorres)
363          apply csymbr
364          apply (simp add: storePDPTE_def' swp_def)
365          apply clarsimp
366          apply(simp only: dc_def[symmetric] bit_simps_corres[symmetric])
367          apply (ctac add: clearMemory_setObject_PDPTE_ccorres)
368         apply wp
369        apply (simp del: Collect_const)
370        apply (vcg exspec=unmapPDPT_modifies)
371       apply (simp add: to_bool_def)
372       apply (rule ccorres_return_Skip')
373      apply (simp add: cap_get_tag_isCap_ArchObject[symmetric])
374      apply (clarsimp simp: cap_lift_pdpt_cap cap_to_H_def
375                            cap_pdpt_cap_lift_def
376                     elim!: ccap_relationE cong: if_cong)
377     apply (simp add: liftM_def getSlotCap_def updateCap_def
378                 del: Collect_const)
379     apply (rule ccorres_move_c_guard_cte)
380     apply (rule ccorres_getCTE)+
381     apply (rule_tac P="cte_wp_at' ((=) rv) ctSlot
382                          and (\<lambda>_. rv = rva \<and> isArchCap isPDPointerTableCap (cteCap rv))"
383                in ccorres_from_vcg_throws [where P'=UNIV])
384     apply (rule allI, rule conseqPre, vcg)
385     apply (clarsimp simp: cte_wp_at_ctes_of cap_get_tag_isCap_ArchObject)
386     apply (rule cmap_relationE1[OF cmap_relation_cte], assumption+)
387     apply (frule ccte_relation_ccap_relation)
388     apply (clarsimp simp: typ_heap_simps cap_get_tag_isCap_ArchObject)
389     apply (rule fst_setCTE [OF ctes_of_cte_at], assumption)
390     apply (erule rev_bexI)
391     apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
392                           typ_heap_simps')
393     apply (rule conjI)
394      apply (clarsimp simp: cpspace_relation_def typ_heap_simps')
395      apply (subst setCTE_tcb_case, assumption+)
396      apply (clarsimp dest!: ksPSpace_update_eq_ExD)
397      apply (erule cmap_relation_updI, assumption)
398       apply (simp add: cap_get_tag_isCap_ArchObject[symmetric])
399       apply (clarsimp simp: ccte_relation_def c_valid_cte_def
400                      elim!: ccap_relationE)
401       apply (subst cteCap_update_cte_to_H)
402         apply (clarsimp simp: map_option_Some_eq2)
403         apply (rule trans, rule sym, rule option.sel, rule sym, erule arg_cong)
404        apply (erule iffD1[OF cap_pdpt_cap_lift])
405       apply (clarsimp simp: map_option_Some_eq2 cap_get_tag_isCap_ArchObject[symmetric]
406                             cap_lift_pdpt_cap cap_to_H_def
407                             cap_pdpt_cap_lift_def)
408      apply simp
409     apply (clarsimp simp: carch_state_relation_def cmachine_state_relation_def
410                           fpu_null_state_heap_update_tag_disj_simps
411                           global_ioport_bitmap_heap_update_tag_disj_simps
412                           cvariable_array_map_const_add_map_option[where f="tcb_no_ctes_proj"]
413                    dest!: ksPSpace_update_eq_ExD)
414    apply (simp add: cte_wp_at_ctes_of)
415    apply (wp mapM_x_wp' | wpc | simp)+
416   apply (simp add: guard_is_UNIV_def)
417  apply (clarsimp simp: cap_get_tag_isCap_ArchObject[symmetric] cte_wp_at_ctes_of)
418  apply (frule ctes_of_valid', clarsimp)
419  apply (frule_tac x=s in fun_cong[OF diminished_valid'])
420  apply (frule valid_global_refsD_with_objSize, clarsimp)
421  apply (clarsimp simp: cap_lift_pdpt_cap cap_to_H_def
422                        cap_pdpt_cap_lift_def isCap_simps
423                        valid_cap'_def get_capSizeBits_CL_def
424                        bit_simps capAligned_def
425                        to_bool_def mask_def pd_pointer_table_at'_def
426                        capRange_def Int_commute asid_bits_def
427                 elim!: ccap_relationE cong: if_cong
428                 dest!: diminished_capMaster)
429  apply (drule spec[where x=0])
430  apply (auto simp add: word_and_le1)
431  done
432
433lemma cap_case_PML4Cap2:
434  "(case cap of ArchObjectCap (PML4Cap pd mapdata)
435                   \<Rightarrow> f pd mapdata | _ \<Rightarrow> g)
436   = (if isArchObjectCap cap \<and> isPML4Cap (capCap cap)
437      then f (capPML4BasePtr (capCap cap)) (capPML4MappedASID (capCap cap))
438      else g)"
439  by (simp add: isCap_simps
440         split: capability.split arch_capability.split)
441
442lemma ap_eq_D:
443  "x \<lparr>array_C := arr'\<rparr> = asid_pool_C arr \<Longrightarrow> arr' = arr"
444  by (cases x) simp
445
446declare Kernel_C.asid_pool_C_size [simp del]
447
448lemma createObjects_asidpool_ccorres:
449  shows "ccorres dc xfdc
450   ((\<lambda>s. \<exists>p. cte_wp_at' (\<lambda>cte. cteCap cte = UntypedCap isdev frame pageBits idx ) p s)
451    and pspace_aligned' and pspace_distinct' and valid_objs'
452    and ret_zero frame (2 ^ pageBits)
453    and valid_global_refs' and pspace_no_overlap' frame pageBits)
454   ({s. region_actually_is_bytes frame (2^pageBits) s})
455   hs
456   (placeNewObject frame (makeObject::asidpool) 0)
457   (CALL memzero(Ptr frame, (2 ^ pageBits));;
458   (global_htd_update (\<lambda>_. ptr_retyp (ap_Ptr frame))))"
459proof -
460  have helper: "\<forall>\<sigma> x. (\<sigma>, x) \<in> rf_sr \<and> is_aligned frame pageBits \<and> frame \<noteq> 0
461  \<and> pspace_aligned' \<sigma> \<and> pspace_distinct' \<sigma>
462  \<and> pspace_no_overlap' frame pageBits \<sigma>
463  \<and> ret_zero frame (2 ^ pageBits) \<sigma>
464  \<and> region_actually_is_bytes frame (2 ^ pageBits) x
465  \<and> {frame ..+ 2 ^ pageBits} \<inter> kernel_data_refs = {}
466 \<longrightarrow>
467  (\<sigma>\<lparr>ksPSpace := foldr (\<lambda>addr. data_map_insert addr (KOArch (KOASIDPool makeObject))) (new_cap_addrs (Suc 0) frame (KOArch (KOASIDPool makeObject))) (ksPSpace \<sigma>)\<rparr>,
468   x\<lparr>globals := globals x
469                 \<lparr>t_hrs_' := hrs_htd_update (ptr_retyps_gen 1 (ap_Ptr frame) False)
470                       (hrs_mem_update
471                         (heap_update_list frame (replicate (2 ^ pageBits) 0))
472                         (t_hrs_' (globals x)))\<rparr>\<rparr>) \<in> rf_sr"
473    (is "\<forall>\<sigma> x. ?P \<sigma> x \<longrightarrow>
474        (\<sigma>\<lparr>ksPSpace := ?ks \<sigma>\<rparr>, x\<lparr>globals := globals x\<lparr>t_hrs_' := ?ks' x\<rparr>\<rparr>) \<in> rf_sr")
475  proof (intro impI allI)
476  fix \<sigma> x
477  let ?thesis = "(\<sigma>\<lparr>ksPSpace := ?ks \<sigma>\<rparr>, x\<lparr>globals := globals x\<lparr>t_hrs_' := ?ks' x\<rparr>\<rparr>) \<in> rf_sr"
478  let ?ks = "?ks \<sigma>"
479  let ?ks' = "?ks' x"
480  let ?ptr = "ap_Ptr frame"
481
482  assume "?P \<sigma> x"
483  hence rf: "(\<sigma>, x) \<in> rf_sr" and al: "is_aligned frame pageBits" and ptr0: "frame \<noteq> 0"
484    and pal: "pspace_aligned' \<sigma>" and pdst: "pspace_distinct' \<sigma>"
485    and pno: "pspace_no_overlap' frame pageBits \<sigma>"
486    and zro: "ret_zero frame (2 ^ pageBits) \<sigma>"
487    and actually: "region_actually_is_bytes frame (2 ^ pageBits) x"
488    and kdr: "{frame ..+ 2 ^ pageBits} \<inter> kernel_data_refs = {}"
489    by simp_all
490
491  note empty = region_actually_is_bytes[OF actually]
492
493  have relrl:
494    "casid_pool_relation makeObject (from_bytes (replicate (size_of TYPE(asid_pool_C)) 0))"
495    unfolding casid_pool_relation_def
496    apply (clarsimp simp: makeObject_asidpool split: asid_pool_C.splits)
497    apply (clarsimp simp: array_relation_def option_to_ptr_def asid_map_relation_def)
498    apply (rename_tac words i)
499    apply (subgoal_tac "asid_map_C.words_C (words.[unat i]).[0] = 0")
500    apply (clarsimp simp: asid_map_lift_def asid_map_get_tag_def asid_map_tag_defs)
501    apply (simp add: from_bytes_def)
502    apply (simp add: typ_info_simps asid_pool_C_tag_def
503                     size_td_lt_final_pad size_td_lt_ti_typ_pad_combine Let_def size_of_def)
504    apply (simp add: final_pad_def Let_def size_td_lt_ti_typ_pad_combine)
505    apply (simp add: padup_def align_td_array')
506    apply (subst (asm) size_td_array)
507    apply simp
508    apply (simp add: size_td_array ti_typ_pad_combine_def ti_typ_combine_def
509                     Let_def empty_typ_info_def update_ti_adjust_ti
510                del: replicate_numeral Kernel_C.pde_C_size)
511    apply (simp add: typ_info_array array_tag_def
512                del: replicate_numeral)
513    apply (drule ap_eq_D)
514    apply (clarsimp simp del: replicate_numeral)
515    apply (subst update_ti_t_array_tag_n_rep)
516      apply (simp del: replicate_numeral)
517     apply simp
518    apply (simp del: replicate_numeral)
519    apply (subst index_fold_update)
520       apply simp
521      apply auto[1]
522     apply (simp add: asid_low_bits_def word_le_nat_alt)
523    apply (simp split: if_split)
524    apply (rule conjI[rotated], simp add: asid_low_bits_def word_le_nat_alt)
525    apply (simp add: asid_map_C_typ_info asid_map_C_tag_def
526                     final_pad_def Let_def size_td_lt_ti_typ_pad_combine
527                     padup_def align_td_array' size_td_array)
528    apply (simp add: ti_typ_pad_combine_def ti_typ_combine_def empty_typ_info_def
529                     update_ti_adjust_ti size_td_array typ_info_array array_tag_def
530                     update_ti_t_array_tag_n_rep update_ti_t_machine_word_0s)
531    done
532
533  define ko where "ko \<equiv> KOArch (KOASIDPool makeObject)"
534
535  have rc :"range_cover frame (objBitsKO ko) (objBitsKO ko) (Suc 0)"
536    by (simp add:objBits_simps ko_def archObjSize_def al range_cover_full)
537
538  have rc' :"range_cover frame (objBitsKO ko) (objBitsKO ko) (2 ^ 0)"
539    by (simp add:objBits_simps ko_def archObjSize_def al range_cover_full range_cover_rel)
540
541  have pno': "pspace_no_overlap' frame (objBitsKO ko) \<sigma>"
542    by (simp add:objBits_simps pno ko_def archObjSize_def al)
543
544  have al': "is_aligned frame (objBitsKO (ko::kernel_object))"
545    by (simp add:objBits_simps ko_def archObjSize_def al)
546
547  (* s/obj/obj'/ *)
548  have szo: "size_of TYPE(asid_pool_C) = 2 ^ objBitsKO ko"
549    by (simp add: size_of_def objBits_simps ko_def archObjSize_def pageBits_def)
550  have szko: "objBitsKO ko = pageBits"
551    by (simp add: objBits_simps ko_def archObjSize_def)
552  hence sz: "objBitsKO ko \<le> pageBits" by simp
553  have szo': "2 ^ pageBits = 2 ^ (pageBits - objBitsKO ko) * size_of TYPE(asid_pool_C)" using szko
554    apply (subst szo)
555    apply (simp add: power_add [symmetric])
556    done
557
558  have [simp]: "(2::nat) ^ (pageBits - objBitsKO ko) * 2 ^ objBitsKO ko = 2 ^ pageBits"
559     by (clarsimp simp:pageBits_def objBits_simps ko_def archObjSize_def)
560
561  have ptr_retyp:
562    "hrs_htd_update (ptr_retyps (2 ^ (pageBits - objBitsKO ko)) (ap_Ptr frame)) = hrs_htd_update (ptr_retyp (ap_Ptr frame))"
563    apply (simp add: szko hrs_htd_update_def)
564    done
565
566  note rl' = cslift_ptr_retyp_memset_other_inst [OF _ rc' _ szo,
567     simplified, OF empty[folded szko] szo[symmetric], unfolded szko]
568
569  have szb: "pageBits < word_bits" by simp
570  have mko: "\<And>dev. makeObjectKO dev (Inl (KOArch (KOASIDPool f))) = Some ko"
571    by (simp add: ko_def makeObjectKO_def)
572
573
574  note rl = projectKO_opt_retyp_other [OF rc pal pno' ko_def]
575
576  note cterl = retype_ctes_helper
577                 [OF pal pdst pno' al' le_refl
578                     range_cover_sz'[where 'a=machine_word_len,
579                                     folded word_bits_def, OF rc]
580                     mko rc, simplified]
581
582  note ht_rl = clift_eq_h_t_valid_eq[OF rl', OF tag_disj_via_td_name, simplified]
583    uinfo_array_tag_n_m_not_le_typ_name
584
585  have guard:
586    "\<forall>n<2 ^ (pageBits - objBitsKO ko). c_guard (CTypesDefs.ptr_add ?ptr (of_nat n))"
587    apply (rule retype_guard_helper[where m=3])
588        apply (rule range_cover_rel[OF rc])
589         apply fastforce
590        apply simp
591       apply (clarsimp simp:objBits_simps ko_def archObjSize_def)
592      apply (simp add:ptr0)
593     apply (simp add:szo)
594   apply (simp add:align_of_def objBits_simps pageBits_def ko_def archObjSize_def)+
595   done
596
597  have cslift_ptr_retyp_helper:
598   "\<forall>x::asid_pool_C ptr\<in>dom (cslift x). is_aligned (ptr_val x) (objBitsKO ko)
599   \<Longrightarrow> clift (hrs_htd_update (ptr_retyps_gen 1 (ap_Ptr frame) False)
600           (hrs_mem_update (heap_update_list frame (replicate ((2::nat) ^ pageBits) (0::word8)))
601             (t_hrs_' (globals x)))) =
602   (\<lambda>y::asid_pool_C ptr.
603       if y \<in> (CTypesDefs.ptr_add (ap_Ptr frame) \<circ> of_nat) ` {k::nat. k < (2::nat) ^ (pageBits - objBitsKO ko)}
604       then Some (from_bytes (replicate (size_of TYPE(asid_pool_C)) (0::word8))) else cslift x y)"
605    using guard
606    apply (subst clift_ptr_retyps_gen_memset_same, simp_all add: szo szko)
607    apply (simp add: szo empty szko)
608    done
609
610  from rf have "cpspace_relation (ksPSpace \<sigma>) (underlying_memory (ksMachineState \<sigma>)) (t_hrs_' (globals x))"
611    unfolding rf_sr_def cstate_relation_def by (simp add: Let_def)
612  hence "cpspace_relation ?ks (underlying_memory (ksMachineState \<sigma>)) ?ks'"
613    unfolding cpspace_relation_def
614    apply -
615    apply (clarsimp simp: rl' cterl[unfolded ko_def] tag_disj_via_td_name
616                 foldr_upd_app_if [folded data_map_insert_def] cte_C_size tcb_C_size)
617    apply (subst cslift_ptr_retyp_helper[simplified])
618     apply (erule pspace_aligned_to_C [OF pal])
619      apply (simp add: projectKOs ko_def)
620     apply (simp add: ko_def projectKOs objBits_simps archObjSize_def)
621    apply (simp add: ptr_add_to_new_cap_addrs [OF szo] ht_rl)
622    apply (simp add: rl[unfolded ko_def] projectKO_opt_retyp_same ko_def projectKOs cong: if_cong)
623    apply (simp add:objBits_simps archObjSize_def)
624    apply (erule cmap_relation_retype)
625    apply (rule relrl)
626    done
627
628  thus ?thesis using rf empty kdr zro
629    apply (simp add: rf_sr_def cstate_relation_def Let_def rl' tag_disj_via_td_name
630                     ko_def[symmetric])
631    apply (simp add: carch_state_relation_def cmachine_state_relation_def
632                     fpu_null_state_relation_def)
633    apply (simp add: rl' cterl tag_disj_via_td_name h_t_valid_clift_Some_iff tcb_C_size)
634    apply (clarsimp simp: hrs_htd_update ptr_retyps_htd_safe_neg szo szko
635                          kernel_data_refs_domain_eq_rotate
636                          cvariable_array_ptr_retyps[OF szo]
637                          foldr_upd_app_if [folded data_map_insert_def]
638                          zero_ranges_ptr_retyps
639                          rl empty projectKOs)
640    done
641  qed
642
643  have [simp]:
644    "of_nat pageBits < (4::word32) = False" by (simp add: pageBits_def)
645
646  show ?thesis
647  apply (rule ccorres_from_vcg_nofail2, rule allI)
648  apply (rule conseqPre)
649   apply vcg
650  apply (clarsimp simp: cte_wp_at_ctes_of split: if_split_asm)
651  apply (frule(1) ctes_of_valid', clarsimp)
652  apply (subst ghost_assertion_size_logic[unfolded o_def, rotated], assumption)
653   apply (drule(1) valid_global_refsD_with_objSize)
654   apply (simp add: X64SmallPageBits_def pageBits_def)
655  apply (erule valid_untyped_capE)
656  apply (subst simpler_placeNewObject_def)
657      apply ((simp add: word_bits_conv objBits_simps archObjSize_def
658                        capAligned_def)+)[4]
659  apply (simp add: simpler_modify_def rf_sr_htd_safe)
660  apply (subgoal_tac "{frame ..+ 2 ^ pageBits} \<inter> kernel_data_refs = {}")
661   prefer 2
662   apply (drule(1) valid_global_refsD')
663   apply (clarsimp simp: Int_commute pageBits_def X64SmallPageBits_def
664                         intvl_range_conv[where bits=pageBits, unfolded pageBits_def word_bits_def,
665                                          simplified])
666  apply (intro conjI impI)
667       apply (erule is_aligned_no_wrap')
668       apply (clarsimp simp: X64SmallPageBits_def pageBits_def)
669      apply (erule is_aligned_weaken, simp add:pageBits_def)
670     apply (simp add: is_aligned_def X64SmallPageBits_def bit_simps)
671    apply (simp add: region_actually_is_bytes_dom_s pageBits_def X64SmallPageBits_def)
672   apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def
673                         kernel_data_refs_domain_eq_rotate
674                         size_of_def pageBits_def
675                  elim!: ptr_retyp_htd_safe_neg)
676  apply clarsimp
677  apply (cut_tac helper[rule_format])
678   prefer 2
679   apply fastforce
680  apply (subst data_map_insert_def[symmetric])
681  apply (erule iffD1[OF rf_sr_upd, rotated -1])
682   apply simp_all
683  apply (simp add: hrs_htd_update_def hrs_mem_update_def split_def)
684  apply (simp add: pageBits_def X64SmallPageBits_def ptr_retyps_gen_def
685              del: replicate_numeral)
686  done
687qed
688
689lemma cmap_relation_ccap_relation:
690  "\<lbrakk>cmap_relation (ctes_of s) (cslift s') cte_Ptr ccte_relation;ctes_of s p = Some cte; cteCap cte = cap\<rbrakk>
691    \<Longrightarrow> ccap_relation cap
692    (h_val (hrs_mem (t_hrs_' (globals s'))) (cap_Ptr &(cte_Ptr p\<rightarrow>[''cap_C''])))"
693  apply (erule(1) cmap_relationE1)
694  apply (clarsimp simp add: typ_heap_simps' ccte_relation_ccap_relation)
695  done
696
697lemma ccorres_move_Guard_Seq_strong:
698  "\<lbrakk>\<forall>s s'. (s, s') \<in> sr \<and> P s \<and> P' s' \<longrightarrow> G' s';
699    ccorres_underlying sr \<Gamma> r xf arrel axf A C' hs a (c;;d) \<rbrakk>
700  \<Longrightarrow> ccorres_underlying sr \<Gamma> r xf arrel axf (A and P) {s. P' s \<and> (G' s \<longrightarrow> s \<in> C')} hs a
701    (Guard F (Collect G') c;;
702     d)"
703  apply (rule ccorres_guard_imp2, erule ccorres_move_Guard_Seq)
704   apply assumption
705  apply auto
706  done
707
708lemma ghost_assertion_data_get_gs_clear_region:
709  "gs_get_assn proc (gs_clear_region addr n gs) = gs_get_assn proc gs"
710  by (clarsimp simp: ghost_assertion_data_get_def gs_clear_region_def)
711
712lemma ghost_assertion_size_logic_flex:
713  "unat (sz :: machine_word) \<le> gsMaxObjectSize s
714    \<Longrightarrow> (s, \<sigma>') \<in> rf_sr
715    \<Longrightarrow> gs_get_assn cap_get_capSizeBits_'proc (ghost'state_' (globals \<sigma>'))
716        = gs_get_assn cap_get_capSizeBits_'proc gs
717    \<Longrightarrow> gs_get_assn cap_get_capSizeBits_'proc gs = 0 \<or>
718            sz \<le> gs_get_assn cap_get_capSizeBits_'proc gs"
719  by (metis ghost_assertion_size_logic)
720
721lemma canonical_pspace_strg:
722  "valid_pspace' s \<longrightarrow> pspace_canonical' s"
723  by (simp add: valid_pspace'_def)
724
725lemma performASIDControlInvocation_ccorres:
726notes replicate_numeral[simp del]
727shows
728  "ccorres (K (K \<bottom>) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
729             (invs'
730               and ct_active'
731               and sch_act_simple
732               and cte_wp_at' (\<lambda>cte. cteCap cte = capability.UntypedCap isdev frame pageBits idx) parent
733               and (\<lambda>s. descendants_of' parent (ctes_of s) = {})
734               and ex_cte_cap_to' parent
735               and (\<lambda>_. base \<le> mask asid_bits \<and> is_aligned base asid_low_bits))
736             (UNIV \<inter> {s. frame_' s = Ptr frame}
737                   \<inter> {s. slot_' s = cte_Ptr slot}
738                   \<inter> {s. parent_' s = cte_Ptr parent}
739                   \<inter> {s. asid_base_' s = base}) []
740       (liftE (performASIDControlInvocation (MakePool frame slot parent base)))
741       (Call performASIDControlInvocation_'proc)"
742  apply (rule ccorres_gen_asm)
743  apply (simp only: liftE_liftM ccorres_liftM_simp)
744  apply (cinit lift: frame_' slot_' parent_' asid_base_')
745   apply (rule_tac P="is_aligned frame pageBits \<and> canonical_address frame \<and> frame \<in> kernel_mappings" in ccorres_gen_asm)
746   apply (rule ccorres_rhs_assoc2)
747   apply (rule ccorres_split_nothrow[where c="Seq c c'" for c c'])
748       apply (fold pageBits_def)[1]
749       apply (simp add: hrs_htd_update)
750       apply (rule deleteObjects_ccorres)
751      apply ceqv
752      apply (rule ccorres_rhs_assoc2)
753      apply (rule ccorres_abstract_cleanup)
754      apply (rule ccorres_symb_exec_l)
755        apply (rule_tac P = "rva = (capability.UntypedCap isdev frame pageBits idx)" in ccorres_gen_asm)
756        apply (simp add: hrs_htd_update del:fun_upd_apply)
757        apply (rule ccorres_split_nothrow)
758
759            apply (rule_tac cap'="UntypedCap isdev frame pageBits idx" in updateFreeIndex_ccorres)
760            apply (rule allI, rule conseqPre, vcg)
761            apply (rule subsetI, clarsimp simp: typ_heap_simps' pageBits_def isCap_simps)
762            apply (frule ccte_relation_ccap_relation, clarsimp)
763            apply (frule cap_get_tag_isCap_unfolded_H_cap)
764            apply (clarsimp simp: isCap_simps cap_lift_untyped_cap
765                                  cap_to_H_simps cap_untyped_cap_lift_def
766                                  ccap_relation_def modify_map_def
767                                  fun_eq_iff
768                          dest!: word_unat.Rep_inverse' split: if_split)
769            apply (rule exI, strengthen refl)
770            apply (case_tac cte', simp add: cap_lift_untyped_cap max_free_index_def mask_def)
771            apply (simp add: mex_def meq_def del: split_paired_Ex)
772            apply blast
773           apply ceqv
774          apply ccorres_remove_UNIV_guard
775          apply csymbr
776          apply (rule ccorres_Guard_Seq[where F=ShiftError])+
777          apply (simp del: Collect_const, simp add: framesize_to_H_def)
778          apply (ctac (c_lines 2) add:createObjects_asidpool_ccorres[where idx="max_free_index pageBits"]
779                              pre del: ccorres_Guard_Seq)
780            apply csymbr
781             apply (ctac (no_vcg) add: cteInsert_ccorres)
782              apply (simp add: ccorres_seq_skip del: fun_upd_apply)
783              apply (rule ccorres_assert)
784              apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
785              apply (rule allI, rule conseqPre, vcg)
786              apply (clarsimp simp: gets_def modify_def return_def put_def get_def bind_def
787                             asid_shiftr_low_bits_less
788                   simp del: fun_upd_apply Collect_const)
789              apply (clarsimp simp: word_sle_def word_sless_def Kernel_C.asidLowBits_def simp del: fun_upd_apply)
790              apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def cmachine_state_relation_def  simp del: fun_upd_apply)
791              apply (clarsimp simp: carch_state_relation_def carch_globals_def simp del: fun_upd_apply)
792              apply (simp add: asid_high_bits_of_def fun_upd_def[symmetric]
793                          del: fun_upd_apply)
794              apply (erule array_relation_update)
795                apply (simp add: unat_ucast)
796                apply (subst Divides.mod_less, simp)
797                 apply (drule leq_asid_bits_shift)
798                 apply (simp add: asid_high_bits_def mask_def word_le_nat_alt)
799                apply simp
800               apply (clarsimp simp: option_to_ptr_def option_to_0_def)
801              apply (clarsimp simp: asid_high_bits_def)
802             apply wp+
803            apply (strengthen valid_pspace_mdb' vp_strgs' valid_pspace_valid_objs' canonical_pspace_strg)
804            apply (clarsimp simp: is_simple_cap'_def isCap_simps conj_comms placeNewObject_def2)
805            apply (wp createObjects_valid_pspace'[where ty="Inl (KOArch (KOASIDPool f))" and sz = pageBits for f]
806                      createObjects_cte_wp_at'[where sz = pageBits]
807                   | simp add:makeObjectKO_def objBits_simps archObjSize_def range_cover_full
808                   | simp add: bit_simps untypedBits_defs)+
809            apply (clarsimp simp:valid_cap'_def capAligned_def)
810            apply (wp createObject_typ_at')
811           apply clarsimp
812           apply vcg
813          apply (clarsimp simp:conj_comms objBits_simps archObjSize_def |
814                 strengthen valid_pspace_mdb' vp_strgs' invs_valid_pspace'
815                 valid_pspace_valid_objs' invs_valid_global'
816                 invs_urz)+
817          apply (wp updateFreeIndex_forward_invs'
818                    updateFreeIndex_caps_no_overlap''[where sz=pageBits]
819                    updateFreeIndex_pspace_no_overlap'[where sz=pageBits]
820                    updateFreeIndex_caps_overlap_reserved
821                    updateFreeIndex_cte_wp_at
822                    )
823          apply (strengthen exI[where x=parent])
824          apply (wp updateFreeIndex_cte_wp_at)
825         apply clarsimp
826         apply vcg
827        apply wp
828       apply clarsimp
829       apply (wp getSlotCap_wp)
830      apply clarsimp
831     apply (rule_tac Q="\<lambda>_. cte_wp_at' ((=) (UntypedCap isdev frame pageBits idx) o cteCap) parent
832                           and (\<lambda>s. descendants_range_in' {frame..frame + (2::machine_word) ^ pageBits - (1::machine_word)} parent (ctes_of s))
833                           and pspace_no_overlap' frame pageBits
834                           and invs'
835                           and ct_active'"
836                 in hoare_post_imp)
837     apply (clarsimp simp: cte_wp_at_ctes_of invs_valid_objs' range_cover_full word_bits_conv
838            pageBits_def max_free_index_def asid_low_bits_def)
839     apply (case_tac cte,clarsimp simp:invs_valid_pspace')
840     apply (frule(1) ctes_of_valid_cap'[OF _ invs_valid_objs'])
841     apply (clarsimp simp:valid_cap'_def asid_low_bits_def invs_urz)
842     apply (strengthen descendants_range_in_subseteq'[mk_strg I E] refl)
843     apply (simp add: untypedBits_defs word_size_bits_def asid_wf_def)
844     apply (intro context_conjI)
845         apply (simp add:is_aligned_def)
846        apply (rule le_m1_iff_lt[THEN iffD1,THEN iffD1])
847         apply (simp add:asid_bits_def)
848        apply (simp add:mask_def)
849       apply simp
850      apply (rule descendants_range_caps_no_overlapI'[where d=isdev and cref = parent])
851        apply simp
852        apply (fastforce simp:cte_wp_at_ctes_of is_aligned_neg_mask_eq)
853       apply (clarsimp simp:is_aligned_neg_mask_eq)
854     apply (clarsimp dest!: upto_intvl_eq)
855    apply (wp deleteObjects_cte_wp_at'[where d=isdev and idx = idx and p = parent]
856              deleteObjects_descendants[where d=isdev and p = parent and idx = idx]
857              deleteObjects_invs'[where d=isdev and p = parent and idx = idx]
858              Detype_R.deleteObjects_descendants[where p = parent and idx = idx]
859              deleteObjects_ct_active'[where d=isdev and cref = parent and idx = idx])
860   apply clarsimp
861   apply vcg
862  apply (clarsimp simp: conj_comms invs_valid_pspace')
863  apply (frule cte_wp_at_valid_objs_valid_cap', fastforce)
864  apply (clarsimp simp: valid_cap'_def capAligned_def cte_wp_at_ctes_of untypedBits_defs
865                        descendants_range'_def2 empty_descendants_range_in')
866  apply (intro conjI; (rule refl)?)
867        apply clarsimp
868        apply (drule(1) cte_cap_in_untyped_range[where ptr = frame])
869             apply (fastforce simp: cte_wp_at_ctes_of)
870            apply assumption+
871         apply fastforce
872        apply simp
873       apply assumption
874      apply simp
875     apply simp
876    apply (erule empty_descendants_range_in')
877   apply (fastforce)
878  apply (erule(1) cmap_relationE1[OF cmap_relation_cte])
879  apply (clarsimp simp: typ_heap_simps cap_get_tag_isCap dest!: ccte_relation_ccap_relation)
880  apply (clarsimp simp: is_aligned_mask max_free_index_def pageBits_def)
881  apply (rule conjI, rule UNIV_I)?
882  apply clarsimp?
883  apply (erule_tac s = sa in rf_sr_ctes_of_cliftE)
884   apply assumption
885  apply (frule_tac s = sa in rf_sr_cte_relation)
886   apply simp+
887  apply (clarsimp simp:typ_heap_simps' region_is_bytes'_def[where sz=0])
888  apply (frule ccte_relation_ccap_relation)
889  apply (clarsimp simp: cap_get_tag_isCap hrs_htd_update)
890  apply (clarsimp simp: hrs_htd_update_def split_def
891                         pageBits_def
892                   split: if_split)
893  apply (clarsimp simp: X64SmallPageBits_def word_sle_def is_aligned_mask[symmetric]
894                        ghost_assertion_data_get_gs_clear_region[unfolded o_def])
895  apply (subst ghost_assertion_size_logic_flex[unfolded o_def, rotated])
896     apply assumption
897    apply (simp add: ghost_assertion_data_get_gs_clear_region[unfolded o_def])
898   apply (drule valid_global_refsD_with_objSize, clarsimp)+
899   apply (clarsimp simp: isCap_simps dest!: ccte_relation_ccap_relation)
900  apply (cut_tac ptr=frame and bits=12
901                   and htd="typ_region_bytes frame 12 (hrs_htd (t_hrs_' (globals s')))"
902                          in typ_region_bytes_actually_is_bytes)
903   apply (simp add: hrs_htd_update)
904  apply (clarsimp simp: region_actually_is_bytes'_def[where len=0])
905  apply (intro conjI)
906         apply (clarsimp elim!:is_aligned_weaken)
907        apply (clarsimp simp: vm_page_size_defs)
908       apply (simp add:is_aligned_def)
909      apply (erule is_aligned_no_wrap',simp)
910     apply (simp add: hrs_htd_def)
911    apply (clarsimp simp: framesize_to_H_def pageBits_def)
912   apply (drule region_actually_is_bytes_dom_s[OF _ order_refl])
913   apply (simp add: hrs_htd_def split_def)
914  apply (clarsimp simp: ccap_relation_def)
915  apply (clarsimp simp: cap_asid_pool_cap_lift)
916  apply (clarsimp simp: cap_to_H_def)
917  apply (clarsimp simp: asid_bits_def)
918  apply (drule word_le_mask_eq, simp)
919  apply (simp add: asid_bits_def sign_extend_canonical_address)
920  done
921
922lemmas performX64MMUInvocations
923    = ccorres_invocationCatch_Inr performInvocation_def
924      X64_H.performInvocation_def performX64MMUInvocation_def
925      liftE_bind_return_bindE_returnOk
926
927lemma slotcap_in_mem_PML4:
928  "\<lbrakk> slotcap_in_mem cap slot (ctes_of s); (s, s') \<in> rf_sr \<rbrakk>
929       \<Longrightarrow> \<exists>v. cslift s' (cte_Ptr slot) = Some v
930           \<and> (cap_get_tag (cte_C.cap_C v) = scast cap_pml4_cap)
931              = (isArchObjectCap cap \<and> isPML4Cap (capCap cap))
932           \<and> ccap_relation cap (cte_C.cap_C v)"
933  apply (clarsimp simp: slotcap_in_mem_def)
934  apply (erule(1) cmap_relationE1[OF cmap_relation_cte])
935  apply (clarsimp dest!: ccte_relation_ccap_relation)
936  apply (simp add: cap_get_tag_isCap_ArchObject2)
937  done
938
939declare if_split [split del]
940
941lemma isValidNativeRoot_spec:
942  "\<forall>s.
943    \<Gamma> \<turnstile> {\<sigma>. s = \<sigma> \<and> True}
944            Call isValidNativeRoot_'proc
945         {t. \<forall>cap. ccap_relation cap (cap_' s) \<longrightarrow> ret__unsigned_long_' t = from_bool (isArchObjectCap cap \<and> isPML4Cap (capCap cap) \<and>
946                                        capPML4MappedASID (capCap cap) \<noteq> None)}"
947  apply (vcg, clarsimp)
948  apply (rule conjI, clarsimp simp: from_bool_def case_bool_If if_1_0_0
949                             split: if_split)
950   apply (rule conjI; clarsimp simp: cap_pml4_cap_lift)
951   apply (erule ccap_relationE, clarsimp simp: cap_to_H_def isCap_simps to_bool_def
952                                        split: if_split_asm)
953   apply (erule ccap_relationE, clarsimp simp: isCap_simps cap_to_H_def)
954  by (clarsimp simp: from_bool_def case_bool_If isCap_simps
955              dest!: cap_get_tag_isCap_unfolded_H_cap
956              split: if_split)
957
958definition
959  get_capMappedASID_CL :: "cap_CL option \<Rightarrow> machine_word"
960where
961  "get_capMappedASID_CL \<equiv> \<lambda>cap. case cap of
962     Some (Cap_pml4_cap c) \<Rightarrow> capPML4MappedASID_CL c
963   | Some (Cap_pdpt_cap c) \<Rightarrow> capPDPTMappedASID_CL c
964   | Some (Cap_page_directory_cap c) \<Rightarrow> capPDMappedASID_CL c"
965
966lemma cap_get_capMappedASID_spec:
967  "\<forall>s. \<Gamma> \<turnstile> \<lbrace>s. cap_get_tag \<acute>cap = scast cap_pml4_cap \<or> cap_get_tag \<acute>cap = scast cap_pdpt_cap
968               \<or> cap_get_tag \<acute>cap = scast cap_page_directory_cap\<rbrace>
969       \<acute>ret__unsigned_long :== PROC cap_get_capMappedASID(\<acute>cap)
970       \<lbrace>\<acute>ret__unsigned_long = (get_capMappedASID_CL (cap_lift \<^bsup>s\<^esup>cap))\<rbrace>"
971  apply vcg
972  apply (clarsimp simp: get_capMappedASID_CL_def)
973  apply (intro conjI impI;
974         clarsimp simp: cap_lifts
975                        cap_lift_asid_control_cap
976                        cap_lift_irq_control_cap cap_lift_null_cap
977                        Kernel_C.asidLowBits_def asid_low_bits_def
978                        word_sle_def Let_def mask_def
979                        isZombieTCB_C_def ZombieTCB_C_def
980                        cap_lift_domain_cap cap_get_tag_scast
981                        objBits_defs wordRadix_def
982                        c_valid_cap_def cl_valid_cap_def
983                 dest!: sym [where t = "ucast (cap_get_tag cap)" for cap])
984  apply (auto simp: cap_pml4_cap_lift[symmetric] cap_pdpt_cap_lift[symmetric]
985                    cap_page_directory_cap_lift[symmetric])
986  done
987
988lemma addrFromPPtr_mask_middle_pml4ShiftBits:
989  "\<lbrakk>is_aligned p pageBits; p \<in> kernel_mappings\<rbrakk> \<Longrightarrow>
990   addrFromPPtr p && (mask pml4ShiftBits << pageBits) = addrFromPPtr p"
991  apply (clarsimp simp: mask_shiftl_decompose kernel_mappings_def)
992  apply (subst word_bool_alg.conj.assoc[symmetric])
993  apply (subst is_aligned_neg_mask_eq)
994   apply (rule aligned_already_mask)
995   apply (erule is_aligned_addrFromPPtr)
996  apply (clarsimp simp: addrFromPPtr_def X64.pptrBase_def pptr_base_def bit_simps)
997  apply (clarsimp simp: mask_def)
998  apply (word_bitwise, clarsimp)
999  done
1000
1001(* FIXME: move *)
1002lemma obj_at_kernel_mappings':
1003  "\<lbrakk>pspace_in_kernel_mappings' s; obj_at' P p s\<rbrakk> \<Longrightarrow>
1004   p \<in> kernel_mappings"
1005  by (clarsimp simp: pspace_in_kernel_mappings'_def obj_at'_def dom_def)
1006
1007lemma decodeX64PageTableInvocation_ccorres:
1008  "\<lbrakk>interpret_excaps extraCaps' = excaps_map extraCaps; isPageTableCap cp\<rbrakk> \<Longrightarrow>
1009   ccorres
1010       (intr_and_se_rel \<currency> dc)
1011       (liftxf errstate id (K ()) ret__unsigned_long_')
1012       (invs' and (\<lambda>s. ksCurThread s = thread) and ct_active' and sch_act_simple
1013              and (excaps_in_mem extraCaps \<circ> ctes_of)
1014              and cte_wp_at' (diminished' (ArchObjectCap cp) \<circ> cteCap) slot
1015              and valid_cap' (ArchObjectCap cp)
1016              and (\<lambda>s. \<forall>v \<in> set extraCaps. ex_cte_cap_wp_to' isCNodeCap (snd v) s)
1017              and sysargs_rel args buffer)
1018       (UNIV \<inter> {s. invLabel_' s = label}
1019             \<inter> {s. unat (length___unsigned_long_' s) = length args}
1020             \<inter> {s. cte_' s = cte_Ptr slot}
1021             \<inter> {s. excaps_' s = extraCaps'}
1022             \<inter> {s. ccap_relation (ArchObjectCap cp) (cap_' s)}
1023             \<inter> {s. buffer_' s = option_to_ptr buffer})
1024       hs
1025       (decodeX64MMUInvocation label args cptr slot cp extraCaps
1026              >>= invocationCatch thread isBlocking isCall InvokeArchObject)
1027       (Call decodeX86PageTableInvocation_'proc)"
1028   (is "_ \<Longrightarrow> _ \<Longrightarrow> ccorres _ _ ?pre _ _ _ _")
1029  supply Collect_const[simp del] if_cong[cong]
1030  apply (clarsimp simp only: isCap_simps)
1031  apply (cinit' lift: invLabel_' length___unsigned_long_' cte_' excaps_' cap_' buffer_'
1032                simp: decodeX64MMUInvocation_def invocation_eq_use_types decodeX64PageTableInvocation_def)
1033   apply (simp add: Let_def isCap_simps if_to_top_of_bind
1034               cong: StateSpace.state.fold_congs globals.fold_congs)
1035   apply (rule ccorres_Cond_rhs_Seq)
1036    apply (rule ccorres_split_throws)
1037     apply (simp add: liftE_bindE bind_assoc)
1038     apply (rule ccorres_symb_exec_l[OF _ getCTE_inv _ empty_fail_getCTE])
1039      apply (rule ccorres_rhs_assoc)+
1040      apply (ctac add: isFinalCapability_ccorres)
1041        apply (simp add: unlessE_def if_to_top_of_bind if_to_top_of_bindE ccorres_seq_cond_raise)
1042        apply (rule ccorres_cond2'[where R=\<top>])
1043          apply (clarsimp simp: from_bool_0)
1044         apply (simp add: throwError_bind invocationCatch_def)
1045         apply (rule syscall_error_throwError_ccorres_n)
1046         apply (simp add: syscall_error_to_H_cases)
1047        apply (simp add: returnOk_bind bindE_assoc performX64MMUInvocations)
1048        apply (ctac add: setThreadState_ccorres)
1049          apply (ctac add: performPageTableInvocationUnmap_ccorres)
1050             apply (rule ccorres_alternative2)
1051             apply (rule ccorres_return_CE, simp+)[1]
1052            apply (rule ccorres_inst[where P=\<top> and P'=UNIV], simp)
1053           apply wp
1054          apply simp
1055          apply (vcg exspec=performX86PageTableInvocationUnmap_modifies)
1056         apply (wp sts_invs_minor')
1057        apply simp
1058        apply (vcg exspec=setThreadState_modifies)
1059       apply (wp hoare_drop_imps isFinalCapability_inv)
1060      apply simp
1061      apply (vcg exspec=isFinalCapability_modifies)
1062     apply (wp getCTE_wp)
1063    apply (vcg exspec=performX86PageTableInvocationUnmap_modifies exspec=isFinalCapability_modifies)
1064   apply simp
1065   apply (rule ccorres_Cond_rhs_Seq)
1066    apply (rule ccorres_equals_throwError)
1067     apply (simp split: invocation_label.split arch_invocation_label.split
1068                   add: throwError_bind invocationCatch_def)
1069     apply fastforce
1070    apply (rule syscall_error_throwError_ccorres_n)
1071    apply (simp add: syscall_error_to_H_cases)
1072   apply simp
1073   apply csymbr
1074   apply (rule ccorres_Cond_rhs_Seq)
1075    apply (simp add: word_less_nat_alt throwError_bind invocationCatch_def)
1076    apply (rule ccorres_cond_true_seq)
1077    apply (rule ccorres_equals_throwError)
1078     apply (simp add: throwError_bind split: list.split)
1079     apply fastforce
1080    apply (rule syscall_error_throwError_ccorres_n)
1081    apply (simp add: syscall_error_to_H_cases)
1082   apply csymbr
1083   apply (simp add: interpret_excaps_test_null excaps_map_def)
1084   apply (rule ccorres_Cond_rhs_Seq)
1085    apply (simp add: throwError_bind invocationCatch_def)
1086    apply (rule ccorres_equals_throwError)
1087     apply (fastforce simp: throwError_bind split: list.split)
1088    apply (rule syscall_error_throwError_ccorres_n)
1089    apply (simp add: syscall_error_to_H_cases)
1090   apply (simp add: list_case_If2 split_def
1091                    word_less_nat_alt length_ineq_not_Nil Let_def
1092                    whenE_bindE_throwError_to_if if_to_top_of_bind)
1093   apply csymbr
1094   apply (rule_tac Q=\<top> and Q'=\<top> in ccorres_if_cond_throws[rotated -1])
1095      apply vcg
1096     apply clarsimp
1097     apply (frule cap_get_tag_isCap_unfolded_H_cap)
1098     apply (clarsimp simp: cap_lift_page_table_cap cap_page_table_cap_lift_def
1099                           cap_to_H_def
1100                    elim!: ccap_relationE split: if_split)
1101     apply (simp add: to_bool_def)
1102    apply (simp add: throwError_bind invocationCatch_def)
1103    apply (rule syscall_error_throwError_ccorres_n)
1104    apply (simp add: syscall_error_to_H_cases)
1105   apply (simp add: cap_case_PML4Cap2 split_def)
1106   apply (rule ccorres_add_return)
1107   apply (ctac add: getSyscallArg_ccorres_foo[where args=args and n=0 and buffer=buffer])
1108     apply csymbr
1109     apply (rule ccorres_add_return)
1110     apply (ctac add: getSyscallArg_ccorres_foo[where args=args and n=1 and buffer=buffer])
1111       apply csymbr
1112       apply (rule ccorres_add_return)
1113       apply (rule_tac r'="(\<lambda>rv _ rv'. ((cap_get_tag rv' = scast cap_pml4_cap)
1114                                          = (isArchObjectCap rv \<and> isPML4Cap (capCap rv)))
1115                                     \<and> (ccap_relation rv rv')) (fst (extraCaps ! 0))"
1116                and xf'=vspaceCap_' in ccorres_split_nothrow)
1117           apply (rule ccorres_from_vcg[where P="excaps_in_mem extraCaps \<circ> ctes_of" and P'=UNIV])
1118           apply (rule allI, rule conseqPre, vcg)
1119           apply (clarsimp simp: excaps_in_mem_def return_def neq_Nil_conv)
1120           apply (drule(1) slotcap_in_mem_PML4)
1121           apply (frule interpret_excaps_eq[rule_format, where n=0], simp)
1122           apply (clarsimp simp: typ_heap_simps' mask_def)
1123           apply (rename_tac rv' t t')
1124          apply (simp add: word_sless_def word_sle_def)
1125          apply ceqv
1126         apply csymbr
1127         apply clarsimp
1128         apply (rule ccorres_Cond_rhs_Seq)
1129          apply ccorres_rewrite
1130          apply (clarsimp simp: from_bool_0)
1131          apply (rule_tac P="isArchObjectCap (fst (extraCaps ! 0)) \<and>
1132                             isPML4Cap (capCap (fst (extraCaps ! 0)))"
1133                          in ccorres_cases)
1134           apply (clarsimp simp: hd_conv_nth throwError_bind invocationCatch_def cong: if_cong)
1135           apply (rule syscall_error_throwError_ccorres_n[simplified dc_def id_def o_def])
1136           apply (simp add: syscall_error_to_H_cases)
1137          apply (clarsimp simp: hd_conv_nth throwError_bind invocationCatch_def cong: if_cong)
1138          apply (rule syscall_error_throwError_ccorres_n[simplified dc_def id_def o_def])
1139          apply (simp add: syscall_error_to_H_cases)
1140         apply (simp add: hd_conv_nth)
1141         apply csymbr
1142         apply csymbr
1143         apply csymbr
1144         apply (simp add: case_option_If2 if_to_top_of_bind if_to_top_of_bindE hd_conv_nth)
1145         apply (rule ccorres_if_cond_throws[rotated -1, where Q=\<top> and Q'=\<top>])
1146            apply vcg
1147           apply (fastforce simp: user_vtop_def X64.pptrUserTop_def bit_simps
1148                                  word_le_nat_alt mask_def hd_conv_nth
1149                                  length_ineq_not_Nil)
1150          apply (simp add: throwError_bind invocationCatch_def)
1151          apply (rule syscall_error_throwError_ccorres_n)
1152          apply (simp add: syscall_error_to_H_cases)
1153         apply (simp add: lookupError_injection invocationCatch_use_injection_handler
1154                          injection_bindE[OF refl refl] injection_handler_If
1155                          injection_handler_throwError injection_liftE[OF refl]
1156                          injection_handler_returnOk bindE_assoc cap_pml4_cap_lift get_capMappedASID_CL_def
1157                    cong: if_cong)
1158         apply csymbr
1159         apply (ctac add: ccorres_injection_handler_csum1 [OF ccorres_injection_handler_csum1,
1160                                                           OF findVSpaceForASID_ccorres])
1161            apply (simp add: Collect_False if_to_top_of_bindE)
1162            apply (rule ccorres_if_cond_throws[rotated -1, where Q=\<top> and Q'=\<top>])
1163               apply vcg
1164              apply (clarsimp simp: isCap_simps get_capPtr_CL_def)
1165              apply (frule cap_get_tag_isCap_unfolded_H_cap)
1166              apply (erule_tac v="rv'" in ccap_relationE, clarsimp simp: cap_to_H_def)
1167             apply (rule syscall_error_throwError_ccorres_n)
1168             apply (simp add: syscall_error_to_H_cases)
1169            apply (simp add: bindE_assoc,
1170                   simp add: liftE_bindE)
1171            apply (ctac add: ccorres_injection_handler_csum1 [OF ccorres_injection_handler_csum1,
1172                                                              OF lookupPDSlot_ccorres])
1173            apply (rename_tac pd_slot lookup_pd_res)
1174               apply (rule ccorres_pre_getObject_pde)
1175               apply (rename_tac pde)
1176               apply (rule ccorres_cond_false_seq)
1177               apply ccorres_rewrite
1178               apply clarsimp
1179               apply (rename_tac pml4_mapped_asid)
1180               apply (rule ccorres_rhs_assoc2, rule ccorres_rhs_assoc2, rule ccorres_rhs_assoc2)
1181               apply (rule_tac xf'=ret__int_' and R="ko_at' pde pd_slot" and R'=UNIV and
1182                               val="from_bool (pde \<noteq> InvalidPDE)"
1183                               in ccorres_symb_exec_r_known_rv)
1184                  apply vcg
1185                  apply (clarsimp simp: from_bool_0)
1186                  apply (erule cmap_relationE1[OF rf_sr_cpde_relation], erule ko_at_projectKO_opt)
1187                  apply (clarsimp simp: typ_heap_simps from_bool_eq_if)
1188                  apply (auto simp: cpde_relation_def Let_def pde_pde_pt_lift_def
1189                                    pde_pde_pt_lift pde_tag_defs pde_pde_large_lift_def
1190                                    pde_lift_def from_bool_def case_bool_If
1191                             split: pde.split_asm if_splits)[1]
1192                 apply ceqv
1193                apply clarsimp
1194                apply (rule ccorres_Cond_rhs_Seq)
1195                 apply (clarsimp simp: throwError_bind invocationCatch_def
1196                                       injection_handler_throwError)
1197                 apply ccorres_rewrite
1198                 apply (rule syscall_error_throwError_ccorres_n)
1199                 apply (clarsimp simp: syscall_error_to_H_cases)
1200                apply (clarsimp simp: injection_handler_returnOk)
1201                apply (simp add: injection_handler_returnOk performX64MMUInvocations bindE_assoc)
1202                apply csymbr
1203                apply csymbr
1204                apply csymbr
1205                apply csymbr
1206                apply csymbr
1207                apply csymbr
1208                apply (ctac add: setThreadState_ccorres)
1209                  apply (ctac add: performPageTableInvocationMap_ccorres)
1210                     apply (rule ccorres_alternative2)
1211                     apply (rule ccorres_return_CE, simp+)[1]
1212                    apply (rule ccorres_inst[where P=\<top> and P'=UNIV], simp)
1213                   apply wpsimp+
1214                  apply (vcg exspec=performX86PageTableInvocationMap_modifies)
1215                 apply wpsimp+
1216                apply (vcg exspec=setThreadState_modifies)
1217               apply (simp add: get_capPtr_CL_def)
1218              apply (rule_tac s=pml4_mapped_asid and
1219                              t="the (capPML4MappedASID (capCap (fst (extraCaps ! 0))))"
1220                              in subst,
1221                     fastforce)
1222               apply vcg
1223              apply (rename_tac lookup_pd_ret)
1224              apply clarsimp
1225              apply (rule_tac P'="{s. errstate s = lookup_pd_ret}" in ccorres_from_vcg_split_throws[where P=\<top>])
1226               apply vcg
1227              apply (rule conseqPre, vcg)
1228              apply (clarsimp simp: throwError_def return_def syscall_error_rel_def
1229                                    syscall_error_to_H_cases exception_defs false_def)
1230              apply (erule lookup_failure_rel_fault_lift[rotated])
1231              apply (clarsimp simp: exception_defs)
1232             apply clarsimp
1233             apply (wp injection_wp[OF refl] hoare_drop_imps)
1234            apply (clarsimp simp: get_capPtr_CL_def)
1235            apply (rename_tac pml4_mapped_asid)
1236            apply (rule_tac s=pml4_mapped_asid and
1237                            t="the (capPML4MappedASID (capCap (fst (extraCaps ! 0))))"
1238                            in subst,
1239                   fastforce)
1240            apply (vcg exspec=lookupPDSlot_modifies)
1241           apply clarsimp
1242           apply (rule_tac P'="{s. errstate s = find_ret}" in ccorres_from_vcg_split_throws[where P=\<top>])
1243            apply vcg
1244           apply simp
1245           apply (rule conseqPre, vcg)
1246           apply (clarsimp simp: throwError_def return_def syscall_error_rel_def
1247                                 syscall_error_to_H_cases exception_defs false_def)
1248           apply (erule lookup_failure_rel_fault_lift[rotated])
1249           apply (simp add: exception_defs)
1250          apply simp
1251          apply (wp injection_wp[OF refl] hoare_drop_imps)
1252         apply simp
1253         apply (vcg exspec=findVSpaceForASID_modifies)
1254        apply simp
1255        apply (rule_tac Q="\<lambda>a b. invs' b \<and> valid_cap' (fst (extraCaps ! 0)) b \<and> tcb_at' thread b \<and>
1256                                 sch_act_wf (ksSchedulerAction b) b \<and> cte_wp_at' (\<lambda>_. True) slot b"
1257                                 in hoare_strengthen_post)
1258         apply wp
1259        apply (clarsimp simp: isCap_simps invs_valid_objs' valid_cap'_def valid_tcb_state'_def
1260                              invs_arch_state' invs_no_0_obj')
1261       apply vcg
1262      apply wp
1263     apply simp
1264     apply (vcg exspec=getSyscallArg_modifies)
1265    apply wpsimp+
1266  apply (vcg exspec=getSyscallArg_modifies)
1267  apply (clarsimp simp: cte_wp_at_ctes_of excaps_map_def
1268                        word_sle_def word_sless_def bit_simps)
1269  (* X64PageTableUnmap *)
1270  apply (rule conjI)
1271   apply (auto simp: ct_in_state'_def isCap_simps valid_tcb_state'_def
1272              elim!: pred_tcb'_weakenE dest!: st_tcb_at_idle_thread')[1]
1273  (* X64PageTableMap *)
1274  apply (rule conjI)
1275   apply (clarsimp simp: isCap_simps)
1276   apply (subgoal_tac "s \<turnstile>' fst (extraCaps ! 0)")
1277    apply (clarsimp simp: sysargs_rel_to_n word_le_nat_alt
1278                          linorder_not_less)
1279    apply (clarsimp | drule length_le_helper)+
1280    apply (clarsimp simp: valid_cap'_def neq_Nil_conv
1281                          mask_add_aligned page_directory_at'_def
1282                          word_le_nat_alt[symmetric] bit_simps)
1283    apply (auto simp: ct_in_state'_def pred_tcb_at' mask_def valid_tcb_state'_def
1284               elim!: pred_tcb'_weakenE dest!: st_tcb_at_idle_thread')[1]
1285   apply (clarsimp simp: neq_Nil_conv excaps_in_mem_def
1286                         slotcap_in_mem_def)
1287   apply (auto dest: ctes_of_valid')[1]
1288  (* X64PageTableUnmap *)
1289  apply (rule conjI)
1290   apply (fastforce simp: rf_sr_ksCurThread "StrictC'_thread_state_defs"
1291                          mask_eq_iff_w2p word_size
1292                          ct_in_state'_def st_tcb_at'_def
1293                          word_sle_def word_sless_def
1294                          typ_heap_simps' bit_simps)
1295  (* X64PageTableMap *)
1296  apply (clarsimp simp: cap_get_tag_isCap_ArchObject isCap_simps
1297                        word_sle_def word_sless_def
1298                        word_less_nat_alt)
1299  apply (frule length_ineq_not_Nil)
1300  apply (frule cap_get_tag_isCap_unfolded_H_cap(14))
1301  apply (clarsimp simp: cap_lift_page_directory_cap hd_conv_nth
1302                        cap_lift_page_table_cap bit_simps
1303                        cap_page_directory_cap_lift_def
1304                        to_bool_def
1305                        typ_heap_simps' shiftl_t2n[where n=3] field_simps
1306                 elim!: ccap_relationE)
1307  apply (clarsimp simp: neq_Nil_conv[where xs=extraCaps]
1308                        excaps_in_mem_def slotcap_in_mem_def
1309                 dest!: sym[where s="ArchObjectCap cp" for cp])
1310  apply (cut_tac p="snd (hd extraCaps)" in ctes_of_valid', simp, clarsimp)
1311  apply (rule conjI[rotated], clarsimp simp: cap_tag_defs)
1312  apply (clarsimp simp: valid_cap_simps' rf_sr_ksCurThread)
1313  apply (rule context_conjI)
1314   subgoal by (clarsimp simp: get_capMappedASID_CL_def dest!: cap_to_H_PML4Cap)
1315  apply clarsimp
1316  apply (rule conjI, clarsimp simp: mask_def)
1317  apply clarsimp
1318  apply (rule conjI, clarsimp, rule conjI, clarsimp simp: pde_tag_defs, clarsimp)
1319   apply (rule conjI[rotated])
1320    apply (fastforce dest!: cap_lift_page_table_cap
1321                     intro!: is_aligned_addrFromPPtr[simplified bit_simps, simplified]
1322                     simp: vmsz_aligned'_def cap_to_H_simps cap_page_table_cap_lift_def bit_simps capAligned_def)
1323   apply clarsimp
1324   apply (rule conjI, clarsimp simp: ThreadState_Restart_def mask_def)
1325   apply (rule conjI)
1326    (* ccap_relation *)
1327    apply (clarsimp simp: ccap_relation_def map_option_Some_eq2 cap_page_table_cap_lift[THEN iffD1]
1328                          cap_to_H_simps asid_wf_def3[simplified asid_bits_def, simplified])
1329    apply (rule conjI[rotated],
1330           fastforce simp: sign_extend_canonical_address le_def[symmetric] mask_def word_bw_assocs
1331                           le_user_vtop_canonical_address[simplified user_vtop_def X64.pptrUserTop_def word_le_nat_alt])
1332    subgoal by (clarsimp dest!: cap_lift_page_table_cap simp: cap_page_table_cap_lift_def)
1333   (* cpde_relation *)
1334   apply (rule conjI, clarsimp simp: cpde_relation_def)
1335    apply (clarsimp simp: superuser_from_vm_rights_def writable_from_vm_rights_def
1336                          x64CacheDisabled_def attribsFromWord_def word_and_1 nth_shiftr
1337                   split: if_split)
1338    apply (clarsimp dest!: cap_lift_page_table_cap simp: cap_to_H_simps cap_page_table_cap_lift_def)
1339    apply (rule addrFromPPtr_mask_middle_pml4ShiftBits[simplified pml4ShiftBits_def bit_simps, simplified])
1340     subgoal by (fastforce simp: valid_cap_simps' capAligned_def bit_simps)
1341    apply (fastforce dest!: page_table_pte_atI'[where x=0, simplified bit_simps, simplified]
1342                     intro!: obj_at_kernel_mappings'
1343                     simp: typ_at_to_obj_at_arches)
1344   apply (frule (1) cap_lift_PML4Cap_Base)
1345   subgoal by (auto simp: cap_to_H_simps cap_pml4_cap_lift_def get_capPtr_CL_def
1346                    dest!: cap_to_H_PML4Cap_tag cap_lift_pml4_cap)
1347  (* the below proof duplicates some of the sections above *)
1348  apply (clarsimp simp: pde_tag_defs pde_get_tag_def word_and_1)
1349  apply safe
1350     apply (clarsimp simp: ThreadState_Restart_def mask_def)
1351    (* ccap_relation *)
1352    apply (clarsimp simp: ccap_relation_def map_option_Some_eq2 cap_page_table_cap_lift[THEN iffD1]
1353                          cap_to_H_simps asid_wf_def3[simplified asid_bits_def, simplified])
1354     apply (rule conjI[rotated],
1355            fastforce simp: sign_extend_canonical_address le_def[symmetric] mask_def word_bw_assocs
1356                            le_user_vtop_canonical_address[simplified user_vtop_def X64.pptrUserTop_def word_le_nat_alt])
1357     subgoal by (clarsimp dest!: cap_lift_page_table_cap simp: cap_page_table_cap_lift_def)
1358    (* cpde_relation *)
1359    apply (clarsimp simp: cpde_relation_def superuser_from_vm_rights_def writable_from_vm_rights_def
1360                          x64CacheDisabled_def attribsFromWord_def word_and_1 nth_shiftr
1361                    split: if_split)
1362    apply (clarsimp dest!: cap_lift_page_table_cap simp: cap_to_H_simps cap_page_table_cap_lift_def)
1363    apply (rule addrFromPPtr_mask_middle_pml4ShiftBits[simplified pml4ShiftBits_def bit_simps, simplified])
1364     subgoal by (fastforce simp: valid_cap_simps' capAligned_def bit_simps)
1365    apply (fastforce dest!: page_table_pte_atI'[where x=0, simplified bit_simps, simplified]
1366                     intro!: obj_at_kernel_mappings'
1367                     simp: typ_at_to_obj_at_arches)
1368   apply (frule (1) cap_lift_PML4Cap_Base)
1369   subgoal by (auto simp: cap_to_H_simps cap_pml4_cap_lift_def get_capPtr_CL_def
1370                    dest!: cap_to_H_PML4Cap_tag cap_lift_pml4_cap)
1371  apply (fastforce dest!: cap_lift_page_table_cap
1372                   intro!: is_aligned_addrFromPPtr[simplified bit_simps, simplified]
1373                   simp: vmsz_aligned'_def cap_to_H_simps cap_page_table_cap_lift_def bit_simps capAligned_def)
1374  done
1375
1376lemma checkVPAlignment_spec:
1377  "\<forall>s. \<Gamma>\<turnstile> \<lbrace>s. \<acute>sz < 3\<rbrace> Call checkVPAlignment_'proc
1378          {t. ret__unsigned_long_' t = from_bool
1379               (vmsz_aligned' (w_' s) (framesize_to_H (sz_' s)))}"
1380  apply (rule allI, rule conseqPre, vcg)
1381  apply (clarsimp simp: mask_eq_iff_w2p word_size)
1382  apply (rule conjI)
1383   apply (simp add: pageBitsForSize_def bit_simps split: vmpage_size.split)
1384  apply (simp add: from_bool_def vmsz_aligned'_def is_aligned_mask
1385                   mask_def split: if_split)
1386  done
1387
1388definition
1389  "createSafeMappingEntries ptr vaddr vsz vrights attr pd = doE
1390     entries \<leftarrow> createMappingEntries ptr vaddr vsz vrights attr pd;
1391     ensureSafeMapping entries;
1392     returnOk entries
1393   odE"
1394
1395lemma createSafeMappingEntries_fold:
1396  "(doE
1397     entries \<leftarrow> createMappingEntries ptr vaddr vsz vrights attr pd;
1398     _ \<leftarrow> ensureSafeMapping entries;
1399     f entries
1400   odE) = (createSafeMappingEntries ptr vaddr vsz vrights attr pd >>=E f)"
1401  by (simp add: createSafeMappingEntries_def bindE_assoc)
1402
1403definition
1404  ptr_range_to_list :: "('a :: c_type) ptr \<Rightarrow> machine_word \<Rightarrow> 'a ptr list"
1405where
1406 "ptr_range_to_list ptr lenV \<equiv>
1407    map (\<lambda>x. CTypesDefs.ptr_add ptr (of_nat x)) [0 ..< unat lenV]"
1408
1409definition
1410 "pte_range_relation xs pte_ran \<equiv>
1411     xs = map ptr_val (ptr_range_to_list (pte_range_C.base_C pte_ran)
1412                            (pte_range_C.length_C pte_ran))
1413          \<and> 1 \<le> pte_range_C.length_C pte_ran"
1414
1415definition
1416 "pde_range_relation xs pde_ran \<equiv>
1417     xs = map ptr_val (ptr_range_to_list (pde_range_C.base_C pde_ran)
1418                            (pde_range_C.length_C pde_ran))
1419          \<and> 1 \<le> pde_range_C.length_C pde_ran"
1420
1421definition
1422 "vm_attribs_relation attr attr' \<equiv>
1423       x86PATBit_CL (vm_attributes_lift attr') = from_bool (x64PAT attr)
1424     \<and> x86PCDBit_CL (vm_attributes_lift attr') = from_bool (x64CacheDisabled attr)
1425     \<and> x86PWTBit_CL (vm_attributes_lift attr') = from_bool (x64WriteThrough attr)"
1426
1427lemma framesize_from_H_eqs:
1428  "(framesize_from_H vsz = scast Kernel_C.X86_SmallPage) = (vsz = X64SmallPage)"
1429  "(framesize_from_H vsz = scast Kernel_C.X86_LargePage) = (vsz = X64LargePage)"
1430  "(framesize_from_H vsz = scast Kernel_C.X64_HugePage) = (vsz = X64HugePage)"
1431  by (simp add: framesize_from_H_def vm_page_size_defs split: vmpage_size.split)+
1432
1433lemma pde_get_tag_alt:
1434  "pde_lift v = Some pdeC
1435    \<Longrightarrow> pde_get_tag v = (case pdeC of Pde_pde_pt _ \<Rightarrow> scast pde_pde_pt
1436          | Pde_pde_large _ \<Rightarrow> scast pde_pde_large)"
1437  by (auto simp add: pde_lift_def Let_def split: if_split_asm)
1438
1439thm cpde_relation_def
1440
1441(* FIXME x64: this is going to be annoying because
1442 * you cannot determine validity just from pde_get_tag *)
1443lemma cpde_relation_pde_case:
1444  "cpde_relation pde cpde
1445     \<Longrightarrow> (case pde of InvalidPDE \<Rightarrow> P
1446             | PageTablePDE _ _ _ _ _ _ \<Rightarrow> Q
1447             | LargePagePDE _ _ _ _ _ _ _ _ _ \<Rightarrow> R)
1448         = (if pde_get_tag cpde = scast pde_pde_pt then
1449                 if (pde_pde_pt_CL.present_CL (pde_pde_pt_lift cpde) = 0) then P else Q
1450            else R)"
1451  by (clarsimp simp: cpde_relation_def Let_def pde_get_tag_alt
1452                     pde_tag_defs pde_pde_pt_lift_def
1453              split: X64_H.pde.split_asm)
1454
1455lemma ccorres_pre_getObject_pte:
1456  "(\<And>rv. ccorresG rf_sr \<Gamma> r xf (P rv) (P' rv) hs (f rv) c) \<Longrightarrow>
1457   ccorresG rf_sr \<Gamma> r xf (\<lambda>s. \<forall>pte. ko_at' pte p s \<longrightarrow> P pte s)
1458            {s. \<forall>pte pte'. cslift s (pte_Ptr p) = Some pte' \<and> cpte_relation pte pte' \<longrightarrow> s \<in> P' pte} hs
1459     (getObject p >>= f) c"
1460  apply (rule ccorres_guard_imp2)
1461   apply (rule ccorres_symb_exec_l)
1462      apply (rule_tac P="ko_at' rv p" in ccorres_cross_over_guard)
1463      apply assumption
1464     apply (wp getObject_inv loadObject_default_inv
1465               getPTE_wp empty_fail_getObject | simp)+
1466  apply clarsimp
1467  apply (erule cmap_relationE1[OF rf_sr_cpte_relation], erule ko_at_projectKO_opt)
1468  apply clarsimp
1469  done
1470
1471lemma ptr_add_uint_of_nat [simp]:
1472    "a  +\<^sub>p uint (of_nat b :: machine_word) = a  +\<^sub>p (int b)"
1473  by (clarsimp simp: CTypesDefs.ptr_add_def)
1474
1475declare int_unat[simp]
1476
1477lemma obj_at_pte_aligned:
1478  "obj_at' (\<lambda>a::X64_H.pte. True) ptr s ==> is_aligned ptr word_size_bits"
1479  apply (drule obj_at_ko_at')
1480  apply (clarsimp dest!:ko_at_is_aligned'
1481                  simp: objBits_simps archObjSize_def bit_simps
1482                  elim!: is_aligned_weaken)
1483  done
1484
1485(* FIXME x64: dont know what these are for yet *)
1486lemma addrFromPPtr_mask_5:
1487  "addrFromPPtr ptr && mask (5::nat) = ptr && mask (5::nat)"
1488  apply (simp add:addrFromPPtr_def X64.pptrBase_def)
1489  apply word_bitwise
1490  apply (simp add:mask_def)
1491  done
1492
1493lemma addrFromPPtr_mask_6:
1494  "addrFromPPtr ptr && mask (6::nat) = ptr && mask (6::nat)"
1495  apply (simp add:addrFromPPtr_def X64.pptrBase_def)
1496  apply word_bitwise
1497  apply (simp add:mask_def)
1498  done
1499
1500lemma cpde_relation_invalid:
1501 "cpde_relation pdea pde \<Longrightarrow> (pde_get_tag pde = scast pde_pde_pt \<and> pde_pde_pt_CL.present_CL (pde_pde_pt_lift pde) = 0) = isInvalidPDE pdea"
1502  apply (simp add: cpde_relation_def Let_def)
1503  apply (simp add: pde_pde_pt_lift_def)
1504  apply (case_tac pdea, simp_all add: isInvalidPDE_def) [1]
1505   apply (clarsimp simp: pde_pde_pt_lift pde_pde_pt_lift_def)
1506  apply (simp add: pde_pde_pt_lift_def pde_pde_pt_lift)
1507  done
1508
1509lemma storePTE_Basic_ccorres'':
1510  "ccorres dc xfdc \<top> {s. ptr_val (f s) = p \<and> cpte_relation pte pte'} hs
1511     (storePTE p pte)
1512     (Guard C_Guard {s. s \<Turnstile>\<^sub>c f s}
1513        (Basic (\<lambda>s. globals_update( t_hrs_'_update
1514            (hrs_mem_update (heap_update (f s) pte'))) s)))"
1515  apply (rule ccorres_guard_imp2)
1516   apply (rule ccorres_gen_asm2, erule storePTE_Basic_ccorres')
1517  apply simp
1518  done
1519
1520lemma pageBitsForSize_le_64: "of_nat (pageBitsForSize x) < (64::machine_word)"
1521  by (cases x; simp add: bit_simps)
1522
1523lemma pte_sadness:
1524  "\<lbrakk> pte' = pte\<lparr>pte_CL.page_base_address_CL := pte_CL.page_base_address_CL pte'\<rparr>;
1525     f (pte_CL.page_base_address_CL pte) = pte_CL.page_base_address_CL pte \<rbrakk>
1526   \<Longrightarrow> pte'\<lparr>pte_CL.page_base_address_CL := f (pte_CL.page_base_address_CL pte)\<rparr> = pte"
1527  apply (cases pte', cases pte, simp)
1528  done
1529
1530lemma hd_in_zip_set:
1531   "slots \<noteq> [] \<Longrightarrow> (hd slots, 0) \<in> set (zip slots [0.e.of_nat (length slots - Suc 0)::machine_word])"
1532   by (cases slots; simp add: upto_enum_word upto_0_to_n2 del: upt_Suc)
1533
1534lemma last_in_zip_set:
1535  "\<lbrakk> slots \<noteq> []; length js = length slots \<rbrakk> \<Longrightarrow> (last slots, last js) \<in> set (zip slots js)"
1536   apply (simp add: in_set_zip last_conv_nth)
1537   apply (rule_tac x="length slots - 1" in exI)
1538   apply clarsimp
1539   apply (subst last_conv_nth)
1540    apply (cases js; simp)
1541   apply simp
1542   done
1543
1544definition
1545  "isVMPTE entry \<equiv> \<exists>x y. entry = (VMPTE x, VMPTEPtr y)"
1546
1547primrec (nonexhaustive)
1548  thePTE :: "vmpage_entry \<Rightarrow> pte" where
1549  "thePTE (VMPTE pte) = pte"
1550
1551primrec (nonexhaustive)
1552  thePTEPtr :: "vmpage_entry_ptr \<Rightarrow> machine_word" where
1553  "thePTEPtr (VMPTEPtr p) = p"
1554
1555lemma performPageInvocationMapPTE_ccorres:
1556  "ccorres (K (K \<bottom>) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
1557       (invs' and cte_at' slot and (\<lambda>s. 7 \<le> gsMaxObjectSize s)
1558           and (\<lambda>_. page_entry_map_corres mapping \<and>
1559                  (isArchPageCap cap \<and> capVPMappedAddress (capCap cap) \<noteq> None)))
1560       (UNIV \<inter> {s. cpte_relation (thePTE (fst mapping)) (pte_' s)}
1561             \<inter> {s. ptSlot_' s = pte_Ptr (thePTEPtr (snd mapping))}
1562             \<inter> {s. ccap_relation cap (cap_' s)}
1563             \<inter> {s. ctSlot_' s = cte_Ptr slot}
1564             \<inter> {s. vspace_' s = pml4e_Ptr vspace}
1565             \<inter> {s. isVMPTE mapping}) []
1566       (liftE (performPageInvocation (PageMap cap slot mapping vspace)))
1567       (Call performX86PageInvocationMapPTE_'proc)"
1568  supply pageBitsForSize_le_64 [simp]
1569  apply (rule ccorres_gen_asm2)+
1570  apply (rule ccorres_gen_asm)
1571  apply (simp only: liftE_liftM ccorres_liftM_simp)
1572  apply (cinit lift:  pte_' ptSlot_' cap_' ctSlot_' vspace_')
1573   apply (clarsimp simp: isVMPTE_def page_entry_map_corres_def)
1574   apply (ctac)
1575     apply (rule ccorres_split_nothrow)
1576         apply (rule storePTE_Basic_ccorres'')
1577        apply ceqv
1578       apply csymbr
1579       apply csymbr
1580       apply (clarsimp simp: isCap_simps)
1581       apply (frule cap_get_tag_isCap_unfolded_H_cap,
1582                clarsimp simp: cap_frame_cap_lift cap_to_H_def
1583                        elim!: ccap_relationE
1584                        split: if_split_asm)
1585       apply (rule ccorres_add_return2)
1586       apply (rule ccorres_split_nothrow)
1587           apply (rule ccorres_call[where xf'=xfdc])
1588              apply (ctac add: invalidatePageStructureCacheASID_ccorres)
1589             apply clarsimp
1590            apply clarsimp
1591           apply clarsimp
1592          apply ceqv
1593         apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
1594         apply (rule allI, rule conseqPre, vcg)
1595         apply (clarsimp simp: return_def)
1596        apply wp
1597       apply (vcg exspec=invalidatePageStructureCacheASID_modifies)
1598      apply wp
1599     apply vcg
1600    apply wp
1601   apply (rule conseqPre[where P'="UNIV" and P=UNIV, simplified])
1602   apply vcg
1603   apply (clarsimp simp: isCap_simps dest!: cap_get_tag_isCap_unfolded_H_cap)
1604  apply clarsimp
1605  done
1606
1607lemma pde_align_ptBits:
1608  "\<lbrakk> ko_at' (X64_H.PageTablePDE x y z w u v) slot s ;valid_objs' s\<rbrakk>
1609  \<Longrightarrow> is_aligned (ptrFromPAddr x) ptBits"
1610  apply (drule ko_at_valid_objs')
1611    apply simp
1612   apply (simp add:projectKO_opts_defs injectKO_defs
1613    split:Structures_H.kernel_object.splits
1614    arch_kernel_object.splits)+
1615  apply (simp add:valid_obj'_def)
1616  apply (rule is_aligned_ptrFromPAddr_n)
1617   apply simp
1618  apply (simp add: bit_simps)
1619  done
1620
1621lemma vaddr_segment_nonsense3_folded:
1622  "is_aligned (p :: machine_word) pageBits \<Longrightarrow>
1623   (p + ((vaddr >> pageBits) && mask (pt_bits - word_size_bits) << word_size_bits) && ~~ mask pt_bits) = p"
1624  apply (rule is_aligned_add_helper[THEN conjunct2])
1625   apply (simp add: bit_simps mask_def)+
1626  apply (rule shiftl_less_t2n[where m=12 and n=3, simplified, OF and_mask_less'[where n=9, unfolded mask_def, simplified]])
1627   apply simp+
1628  done
1629
1630lemma storePDE_Basic_ccorres'':
1631  "ccorres dc xfdc
1632     (\<lambda>_. True)
1633     {s. ptr_val (f s) = p \<and> cpde_relation pde pde'} hs
1634     (storePDE p pde)
1635     (Guard C_Guard {s. s \<Turnstile>\<^sub>c f s}
1636        (Basic (\<lambda>s. globals_update( t_hrs_'_update
1637            (hrs_mem_update (heap_update (f s) pde'))) s)))"
1638  apply (rule ccorres_guard_imp2)
1639   apply (rule ccorres_gen_asm2, erule storePDE_Basic_ccorres')
1640  apply simp
1641  done
1642
1643lemma pde_lift_to_large:
1644  "pde_lift pde = Some (Pde_pde_large pde') \<Longrightarrow> pde_pde_large_lift pde = pde'"
1645  by (simp add: pde_pde_large_lift_def)
1646
1647lemma pde_lift_to_tag:
1648  "pde_lift pde = Some (Pde_pde_large pde') \<Longrightarrow> pde_get_tag pde = scast pde_pde_large"
1649  by (simp add: pde_lift_def Let_def split: if_split_asm)
1650
1651lemmas pde_lift_section = pde_lift_to_large pde_lift_to_tag
1652
1653definition
1654  "isVMPDE entry \<equiv> \<exists>x y. entry = (VMPDE x, VMPDEPtr y)"
1655
1656primrec (nonexhaustive)
1657  thePDE :: "vmpage_entry \<Rightarrow> pde" where
1658  "thePDE (VMPDE pde) = pde"
1659
1660primrec (nonexhaustive)
1661  thePDEPtr :: "vmpage_entry_ptr \<Rightarrow> machine_word" where
1662  "thePDEPtr (VMPDEPtr p) = p"
1663
1664lemma performPageInvocationMapPDE_ccorres:
1665  "ccorres (K (K \<bottom>) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
1666       (invs' and cte_at' slot and (\<lambda>s. 7 \<le> gsMaxObjectSize s)
1667              and (\<lambda>s. page_entry_map_corres mapping \<and>
1668                       isArchPageCap cap \<and> capVPMappedAddress (capCap cap) \<noteq> None))
1669       (UNIV \<inter> {s. cpde_relation (thePDE (fst mapping)) (pde_' s)}
1670             \<inter> {s. pdSlot_' s = pde_Ptr (thePDEPtr (snd mapping))}
1671             \<inter> {s. ccap_relation cap (cap_' s)}
1672             \<inter> {s. ctSlot_' s = cte_Ptr slot}
1673             \<inter> {s. vspace_' s = pml4e_Ptr vspace}
1674             \<inter> {s. isVMPDE mapping}) []
1675       (liftE (performPageInvocation (PageMap cap slot mapping vspace)))
1676       (Call performX86PageInvocationMapPDE_'proc)"
1677  supply pageBitsForSize_le_64 [simp]
1678  apply (rule ccorres_gen_asm2)
1679  apply (rule ccorres_gen_asm)
1680  apply (simp only: liftE_liftM ccorres_liftM_simp)
1681  apply (cinit lift:  pde_' pdSlot_' cap_' ctSlot_' vspace_')
1682   apply (clarsimp simp: isVMPDE_def page_entry_map_corres_def)
1683   apply (ctac)
1684     apply (rule ccorres_split_nothrow)
1685         apply (rule storePDE_Basic_ccorres'')
1686        apply ceqv
1687       apply csymbr
1688       apply csymbr
1689       apply (clarsimp simp: isCap_simps)
1690       apply (frule cap_get_tag_isCap_unfolded_H_cap,
1691                clarsimp simp: cap_frame_cap_lift cap_to_H_def
1692                        elim!: ccap_relationE
1693                        split: if_split_asm)
1694       apply (rule ccorres_add_return2)
1695       apply (rule ccorres_split_nothrow)
1696           apply (rule ccorres_call[where xf'=xfdc])
1697              apply (ctac add: invalidatePageStructureCacheASID_ccorres)
1698             apply clarsimp
1699            apply clarsimp
1700           apply clarsimp
1701          apply ceqv
1702         apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
1703         apply (rule allI, rule conseqPre, vcg)
1704         apply (clarsimp simp: return_def)
1705        apply wp
1706       apply (vcg exspec=invalidatePageStructureCacheASID_modifies)
1707      apply wp
1708     apply vcg
1709    apply wp
1710   apply (rule conseqPre[where P'="UNIV" and P=UNIV, simplified])
1711   apply vcg
1712   apply (clarsimp simp: isCap_simps dest!: cap_get_tag_isCap_unfolded_H_cap)
1713  apply clarsimp
1714  done
1715
1716lemma performPageInvocationRemapPDE_ccorres:
1717  "ccorres (K (K \<bottom>) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
1718       (invs' and (\<lambda>s. 7 \<le> gsMaxObjectSize s)
1719           and (\<lambda>_. asid \<le> mask asid_bits \<and> page_entry_map_corres mapping))
1720       (UNIV \<inter> {s. cpde_relation (thePDE (fst mapping)) (pde_' s)}
1721             \<inter> {s. pdSlot_' s = pde_Ptr (thePDEPtr (snd mapping))}
1722             \<inter> {s. asid_' s = asid}
1723             \<inter> {s. vspace_' s = pml4e_Ptr vspace}
1724             \<inter> {s. isVMPDE mapping}) []
1725       (liftE (performPageInvocation (PageRemap mapping asid vspace)))
1726       (Call performX86PageInvocationRemapPDE_'proc)"
1727  supply pageBitsForSize_le_64 [simp]
1728  apply (rule ccorres_gen_asm2)
1729  apply (rule ccorres_gen_asm)
1730  apply (simp only: liftE_liftM ccorres_liftM_simp)
1731  apply (cinit lift:  pde_' pdSlot_' asid_' vspace_')
1732   apply (clarsimp simp: page_entry_map_corres_def isVMPDE_def)
1733   (* FIXME x64: UGLY HACK *)
1734   apply (rule ccorres_seq_skip[THEN iffD1],
1735             rule ccorres_add_return,
1736              ctac (no_vcg) add: ccorres_return_Skip)
1737     apply (rule ccorres_split_nothrow)
1738         apply (rule storePDE_Basic_ccorres'')
1739        apply ceqv
1740       apply csymbr
1741       apply (rule ccorres_add_return2)
1742       apply (rule ccorres_split_nothrow)
1743           apply (rule ccorres_call[where xf'=xfdc])
1744              apply (ctac add: invalidatePageStructureCacheASID_ccorres)
1745             apply clarsimp
1746            apply clarsimp
1747           apply clarsimp
1748          apply ceqv
1749         apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
1750         apply (rule allI, rule conseqPre, vcg)
1751         apply (clarsimp simp: return_def)
1752        apply wp
1753       apply (vcg exspec=invalidatePageStructureCacheASID_modifies)
1754      apply wp
1755     apply vcg
1756    apply wp
1757   apply (clarsimp simp: isCap_simps dest!: cap_get_tag_isCap_unfolded_H_cap)
1758  apply clarsimp
1759  done
1760
1761lemma performPageInvocationRemapPTE_ccorres:
1762  "ccorres (K (K \<bottom>) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
1763       (invs' and (\<lambda>s. 7 \<le> gsMaxObjectSize s)
1764           and (\<lambda>_. asid \<le> mask asid_bits \<and> page_entry_map_corres mapping))
1765       (UNIV \<inter> {s. cpte_relation (thePTE (fst mapping)) (pte_' s)}
1766             \<inter> {s. ptSlot_' s = pte_Ptr (thePTEPtr (snd mapping))}
1767             \<inter> {s. asid_' s = asid}
1768             \<inter> {s. vspace_' s = pml4e_Ptr vspace}
1769             \<inter> {s. isVMPTE mapping}) []
1770       (liftE (performPageInvocation (PageRemap mapping asid vspace)))
1771       (Call performX86PageInvocationRemapPTE_'proc)"
1772  supply pageBitsForSize_le_64 [simp]
1773  apply (rule ccorres_gen_asm2)
1774  apply (rule ccorres_gen_asm)
1775  apply (simp only: liftE_liftM ccorres_liftM_simp)
1776  apply (cinit lift:  pte_' ptSlot_' asid_' vspace_')
1777   apply (clarsimp simp: page_entry_map_corres_def isVMPTE_def)
1778   (* FIXME x64: UGLY HACK *)
1779   apply (rule ccorres_seq_skip[THEN iffD1],
1780             rule ccorres_add_return,
1781              ctac (no_vcg) add: ccorres_return_Skip)
1782     apply (rule ccorres_split_nothrow)
1783         apply (rule storePTE_Basic_ccorres'')
1784        apply ceqv
1785       apply csymbr
1786       apply (rule ccorres_add_return2)
1787       apply (rule ccorres_split_nothrow)
1788           apply (rule ccorres_call[where xf'=xfdc])
1789              apply (ctac add: invalidatePageStructureCacheASID_ccorres)
1790             apply clarsimp
1791            apply clarsimp
1792           apply clarsimp
1793          apply ceqv
1794         apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
1795         apply (rule allI, rule conseqPre, vcg)
1796         apply (clarsimp simp: return_def)
1797        apply wp
1798       apply (vcg exspec=invalidatePageStructureCacheASID_modifies)
1799      apply wp
1800     apply vcg
1801    apply wp
1802   apply (clarsimp simp: isCap_simps dest!: cap_get_tag_isCap_unfolded_H_cap)
1803  apply clarsimp
1804  done
1805
1806definition
1807  "isVMPDPTE entry \<equiv> \<exists>x y. entry = (VMPDPTE x, VMPDPTEPtr y)"
1808
1809primrec (nonexhaustive)
1810  thePDPTE :: "vmpage_entry \<Rightarrow> pdpte" where
1811  "thePDPTE (VMPDPTE pdpte) = pdpte"
1812
1813primrec (nonexhaustive)
1814  thePDPTEPtr :: "vmpage_entry_ptr \<Rightarrow> machine_word" where
1815  "thePDPTEPtr (VMPDPTEPtr p) = p"
1816
1817lemma storePDPTE_Basic_ccorres'':
1818  "ccorres dc xfdc \<top> {s. ptr_val (f s) = p \<and> cpdpte_relation pte pte'} hs
1819     (storePDPTE p pte)
1820     (Guard C_Guard {s. s \<Turnstile>\<^sub>c f s}
1821        (Basic (\<lambda>s. globals_update( t_hrs_'_update
1822            (hrs_mem_update (heap_update (f s) pte'))) s)))"
1823  apply (rule ccorres_guard_imp2)
1824   apply (rule ccorres_gen_asm2, erule storePDPTE_Basic_ccorres')
1825  apply simp
1826  done
1827
1828lemma updatePDPTE_ccorres:
1829  "ccorres
1830            (K (K \<bottom>) \<currency> dc)
1831            (liftxf errstate id (\<lambda>y. ()) ret__unsigned_long_')
1832            \<top>
1833            (UNIV \<inter> {s. asid_' s = asid}
1834                  \<inter> {s. cpdpte_relation pdpte (pdpte_' s)}
1835                  \<inter> {s. pdptSlot_' s = pdpte_Ptr pdptPtr}
1836                  \<inter> {s. vspace_' s = pml4e_Ptr vspace})
1837            hs
1838            (doE y <- liftE (storePDPTE pdptPtr pdpte);
1839                liftE (invalidatePageStructureCacheASID
1840                 (addrFromPPtr vspace) asid)
1841             odE)
1842            (Call updatePDPTE_'proc)"
1843  supply Collect_const[simp del]
1844  apply (cinit' lift: asid_' pdpte_' pdptSlot_' vspace_')
1845   apply (clarsimp simp: bind_liftE_distrib[symmetric])
1846   apply (simp only: liftE_liftM ccorres_liftM_simp)
1847   apply (rule ccorres_split_nothrow)
1848       apply (rule storePDPTE_Basic_ccorres'')
1849      apply ceqv
1850     apply csymbr
1851     apply (rule ccorres_add_return2)
1852     apply (rule ccorres_split_nothrow)
1853         apply (rule ccorres_call[where xf'=xfdc])
1854            apply (ctac add: invalidatePageStructureCacheASID_ccorres)
1855           apply clarsimp
1856          apply clarsimp
1857         apply clarsimp
1858        apply ceqv
1859       apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
1860       apply (rule allI, rule conseqPre, vcg)
1861       apply (clarsimp simp: return_def)
1862      apply wp
1863     apply (vcg exspec=invalidatePageStructureCacheASID_modifies)
1864    apply wp
1865   apply vcg
1866  apply clarsimp
1867  done
1868
1869lemma performPageInvocationRemapPDPTE_ccorres:
1870  "ccorres (K (K \<bottom>) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
1871       ((\<lambda>s. page_entry_map_corres mapping))
1872       (UNIV \<inter> {s. cpdpte_relation (thePDPTE (fst mapping)) (pdpte_' s)}
1873             \<inter> {s. pdptSlot_' s = pdpte_Ptr (thePDPTEPtr (snd mapping))}
1874             \<inter> {s. asid_' s = asid}
1875             \<inter> {s. vspace_' s = pml4e_Ptr vspace}
1876             \<inter> {s. isVMPDPTE mapping}) hs
1877       (liftE (performPageInvocation (PageRemap mapping asid vspace)))
1878       (Call performX64ModeRemap_'proc)"
1879  supply pageBitsForSize_le_64 [simp]
1880  apply (rule ccorres_gen_asm2)
1881  apply (simp add: performPageInvocation_def bind_liftE_distrib)
1882  apply (cinit' lift:  pdpte_' pdptSlot_' asid_' vspace_')
1883   apply (clarsimp simp: page_entry_map_corres_def isVMPDPTE_def)
1884   (* FIXME x64: UGLY HACK *)
1885   apply (rule ccorres_seq_skip[THEN iffD1],
1886             rule ccorres_add_return,
1887              ctac (no_vcg) add: ccorres_return_Skip)
1888     apply (rule ccorres_add_returnOk)
1889     apply (ctac add: updatePDPTE_ccorres)
1890        apply (rule ccorres_return_CE)
1891          apply clarsimp
1892         apply clarsimp
1893        apply clarsimp
1894       apply (rule ccorres_inst[where P=\<top> and P'=UNIV])
1895       apply clarsimp
1896      apply wp
1897     apply vcg
1898    apply wp
1899   apply clarsimp
1900  apply clarsimp
1901  done
1902
1903lemma performPageInvocationMapPDPTE_ccorres:
1904  "ccorres (K (K \<bottom>) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
1905       (cte_at' slot and (\<lambda>s. page_entry_map_corres mapping \<and> isArchPageCap cap
1906                            \<and> capVPMappedAddress (capCap cap) \<noteq> None))
1907       (UNIV \<inter> {s. cpdpte_relation (thePDPTE (fst mapping)) (pdpte_' s)}
1908             \<inter> {s. pdptSlot_' s = pdpte_Ptr (thePDPTEPtr (snd mapping))}
1909             \<inter> {s. ccap_relation cap (cap_' s)}
1910             \<inter> {s. ctSlot_' s = cte_Ptr slot}
1911             \<inter> {s. vspace_' s = pml4e_Ptr vspace}
1912             \<inter> {s. isVMPDPTE mapping}) hs
1913       (liftE (performPageInvocation (PageMap cap slot mapping vspace)))
1914       (Call performX64ModeMap_'proc)"
1915  supply pageBitsForSize_le_64 [simp]
1916  apply (rule ccorres_gen_asm2)
1917  apply (rule ccorres_gen_asm)
1918  apply (simp add: performPageInvocation_def)
1919  apply (cinit' lift:  pdpte_' pdptSlot_' cap_' ctSlot_' vspace_')
1920   apply (clarsimp simp: page_entry_map_corres_def isVMPDPTE_def)
1921   apply (rename_tac apdpte apdpteptr a b)
1922   apply (simp only: bind_liftE_distrib)
1923   apply (subst liftE_bindE)
1924   apply ctac
1925     apply (clarsimp simp: isCap_simps)
1926     apply csymbr
1927     apply (clarsimp simp: liftE_foo_to_fooE)
1928     apply (rule ccorres_add_returnOk)
1929     apply (ctac add: updatePDPTE_ccorres)
1930        apply (rule ccorres_return_CE)
1931          apply clarsimp
1932         apply clarsimp
1933        apply clarsimp
1934       apply (rule ccorres_inst[where P=\<top> and P'=UNIV])
1935       apply clarsimp
1936      apply wp
1937     apply (vcg exspec=updatePDPTE_modifies)
1938    apply wp
1939   apply (clarsimp simp: isCap_simps)
1940   apply (clarsimp simp: ccap_relation_PageCap_MappedASID)
1941   apply (frule cap_get_tag_isCap_unfolded_H_cap)
1942   apply clarsimp
1943   apply vcg
1944  apply clarsimp
1945  done
1946
1947lemma vmsz_aligned_addrFromPPtr':
1948  "vmsz_aligned' (addrFromPPtr p) sz
1949       = vmsz_aligned' p sz"
1950  apply (simp add: vmsz_aligned'_def addrFromPPtr_def
1951                   X64.addrFromPPtr_def)
1952  apply (subgoal_tac "is_aligned X64.pptrBase (pageBitsForSize sz)")
1953   apply (rule iffI)
1954    apply (drule(1) aligned_add_aligned)
1955      apply (simp add: pageBitsForSize_def word_bits_def split: vmpage_size.split)
1956     apply simp
1957   apply (erule(1) aligned_sub_aligned)
1958    apply (simp add: pageBitsForSize_def word_bits_def split: vmpage_size.split)
1959  apply (simp add: pageBitsForSize_def X64.pptrBase_def is_aligned_def bit_simps
1960            split: vmpage_size.split)+
1961  done
1962
1963lemmas vmsz_aligned_addrFromPPtr
1964    = vmsz_aligned_addrFromPPtr'
1965      vmsz_aligned_addrFromPPtr'[unfolded addrFromPPtr_def]
1966      vmsz_aligned_addrFromPPtr'[unfolded vmsz_aligned'_def]
1967      vmsz_aligned_addrFromPPtr'[unfolded addrFromPPtr_def vmsz_aligned'_def]
1968
1969lemmas framesize_from_H_simps
1970    = framesize_from_H_def[split_simps vmpage_size.split]
1971
1972lemma shiftr_asid_low_bits_mask_asid_high_bits:
1973  "(asid :: machine_word) \<le> mask asid_bits
1974      \<Longrightarrow> (asid >> asid_low_bits) && mask asid_high_bits = asid >> asid_low_bits"
1975  apply (rule iffD2 [OF mask_eq_iff_w2p])
1976   apply (simp add: asid_high_bits_def word_size)
1977  apply (rule shiftr_less_t2n)
1978  apply (simp add: asid_low_bits_def asid_high_bits_def mask_def)
1979  apply (simp add: asid_bits_def)
1980  done
1981
1982lemma shiftr_asid_low_bits_mask_eq_0:
1983  "\<lbrakk> (asid :: machine_word) \<le> mask asid_bits; asid >> asid_low_bits = 0 \<rbrakk>
1984        \<Longrightarrow> (asid && mask asid_low_bits = 0) = (asid = 0)"
1985  apply (rule iffI[rotated])
1986   apply simp
1987  apply (rule asid_low_high_bits)
1988     apply (rule upcast_ucast_id[where 'b=machine_word_len]; simp add: asid_low_bits_of_mask_eq)
1989    apply (simp add: ucast_asid_high_bits_is_shift)
1990   apply (simp add: asid_wf_def mask_def)
1991  apply (rule asid_wf_0)
1992  done
1993
1994lemma slotcap_in_mem_valid:
1995  "\<lbrakk> slotcap_in_mem cap slot (ctes_of s); valid_objs' s \<rbrakk>
1996            \<Longrightarrow> s \<turnstile>' cap"
1997  apply (clarsimp simp: slotcap_in_mem_def)
1998  apply (erule(1) ctes_of_valid')
1999  done
2000
2001lemma list_length_less:
2002  "(args = [] \<or> length args \<le> Suc 0) = (length args < 2)"
2003  by (case_tac args,simp_all)
2004
2005lemma unat_less_iff64:
2006  "\<lbrakk>unat (a::machine_word) = b;c < 2^word_bits\<rbrakk>
2007   \<Longrightarrow> (a < of_nat c) = (b < c)"
2008  apply (rule iffI)
2009    apply (drule unat_less_helper)
2010    apply simp
2011  apply (simp add:unat64_eq_of_nat)
2012  apply (rule of_nat_mono_maybe)
2013   apply (simp add:word_bits_def)
2014  apply simp
2015  done
2016
2017lemma injection_handler_if_returnOk:
2018  "injection_handler Inl (if a then b else returnOk c)
2019  = (if a then (injection_handler Inl b) else returnOk c)"
2020  apply (clarsimp simp:whenE_def injection_handler_def)
2021  apply (clarsimp simp:injection_handler_def
2022    throwError_def return_def bind_def returnOk_def
2023    handleE'_def split:if_splits)
2024  done
2025
2026lemma pbfs_less: "pageBitsForSize sz < 31"
2027  by (case_tac sz,simp_all add: bit_simps)
2028
2029lemma at_least_2_args:
2030  "\<not>  length args < 2 \<Longrightarrow> \<exists>a b c. args = a#b#c"
2031  apply (case_tac args)
2032   apply simp
2033  apply (case_tac list)
2034   apply simp
2035  apply (case_tac lista)
2036   apply simp
2037  apply simp
2038  done
2039
2040definition
2041  to_option :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a option"
2042where
2043  "to_option f x \<equiv> if f x then Some x else None"
2044
2045(* FIXME: move *)
2046lemma valid_objs_valid_pte': "\<lbrakk> valid_objs' s ; ko_at' (ko :: pte) p s \<rbrakk> \<Longrightarrow> valid_pte' ko s"
2047  by (fastforce simp add: obj_at'_def ran_def valid_obj'_def projectKOs valid_objs'_def)
2048
2049lemma cte_wp_at_diminished_gsMaxObjectSize:
2050  "cte_wp_at' (diminished' cap o cteCap) slot s
2051    \<Longrightarrow> valid_global_refs' s
2052    \<Longrightarrow> 2 ^ capBits cap \<le> gsMaxObjectSize s"
2053  apply (clarsimp simp: cte_wp_at_ctes_of)
2054  apply (drule(1) valid_global_refsD_with_objSize)
2055  apply (clarsimp simp: diminished'_def capMaster_eq_capBits_eq[OF capMasterCap_maskCapRights])
2056  done
2057
2058lemma two_nat_power_pageBitsForSize_le:
2059  "(2 :: nat) ^ pageBits \<le> 2 ^ pageBitsForSize vsz"
2060  by (cases vsz, simp_all add: pageBits_def bit_simps)
2061
2062(* FIXME: move *)
2063lemma is_aligned_pageBitsForSize_minimum:
2064  "\<lbrakk> is_aligned p (pageBitsForSize sz) ; n \<le> pageBits \<rbrakk> \<Longrightarrow> is_aligned p n"
2065  apply (cases sz; clarsimp simp: pageBits_def)
2066  apply (erule is_aligned_weaken, simp)+
2067  done
2068
2069lemma ptrFromPAddr_add_left:
2070  "ptrFromPAddr (x + y) = ptrFromPAddr x + y"
2071  unfolding ptrFromPAddr_def by simp
2072
2073lemma at_least_3_args:
2074  "\<not>  length args < 3 \<Longrightarrow> \<exists>a b c d. args = a#b#c#d"
2075  apply (case_tac args; simp)
2076  apply (rename_tac list, case_tac list; simp)+
2077  done
2078
2079lemma createSafeMappingEntries_PTE_ccorres:
2080  "ccorres (syscall_error_rel \<currency> (\<lambda>rv rv'. isVMPTE rv \<and> cpte_relation (thePTE (fst rv)) (fst rv')
2081                                         \<and> (snd rv' = pte_Ptr (thePTEPtr (snd rv)))))
2082     (liftxf errstate create_mapping_pte_return_C.status_C
2083             (\<lambda>v. (create_mapping_pte_return_C.pte_C v,
2084                   create_mapping_pte_return_C.ptSlot_C v))
2085             ret__struct_create_mapping_pte_return_C_')
2086     (valid_objs' and page_map_l4_at' pml4 and
2087         (\<lambda>_. vsz = X64SmallPage \<and> vmsz_aligned' (addrFromPPtr base) vsz
2088              \<and> base \<in> kernel_mappings))
2089     (UNIV \<inter> {s. base_' s = (addrFromPPtr base)} \<inter> {s. vaddr___unsigned_long_' s = vaddr}
2090           \<inter> {s. vmrights_to_H (vmRights_' s) = vrights \<and> vmRights_' s < 4 \<and> vmRights_' s \<noteq> 0}
2091           \<inter> {s. vm_attribs_relation attr (attr_' s)}
2092           \<inter> {s. vspace_' s = pml4e_Ptr pml4}) hs
2093     (createSafeMappingEntries (addrFromPPtr base) vaddr vsz vrights attr pml4)
2094     (Call createSafeMappingEntries_PTE_'proc)"
2095  supply Collect_const[simp del]
2096    apply (rule ccorres_gen_asm)
2097  apply (cinit lift: base_' vaddr___unsigned_long_' vmRights_' attr_' vspace_')
2098   apply (simp add: createSafeMappingEntries_def createMappingEntries_def
2099                    ensureSafeMapping_def framesize_from_H_eqs)
2100   apply (simp add: lookupError_injection bindE_assoc)
2101   apply (ctac add: ccorres_injection_handler_csum1[OF lookupPTSlot_ccorres])
2102      apply (simp add: Collect_False liftE_bindE)
2103      apply csymbr+
2104      apply (rule ccorres_return_CE, simp+)[1]
2105     apply (simp, ccorres_rewrite)
2106     apply (rule_tac P'="{s. lu_ret___struct_lookupPTSlot_ret_C = errstate s}" in ccorres_from_vcg_throws[where P=\<top>])
2107     apply (rule allI, rule conseqPre, vcg)
2108     apply (clarsimp simp: fst_throwError_returnOk syscall_error_to_H_cases
2109                           syscall_error_rel_def exception_defs false_def)
2110     apply (erule lookup_failure_rel_fault_lift[rotated])
2111     apply (simp add: exception_defs)
2112    apply wpsimp
2113   apply (clarsimp simp: isVMPTE_def)
2114   apply (vcg exspec=lookupPTSlot_modifies)
2115  apply clarsimp
2116  apply (clarsimp simp: cpte_relation_def Let_def vm_attribs_relation_def from_bool_def)
2117  apply (rule addrFromPPtr_mask_middle_pml4ShiftBits[simplified, simplified bit_simps])
2118   apply (clarsimp simp: vmsz_aligned_addrFromPPtr' vmsz_aligned_aligned_pageBits[simplified bit_simps])
2119  apply clarsimp
2120  done
2121
2122lemma pde_case_isPageTablePDE:
2123  "(case pde of PageTablePDE _ _ _ _ _ _ \<Rightarrow> P | _ \<Rightarrow> Q)
2124       = (if isPageTablePDE pde then P else Q)"
2125  by (clarsimp simp: isPageTablePDE_def split: pde.splits)
2126
2127primrec
2128  shiftBitsForSize :: "vmpage_size \<Rightarrow> nat"
2129where
2130  "shiftBitsForSize X64SmallPage = pml4ShiftBits"
2131| "shiftBitsForSize X64LargePage = pdptShiftBits"
2132| "shiftBitsForSize X64HugePage = pdShiftBits"
2133
2134lemma shiftBits_pageBits_add[simp]:
2135  "shiftBitsForSize sz + pageBitsForSize sz = 51"
2136  by (case_tac sz; clarsimp simp: bit_simps)
2137
2138lemma addrFromPPtr_mask_middle_shiftBits:
2139  "\<lbrakk>is_aligned p (pageBitsForSize sz); p \<in> kernel_mappings\<rbrakk> \<Longrightarrow>
2140   addrFromPPtr p && (mask (shiftBitsForSize sz) << (pageBitsForSize sz)) = addrFromPPtr p"
2141  apply (clarsimp simp: mask_shiftl_decompose kernel_mappings_def)
2142  apply (subst word_bool_alg.conj.assoc[symmetric])
2143  apply (subst is_aligned_neg_mask_eq)
2144   apply (rule aligned_already_mask)
2145   apply (clarsimp simp: is_aligned_addrFromPPtr_pageBitsForSize)
2146  apply (clarsimp simp: addrFromPPtr_def X64.pptrBase_def pptr_base_def bit_simps)
2147  apply (clarsimp simp: mask_def)
2148  apply (word_bitwise, clarsimp)
2149  done
2150
2151lemma createSafeMappingEntries_PDE_ccorres:
2152  "ccorres (syscall_error_rel \<currency> (\<lambda>rv rv'. isVMPDE rv \<and> cpde_relation (thePDE (fst rv)) (fst rv')
2153                                         \<and> (snd rv' = pde_Ptr (thePDEPtr (snd rv)))))
2154     (liftxf errstate create_mapping_pde_return_C.status_C
2155             (\<lambda>v. (create_mapping_pde_return_C.pde_C v,
2156                   create_mapping_pde_return_C.pdSlot_C v))
2157             ret__struct_create_mapping_pde_return_C_')
2158     (valid_objs' and page_map_l4_at' pml4 and
2159         (\<lambda>_. vsz = X64LargePage \<and> vmsz_aligned' (addrFromPPtr base) vsz
2160              \<and> base \<in> kernel_mappings))
2161     (UNIV \<inter> {s. base_' s = (addrFromPPtr base)} \<inter> {s. vaddr___unsigned_long_' s = vaddr}
2162           \<inter> {s. vmrights_to_H (vmRights_' s) = vrights \<and> vmRights_' s < 4 \<and> vmRights_' s \<noteq> 0}
2163           \<inter> {s. vm_attribs_relation attr (attr_' s)}
2164           \<inter> {s. vspace_' s = pml4e_Ptr pml4}) hs
2165     (createSafeMappingEntries (addrFromPPtr base) vaddr vsz vrights attr pml4)
2166     (Call createSafeMappingEntries_PDE_'proc)"
2167  supply Collect_const[simp del]
2168  apply (rule ccorres_gen_asm)
2169  apply (cinit lift: base_' vaddr___unsigned_long_' vmRights_' attr_' vspace_')
2170   apply (simp add: createSafeMappingEntries_def createMappingEntries_def
2171                    ensureSafeMapping_def framesize_from_H_eqs)
2172   apply (simp add: lookupError_injection bindE_assoc)
2173   apply (ctac add: ccorres_injection_handler_csum1[OF lookupPDSlot_ccorres])
2174      apply (simp add: Collect_False liftE_bindE)
2175      apply (rule ccorres_pre_getObject_pde)
2176      apply csymbr
2177      apply (rule ccorres_rhs_assoc2)
2178      apply (rule ccorres_rhs_assoc2)
2179      apply (rule_tac xf'="ret__int_'" and val="from_bool (isPageTablePDE x)"
2180               and R="ko_at' x rv"
2181               in ccorres_symb_exec_r_known_rv[where R'=UNIV])
2182         apply (rule conseqPre, vcg)
2183         apply clarsimp
2184         apply (erule cmap_relationE1[OF rf_sr_cpde_relation], erule ko_at_projectKO_opt)
2185         apply (clarsimp simp: typ_heap_simps cpde_relation_def Let_def)
2186         apply (case_tac x; fastforce simp: if_1_0_0 pde_lifts isPageTablePDE_def false_def true_def
2187                                            pde_pde_pt_lift_def)
2188        apply ceqv
2189       apply (clarsimp simp: pde_case_isPageTablePDE)
2190       apply (rule ccorres_Cond_rhs_Seq, clarsimp)
2191        apply ccorres_rewrite
2192        apply (rule ccorres_from_vcg_throws[where P=\<top> and P'=UNIV])
2193        apply (rule allI, rule conseqPre, vcg)
2194        apply (clarsimp simp: fst_throwError_returnOk exception_defs syscall_error_to_H_cases syscall_error_rel_def)
2195       apply clarsimp
2196       apply csymbr+
2197       apply (rule ccorres_return_CE, simp+)[1]
2198      apply (clarsimp simp: isVMPDE_def)
2199      apply vcg
2200     apply (simp, ccorres_rewrite)
2201     apply (rule_tac P'="{s. lu_ret___struct_lookupPDSlot_ret_C = errstate s}" in ccorres_from_vcg_throws[where P=\<top>])
2202     apply (rule allI, rule conseqPre, vcg)
2203     apply (clarsimp simp: fst_throwError_returnOk syscall_error_to_H_cases
2204                           syscall_error_rel_def exception_defs false_def)
2205     apply (erule lookup_failure_rel_fault_lift[rotated])
2206     apply (simp add: exception_defs)
2207    apply clarsimp
2208    apply wpsimp
2209   apply (clarsimp simp: if_P_1_0_neq_0)
2210   apply (vcg exspec=lookupPDSlot_modifies)
2211  apply (clarsimp simp: vm_attribs_relation_def typ_heap_simps from_bool_eq_if'
2212                        is_aligned_addrFromPPtr_pageBitsForSize[where sz=X64LargePage, simplified]
2213                        vmsz_aligned'_def
2214                  dest!: addrFromPPtr_mask_middle_shiftBits[where sz=X64LargePage, simplified])
2215  apply (clarsimp simp: bit_simps cpde_relation_def Let_def true_def false_def
2216                        isPageTablePDE_def of_bool_from_bool
2217                split: pde.splits)
2218  done
2219
2220lemma createSafeMappingEntries_PDPTE_ccorres:
2221  "ccorres (syscall_error_rel \<currency> (\<lambda>rv rv'. isVMPDPTE rv \<and> cpdpte_relation (thePDPTE (fst rv)) (fst rv')
2222                                         \<and> (snd rv' = pdpte_Ptr (thePDPTEPtr (snd rv)))))
2223     (liftxf errstate create_mapping_pdpte_return_C.status_C
2224             (\<lambda>v. (create_mapping_pdpte_return_C.pdpte_C v,
2225                   create_mapping_pdpte_return_C.pdptSlot_C v))
2226             ret__struct_create_mapping_pdpte_return_C_')
2227     (valid_objs' and page_map_l4_at' pml4 and
2228         (\<lambda>_. vsz = X64HugePage \<and> vmsz_aligned' (addrFromPPtr base) vsz
2229              \<and> base \<in> kernel_mappings))
2230     (UNIV \<inter> {s. base_' s = (addrFromPPtr base)} \<inter> {s. vaddr___unsigned_long_' s = vaddr}
2231           \<inter> {s. vmrights_to_H (vmRights_' s) = vrights \<and> vmRights_' s < 4 \<and> vmRights_' s \<noteq> 0}
2232           \<inter> {s. vm_attribs_relation attr (attr_' s)}
2233           \<inter> {s. vspace_' s = pml4e_Ptr pml4}) hs
2234     (createSafeMappingEntries (addrFromPPtr base) vaddr vsz vrights attr pml4)
2235     (Call createSafeMappingEntries_PDPTE_'proc)"
2236  supply Collect_const[simp del]
2237  apply (rule ccorres_gen_asm)
2238  apply (cinit lift: base_' vaddr___unsigned_long_' vmRights_' attr_' vspace_')
2239   apply (simp add: createSafeMappingEntries_def createMappingEntries_def
2240                    ensureSafeMapping_def framesize_from_H_eqs)
2241   apply (simp add: lookupError_injection bindE_assoc)
2242   apply (ctac add: ccorres_injection_handler_csum1[OF lookupPDPTSlot_ccorres])
2243      apply (simp add: Collect_False liftE_bindE)
2244      apply (rule ccorres_pre_getObject_pdpte)
2245      apply csymbr
2246      apply (rule ccorres_rhs_assoc2)
2247      apply (rule ccorres_rhs_assoc2)
2248      apply (rule_tac xf'="ret__int_'" and val="from_bool (isPageDirectoryPDPTE x)"
2249               and R="ko_at' x rv"
2250               in ccorres_symb_exec_r_known_rv[where R'=UNIV])
2251         apply (rule conseqPre, vcg)
2252         apply clarsimp
2253         apply (erule cmap_relationE1[OF rf_sr_cpdpte_relation], erule ko_at_projectKO_opt)
2254         apply (clarsimp simp: typ_heap_simps cpdpte_relation_def Let_def)
2255         apply (case_tac x; fastforce simp: if_1_0_0 pdpte_lifts isPageDirectoryPDPTE_def false_def true_def
2256                                            pdpte_pdpte_pd_lift_def)
2257        apply ceqv
2258       apply (clarsimp simp: pdpte_case_isPageDirectoryPDPTE)
2259       apply (rule ccorres_Cond_rhs_Seq, clarsimp)
2260        apply ccorres_rewrite
2261        apply (rule ccorres_from_vcg_throws[where P=\<top> and P'=UNIV])
2262        apply (rule allI, rule conseqPre, vcg)
2263        apply (clarsimp simp: fst_throwError_returnOk exception_defs syscall_error_to_H_cases syscall_error_rel_def)
2264       apply clarsimp
2265       apply csymbr+
2266       apply (rule ccorres_return_CE, simp+)[1]
2267      apply (clarsimp simp: isVMPDPTE_def)
2268      apply vcg
2269     apply (simp, ccorres_rewrite)
2270     apply (rule_tac P'="{s. lu_ret___struct_lookupPDPTSlot_ret_C = errstate s}" in ccorres_from_vcg_throws[where P=\<top>])
2271     apply (rule allI, rule conseqPre, vcg)
2272     apply (clarsimp simp: fst_throwError_returnOk syscall_error_to_H_cases
2273                           syscall_error_rel_def exception_defs false_def)
2274     apply (erule lookup_failure_rel_fault_lift[rotated])
2275     apply (simp add: exception_defs)
2276    apply clarsimp
2277    apply wpsimp
2278   apply (clarsimp simp: if_P_1_0_neq_0)
2279   apply (vcg exspec=lookupPDPTSlot_modifies)
2280  apply (clarsimp simp: vm_attribs_relation_def typ_heap_simps from_bool_eq_if'
2281                        is_aligned_addrFromPPtr_pageBitsForSize[where sz=X64HugePage, simplified]
2282                        vmsz_aligned'_def
2283                  dest!: addrFromPPtr_mask_middle_shiftBits[where sz=X64HugePage, simplified])
2284  apply (clarsimp simp: bit_simps cpdpte_relation_def Let_def true_def false_def
2285                        isPageDirectoryPDPTE_def of_bool_from_bool
2286                split: pdpte.splits)
2287  done
2288
2289lemma decodeX86ModeRemapPage_ccorres:
2290  "ccorres (intr_and_se_rel \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
2291            (invs' and page_map_l4_at' pml4 and (\<lambda>s. thread = ksCurThread s)
2292                   and K (vmsz_aligned' base vsz \<and> base \<in> kernel_mappings \<and> vsz = X64HugePage
2293                          \<and> isPageCap cap \<and> (\<exists>addr. capVPMappedAddress cap = Some (asid, addr))
2294                          \<and> capVPBasePtr cap = base))
2295            (UNIV \<inter> \<lbrace>\<acute>label___unsigned_long = scast X86PageRemap\<rbrace>
2296                  \<inter> \<lbrace>\<acute>page_size = framesize_from_H vsz\<rbrace>
2297                  \<inter> \<lbrace>\<acute>vroot = pml4e_Ptr pml4\<rbrace>
2298                  \<inter> \<lbrace>ccap_relation (ArchObjectCap cap) \<acute>cap\<rbrace>
2299                  \<inter> \<lbrace>\<acute>paddr = addrFromPPtr base\<rbrace>
2300                  \<inter> \<lbrace>vmrights_to_H \<acute>vm_rights = rights \<and> \<acute>vm_rights < 4 \<and> \<acute>vm_rights \<noteq> 0\<rbrace>
2301                  \<inter> \<lbrace>vm_attribs_relation attribs \<acute>vm_attr\<rbrace>
2302                  \<inter> \<lbrace>\<acute>vaddr___unsigned_long = vaddr\<rbrace>)
2303            hs
2304            (doE x <- injection_handler Inl
2305                  (createSafeMappingEntries (addrFromPPtr base) vaddr vsz rights attribs pml4);
2306                 invocationCatch thread isBlocking isCall Invocations_H.invocation.InvokeArchObject
2307                    (Inr (invocation.InvokePage (PageRemap x asid pml4)))
2308             odE)
2309            (Call decodeX86ModeMapRemapPage_'proc)"
2310  supply if_cong[cong] tl_drop_1[simp] Collect_const[simp del] dc_simp[simp del]
2311  apply (simp add: K_def)
2312  apply (rule ccorres_gen_asm)
2313  apply (cinit' lift: label___unsigned_long_' page_size_' vroot_' cap_' paddr_' vm_rights_' vm_attr_'
2314                      vaddr___unsigned_long_'
2315                simp: bindE_assoc injection_handler_bindE)
2316   apply (simp cong: StateSpace.state.fold_congs globals.fold_congs)
2317   apply (clarsimp simp: framesize_from_H_simps, ccorres_rewrite)
2318   apply (rule ccorres_rhs_assoc)+
2319   apply csymbr
2320   apply (ctac add: ccorres_injection_handler_csum1 [OF createSafeMappingEntries_PDPTE_ccorres])
2321      apply (simp add: Collect_False performX64MMUInvocations bindE_assoc bind_assoc)
2322      apply (ctac add: setThreadState_ccorres)
2323        apply (rule ccorres_cond_false)
2324        apply (rule ccorres_rhs_assoc)+
2325        apply csymbr
2326        apply (ctac(no_vcg) add: performPageInvocationRemapPDPTE_ccorres)
2327          apply (rule ccorres_alternative2)
2328          apply (rule ccorres_return_CE, simp+)[1]
2329         apply (rule ccorres_inst[where P=\<top> and P'=UNIV], simp)
2330        apply wp
2331       apply (wp sts_invs_minor')
2332      apply simp
2333      apply (vcg exspec=setThreadState_modifies)
2334     apply (simp, ccorres_rewrite)
2335     apply (rule ccorres_return_C_errorE, simp+)[1]
2336    apply (simp add: createSafeMappingEntries_def)
2337    apply (wp injection_wp[OF refl] createMappingEntries_wf)
2338   apply (simp add: all_ex_eq_helper)
2339   apply (vcg exspec=createSafeMappingEntries_PDPTE_modifies)
2340  apply (clarsimp simp: invs_valid_objs' tcb_at_invs' vmsz_aligned_addrFromPPtr' invs_queues
2341                        valid_tcb_state'_def invs_sch_act_wf' ThreadState_Restart_def rf_sr_ksCurThread
2342                        arch_invocation_label_defs mask_def isCap_simps)
2343  apply (frule cap_get_tag_isCap_unfolded_H_cap)
2344  by (clarsimp dest!: ccap_relation_PageCap_MappedASID)
2345
2346lemma ccap_relation_PML4Cap_BasePtr:
2347  "ccap_relation (ArchObjectCap (PML4Cap p r)) ccap
2348    \<Longrightarrow> capPML4BasePtr_CL (cap_pml4_cap_lift ccap) = p"
2349  apply (clarsimp simp: ccap_relation_def map_option_Some_eq2)
2350  by (drule (1) cap_lift_PML4Cap_Base, clarsimp)
2351
2352lemma decodeX86ModeMapPage_ccorres:
2353  "ccorres (intr_and_se_rel \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
2354            (invs' and page_map_l4_at' pml4
2355                   and cte_wp_at' \<top> slot and (\<lambda>s. thread = ksCurThread s)
2356                   and K (vmsz_aligned' base vsz \<and> base \<in> kernel_mappings \<and> vsz = X64HugePage
2357                          \<and> isPageCap cap \<and> capVPBasePtr cap = base \<and> capVPMappedAddress cap \<noteq> None))
2358            (UNIV \<inter> \<lbrace>\<acute>label___unsigned_long = scast X86PageMap\<rbrace>
2359                  \<inter> \<lbrace>\<acute>page_size = framesize_from_H vsz\<rbrace>
2360                  \<inter> \<lbrace>\<acute>vroot = pml4e_Ptr pml4\<rbrace>
2361                  \<inter> \<lbrace>ccap_relation (ArchObjectCap cap) \<acute>cap\<rbrace>
2362                  \<inter> \<lbrace>\<acute>paddr = addrFromPPtr base\<rbrace>
2363                  \<inter> \<lbrace>vmrights_to_H \<acute>vm_rights = rights \<and> \<acute>vm_rights < 4 \<and> \<acute>vm_rights \<noteq> 0\<rbrace>
2364                  \<inter> \<lbrace>vm_attribs_relation attribs \<acute>vm_attr\<rbrace>
2365                  \<inter> \<lbrace>\<acute>vaddr___unsigned_long = vaddr\<rbrace>
2366                  \<inter> \<lbrace>\<acute>cte = cte_Ptr slot\<rbrace>)
2367            hs
2368            (doE x <- injection_handler Inl
2369                  (createSafeMappingEntries (addrFromPPtr base) vaddr vsz rights attribs pml4);
2370                 invocationCatch thread isBlocking isCall Invocations_H.invocation.InvokeArchObject
2371                    (Inr (invocation.InvokePage (PageMap (ArchObjectCap cap) slot x pml4)))
2372             odE)
2373            (Call decodeX86ModeMapRemapPage_'proc)"
2374  supply if_cong[cong] tl_drop_1[simp] Collect_const[simp del] dc_simp[simp del]
2375  apply (simp add: K_def)
2376  apply (rule ccorres_gen_asm)
2377  apply (cinit' lift: label___unsigned_long_' page_size_' vroot_' cap_' paddr_' vm_rights_' vm_attr_'
2378                      vaddr___unsigned_long_' cte_'
2379                simp: bindE_assoc injection_handler_bindE)
2380   apply (simp cong: StateSpace.state.fold_congs globals.fold_congs)
2381   apply (clarsimp simp: framesize_from_H_simps, ccorres_rewrite)
2382   apply (rule ccorres_rhs_assoc)+
2383   apply csymbr
2384   apply (ctac add: ccorres_injection_handler_csum1 [OF createSafeMappingEntries_PDPTE_ccorres])
2385      apply (simp add: Collect_False performX64MMUInvocations bindE_assoc bind_assoc)
2386      apply (ctac add: setThreadState_ccorres)
2387        apply (ctac(no_vcg) add: performPageInvocationMapPDPTE_ccorres)
2388          apply (rule ccorres_alternative2)
2389          apply (rule ccorres_return_CE, simp+)[1]
2390         apply (rule ccorres_inst[where P=\<top> and P'=UNIV], simp)
2391        apply wp
2392       apply (wp sts_invs_minor')
2393      apply simp
2394      apply (vcg exspec=setThreadState_modifies)
2395     apply (simp, ccorres_rewrite)
2396     apply (rule ccorres_return_C_errorE, simp+)[1]
2397    apply (simp add: createSafeMappingEntries_def)
2398    apply (wp injection_wp[OF refl] createMappingEntries_wf)
2399   apply (simp add: all_ex_eq_helper)
2400   apply (vcg exspec=createSafeMappingEntries_PDPTE_modifies)
2401  by (clarsimp simp: invs_valid_objs' tcb_at_invs' vmsz_aligned_addrFromPPtr' invs_queues
2402                     valid_tcb_state'_def invs_sch_act_wf' ThreadState_Restart_def rf_sr_ksCurThread
2403                     arch_invocation_label_defs mask_def isCap_simps)
2404
2405lemma valid_cap'_PageCap_kernel_mappings:
2406  "\<lbrakk>pspace_in_kernel_mappings' s; isPageCap cap; valid_cap' (ArchObjectCap cap) s\<rbrakk>
2407     \<Longrightarrow> capVPBasePtr cap \<in> kernel_mappings"
2408  apply (clarsimp simp: valid_cap'_def isCap_simps)
2409  apply (drule_tac x=0 in spec)
2410  apply (case_tac v3; fastforce simp: bit_simps typ_at_to_obj_at_arches obj_at_kernel_mappings'
2411                              split: if_splits)
2412  done
2413
2414lemma ccap_relation_PageCap_MappedAddress_update:
2415  "\<lbrakk>cap_lift cap = Some (Cap_frame_cap (cap_frame_cap_lift cap));
2416    cp = PageCap f_base f_vmr mt sz f_dev f_addr;
2417    cap_to_H (Cap_frame_cap (cap_frame_cap_lift cap)) = ArchObjectCap cp;
2418    c_valid_cap cap; 0 < asid; asid_wf asid;
2419    cap_frame_cap_lift cap' = cap_frame_cap_lift cap
2420        \<lparr>capFMappedASID_CL := asid && mask 16,
2421           capFMappedAddress_CL := sign_extend 47 vaddr,
2422           capFMapType_CL := newmt\<rparr>;
2423    cap_lift cap' = Some (Cap_frame_cap (cap_frame_cap_lift cap
2424            \<lparr>capFMappedASID_CL := asid && mask 16,
2425               capFMappedAddress_CL := sign_extend 47 vaddr,
2426               capFMapType_CL := newmt\<rparr>));
2427    newmt = scast X86_MappingVSpace; canonical_address vaddr\<rbrakk>
2428    \<Longrightarrow> ccap_relation (ArchObjectCap (PageCap f_base f_vmr VMVSpaceMap sz f_dev (Some (asid, vaddr)))) cap'"
2429   apply (clarsimp simp: ccap_relation_def map_option_Some_eq2 cap_frame_cap_lift[THEN iffD1]
2430                         cap_to_H_simps asid_wf_def3[simplified asid_bits_def, simplified])
2431   apply (rule conjI, clarsimp simp: maptype_to_H_def vm_page_map_type_defs mask_def)
2432   apply (rule conjI,
2433           clarsimp simp: sign_extend_canonical_address le_def[symmetric] mask_def word_bw_assocs
2434                          le_user_vtop_canonical_address[simplified user_vtop_def
2435                          X64.pptrUserTop_def word_le_nat_alt] vm_page_map_type_defs
2436                          canonical_address_user_vtop[simplified user_vtop_def X64.pptrUserTop_def]
2437                   split: if_split)
2438    apply (rule conjI; word_bitwise, clarsimp)
2439   by (clarsimp simp: c_valid_cap_def cl_valid_cap_def vm_page_map_type_defs mask_def)
2440
2441lemma framesize_to_from_H:
2442  "sz < 3 \<Longrightarrow> framesize_from_H (framesize_to_H sz) = sz"
2443   apply (clarsimp simp: framesize_to_H_def framesize_from_H_def
2444                Kernel_C.X86_SmallPage_def Kernel_C.X86_LargePage_def
2445                Kernel_C.X64_HugePage_def
2446           split: if_split vmpage_size.splits)
2447  by (word_bitwise, auto)
2448
2449lemma decodeX64FrameInvocation_ccorres:
2450  notes if_cong[cong] tl_drop_1[simp] Collect_const[simp del]
2451  shows
2452  "\<lbrakk> interpret_excaps extraCaps' = excaps_map extraCaps;
2453          isPageCap cp \<rbrakk>
2454     \<Longrightarrow>
2455   ccorres (intr_and_se_rel \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
2456       (invs' and (\<lambda>s. ksCurThread s = thread) and ct_active' and sch_act_simple
2457              and (excaps_in_mem extraCaps \<circ> ctes_of)
2458              and cte_wp_at' (diminished' (ArchObjectCap cp) \<circ> cteCap) slot
2459              and (\<lambda>s. \<forall>v \<in> set extraCaps. ex_cte_cap_wp_to' isCNodeCap (snd v) s)
2460              and sysargs_rel args buffer and valid_objs')
2461       (UNIV \<inter> {s. invLabel_' s = label}
2462             \<inter> {s. unat (length___unsigned_long_' s) = length args}
2463             \<inter> {s. cte_' s = cte_Ptr slot}
2464             \<inter> {s. excaps_' s = extraCaps'}
2465             \<inter> {s. ccap_relation (ArchObjectCap cp) (cap_' s)}
2466             \<inter> {s. buffer_' s = option_to_ptr buffer}) []
2467       (decodeX64MMUInvocation label args cptr slot cp extraCaps
2468              >>= invocationCatch thread isBlocking isCall InvokeArchObject)
2469       (Call decodeX86FrameInvocation_'proc)"
2470  apply (clarsimp simp only: isCap_simps) using [[goals_limit=30]]
2471  apply (cinit' lift: invLabel_' length___unsigned_long_' cte_' excaps_' cap_' buffer_'
2472                simp: decodeX64MMUInvocation_def )
2473   apply (simp add: Let_def isCap_simps invocation_eq_use_types split_def decodeX64FrameInvocation_def
2474               del: Collect_const
2475              cong: StateSpace.state.fold_congs globals.fold_congs
2476                    if_cong invocation_label.case_cong arch_invocation_label.case_cong list.case_cong)
2477   apply (rule ccorres_Cond_rhs[rotated])+
2478         apply (rule ccorres_inst[where P=\<top> and P'=UNIV])
2479        apply (rule ccorres_equals_throwError)
2480         apply (fastforce simp: throwError_bind invocationCatch_def
2481                        split: invocation_label.split arch_invocation_label.split)
2482        apply (rule syscall_error_throwError_ccorres_n)
2483        apply (simp add: syscall_error_to_H_cases)
2484       (* PageGetAddress *)
2485       apply (simp add: returnOk_bind bindE_assoc performX64MMUInvocations)
2486       apply (rule ccorres_rhs_assoc)+
2487       apply (ctac add: setThreadState_ccorres)
2488         apply csymbr
2489         apply (ctac(no_vcg) add: performPageGetAddress_ccorres)
2490           apply (rule ccorres_alternative2)
2491           apply (rule ccorres_return_CE, simp+)[1]
2492          apply (rule ccorres_inst[where P=\<top> and P'=UNIV], simp)
2493         apply wp+
2494       apply (vcg exspec=setThreadState_modifies)
2495      apply (rule ccorres_rhs_assoc)+
2496
2497(* X64PageUnmap *)
2498     apply (simp add: returnOk_bind bindE_assoc
2499                      performX64MMUInvocations)
2500     apply (ctac add: setThreadState_ccorres)
2501       apply (ctac(no_vcg) add: performPageInvocationUnmap_ccorres)
2502         apply (rule ccorres_alternative2)
2503         apply (rule ccorres_return_CE, simp+)[1]
2504        apply (rule ccorres_inst[where P=\<top> and P'=UNIV], simp)
2505       apply wp
2506      apply (wp sts_invs_minor')
2507     apply simp
2508     apply (vcg exspec=setThreadState_modifies)
2509(* Remap*)
2510    apply (rule ccorres_rhs_assoc)+
2511    apply csymbr+
2512    apply (simp add: if_1_0_0 word_less_nat_alt del: Collect_const)
2513    apply (rule ccorres_Cond_rhs_Seq)
2514     apply (rule ccorres_equals_throwError)
2515      apply (fastforce simp: throwError_bind invocationCatch_def
2516                     split: list.split)
2517     apply (simp del: Collect_const)
2518     apply (rule ccorres_cond_true_seq)
2519     apply (rule syscall_error_throwError_ccorres_n)
2520     apply (simp add: syscall_error_to_H_cases)
2521    apply csymbr
2522    apply (simp add: if_1_0_0 interpret_excaps_test_null
2523                     excaps_map_def del: Collect_const)
2524    apply (rule ccorres_Cond_rhs_Seq)
2525     apply (rule ccorres_equals_throwError)
2526      apply (fastforce simp: throwError_bind invocationCatch_def
2527                     split: list.split)
2528     apply (rule syscall_error_throwError_ccorres_n)
2529     apply (simp add: syscall_error_to_H_cases)
2530    apply (clarsimp simp: list_case_If2 dest!: at_least_2_args)
2531    apply csymbr
2532    apply (frule ccap_relation_PageCap_MapType)
2533    apply clarsimp
2534    apply (rule ccorres_Cond_rhs_Seq)
2535     apply ccorres_rewrite
2536     apply (clarsimp simp: maptype_from_H_def throwError_bind invocationCatch_def
2537                    split: vmmap_type.split_asm)
2538     apply (rule syscall_error_throwError_ccorres_n[simplified id_def o_def dc_def])
2539     apply (clarsimp simp: syscall_error_to_H_cases)
2540    apply (clarsimp simp: maptype_from_H_def vm_page_map_type_defs
2541                   split: vmmap_type.split_asm)
2542    apply (rule ccorres_add_return)
2543    apply (ctac add: getSyscallArg_ccorres_foo[where args=args and n=0 and buffer=buffer])
2544      apply (rule ccorres_add_return)
2545      apply (ctac add: getSyscallArg_ccorres_foo[where args=args and n=1 and buffer=buffer])
2546        apply csymbr
2547        apply (rule getSlotCap_ccorres_fudge_n[where vals=extraCaps and n=0])
2548        apply (rule ccorres_move_c_guard_cte)
2549        apply (rule_tac r'="\<lambda>rv rv'. ((cap_get_tag rv' = scast cap_pml4_cap)
2550                                             = (isArchObjectCap rv \<and> isPML4Cap (capCap rv)))
2551                                     \<and> (ccap_relation rv rv')"
2552                   and xf'=vspaceCap_' in ccorres_split_nothrow[where F=UNIV])
2553            apply (simp add: getSlotCap_def del: Collect_const)
2554            apply (rule ccorres_symb_exec_l[OF _ getCTE_inv getCTE_sp[where P=\<top>]
2555                        empty_fail_getCTE])
2556            apply (rule ccorres_from_vcg[where P'=UNIV])
2557            apply (rule allI, rule conseqPre, vcg)
2558            apply (clarsimp simp: return_def cte_wp_at_ctes_of)
2559            apply (erule(1) cmap_relationE1[OF cmap_relation_cte])
2560            apply (frule interpret_excaps_eq[rule_format, where n=0], simp)
2561            apply (clarsimp simp: typ_heap_simps' mask_def split_def
2562                                  cap_get_tag_isCap_ArchObject2
2563                                  word_sless_def word_sle_def
2564                           dest!: ccte_relation_ccap_relation)
2565           apply (simp add:sless_positive)
2566           apply ceqv
2567          apply (rule ccorres_assert2)
2568          apply csymbr+
2569          apply (clarsimp simp: split_def cap_case_PML4Cap2 if_1_0_0)
2570          apply (rule ccorres_Cond_rhs_Seq)
2571           apply (rule ccorres_equals_throwError)
2572            apply (fastforce simp: invocationCatch_def throwError_bind
2573                                   hd_conv_nth from_bool_0
2574                             cong: conj_cong)
2575           apply ccorres_rewrite
2576           apply (rule syscall_error_throwError_ccorres_n)
2577           apply (clarsimp simp: syscall_error_to_H_cases)
2578          apply (clarsimp simp: hd_conv_nth)
2579          apply csymbr+
2580          apply (clarsimp simp add: case_option_If2 if_to_top_of_bind if_to_top_of_bindE
2581                           hd_conv_nth cong: conj_cong)
2582          apply (rule ccorres_if_cond_throws2[rotated -1, where Q=\<top> and Q'=\<top>])
2583             apply vcg
2584            apply (clarsimp simp: if_1_0_0)
2585            apply (frule ccap_relation_mapped_asid_0, clarsimp simp: asidInvalid_def)
2586            apply (erule_tac ccap_relationE,
2587                     clarsimp simp: cap_to_H_def Let_def split: cap_CL.split_asm if_split_asm)
2588           apply (simp add: throwError_bind invocationCatch_def)
2589           apply (rule syscall_error_throwError_ccorres_n)
2590           apply (simp add: syscall_error_to_H_cases)
2591          apply csymbr
2592          apply csymbr
2593          apply csymbr
2594          apply csymbr
2595          apply (simp add: createSafeMappingEntries_fold)
2596          apply (simp add: whenE_bindE_throwError_to_if
2597                           invocationCatch_use_injection_handler
2598                           injection_bindE[OF refl refl] bindE_assoc
2599                           injection_handler_returnOk
2600                           injection_handler_If if_to_top_of_bindE
2601                           lookupError_injection)
2602          apply (ctac add: ccorres_injection_handler_csum1
2603                                [OF ccorres_injection_handler_csum1,
2604                                 OF findVSpaceForASID_ccorres])
2605             apply (simp add: Collect_False del: Collect_const)
2606             apply (rule ccorres_Cond_rhs_Seq)
2607              apply (clarsimp simp: isCap_simps ccap_relation_PageCap_MappedASID)
2608              apply (clarsimp simp: cap_pml4_cap_lift get_capMappedASID_CL_def get_capPtr_CL_def)
2609              apply (erule_tac v=rv' in ccap_relationE,
2610                        clarsimp simp: cap_to_H_def Let_def split: cap_CL.split_asm if_split_asm)
2611              apply (clarsimp simp: invocationCatch_def throwError_bind
2612                                    injection_handler_throwError)
2613              apply (rule syscall_error_throwError_ccorres_n)
2614              apply (simp add: syscall_error_to_H_cases)
2615             apply (clarsimp simp: isCap_simps ccap_relation_PageCap_MappedASID)
2616             apply (clarsimp simp: cap_pml4_cap_lift get_capMappedASID_CL_def get_capPtr_CL_def)
2617             apply (erule_tac v=rv' in ccap_relationE,
2618                      clarsimp simp: cap_to_H_def Let_def split: cap_CL.split_asm if_split_asm)
2619             apply csymbr
2620             apply csymbr
2621             apply csymbr
2622             apply csymbr
2623             apply csymbr
2624             apply csymbr
2625             apply csymbr
2626             apply csymbr
2627             apply csymbr
2628             apply (rule ccorres_symb_exec_r)
2629               apply csymbr
2630               apply (clarsimp simp: checkVPAlignment_def unlessE_def
2631                                     injection_handler_If if_to_top_of_bindE)
2632               apply wpfix
2633               apply (rule ccorres_if_cond_throws2[rotated -1, where Q=\<top> and Q'=\<top>])
2634                  apply vcg
2635                 apply (clarsimp simp: from_bool_0 vmsz_aligned'_def is_aligned_mask)
2636                 apply (frule cap_get_tag_isCap_unfolded_H_cap(18) (* PageCap *))
2637                 apply (drule cap_get_tag_PageCap_frame)
2638                 apply (clarsimp split: if_split_asm)
2639                apply (simp add: injection_handler_throwError throwError_bind
2640                                 invocationCatch_def)
2641                apply (rule syscall_error_throwError_ccorres_n)
2642                apply (simp add: syscall_error_to_H_cases)
2643               apply (rule ccorres_Cond_rhs)
2644                apply (simp add: injection_handler_returnOk)
2645                apply (rule ccorres_rhs_assoc)+
2646                apply csymbr
2647                apply (ctac add: ccorres_injection_handler_csum1[OF createSafeMappingEntries_PTE_ccorres])
2648                   apply (simp add: Collect_False performX64MMUInvocations bindE_assoc)
2649                   apply (rule_tac P="\<lambda>s. thread = ksCurThread s" in ccorres_cross_over_guard)
2650                   apply (ctac add: setThreadState_ccorres)
2651                     apply (ctac(no_vcg) add: performPageInvocationRemapPTE_ccorres)
2652                      apply (rule ccorres_alternative2)
2653                      apply (rule ccorres_return_CE, simp+)[1]
2654                      apply (rule ccorres_inst[where P=\<top> and P'=UNIV], simp)
2655                     apply wp
2656                    apply (wp sts_invs_minor')
2657                   apply simp
2658                   apply (vcg exspec=setThreadState_modifies)
2659                  apply (simp, ccorres_rewrite)
2660                  apply (rule ccorres_return_C_errorE, simp+)[1]
2661                 apply (simp add: createSafeMappingEntries_def)
2662                 apply (wp injection_wp[OF refl] createMappingEntries_wf)
2663                apply (simp add: all_ex_eq_helper)
2664                apply (rule_tac Q'=UNIV and A'="{}" in conseqPost)
2665                  apply (vcg exspec=createSafeMappingEntries_PTE_modifies)
2666                 apply (clarsimp simp: ThreadState_Restart_def mask_def rf_sr_ksCurThread)
2667                apply clarsimp
2668               apply (simp add: injection_handler_returnOk)
2669               apply (rule ccorres_Cond_rhs)
2670                apply (rule ccorres_rhs_assoc)+
2671                apply csymbr
2672                apply (ctac add: ccorres_injection_handler_csum1[OF createSafeMappingEntries_PDE_ccorres])
2673                   apply (simp add: performX64MMUInvocations bindE_assoc)
2674                   apply ccorres_rewrite
2675                   apply (rule_tac P="\<lambda>s. thread = ksCurThread s" in ccorres_cross_over_guard)
2676                   apply (ctac add: setThreadState_ccorres)
2677                     apply (ctac(no_vcg) add: performPageInvocationRemapPDE_ccorres)
2678                      apply (rule ccorres_alternative2)
2679                      apply (rule ccorres_return_CE, simp+)[1]
2680                      apply (rule ccorres_inst[where P=\<top> and P'=UNIV], simp)
2681                     apply wp
2682                    apply (wp sts_invs_minor')
2683                   apply simp
2684                  apply (vcg exspec=setThreadState_modifies)
2685                 apply (simp, ccorres_rewrite)
2686                  apply (rule ccorres_return_C_errorE, simp+)[1]
2687                apply (simp add: createSafeMappingEntries_def)
2688                apply (wp injection_wp[OF refl] createMappingEntries_wf)
2689               apply (simp add: all_ex_eq_helper)
2690                apply (rule_tac Q'=UNIV and A'="{}" in conseqPost)
2691                  apply (vcg exspec=createSafeMappingEntries_PDE_modifies)
2692                 apply (clarsimp simp: ThreadState_Restart_def mask_def rf_sr_ksCurThread)
2693                apply clarsimp
2694               apply (rule ccorres_add_returnOk)
2695               apply (ctac add: decodeX86ModeRemapPage_ccorres[where cap=cp])
2696                  apply (rule ccorres_return_CE, simp+)[1]
2697                 apply (rule ccorres_return_C_errorE, simp+)[1]
2698                apply wpsimp
2699               apply clarsimp
2700               apply (vcg exspec=decodeX86ModeMapRemapPage_modifies)
2701              apply clarsimp
2702              apply (rule_tac Q' = "{x. vmrights_to_H (vmRights_' x) = maskVMRights v1 (rightsFromWord a)
2703                                    \<and> vmRights_' x < 4 \<and> vmRights_' x \<noteq> 0}"
2704                          and A' = "{}"
2705                          in conseqPost[rotated])
2706                apply (clarsimp simp: Collect_const)
2707                apply (frule cap_get_tag_isCap_unfolded_H_cap)
2708                apply (rule conjI, clarsimp)
2709                 apply (clarsimp simp: ccap_relation_PageCap_BasePtr ccap_relation_PageCap_MappedAddress
2710                                       ccap_relation_PageCap_Size attribsFromWord_def
2711                                       vm_attribs_relation_def invocation_eq_use_types
2712                                       of_bool_nth[simplified of_bool_from_bool])
2713                apply (clarsimp dest!: ccap_relation_c_valid_cap
2714                                 simp: c_valid_cap_def cl_valid_cap_def cap_frame_cap_lift)
2715               apply clarsimp
2716              apply vcg
2717             apply (rule conseqPre, vcg, clarsimp)
2718            apply ccorres_rewrite
2719            apply (rule_tac P'="{s. find_ret = errstate s}" in ccorres_from_vcg_throws[where P=\<top>])
2720            apply (rule allI, rule conseqPre, vcg)
2721            apply (clarsimp simp: fst_throwError_returnOk exception_defs
2722                                  syscall_error_rel_def syscall_error_to_H_cases
2723                                  false_def)
2724            apply (erule lookup_failure_rel_fault_lift[rotated])
2725            apply (simp add: exception_defs)
2726           apply (frule ccap_relation_PageCap_Size, clarsimp simp: framesize_from_H_eqs)
2727           apply ((wp injection_wp[OF refl] static_imp_wp hoare_vcg_const_imp_lift_R
2728                       | simp | wp_once hoare_drop_imps)+)[1]
2729          apply simp
2730          apply (vcg exspec=findVSpaceForASID_modifies)
2731         apply (wp getSlotCap_wp)
2732        apply simp
2733        apply (vcg exspec=getSyscallArg_modifies)
2734       apply simp
2735       apply wp
2736      apply simp
2737      apply (vcg exspec=getSyscallArg_modifies)
2738     apply simp
2739     apply (rule_tac s=args in ssubst, simp)
2740     apply wp
2741    apply simp
2742    apply (rule_tac t=a and s="hd args" in ssubst, fastforce)
2743    apply (rule_tac t=b and s="hd (tl args)" in ssubst, fastforce)
2744    apply (vcg exspec=getSyscallArg_modifies)
2745   \<comment> \<open>PageMap\<close>
2746   apply (rule ccorres_rhs_assoc)+
2747   apply csymbr+
2748   apply (simp add: if_1_0_0 word_less_nat_alt del: Collect_const)
2749   apply (rule ccorres_Cond_rhs_Seq)
2750    apply (rule ccorres_equals_throwError)
2751     apply (fastforce simp: throwError_bind invocationCatch_def
2752                    split: list.split)
2753    apply (simp del: Collect_const)
2754    apply (rule ccorres_cond_true_seq)
2755    apply (rule syscall_error_throwError_ccorres_n)
2756    apply (simp add: syscall_error_to_H_cases)
2757   apply csymbr
2758   apply (simp add: if_1_0_0 interpret_excaps_test_null
2759                    excaps_map_def del: Collect_const)
2760   apply (rule ccorres_Cond_rhs_Seq)
2761    apply (rule ccorres_equals_throwError)
2762     apply (fastforce simp: throwError_bind invocationCatch_def
2763                    split: list.split)
2764    apply (rule syscall_error_throwError_ccorres_n)
2765    apply (simp add: syscall_error_to_H_cases)
2766   apply (clarsimp simp: list_case_If2 dest!: at_least_3_args)
2767   apply csymbr
2768   apply csymbr
2769   apply (rule ccorres_add_return)
2770   apply (ctac add: getSyscallArg_ccorres_foo[where args=args and n=0 and buffer=buffer])
2771     apply (rule ccorres_add_return)
2772     apply (ctac add: getSyscallArg_ccorres_foo[where args=args and n=1 and buffer=buffer])
2773       apply (rule ccorres_add_return)
2774       apply (ctac add: getSyscallArg_ccorres_foo[where args=args and n=2 and buffer=buffer])
2775         apply csymbr
2776         apply (rule getSlotCap_ccorres_fudge_n[where vals=extraCaps and n=0])
2777         apply (rule ccorres_move_c_guard_cte)
2778         apply (rule_tac r'="\<lambda>rv rv'. ((cap_get_tag rv' = scast cap_pml4_cap)
2779                                              = (isArchObjectCap rv \<and> isPML4Cap (capCap rv)))
2780                                      \<and> (ccap_relation rv rv')"
2781                    and xf'=vspaceCap_' in ccorres_split_nothrow[where F=UNIV])
2782             apply (simp add: getSlotCap_def del: Collect_const)
2783             apply (rule ccorres_symb_exec_l[OF _ getCTE_inv getCTE_sp[where P=\<top>]
2784                         empty_fail_getCTE])
2785             apply (rule ccorres_from_vcg[where P'=UNIV])
2786             apply (rule allI, rule conseqPre, vcg)
2787             apply (clarsimp simp: return_def cte_wp_at_ctes_of)
2788             apply (erule(1) cmap_relationE1[OF cmap_relation_cte])
2789             apply (frule interpret_excaps_eq[rule_format, where n=0], simp)
2790             apply (clarsimp simp: typ_heap_simps' mask_def split_def
2791                                   cap_get_tag_isCap_ArchObject2
2792                                   word_sle_def word_sless_def
2793                            dest!: ccte_relation_ccap_relation)
2794            apply (simp add:sless_positive)
2795            apply ceqv
2796           apply (rule ccorres_assert2)
2797           apply csymbr+
2798           apply (simp add: if_1_0_0 whenE_bindE_throwError_to_if if_to_top_of_bind)
2799           apply (rule ccorres_if_cond_throws[rotated -1, where Q=\<top> and Q'=\<top>])
2800              apply vcg
2801             apply (drule ccap_relation_mapped_asid_0)
2802             apply (fastforce simp: asidInvalid_def)
2803            apply (simp add: throwError_bind invocationCatch_def)
2804            apply (rule syscall_error_throwError_ccorres_n)
2805            apply (simp add: syscall_error_to_H_cases)
2806           apply csymbr+
2807           apply (clarsimp simp add: split_def cap_case_PML4Cap2 if_1_0_0)
2808           apply (rule ccorres_Cond_rhs_Seq)
2809            apply (rule ccorres_equals_throwError)
2810             apply (fastforce simp: invocationCatch_def throwError_bind
2811                                    hd_conv_nth from_bool_0
2812                              cong: conj_cong)
2813            apply ccorres_rewrite
2814            apply (rule syscall_error_throwError_ccorres_n)
2815            apply (clarsimp simp: syscall_error_to_H_cases)
2816           apply (clarsimp simp: hd_conv_nth)
2817           apply csymbr
2818           apply csymbr
2819           apply csymbr
2820           apply csymbr
2821           apply (simp add: createSafeMappingEntries_fold)
2822           apply (simp add: lookupError_injection invocationCatch_use_injection_handler
2823                            injection_bindE[OF refl refl] injection_handler_returnOk
2824                            injection_handler_If injection_handler_throwError bindE_assoc)
2825           apply (rule_tac t = y
2826                       and s = "the (capPML4MappedASID (capCap (fst (extraCaps ! 0)))) "
2827                       in ssubst, fastforce)
2828           apply (ctac add: ccorres_injection_handler_csum1
2829                                  [OF ccorres_injection_handler_csum1,
2830                                   OF findVSpaceForASID_ccorres])
2831              apply (simp add: Collect_False if_to_top_of_bindE)
2832              apply (rule ccorres_if_cond_throws[rotated -1, where Q=\<top> and Q'=\<top>])
2833                 apply vcg
2834                apply (clarsimp simp: cap_lift_pml4_cap cap_to_H_def get_capPtr_CL_def
2835                                      to_bool_def cap_pml4_cap_lift_def
2836                               elim!: ccap_relationE split: if_split)
2837               apply (rule syscall_error_throwError_ccorres_n[simplified id_def o_def dc_def])
2838               apply (simp add: syscall_error_to_H_cases)
2839              apply csymbr+
2840              apply (rule ccorres_Guard_Seq)
2841              apply csymbr
2842              apply (simp add: if_to_top_of_bindE)
2843              apply (rule ccorres_if_cond_throws[rotated -1, where Q=\<top> and Q'=\<top>])
2844                 apply vcg
2845                apply (frule ccap_relation_PageCap_Size)
2846                apply (clarsimp simp: word_less_nat_alt framesize_from_to_H user_vtop_def
2847                                      X64.pptrUserTop_def)
2848               apply (simp add: throwError_bind invocationCatch_def)?
2849               apply (rule syscall_error_throwError_ccorres_n[simplified id_def o_def dc_def])
2850               apply (simp add: syscall_error_to_H_cases)
2851              apply csymbr
2852              apply (rule ccorres_symb_exec_r)
2853                apply csymbr
2854                apply (simp add: bindE_assoc checkVPAlignment_def unlessE_def
2855                                 injection_handler_If if_to_top_of_bindE)
2856                apply (rule ccorres_if_cond_throws2[rotated -1, where Q=\<top> and Q'=\<top>])
2857                   apply vcg
2858                  apply (clarsimp simp add: from_bool_0 vmsz_aligned'_def is_aligned_mask)
2859                  apply (drule ccap_relation_PageCap_Size)
2860                  apply (clarsimp simp: framesize_from_to_H user_vtop_def X64.pptrUserTop_def)
2861                 apply (simp add: injection_handler_throwError throwError_bind
2862                                  invocationCatch_def)
2863                 apply (rule syscall_error_throwError_ccorres_n[simplified id_def o_def dc_def])
2864                 apply (simp add: syscall_error_to_H_cases)
2865                apply csymbr
2866                apply csymbr
2867                apply csymbr
2868                apply csymbr
2869                apply csymbr
2870                apply (simp add: injection_handler_returnOk bindE_assoc)
2871                apply (rule ccorres_Cond_rhs)
2872                 apply (rule ccorres_rhs_assoc)+
2873                 apply csymbr
2874                 apply (ctac add: ccorres_injection_handler_csum1[OF createSafeMappingEntries_PTE_ccorres])
2875                    apply (simp add: performX64MMUInvocations bindE_assoc)
2876                    apply ccorres_rewrite
2877                    apply (rule_tac P="\<lambda>s. thread = ksCurThread s" in ccorres_cross_over_guard)
2878                    apply (ctac add: setThreadState_ccorres)
2879                      apply (ctac(no_vcg) add: performPageInvocationMapPTE_ccorres)
2880                      apply (rule ccorres_alternative2)
2881                      apply (rule ccorres_return_CE, simp+)[1]
2882                      apply (rule ccorres_inst[where P=\<top> and P'=UNIV], simp)
2883                      apply (wp sts_invs_minor')+
2884                    apply simp
2885                    apply (vcg exspec=setThreadState_modifies)
2886                   apply (simp, ccorres_rewrite)
2887                   apply (rule ccorres_return_C_errorE, simp+)[1]
2888                  apply (simp add: createSafeMappingEntries_def)
2889                  apply (wp injection_wp[OF refl] createMappingEntries_wf)
2890                 apply (simp add: all_ex_eq_helper)
2891                apply (rule_tac Q' = "{s. ccap_relation (ArchObjectCap
2892                                            (PageCap v0 v1 VMVSpaceMap v3 d
2893                                                          (Some (y, a)))) cap}"
2894                            and A' = "{}" in conseqPost)
2895                  apply (vcg exspec=createSafeMappingEntries_PTE_modifies)
2896                 apply (clarsimp simp: ThreadState_Restart_def mask_def rf_sr_ksCurThread
2897                                       isCap_simps cap_pml4_cap_lift
2898                                       get_capPtr_CL_def ccap_relation_PML4Cap_BasePtr)
2899                apply clarsimp
2900                apply (rule ccorres_Cond_rhs)
2901                 apply (rule ccorres_rhs_assoc)+
2902                 apply csymbr
2903                 apply (ctac add: ccorres_injection_handler_csum1
2904                                    [OF createSafeMappingEntries_PDE_ccorres])
2905                    apply (simp add: performX64MMUInvocations bindE_assoc)
2906                    apply ccorres_rewrite
2907                    apply (rule_tac P="\<lambda>s. thread = ksCurThread s" in ccorres_cross_over_guard)
2908                    apply (ctac add: setThreadState_ccorres)
2909                      apply (ctac(no_vcg) add: performPageInvocationMapPDE_ccorres)
2910                      apply (rule ccorres_alternative2)
2911                      apply (rule ccorres_return_CE, simp+)[1]
2912                      apply (rule ccorres_inst[where P=\<top> and P'=UNIV], simp)
2913                      apply wp
2914                     apply (wp sts_invs_minor')
2915                    apply simp
2916                    apply (vcg exspec=setThreadState_modifies)
2917                   apply (simp add: Collect_const)
2918                   apply (rule ccorres_split_throws)
2919                    apply (rule ccorres_return_C_errorE, simp+)[1]
2920                   apply vcg
2921                  apply (simp add: createSafeMappingEntries_def)
2922                  apply (wp injection_wp[OF refl] createMappingEntries_wf)
2923                 apply (simp add: all_ex_eq_helper)
2924                apply (rule_tac Q' = "{s. ccap_relation (ArchObjectCap
2925                                            (PageCap v0 v1 VMVSpaceMap v3 d
2926                                                          (Some (y, a)))) cap}"
2927                            and A' = "{}" in conseqPost)
2928                  apply (vcg exspec=createSafeMappingEntries_PDE_modifies)
2929                 apply (clarsimp simp: ThreadState_Restart_def mask_def rf_sr_ksCurThread
2930                                       isCap_simps cap_pml4_cap_lift
2931                                       get_capPtr_CL_def ccap_relation_PML4Cap_BasePtr)
2932                apply clarsimp
2933                apply (rule ccorres_add_returnOk)
2934                apply (ctac add: decodeX86ModeMapPage_ccorres)
2935                  apply (rule ccorres_return_CE, simp+)[1]
2936                 apply (rule ccorres_return_C_errorE, simp+)[1]
2937                apply wpsimp
2938               apply clarsimp
2939               apply (vcg exspec=decodeX86ModeMapRemapPage_modifies)
2940               apply clarsimp
2941               apply vcg
2942              apply (rule conseqPre, vcg, clarsimp)
2943             apply simp
2944             apply (rule_tac P'="{s. find_ret = errstate s}"
2945                      in ccorres_from_vcg_split_throws[where P=\<top>])
2946              apply vcg
2947             apply (rule conseqPre, vcg)
2948             apply (clarsimp simp: fst_throwError_returnOk exception_defs
2949                                   syscall_error_rel_def syscall_error_to_H_cases
2950                                   false_def)
2951             apply (erule lookup_failure_rel_fault_lift[rotated])
2952             apply (simp add: exception_defs)
2953            apply (simp add: isCap_simps)
2954            apply (wp injection_wp[OF refl]
2955                        | wp_once hoare_drop_imps)+
2956           apply (simp add: all_ex_eq_helper)
2957           apply (rule_tac t = y and
2958                           s = "the (capPML4MappedASID (capCap (fst (extraCaps ! 0))))"
2959                     in ssubst, fastforce)
2960           apply (vcg exspec=findVSpaceForASID_modifies)
2961          apply (wp getSlotCap_wp)
2962         apply (simp add: if_1_0_0 del: Collect_const)
2963         apply vcg
2964        apply simp
2965        apply wp
2966       apply simp
2967       apply (vcg exspec=getSyscallArg_modifies)
2968      apply simp
2969      apply wp
2970     apply simp
2971     apply (vcg exspec=getSyscallArg_modifies)
2972    apply simp
2973    apply (rule_tac s=args in ssubst, simp)
2974    apply (rule_tac t=a and s="hd args" in ssubst, fastforce)
2975    apply wp
2976   apply clarsimp
2977   apply (rule_tac t=a and s="hd args" in ssubst, fastforce)
2978   apply (rule_tac t=b and s="hd (tl args)" in ssubst, fastforce)
2979   apply (rule_tac t=c and s="hd (tl (tl args))" in ssubst, fastforce)
2980   apply (vcg exspec=getSyscallArg_modifies)
2981  apply (rule conjI)
2982   apply clarsimp
2983   apply (frule cte_wp_at_diminished_gsMaxObjectSize, clarsimp)
2984   apply (clarsimp simp: cte_wp_at_ctes_of
2985                         is_aligned_mask[symmetric] vmsz_aligned'_def
2986                         vmsz_aligned_addrFromPPtr)
2987   apply (frule ctes_of_valid', clarsimp+)
2988   apply (simp add: diminished_valid'[symmetric])
2989   apply (frule valid_cap'_PageCap_kernel_mappings[OF invs_pspace_in_kernel_mappings', where cap=cp],
2990             (fastforce simp: isCap_simps)+)[1]
2991   apply (clarsimp simp: valid_cap'_def capAligned_def
2992                         mask_def[where n=asid_bits]
2993                         linorder_not_le simp del: less_1_simp)
2994   apply (subgoal_tac "extraCaps \<noteq> [] \<longrightarrow> (s \<turnstile>' fst (extraCaps ! 0))")
2995    subgoal by (timeit \<open>(
2996             (clarsimp simp: ct_in_state'_def vmsz_aligned'_def isCap_simps
2997                             valid_cap'_def page_map_l4_at'_def tcb_at_invs'
2998                             sysargs_rel_to_n linorder_not_less framesize_from_H_eqs
2999                             excaps_map_def valid_tcb_state'_def split_def is_aligned_addrFromPPtr
3000                             is_aligned_addrFromPPtr_pageBitsForSize[where sz=X64LargePage, simplified]
3001                             is_aligned_addrFromPPtr_pageBitsForSize[where sz=X64HugePage, simplified]
3002                   simp del: less_1_simp
3003           | rule conjI | erule pred_tcb'_weakenE disjE
3004           | (erule order_trans[where x=7, rotated], fastforce simp: bit_simps)
3005           | (frule ccap_relation_capFMappedASID_CL_0 ccap_relation_PageCap_MappedASID, drule ccap_relation_PageCap_Size)
3006           | drule interpret_excaps_eq st_tcb_at_idle_thread'
3007           | solves \<open>clarsimp simp: framesize_from_H_def asid_wf_def split: vmpage_size.splits\<close>
3008       )+)[1]\<close>) (* 55 secs *)
3009   apply (clarsimp simp: neq_Nil_conv excaps_in_mem_def slotcap_in_mem_def
3010                         linorder_not_le)
3011   apply (erule ctes_of_valid', clarsimp)
3012  apply (clarsimp simp: if_1_0_0 rf_sr_ksCurThread "StrictC'_thread_state_defs"
3013                        mask_eq_iff_w2p word_size word_less_nat_alt
3014                        from_bool_0 excaps_map_def cte_wp_at_ctes_of)
3015  apply (frule ctes_of_valid', clarsimp)
3016  apply (simp only: diminished_valid'[symmetric])
3017  apply (clarsimp simp: valid_cap'_def capAligned_def word_sless_def word_sle_def)
3018  apply (frule cap_get_tag_isCap_unfolded_H_cap)
3019  apply clarsimp
3020  apply (frule cap_get_tag_PageCap_frame)
3021 apply (clarsimp simp: asidInvalid_def)
3022  apply (clarsimp simp: word_less_nat_alt vm_attribs_relation_def
3023                        attribsFromWord_def framesize_from_H_eqs
3024                        of_bool_nth[simplified of_bool_from_bool]
3025                        vm_page_size_defs neq_Nil_conv excaps_in_mem_def
3026                        hd_conv_nth numeral_2_eq_2)
3027  apply (rule conjI, clarsimp)
3028   apply (frule interpret_excaps_eq[rule_format, where n=0], simp)
3029   apply (frule(1) slotcap_in_mem_PML4)
3030   apply (clarsimp simp: mask_def[where n=4] typ_heap_simps')
3031   apply (clarsimp simp: isCap_simps)
3032   apply (frule slotcap_in_mem_valid, (clarsimp simp: get_capMappedASID_CL_def)+)
3033   apply (erule_tac c="ArchObjectCap (PML4Cap a b)" for a b in ccap_relationE)
3034   apply (clarsimp simp: cap_lift_pml4_cap to_bool_def
3035                         cap_pml4_cap_lift_def framesize_from_to_H
3036                         cap_to_H_def[split_simps cap_CL.split]
3037                         valid_cap'_def user_vtop_def X64.pptrUserTop_def)
3038   apply (subgoal_tac "(cap_C.words_C (cte_C.cap_C vb).[0] >> 58) && 1 \<noteq> 0")
3039    prefer 2 apply (fastforce split: if_splits)
3040   apply (clarsimp simp: hd_conv_nth cap_lift_frame_cap bit_simps
3041                         to_bool_def cap_frame_cap_lift
3042                         typ_heap_simps' shiftl_t2n[where n=3] field_simps
3043                  elim!: ccap_relationE)
3044   apply (clarsimp simp: neq_Nil_conv[where xs=extraCaps]
3045                         excaps_in_mem_def slotcap_in_mem_def
3046                  dest!: sym[where s="ArchObjectCap cp" for cp])
3047   apply (cut_tac p="snd (hd extraCaps)" in ctes_of_valid', simp, clarsimp)
3048   apply (rule conjI[rotated], clarsimp simp: cap_tag_defs)
3049   apply (clarsimp simp: valid_cap_simps' rf_sr_ksCurThread)
3050   subgoal by (timeit \<open>((clarsimp simp: invocation_eq_use_types framesize_from_to_H
3051              | rule conjI
3052              | solves \<open>clarsimp dest!: ccap_relation_c_valid_cap
3053                           simp: word_less_nat_alt c_valid_cap_def cl_valid_cap_def cap_frame_cap_lift
3054                                 \<close>
3055              | solves \<open>(rule_tac cp=cp in ccap_relation_PageCap_MappedAddress_update;
3056                    fastforce simp: vm_page_map_type_defs mask_def
3057                              le_user_vtop_canonical_address
3058                                            [simplified word_le_not_less user_vtop_def
3059                                                        X64.pptrUserTop_def word_less_nat_alt,
3060                                             simplified])\<close>
3061              | (rule framesize_to_from_H framesize_to_from_H[symmetric], fastforce simp: c_valid_cap_def cl_valid_cap_def)
3062         )+)[1]\<close>) (* 83 sec *)
3063  (* bit of duplicated stuff from previous subgoal here *)
3064  apply clarsimp
3065  apply (frule interpret_excaps_eq[rule_format, where n=0], simp)
3066  apply (frule(1) slotcap_in_mem_PML4)
3067  apply (clarsimp simp: mask_def[where n=4] typ_heap_simps')
3068  apply (clarsimp simp: isCap_simps)
3069  apply (frule slotcap_in_mem_valid, (clarsimp simp: get_capMappedASID_CL_def)+)
3070  apply (erule_tac c="ArchObjectCap (PML4Cap a b)" for a b in ccap_relationE)
3071  apply (clarsimp simp: cap_lift_pml4_cap to_bool_def
3072                        cap_pml4_cap_lift_def framesize_from_to_H
3073                        cap_to_H_def[split_simps cap_CL.split]
3074                        valid_cap'_def user_vtop_def X64.pptrUserTop_def)
3075  apply (subgoal_tac "(cap_C.words_C (cte_C.cap_C vb).[0] >> 58) && 1 \<noteq> 0")
3076   prefer 2 apply (fastforce split: if_splits)
3077  apply (clarsimp simp: hd_conv_nth cap_lift_frame_cap bit_simps
3078                        to_bool_def cap_frame_cap_lift
3079                        typ_heap_simps' shiftl_t2n[where n=3] field_simps
3080                 elim!: ccap_relationE)
3081  by (timeit \<open>((clarsimp simp: cap_tag_defs
3082       | rule conjI
3083       | solves \<open>clarsimp dest!: ccap_relation_PageCap_MappedASID\<close>
3084       | solves \<open>clarsimp dest!: ccap_relation_c_valid_cap
3085                           simp: word_less_nat_alt c_valid_cap_def cl_valid_cap_def cap_frame_cap_lift
3086                                 \<close>)+)[1]\<close> (* 7 sec *))
3087
3088(* FIXME x64: find out some handy ones for X64 *)
3089lemma asidHighBits_handy_convs:
3090  "sint Kernel_C.asidHighBits = 3"
3091  "Kernel_C.asidHighBits \<noteq> 0x20"
3092  "unat Kernel_C.asidHighBits = asid_high_bits"
3093  by (simp add: Kernel_C.asidHighBits_def
3094                asid_high_bits_def)+
3095
3096lemma sts_Restart_ct_active [wp]:
3097  "\<lbrace>\<lambda>s. thread = ksCurThread s\<rbrace> setThreadState Restart thread \<lbrace>\<lambda>_. ct_active'\<rbrace>"
3098  apply (clarsimp simp: ct_in_state'_def)
3099  apply (rule hoare_lift_Pf2 [where f=ksCurThread])
3100   apply (wp sts_st_tcb')
3101   apply (simp split: if_split)
3102  apply wp
3103  done
3104
3105lemma maskCapRights_eq_Untyped [simp]:
3106  "(maskCapRights R cap = UntypedCap d p sz idx) = (cap = UntypedCap d p sz idx)"
3107  apply (cases cap)
3108  apply (auto simp: Let_def isCap_simps maskCapRights_def)
3109  apply (simp add: X64_H.maskCapRights_def isPageCap_def Let_def split: arch_capability.splits)
3110  done
3111
3112
3113lemma le_mask_asid_bits_helper:
3114  "x \<le> 2 ^ asid_high_bits - 1 \<Longrightarrow> (x::machine_word) << asid_low_bits \<le> mask asid_bits"
3115  apply (simp add: mask_def)
3116  apply (drule le2p_bits_unset_64)
3117   apply (simp add: asid_high_bits_def word_bits_def)
3118  apply (subst upper_bits_unset_is_l2p_64 [symmetric])
3119   apply (simp add: asid_bits_def word_bits_def)
3120  apply (clarsimp simp: asid_bits_def asid_low_bits_def asid_high_bits_def nth_shiftl)
3121  done
3122
3123declare Word_Lemmas.from_bool_mask_simp [simp]
3124
3125lemma injection_handler_liftE:
3126  "injection_handler a (liftE f) = liftE f"
3127  by (simp add:injection_handler_def)
3128
3129
3130lemma liftE_case_sum:
3131  "liftE f >>= case_sum (throwError \<circ> Inr) g = f >>= g"
3132  by (simp add:liftE_def)
3133
3134lemma throwError_invocationCatch:
3135  "throwError a >>= invocationCatch b c d e = throwError (Inl a)"
3136  by (simp add:throwError_def invocationCatch_def)
3137
3138lemma framesize_from_H_mask2:
3139  "framesize_from_H a && mask 2 = framesize_from_H a"
3140  apply (rule less_mask_eq)
3141  apply (simp add:framesize_from_H_def
3142    split: vmpage_size.splits)
3143    apply (simp add:Kernel_C.X86_SmallPage_def
3144      Kernel_C.X86_LargePage_def
3145      Kernel_C.X64_HugePage_def)+
3146  done
3147
3148lemma rel_option_alt_def:
3149  "rel_option f a b = (
3150      (a = None \<and>  b = None)
3151      \<or> (\<exists>x y. a = Some x \<and>  b = Some y \<and> f x y))"
3152  apply (case_tac a, case_tac b, simp, simp, case_tac b, auto)
3153  done
3154
3155lemma injection_handler_stateAssert_relocate:
3156  "injection_handler Inl (stateAssert ass xs >>= f) >>=E g
3157    = do v \<leftarrow> stateAssert ass xs; injection_handler Inl (f ()) >>=E g od"
3158  by (simp add: injection_handler_def handleE'_def bind_bindE_assoc bind_assoc)
3159
3160(* FIXME: move to where is_aligned_ptrFromPAddr is *)
3161lemma is_aligned_ptrFromPAddr_pageBitsForSize:
3162  "is_aligned p (pageBitsForSize sz) \<Longrightarrow> is_aligned (ptrFromPAddr p) (pageBitsForSize sz)"
3163  by (cases sz ; simp add: is_aligned_ptrFromPAddr_n pageBits_def bit_simps)
3164
3165lemma makeUserPDPTEPageDirectory_spec:
3166  "\<forall>s. \<Gamma> \<turnstile>
3167  \<lbrace>s. vmsz_aligned' (\<acute>paddr) X64SmallPage\<rbrace>
3168  Call makeUserPDPTEPageDirectory_'proc
3169  \<lbrace> pdpte_lift \<acute>ret__struct_pdpte_C = Some (Pdpte_pdpte_pd \<lparr>
3170       pdpte_pdpte_pd_CL.xd_CL = 0,
3171       pdpte_pdpte_pd_CL.pd_base_address_CL = \<^bsup>s\<^esup>paddr && (mask 39 << 12),
3172       pdpte_pdpte_pd_CL.accessed_CL = 0,
3173       pdpte_pdpte_pd_CL.cache_disabled_CL = x86PCDBit_CL (vm_attributes_lift \<^bsup>s\<^esup>vm_attr),
3174       pdpte_pdpte_pd_CL.write_through_CL = x86PWTBit_CL (vm_attributes_lift \<^bsup>s\<^esup>vm_attr),
3175       pdpte_pdpte_pd_CL.super_user_CL = 1,
3176       pdpte_pdpte_pd_CL.read_write_CL = 1,
3177       pdpte_pdpte_pd_CL.present_CL = 1\<rparr>) \<rbrace>"
3178  apply (rule allI, rule conseqPre, vcg)
3179  apply (clarsimp simp: pdpte_pdpte_pd_lift
3180                        superuser_from_vm_rights_mask writable_from_vm_rights_mask
3181                        vm_attributes_lift_def mask_def)
3182  done
3183
3184lemma ccap_relation_page_directory_mapped_asid:
3185  "ccap_relation (ArchObjectCap (PageDirectoryCap p (Some (asid, vspace)))) cap
3186    \<Longrightarrow> asid = capPDMappedASID_CL (cap_page_directory_cap_lift cap)"
3187  by (frule cap_get_tag_isCap_unfolded_H_cap)
3188     (clarsimp simp: cap_page_directory_cap_lift ccap_relation_def cap_to_H_def split: if_splits)
3189
3190lemma performPageDirectoryInvocationMap_ccorres:
3191  "ccorres (K (K \<bottom>) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
3192       (cte_at' ctSlot)
3193       (UNIV \<inter> \<lbrace>ccap_relation cap \<acute>cap\<rbrace> \<inter> \<lbrace>\<acute>ctSlot = Ptr ctSlot\<rbrace> \<inter> \<lbrace>cpdpte_relation pde \<acute>pdpte\<rbrace> \<inter> \<lbrace>\<acute>pdptSlot = Ptr pdSlot\<rbrace>
3194             \<inter> \<lbrace>\<acute>vspace = Ptr vspace\<rbrace>)
3195       hs
3196       (liftE (performPageDirectoryInvocation (PageDirectoryMap cap ctSlot pde pdSlot vspace)))
3197       (Call performX64PageDirectoryInvocationMap_'proc)"
3198  apply (simp only: liftE_liftM ccorres_liftM_simp)
3199  apply (cinit lift: cap_' ctSlot_' pdpte_' pdptSlot_' vspace_')
3200   apply (ctac (no_vcg))
3201     apply (rule ccorres_split_nothrow)
3202         apply simp
3203         apply (erule storePDPTE_Basic_ccorres)
3204        apply ceqv
3205       apply (rule ccorres_cases[where P="\<exists>p a v. cap = ArchObjectCap (PageDirectoryCap p (Some (a, v)))" and H=\<top> and H'=UNIV];
3206              clarsimp split: capability.splits arch_capability.splits simp: ccorres_fail)
3207       apply csymbr
3208       apply csymbr
3209       apply (rule ccorres_add_return2)
3210       apply (rule ccorres_split_nothrow_novcg)
3211           apply (frule ccap_relation_page_directory_mapped_asid)
3212           apply simp
3213           apply (rule ccorres_call[where xf'=xfdc, OF invalidatePageStructureCacheASID_ccorres]; simp)
3214          apply ceqv
3215         apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
3216         apply (rule allI, rule conseqPre, vcg)
3217         apply (clarsimp simp: return_def)
3218        apply wp
3219       apply (simp add: guard_is_UNIV_def)
3220      apply (clarsimp, wp)
3221     apply vcg
3222    apply wp
3223   apply (clarsimp simp: cap_get_tag_isCap_unfolded_H_cap)
3224  apply simp
3225  done
3226
3227lemma decodeX64PageDirectoryInvocation_ccorres:
3228  "\<lbrakk>interpret_excaps extraCaps' = excaps_map extraCaps; isPageDirectoryCap cp\<rbrakk> \<Longrightarrow>
3229   ccorres
3230       (intr_and_se_rel \<currency> dc)
3231       (liftxf errstate id (K ()) ret__unsigned_long_')
3232       (invs' and (\<lambda>s. ksCurThread s = thread) and ct_active' and sch_act_simple
3233              and (excaps_in_mem extraCaps \<circ> ctes_of)
3234              and cte_wp_at' (diminished' (ArchObjectCap cp) \<circ> cteCap) slot
3235              and valid_cap' (ArchObjectCap cp)
3236              and (\<lambda>s. \<forall>v \<in> set extraCaps. ex_cte_cap_wp_to' isCNodeCap (snd v) s)
3237              and sysargs_rel args buffer)
3238       (UNIV \<inter> {s. label___unsigned_long_' s = label}
3239             \<inter> {s. unat (length___unsigned_long_' s) = length args}
3240             \<inter> {s. cte_' s = cte_Ptr slot}
3241             \<inter> {s. extraCaps___struct_extra_caps_C_' s = extraCaps'}
3242             \<inter> {s. ccap_relation (ArchObjectCap cp) (cap_' s)}
3243             \<inter> {s. buffer_' s = option_to_ptr buffer})
3244       hs
3245       (decodeX64MMUInvocation label args cptr slot cp extraCaps
3246              >>= invocationCatch thread isBlocking isCall InvokeArchObject)
3247       (Call decodeX64PageDirectoryInvocation_'proc)"
3248   (is "_ \<Longrightarrow> _ \<Longrightarrow> ccorres _ _ ?pre _ _ _ _")
3249  supply Collect_const[simp del] if_cong[cong]
3250  apply (clarsimp simp only: isCap_simps)
3251  apply (cinit' lift: label___unsigned_long_' length___unsigned_long_' cte_' extraCaps___struct_extra_caps_C_' cap_' buffer_'
3252                simp: decodeX64MMUInvocation_def invocation_eq_use_types decodeX64PageDirectoryInvocation_def)
3253   apply (simp add: Let_def isCap_simps if_to_top_of_bind
3254               cong: StateSpace.state.fold_congs globals.fold_congs)
3255   apply (rule ccorres_Cond_rhs_Seq)
3256    apply (rule ccorres_split_throws)
3257     apply (simp add: liftE_bindE bind_assoc)
3258     apply (rule ccorres_symb_exec_l[OF _ getCTE_inv _ empty_fail_getCTE])
3259      apply (rule ccorres_rhs_assoc)+
3260      apply (ctac add: isFinalCapability_ccorres)
3261        apply (simp add: unlessE_def if_to_top_of_bind if_to_top_of_bindE ccorres_seq_cond_raise)
3262        apply (rule ccorres_cond2'[where R=\<top>])
3263          apply (clarsimp simp: from_bool_0)
3264         apply (simp add: throwError_bind invocationCatch_def)
3265         apply (rule syscall_error_throwError_ccorres_n)
3266         apply (simp add: syscall_error_to_H_cases)
3267        apply (simp add: returnOk_bind bindE_assoc performX64MMUInvocations)
3268        apply (ctac add: setThreadState_ccorres)
3269          apply (ctac add: performPageDirectoryInvocationUnmap_ccorres)
3270             apply (rule ccorres_alternative2)
3271             apply (rule ccorres_return_CE, simp+)[1]
3272            apply (rule ccorres_inst[where P=\<top> and P'=UNIV], simp)
3273           apply wp
3274          apply simp
3275          apply (vcg exspec=performX64PageDirectoryInvocationUnmap_modifies)
3276         apply (wp sts_invs_minor')
3277        apply simp
3278        apply (vcg exspec=setThreadState_modifies)
3279       apply (wp hoare_drop_imps isFinalCapability_inv)
3280      apply simp
3281      apply (vcg exspec=isFinalCapability_modifies)
3282     apply (wp getCTE_wp)
3283    apply (vcg exspec=performX64PageDirectoryInvocationUnmap_modifies exspec=isFinalCapability_modifies
3284               exspec=setThreadState_modifies)
3285   apply simp
3286   apply (rule ccorres_Cond_rhs_Seq)
3287    apply (rule ccorres_equals_throwError)
3288     apply (simp split: invocation_label.split arch_invocation_label.split
3289                 add: throwError_bind invocationCatch_def)
3290     apply fastforce
3291    apply (rule syscall_error_throwError_ccorres_n)
3292    apply (simp add: syscall_error_to_H_cases)
3293   apply simp
3294   apply csymbr
3295   apply (rule ccorres_Cond_rhs_Seq)
3296    apply (simp add: word_less_nat_alt throwError_bind invocationCatch_def)
3297    apply (rule ccorres_cond_true_seq)
3298    apply (rule ccorres_equals_throwError)
3299     apply (simp add: throwError_bind split: list.split)
3300     apply fastforce
3301    apply (rule syscall_error_throwError_ccorres_n)
3302    apply (simp add: syscall_error_to_H_cases)
3303   apply csymbr
3304   apply (simp add: interpret_excaps_test_null excaps_map_def)
3305   apply (rule ccorres_Cond_rhs_Seq)
3306    apply (simp add: throwError_bind invocationCatch_def)
3307    apply (rule ccorres_equals_throwError)
3308     apply (fastforce simp: throwError_bind split: list.split)
3309    apply (rule syscall_error_throwError_ccorres_n)
3310    apply (simp add: syscall_error_to_H_cases)
3311   apply (simp add: list_case_If2 split_def
3312                    word_less_nat_alt length_ineq_not_Nil Let_def
3313                    whenE_bindE_throwError_to_if if_to_top_of_bind)
3314   apply csymbr
3315   apply (rule_tac Q=\<top> and Q'=\<top> in ccorres_if_cond_throws[rotated -1])
3316      apply vcg
3317     apply clarsimp
3318     apply (frule cap_get_tag_isCap_unfolded_H_cap)
3319     apply (clarsimp simp: cap_lift_page_directory_cap cap_page_directory_cap_lift_def cap_to_H_def
3320                     elim!: ccap_relationE split: if_split)
3321     apply (simp add: to_bool_def)
3322    apply (simp add: throwError_bind invocationCatch_def)
3323    apply (rule syscall_error_throwError_ccorres_n)
3324    apply (simp add: syscall_error_to_H_cases)
3325   apply (simp add: cap_case_PML4Cap2 split_def)
3326   apply (rule ccorres_add_return)
3327   apply (ctac add: getSyscallArg_ccorres_foo[where args=args and n=0 and buffer=buffer])
3328     apply csymbr
3329     apply (rule ccorres_add_return)
3330     apply (ctac add: getSyscallArg_ccorres_foo[where args=args and n=1 and buffer=buffer])
3331       apply csymbr
3332       apply (rule ccorres_add_return)
3333       apply (rule_tac r'="(\<lambda>rv _ rv'. ((cap_get_tag rv' = scast cap_pml4_cap)
3334                                          = (isArchObjectCap rv \<and> isPML4Cap (capCap rv)))
3335                                     \<and> (ccap_relation rv rv')) (fst (extraCaps ! 0))" and
3336                       xf'=vspaceCap_'
3337                       in ccorres_split_nothrow)
3338           apply (rule ccorres_from_vcg[where P="excaps_in_mem extraCaps \<circ> ctes_of" and P'=UNIV])
3339           apply (rule allI, rule conseqPre, vcg)
3340           apply (clarsimp simp: excaps_in_mem_def return_def neq_Nil_conv)
3341           apply (drule(1) slotcap_in_mem_PML4)
3342           apply (frule interpret_excaps_eq[rule_format, where n=0], simp)
3343           apply (clarsimp simp: typ_heap_simps' mask_def)
3344          apply (rename_tac rv' t t')
3345          apply (simp add: word_sless_def word_sle_def)
3346          apply ceqv
3347         apply csymbr
3348         apply clarsimp
3349         apply (rule ccorres_Cond_rhs_Seq)
3350          apply ccorres_rewrite
3351          apply clarsimp
3352          apply (rule_tac P="isArchObjectCap (fst (extraCaps ! 0)) \<and>
3353                             isPML4Cap (capCap (fst (extraCaps ! 0)))"
3354                          in ccorres_cases)
3355           apply (clarsimp simp: hd_conv_nth throwError_bind invocationCatch_def from_bool_0 cong: if_cong)
3356           apply (rule syscall_error_throwError_ccorres_n[simplified dc_def id_def o_def])
3357           apply (simp add: syscall_error_to_H_cases)
3358          apply (clarsimp simp: hd_conv_nth throwError_bind invocationCatch_def from_bool_0 cong: if_cong)
3359          apply (rule syscall_error_throwError_ccorres_n[simplified dc_def id_def o_def])
3360          apply (simp add: syscall_error_to_H_cases)
3361         apply (simp add: hd_conv_nth)
3362         apply csymbr
3363         apply csymbr
3364         apply csymbr
3365         apply (simp add: case_option_If2 if_to_top_of_bind if_to_top_of_bindE hd_conv_nth)
3366         apply (rule ccorres_if_cond_throws[rotated -1, where Q=\<top> and Q'=\<top>])
3367            apply vcg
3368           apply (fastforce simp: user_vtop_def X64.pptrUserTop_def bit_simps
3369                                  word_le_nat_alt mask_def hd_conv_nth length_ineq_not_Nil)
3370          apply (simp add: throwError_bind invocationCatch_def)
3371          apply (rule syscall_error_throwError_ccorres_n)
3372          apply (simp add: syscall_error_to_H_cases)
3373         apply (simp add: lookupError_injection invocationCatch_use_injection_handler
3374                          injection_bindE[OF refl refl] injection_handler_If
3375                          injection_handler_throwError injection_liftE[OF refl]
3376                          injection_handler_returnOk bindE_assoc cap_pml4_cap_lift get_capMappedASID_CL_def
3377                     cong: if_cong)
3378         apply csymbr
3379         apply (ctac add: ccorres_injection_handler_csum1 [OF ccorres_injection_handler_csum1,
3380                                                           OF findVSpaceForASID_ccorres])
3381            apply (simp add: Collect_False if_to_top_of_bindE)
3382            apply (rule ccorres_if_cond_throws[rotated -1, where Q=\<top> and Q'=\<top>])
3383               apply vcg
3384              apply (clarsimp simp: isCap_simps get_capPtr_CL_def)
3385              apply (frule cap_get_tag_isCap_unfolded_H_cap)
3386              apply (erule_tac v="rv'" in ccap_relationE, clarsimp simp: cap_to_H_def)
3387             apply (rule syscall_error_throwError_ccorres_n)
3388             apply (simp add: syscall_error_to_H_cases)
3389            apply (simp add: bindE_assoc,
3390                   simp add: liftE_bindE)
3391            apply (ctac add: ccorres_injection_handler_csum1 [OF ccorres_injection_handler_csum1,
3392                                                              OF lookupPDPTSlot_ccorres])
3393               apply (rule ccorres_pre_getObject_pdpte)
3394               apply (rule ccorres_cond_false_seq)
3395               apply (rename_tac pdpt_slot pdpt_slot_C pdpte)
3396               apply ccorres_rewrite
3397               apply clarsimp
3398               apply (rename_tac pml4_mapped_asid)
3399               apply (rule ccorres_rhs_assoc2, rule ccorres_rhs_assoc2, rule ccorres_rhs_assoc2)
3400               apply (rule_tac xf'=ret__int_' and R="ko_at' pdpte pdpt_slot" and R'=UNIV and
3401                               val="from_bool (pdpte \<noteq> InvalidPDPTE)"
3402                               in ccorres_symb_exec_r_known_rv)
3403                  apply vcg
3404                  apply (clarsimp simp: from_bool_0)
3405                  apply (erule cmap_relationE1[OF rf_sr_cpdpte_relation], erule ko_at_projectKO_opt)
3406                  apply (clarsimp simp: typ_heap_simps from_bool_eq_if)
3407                  apply (auto simp: cpdpte_relation_def Let_def pdpte_pdpte_pd_lift_def
3408                                    pdpte_pdpte_pd_lift pdpte_tag_defs pdpte_pdpte_1g_lift_def
3409                                    pdpte_lift_def from_bool_def case_bool_If
3410                              split: pdpte.split_asm if_splits)[1]
3411                 apply ceqv
3412                apply clarsimp
3413                apply (rule ccorres_Cond_rhs_Seq)
3414                 apply (clarsimp simp: throwError_bind invocationCatch_def injection_handler_throwError)
3415                 apply ccorres_rewrite
3416                 apply (rule syscall_error_throwError_ccorres_n)
3417                 apply (clarsimp simp: syscall_error_to_H_cases)
3418                apply (clarsimp simp: injection_handler_returnOk)
3419                apply (simp add: injection_handler_returnOk performX64MMUInvocations bindE_assoc)
3420                apply csymbr
3421                apply csymbr
3422                apply csymbr
3423                apply csymbr
3424                apply csymbr
3425                apply csymbr
3426                apply (ctac add: setThreadState_ccorres)
3427                  apply (ctac add: performPageDirectoryInvocationMap_ccorres)
3428                     apply (rule ccorres_alternative2)
3429                     apply (rule ccorres_return_CE, simp+)[1]
3430                    apply (rule ccorres_inst[where P=\<top> and P'=UNIV], simp)
3431                   apply wp+
3432                  apply simp
3433                  apply (vcg exspec=performX64PageDirectoryInvocationMap_modifies)
3434                 apply simp
3435                 apply wp
3436                apply simp
3437                apply (vcg exspec=setThreadState_modifies)
3438               apply (simp add: get_capPtr_CL_def)
3439               apply (rule_tac s=pml4_mapped_asid and
3440                               t="the (capPML4MappedASID (capCap (fst (extraCaps ! 0))))"
3441                               in subst, fastforce)
3442               apply vcg
3443              apply (rename_tac lookup_pd_ret)
3444              apply clarsimp
3445              apply (rule_tac P'="{s. errstate s = lookup_pd_ret}" in ccorres_from_vcg_split_throws[where P=\<top>])
3446               apply vcg
3447              apply (rule conseqPre, vcg)
3448              apply (clarsimp simp: throwError_def return_def syscall_error_rel_def syscall_error_to_H_cases exception_defs false_def)
3449              apply (erule lookup_failure_rel_fault_lift[rotated])
3450              apply (clarsimp simp: exception_defs)
3451             apply clarsimp
3452             apply (wp injection_wp[OF refl] hoare_drop_imps)
3453            apply (clarsimp simp: get_capPtr_CL_def)
3454            apply (rename_tac pml4_mapped_asid)
3455            apply (rule_tac s=pml4_mapped_asid and
3456                            t="the (capPML4MappedASID (capCap (fst (extraCaps ! 0))))"
3457                            in subst, fastforce)
3458            apply (vcg exspec=lookupPDPTSlot_modifies)
3459           apply clarsimp
3460           apply (rule_tac P'="{s. errstate s = find_ret}" in ccorres_from_vcg_split_throws[where P=\<top>])
3461            apply vcg
3462           apply simp
3463           apply (rule conseqPre, vcg)
3464           apply (clarsimp simp: throwError_def return_def syscall_error_rel_def
3465                                 syscall_error_to_H_cases exception_defs false_def)
3466           apply (erule lookup_failure_rel_fault_lift[rotated])
3467           apply (simp add: exception_defs)
3468          apply simp
3469          apply (wp injection_wp[OF refl] hoare_drop_imps)
3470         apply simp
3471         apply (vcg exspec=findVSpaceForASID_modifies)
3472        apply simp
3473        apply (rule_tac Q="\<lambda>a b. invs' b \<and> valid_cap' (fst (extraCaps ! 0)) b \<and> tcb_at' thread b \<and>
3474                                 sch_act_wf (ksSchedulerAction b) b \<and> cte_wp_at' (\<lambda>_. True) slot b"
3475                        in hoare_strengthen_post)
3476         apply wp
3477        apply (clarsimp simp: isCap_simps invs_valid_objs' valid_cap'_def valid_tcb_state'_def invs_arch_state' invs_no_0_obj')
3478       apply vcg
3479      apply wp
3480     apply simp
3481     apply (vcg exspec=getSyscallArg_modifies)
3482    apply simp
3483    apply wp
3484   apply simp
3485  apply (vcg exspec=getSyscallArg_modifies)
3486  apply (clarsimp simp: cte_wp_at_ctes_of excaps_map_def
3487                        word_sle_def word_sless_def bit_simps)
3488  apply (rule conjI)
3489   apply (auto simp: ct_in_state'_def isCap_simps valid_tcb_state'_def
3490              elim!: pred_tcb'_weakenE dest!: st_tcb_at_idle_thread')[1]
3491  apply (rule conjI)
3492   apply (clarsimp simp: isCap_simps)
3493   apply (subgoal_tac "s \<turnstile>' fst (extraCaps ! 0)")
3494    apply (clarsimp simp: sysargs_rel_to_n word_le_nat_alt
3495                          linorder_not_less)
3496    apply (clarsimp | drule length_le_helper)+
3497    apply (clarsimp simp: valid_cap'_def neq_Nil_conv
3498                          mask_add_aligned pd_pointer_table_at'_def
3499                          word_le_nat_alt[symmetric] bit_simps)
3500    apply (auto simp: ct_in_state'_def pred_tcb_at' mask_def valid_tcb_state'_def
3501               elim!: pred_tcb'_weakenE dest!: st_tcb_at_idle_thread')[1]
3502   apply (clarsimp simp: neq_Nil_conv excaps_in_mem_def
3503                         slotcap_in_mem_def)
3504   apply (auto dest: ctes_of_valid')[1]
3505  apply (rule conjI)
3506   apply (clarsimp simp: rf_sr_ksCurThread "StrictC'_thread_state_defs"
3507                         mask_eq_iff_w2p word_size
3508                         ct_in_state'_def st_tcb_at'_def
3509                         word_sle_def word_sless_def
3510                         typ_heap_simps' bit_simps)
3511  apply (clarsimp simp: cap_get_tag_isCap_ArchObject isCap_simps
3512                        word_sle_def word_sless_def
3513                        word_less_nat_alt)
3514  apply (frule length_ineq_not_Nil)
3515  apply (frule cap_get_tag_isCap_unfolded_H_cap)
3516  apply (clarsimp simp: cap_lift_pdpt_cap hd_conv_nth
3517                        cap_lift_page_directory_cap bit_simps
3518                        cap_pdpt_cap_lift_def
3519                        to_bool_def
3520                        typ_heap_simps' shiftl_t2n[where n=3] field_simps
3521                 elim!: ccap_relationE)
3522  apply (clarsimp simp: neq_Nil_conv[where xs=extraCaps]
3523                        excaps_in_mem_def slotcap_in_mem_def
3524                 dest!: sym[where s="ArchObjectCap cp" for cp])
3525  apply (cut_tac p="snd (hd extraCaps)" in ctes_of_valid', simp, clarsimp)
3526  apply (rule conjI[rotated], clarsimp simp: cap_tag_defs)
3527  apply (clarsimp simp: valid_cap_simps' rf_sr_ksCurThread)
3528  apply (rule context_conjI)
3529   subgoal by (clarsimp simp: get_capMappedASID_CL_def dest!: cap_to_H_PML4Cap)
3530  apply clarsimp
3531  apply (rule conjI, clarsimp simp: mask_def)
3532  apply clarsimp
3533  apply (rule conjI, clarsimp, rule conjI, clarsimp simp: pdpte_tag_defs, clarsimp)
3534   apply (rule conjI[rotated])
3535    apply (fastforce dest!: cap_lift_page_directory_cap
3536                     intro!: is_aligned_addrFromPPtr[simplified bit_simps, simplified]
3537                     simp: vmsz_aligned'_def cap_to_H_simps cap_page_directory_cap_lift_def bit_simps capAligned_def)
3538   apply clarsimp
3539   apply (rule conjI, clarsimp simp: ThreadState_Restart_def mask_def)
3540   (* ccap_relation *)
3541   apply (rule conjI)
3542    apply (clarsimp simp: ccap_relation_def map_option_Some_eq2 cap_page_directory_cap_lift[THEN iffD1]
3543                          cap_to_H_simps asid_wf_def3[simplified asid_bits_def, simplified])
3544    apply (rule conjI[rotated],
3545           fastforce simp: sign_extend_canonical_address le_def[symmetric] mask_def word_bw_assocs
3546                           le_user_vtop_canonical_address[simplified user_vtop_def X64.pptrUserTop_def word_le_nat_alt])
3547    subgoal by (clarsimp dest!: cap_lift_page_directory_cap simp: cap_page_directory_cap_lift_def)
3548   (* cpdpte_relation *)
3549   apply (rule conjI, clarsimp simp: cpdpte_relation_def)
3550    apply (clarsimp simp: superuser_from_vm_rights_def writable_from_vm_rights_def
3551                          x64CacheDisabled_def attribsFromWord_def word_and_1 nth_shiftr
3552                    split: if_split)
3553    apply (clarsimp dest!: cap_lift_page_directory_cap simp: cap_to_H_simps cap_page_directory_cap_lift_def)
3554    apply (rule addrFromPPtr_mask_middle_pml4ShiftBits[simplified pml4ShiftBits_def bit_simps, simplified])
3555     subgoal by (fastforce simp: valid_cap_simps' capAligned_def bit_simps)
3556    apply (fastforce dest!: page_directory_pde_atI'[where x=0, simplified bit_simps, simplified]
3557                     intro!: obj_at_kernel_mappings'
3558                     simp: typ_at_to_obj_at_arches)
3559   (* basePtr *)
3560   apply (frule (1) cap_lift_PML4Cap_Base)
3561   subgoal by (auto simp: cap_to_H_simps cap_pml4_cap_lift_def get_capPtr_CL_def
3562                    dest!: cap_to_H_PML4Cap_tag cap_lift_pml4_cap)
3563  apply (clarsimp simp: pdpte_tag_defs pdpte_get_tag_def word_and_1)
3564  (* proof below is duplicated from above due to difficulty of reasoning under
3565     so many layers of implications and assumptions
3566     context_conjI creates a mess, separate lemmas would be a bit unwieldy
3567  *)
3568  apply safe
3569      apply (clarsimp simp: ThreadState_Restart_def mask_def)
3570    (* ccap_relation *)
3571     apply (clarsimp simp: ccap_relation_def map_option_Some_eq2 cap_page_directory_cap_lift[THEN iffD1]
3572                           cap_to_H_simps asid_wf_def3[simplified asid_bits_def, simplified])
3573     apply (rule conjI[rotated],
3574            fastforce simp: sign_extend_canonical_address le_def[symmetric] mask_def word_bw_assocs
3575                            le_user_vtop_canonical_address[simplified user_vtop_def X64.pptrUserTop_def word_le_nat_alt])
3576     subgoal by (clarsimp dest!: cap_lift_page_directory_cap simp: cap_page_directory_cap_lift_def)
3577    (* cpdpte_relation *)
3578    apply (clarsimp simp: cpdpte_relation_def)
3579    apply (clarsimp simp: superuser_from_vm_rights_def writable_from_vm_rights_def
3580                          x64CacheDisabled_def attribsFromWord_def word_and_1 nth_shiftr
3581                    split: if_split)
3582    apply (clarsimp dest!: cap_lift_page_directory_cap simp: cap_to_H_simps cap_page_directory_cap_lift_def)
3583    apply (rule addrFromPPtr_mask_middle_pml4ShiftBits[simplified pml4ShiftBits_def bit_simps, simplified])
3584     subgoal by (fastforce simp: valid_cap_simps' capAligned_def bit_simps)
3585    apply (fastforce dest!: page_directory_pde_atI'[where x=0, simplified bit_simps, simplified]
3586                     intro!: obj_at_kernel_mappings'
3587                     simp: typ_at_to_obj_at_arches)
3588   (* basePtr *)
3589   apply (frule (1) cap_lift_PML4Cap_Base)
3590   apply (auto simp: cap_to_H_simps cap_pml4_cap_lift_def get_capPtr_CL_def
3591               dest!: cap_to_H_PML4Cap_tag cap_lift_pml4_cap)[1]
3592  apply (fastforce dest!: cap_lift_page_directory_cap
3593                   intro!: is_aligned_addrFromPPtr[simplified bit_simps, simplified]
3594                   simp: vmsz_aligned'_def cap_to_H_simps cap_page_directory_cap_lift_def bit_simps capAligned_def)
3595  done
3596
3597lemma makeUserPML4E_spec:
3598  "\<forall>s. \<Gamma> \<turnstile>
3599  {s}
3600  Call makeUserPML4E_'proc
3601  \<lbrace> pml4e_lift \<acute>ret__struct_pml4e_C = ( \<lparr>
3602       pml4e_CL.xd_CL = 0,
3603       pml4e_CL.pdpt_base_address_CL = \<^bsup>s\<^esup>paddr && (mask 39 << 12),
3604       pml4e_CL.accessed_CL = 0,
3605       pml4e_CL.cache_disabled_CL = x86PCDBit_CL (vm_attributes_lift \<^bsup>s\<^esup>vm_attr),
3606       pml4e_CL.write_through_CL = x86PWTBit_CL (vm_attributes_lift \<^bsup>s\<^esup>vm_attr),
3607       pml4e_CL.super_user_CL = 1,
3608       pml4e_CL.read_write_CL = 1,
3609       pml4e_CL.present_CL = 1\<rparr>) \<rbrace>"
3610  apply (rule allI, rule conseqPre, vcg)
3611  apply (clarsimp simp: pml4e_lift_def
3612                        superuser_from_vm_rights_mask writable_from_vm_rights_mask
3613                        vm_attributes_lift_def mask_def)
3614  done
3615
3616lemma ccap_relation_pdpt_mapped_asid:
3617  "ccap_relation (ArchObjectCap (PDPointerTableCap p (Some (asid, vspace)))) cap
3618    \<Longrightarrow> asid = capPDPTMappedASID_CL (cap_pdpt_cap_lift cap)"
3619  by (frule cap_get_tag_isCap_unfolded_H_cap)
3620     (clarsimp simp: cap_pdpt_cap_lift ccap_relation_def cap_to_H_def split: if_splits)
3621
3622lemma performPDPTInvocationMap_ccorres:
3623  "ccorres (K (K \<bottom>) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
3624       (cte_at' ctSlot)
3625       (UNIV \<inter> \<lbrace>ccap_relation cap \<acute>cap\<rbrace> \<inter> \<lbrace>\<acute>ctSlot = Ptr ctSlot\<rbrace> \<inter> \<lbrace>cpml4e_relation pml4e \<acute>pml4e\<rbrace> \<inter> \<lbrace>\<acute>pml4Slot = Ptr pml4Slot\<rbrace>
3626             \<inter> \<lbrace>\<acute>vspace = Ptr vspace\<rbrace>)
3627       []
3628       (liftE (performPDPTInvocation (PDPTMap cap ctSlot pml4e pml4Slot vspace)))
3629       (Call performX64PDPTInvocationMap_'proc)"
3630  apply (simp only: liftE_liftM ccorres_liftM_simp)
3631  apply (cinit lift: cap_' ctSlot_' pml4e_' pml4Slot_' vspace_')
3632   apply (ctac (no_vcg))
3633     apply (rule ccorres_split_nothrow)
3634         apply simp
3635         apply (erule storePML4E_Basic_ccorres)
3636        apply ceqv
3637       apply (rule ccorres_cases[where P="\<exists>p a v. cap = ArchObjectCap (PDPointerTableCap p (Some (a, v)))" and H=\<top> and H'=UNIV];
3638              clarsimp split: capability.splits arch_capability.splits simp: ccorres_fail)
3639       apply csymbr
3640       apply csymbr
3641       apply (rule ccorres_add_return2)
3642       apply (rule ccorres_split_nothrow_novcg)
3643           apply (frule ccap_relation_pdpt_mapped_asid)
3644           apply simp
3645           apply (rule ccorres_call[where xf'=xfdc, OF invalidatePageStructureCacheASID_ccorres]; simp)
3646          apply ceqv
3647         apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
3648         apply (rule allI, rule conseqPre, vcg)
3649         apply (clarsimp simp: return_def)
3650        apply wp
3651       apply (simp add: guard_is_UNIV_def)
3652      apply (clarsimp, wp)
3653     apply vcg
3654    apply wp
3655   apply (clarsimp simp: cap_get_tag_isCap_unfolded_H_cap)
3656  apply simp
3657  done
3658
3659lemma performPDPTInvocationMap_ccorres':
3660  "ccorres dc ret__unsigned_long_'
3661       (cte_at' ctSlot)
3662       (UNIV \<inter> \<lbrace>ccap_relation cap \<acute>cap\<rbrace>
3663             \<inter> \<lbrace>\<acute>ctSlot = Ptr ctSlot\<rbrace>
3664             \<inter> \<lbrace>cpml4e_relation pml4e \<acute>pml4e\<rbrace>
3665             \<inter> \<lbrace>\<acute>pml4Slot = Ptr pml4Slot\<rbrace>
3666             \<inter> \<lbrace>\<acute>vspace = Ptr vspace\<rbrace>)
3667       hs
3668       (performPDPTInvocation (PDPTMap cap ctSlot pml4e pml4Slot vspace))
3669       (Call performX64PDPTInvocationMap_'proc)"
3670  apply (cinit lift: cap_' ctSlot_' pml4e_' pml4Slot_' vspace_')
3671   apply (ctac (no_vcg))
3672     apply (rule ccorres_split_nothrow)
3673         apply simp
3674         apply (erule storePML4E_Basic_ccorres)
3675        apply ceqv
3676       apply (rule ccorres_cases[where P="\<exists>p a v. cap = ArchObjectCap (PDPointerTableCap p (Some (a, v)))" and H=\<top> and H'=UNIV];
3677              clarsimp split: capability.splits arch_capability.splits simp: ccorres_fail)
3678       apply csymbr
3679       apply csymbr
3680       apply (rule ccorres_add_return2)
3681       apply (rule ccorres_split_nothrow_novcg)
3682           apply (frule ccap_relation_pdpt_mapped_asid)
3683           apply simp
3684           apply (rule ccorres_call[where xf'=xfdc, OF invalidatePageStructureCacheASID_ccorres]; simp)
3685          apply ceqv
3686         apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
3687         apply (rule allI, rule conseqPre, vcg)
3688         apply (clarsimp simp: return_def)
3689        apply wp
3690       apply (simp add: guard_is_UNIV_def)
3691      apply (clarsimp, wp)
3692     apply vcg
3693    apply wp
3694   apply (clarsimp simp: cap_get_tag_isCap_unfolded_H_cap)
3695  apply simp
3696  done
3697
3698lemmas array_assertion_page_map_l4_at' =
3699  array_assertion_abs_pml4[where n="K 512" and x="K (1 :: machine_word)",
3700                           simplified, rule_format, simplified bit_simps, simplified, OF conjI]
3701
3702lemma decodeX64PDPTInvocation_ccorres:
3703  "\<lbrakk>interpret_excaps extraCaps' = excaps_map extraCaps; isPDPointerTableCap cp\<rbrakk> \<Longrightarrow>
3704   ccorres
3705       (intr_and_se_rel \<currency> dc)
3706       (liftxf errstate id (K ()) ret__unsigned_long_')
3707       (invs' and (\<lambda>s. ksCurThread s = thread) and ct_active' and sch_act_simple
3708              and (excaps_in_mem extraCaps \<circ> ctes_of)
3709              and cte_wp_at' (diminished' (ArchObjectCap cp) \<circ> cteCap) slot
3710              and valid_cap' (ArchObjectCap cp)
3711              and (\<lambda>s. \<forall>v \<in> set extraCaps. ex_cte_cap_wp_to' isCNodeCap (snd v) s)
3712              and sysargs_rel args buffer)
3713       (UNIV \<inter> {s. label___unsigned_long_' s = label}
3714             \<inter> {s. unat (length___unsigned_long_' s) = length args}
3715             \<inter> {s. cte_' s = cte_Ptr slot}
3716             \<inter> {s. extraCaps___struct_extra_caps_C_' s = extraCaps'}
3717             \<inter> {s. ccap_relation (ArchObjectCap cp) (cap_' s)}
3718             \<inter> {s. buffer_' s = option_to_ptr buffer})
3719       hs
3720       (decodeX64MMUInvocation label args cptr slot cp extraCaps
3721              >>= invocationCatch thread isBlocking isCall InvokeArchObject)
3722       (Call decodeX64PDPTInvocation_'proc)"
3723     (is "_ \<Longrightarrow> _ \<Longrightarrow> ccorres _ _ ?pre ?cpre _ _ _")
3724  supply Collect_const[simp del] if_cong[cong] dc_simp[simp del]
3725         from_bool_eq_if[simp] from_bool_eq_if'[simp] from_bool_0[simp] ccorres_IF_True[simp]
3726  apply (clarsimp simp only: isCap_simps)
3727  apply (cinit' lift: label___unsigned_long_' length___unsigned_long_' cte_' extraCaps___struct_extra_caps_C_' cap_' buffer_'
3728                simp: decodeX64MMUInvocation_def invocation_eq_use_types decodeX64PDPointerTableInvocation_def)
3729   apply (simp add: Let_def isCap_simps if_to_top_of_bind
3730               cong: StateSpace.state.fold_congs globals.fold_congs)
3731   apply (rule ccorres_Cond_rhs_Seq)
3732    apply (rule ccorres_split_throws)
3733     apply (simp add: liftE_bindE bind_assoc)
3734     apply (rule ccorres_symb_exec_l[OF _ getCTE_inv _ empty_fail_getCTE])
3735      apply (rule ccorres_rhs_assoc)+
3736      apply (ctac add: isFinalCapability_ccorres)
3737        apply (simp add: unlessE_def if_to_top_of_bind if_to_top_of_bindE
3738                         ccorres_seq_cond_raise)
3739        apply (rule ccorres_cond2'[where R=\<top>])
3740          apply clarsimp
3741         apply (simp add: throwError_bind invocationCatch_def)
3742         apply (rule syscall_error_throwError_ccorres_n)
3743         apply (simp add: syscall_error_to_H_cases)
3744        apply (simp add: returnOk_bind bindE_assoc
3745                         performX64MMUInvocations)
3746        apply (ctac add: setThreadState_ccorres)
3747          apply (ctac add: performPDPTInvocationUnmap_ccorres)
3748             apply (rule ccorres_alternative2)
3749             apply (rule ccorres_return_CE, simp+)[1]
3750            apply (rule ccorres_inst[where P=\<top> and P'=UNIV], simp)
3751           apply wp
3752          apply simp
3753          apply (vcg exspec=performX64PDPTInvocationUnmap_modifies)
3754         apply (wp sts_invs_minor')
3755        apply simp
3756        apply (vcg exspec=setThreadState_modifies)
3757       apply (wp hoare_drop_imps isFinalCapability_inv)
3758      apply simp
3759      apply (vcg exspec=isFinalCapability_modifies)
3760     apply (wp getCTE_wp)
3761    apply (vcg exspec=performX64PDPTInvocationUnmap_modifies exspec=isFinalCapability_modifies
3762               exspec=setThreadState_modifies)
3763   apply simp
3764   apply (rule ccorres_Cond_rhs_Seq)
3765    apply (rule ccorres_equals_throwError)
3766     apply (simp split: invocation_label.split arch_invocation_label.split
3767                   add: throwError_bind invocationCatch_def)
3768     apply fastforce
3769    apply (rule syscall_error_throwError_ccorres_n)
3770    apply (simp add: syscall_error_to_H_cases)
3771   apply simp
3772   apply csymbr
3773   apply (rule ccorres_Cond_rhs_Seq)
3774    apply (simp add: word_less_nat_alt throwError_bind invocationCatch_def)
3775    apply (rule ccorres_cond_true_seq)
3776    apply (rule ccorres_equals_throwError)
3777     apply (simp add: throwError_bind split: list.split)
3778     apply fastforce
3779    apply (rule syscall_error_throwError_ccorres_n)
3780    apply (simp add: syscall_error_to_H_cases)
3781   apply csymbr
3782   apply (simp add: interpret_excaps_test_null excaps_map_def)
3783   apply (rule ccorres_Cond_rhs_Seq)
3784    apply (simp add: throwError_bind invocationCatch_def)
3785    apply (rule ccorres_equals_throwError)
3786     apply (fastforce simp: throwError_bind split: list.split)
3787    apply (rule syscall_error_throwError_ccorres_n)
3788    apply (simp add: syscall_error_to_H_cases)
3789   apply (simp add: list_case_If2 split_def
3790                    word_less_nat_alt length_ineq_not_Nil Let_def
3791                    whenE_bindE_throwError_to_if if_to_top_of_bind)
3792   apply csymbr
3793   apply (rule_tac Q=\<top> and Q'=\<top> in ccorres_if_cond_throws[rotated -1])
3794      apply vcg
3795     apply clarsimp
3796     apply (frule cap_get_tag_isCap_unfolded_H_cap)
3797     apply (clarsimp simp: cap_lift_pdpt_cap cap_pdpt_cap_lift_def
3798                           cap_to_H_def
3799                    elim!: ccap_relationE split: if_split)
3800     apply (simp add: to_bool_def)
3801    apply (simp add: throwError_bind invocationCatch_def)
3802    apply (rule syscall_error_throwError_ccorres_n)
3803    apply (simp add: syscall_error_to_H_cases)
3804   apply (simp add: cap_case_PML4Cap2 split_def)
3805   apply (rule ccorres_add_return)
3806   apply (ctac add: getSyscallArg_ccorres_foo[where args=args and n=0 and buffer=buffer])
3807     apply csymbr
3808     apply (rule ccorres_add_return)
3809     apply (ctac add: getSyscallArg_ccorres_foo[where args=args and n=1 and buffer=buffer])
3810       apply csymbr
3811       apply (rule ccorres_add_return)
3812       apply (rule_tac r'="(\<lambda>rv _ rv'. ((cap_get_tag rv' = scast cap_pml4_cap)
3813                                          = (isArchObjectCap rv \<and> isPML4Cap (capCap rv)))
3814                                     \<and> (ccap_relation rv rv')) (fst (extraCaps ! 0))"
3815                and xf'=vspaceCap_' in ccorres_split_nothrow)
3816           apply (rule ccorres_from_vcg[where P="excaps_in_mem extraCaps \<circ> ctes_of" and P'=UNIV])
3817           apply (rule allI, rule conseqPre, vcg)
3818           apply (clarsimp simp: excaps_in_mem_def return_def neq_Nil_conv)
3819           apply (drule(1) slotcap_in_mem_PML4)
3820           apply (frule interpret_excaps_eq[rule_format, where n=0], simp)
3821           apply (clarsimp simp: typ_heap_simps' mask_def)
3822           apply (rename_tac rv' t t')
3823          apply (simp add: word_sless_def word_sle_def)
3824          apply ceqv
3825         apply csymbr
3826         apply clarsimp
3827         apply (rule ccorres_Cond_rhs_Seq)
3828          apply ccorres_rewrite
3829          apply clarsimp
3830          apply (rule_tac P="isArchObjectCap (fst (extraCaps ! 0)) \<and>
3831                             isPML4Cap (capCap (fst (extraCaps ! 0)))"
3832                          in ccorres_cases)
3833           apply (clarsimp simp: hd_conv_nth throwError_bind invocationCatch_def cong: if_cong)
3834           apply (rule syscall_error_throwError_ccorres_n[simplified id_def o_def])
3835           apply (simp add: syscall_error_to_H_cases)
3836          apply (clarsimp simp: hd_conv_nth throwError_bind invocationCatch_def cong: if_cong)
3837          apply (rule syscall_error_throwError_ccorres_n[simplified id_def o_def])
3838          apply (simp add: syscall_error_to_H_cases)
3839         apply (simp add: hd_conv_nth)
3840         apply csymbr
3841         apply csymbr
3842         apply csymbr
3843         apply (simp add: case_option_If2 if_to_top_of_bind if_to_top_of_bindE hd_conv_nth)
3844         apply (rule ccorres_if_cond_throws[rotated -1, where Q=\<top> and Q'=\<top>])
3845            apply vcg
3846           apply (fastforce simp: user_vtop_def X64.pptrUserTop_def bit_simps
3847                                  word_le_nat_alt mask_def hd_conv_nth
3848                                  length_ineq_not_Nil)
3849          apply (simp add: throwError_bind invocationCatch_def)
3850          apply (rule syscall_error_throwError_ccorres_n)
3851          apply (simp add: syscall_error_to_H_cases)
3852         apply (simp add: lookupError_injection invocationCatch_use_injection_handler
3853                          injection_bindE[OF refl refl] injection_handler_If
3854                          injection_handler_throwError injection_liftE[OF refl]
3855                          injection_handler_returnOk bindE_assoc cap_pml4_cap_lift get_capMappedASID_CL_def
3856                    cong: if_cong)
3857         apply csymbr
3858         apply (ctac add: ccorres_injection_handler_csum1 [OF ccorres_injection_handler_csum1,
3859                                                           OF findVSpaceForASID_ccorres])
3860            apply (simp add: Collect_False if_to_top_of_bindE)
3861            apply (rule ccorres_if_cond_throws[rotated -1, where Q=\<top> and Q'=\<top>])
3862               apply vcg
3863              apply (clarsimp simp: isCap_simps get_capPtr_CL_def)
3864              apply (frule cap_get_tag_isCap_unfolded_H_cap)
3865              apply (erule_tac v="rv'" in ccap_relationE, clarsimp simp: cap_to_H_def)
3866             apply (rule syscall_error_throwError_ccorres_n)
3867             apply (simp add: syscall_error_to_H_cases)
3868            apply (simp add: bindE_assoc,
3869                   simp add: liftE_bindE)
3870            apply (rule ccorres_pre_getObject_pml4e)
3871            apply (clarsimp simp: get_capPtr_CL_def)
3872            apply (rename_tac pml4e pml4_mapped_asid)
3873            apply csymbr
3874            apply csymbr
3875            apply (rename_tac pml4e_present)
3876            apply (rule ccorres_abstract_cleanup)
3877            apply (clarsimp simp add: unlessE_def if_to_top_of_bindE injection_handler_If)
3878            apply (rule_tac Q'="\<lambda>s. \<exists>v. cslift s (pml4e_Ptr (lookup_pml4_slot (capPML4BasePtr (capCap (fst (extraCaps ! 0))))
3879                                                (args ! 0 && 0xFFFFFF8000000000))) = Some v
3880                             \<and> cpml4e_relation pml4e v \<and> pml4e_present = pml4e_CL.present_CL (pml4e_lift v)"
3881                        and Q="valid_cap' (ArchObjectCap cp)"
3882                        in ccorres_if_cond_throws2[rotated -1])
3883               apply vcg
3884              apply (fastforce simp: cpml4e_relation_def Let_def split: pml4e.split_asm)
3885             apply (simp add: injection_handler_throwError throwError_bind invocationCatch_def)
3886             apply (rule syscall_error_throwError_ccorres_n)
3887             apply (simp add: syscall_error_to_H_cases)
3888            apply (simp add: injection_handler_returnOk performX64MMUInvocations bindE_assoc)
3889            apply csymbr
3890            apply csymbr
3891            apply csymbr
3892            apply csymbr
3893            apply csymbr
3894            apply csymbr
3895            apply (ctac add: setThreadState_ccorres)
3896              apply (ctac (no_vcg) add: performPDPTInvocationMap_ccorres)
3897                apply (rule ccorres_alternative2)
3898                apply (rule ccorres_return_CE, simp+)[1]
3899               apply (rule ccorres_inst[where P=\<top> and P'=UNIV], fastforce)
3900              apply wp
3901             apply wp
3902            apply (simp add: bit_simps)
3903            apply clarsimp
3904            apply (rule_tac t=pml4_mapped_asid and
3905                            s="the (capPML4MappedASID (capCap (fst (extraCaps ! 0))))"
3906                            in ssubst, fastforce)
3907            apply (rule conseqPre, vcg exspec=setThreadState_modifies)
3908            apply (rule subset_refl)
3909           apply (rename_tac find_ret)
3910           apply clarsimp
3911           apply (rule_tac P'="{s. errstate s = find_ret}" in ccorres_from_vcg_split_throws[where P=\<top>])
3912            apply vcg
3913           apply (rule conseqPre, vcg)
3914           apply (clarsimp simp: throwError_def return_def syscall_error_rel_def syscall_error_to_H_cases exception_defs false_def)
3915           apply (erule lookup_failure_rel_fault_lift[rotated])
3916           apply (fastforce simp: exception_defs)
3917          apply clarsimp
3918          apply (wp injection_wp[OF refl] hoare_drop_imps)
3919         apply (clarsimp simp: get_capPtr_CL_def)
3920         apply (rename_tac pml4_mapped_asid)
3921         apply (rule_tac t=pml4_mapped_asid and
3922                         s="the (capPML4MappedASID (capCap (fst (extraCaps ! 0))))"
3923                         in ssubst, fastforce)
3924         apply (vcg exspec=findVSpaceForASID_modifies)
3925        apply wpsimp
3926       apply vcg
3927      apply wpsimp
3928     apply (vcg exspec=getSyscallArg_modifies)
3929    apply wpsimp
3930   apply (vcg exspec=getSyscallArg_modifies)
3931  apply (clarsimp simp: cte_wp_at_ctes_of excaps_map_def
3932                        word_sle_def word_sless_def bit_simps)
3933  apply (rule conjI)
3934   apply (auto simp: ct_in_state'_def isCap_simps valid_tcb_state'_def
3935              elim!: pred_tcb'_weakenE dest!: st_tcb_at_idle_thread')[1]
3936  apply (rule conjI)
3937   apply (clarsimp simp: isCap_simps)
3938   apply (subgoal_tac "s \<turnstile>' fst (extraCaps ! 0)")
3939    apply (clarsimp simp: sysargs_rel_to_n word_le_nat_alt
3940                          linorder_not_less)
3941    apply (clarsimp | drule length_le_helper)+
3942    apply (clarsimp simp: valid_cap'_def neq_Nil_conv
3943                          mask_add_aligned page_directory_at'_def
3944                          word_le_nat_alt[symmetric] bit_simps)
3945    apply (auto simp: ct_in_state'_def pred_tcb_at' mask_def valid_tcb_state'_def
3946               elim!: pred_tcb'_weakenE dest!: st_tcb_at_idle_thread')[1]
3947   apply (auto simp: neq_Nil_conv excaps_in_mem_def slotcap_in_mem_def)[1]
3948  apply (rule conjI)
3949   apply (fastforce simp: rf_sr_ksCurThread "StrictC'_thread_state_defs"
3950                          mask_eq_iff_w2p word_size
3951                          ct_in_state'_def st_tcb_at'_def
3952                          word_sle_def word_sless_def
3953                          typ_heap_simps' bit_simps)
3954  apply (frule cap_get_tag_isCap_unfolded_H_cap)
3955  apply clarsimp
3956  apply (rename_tac pml4_mapped_asid)
3957  apply (rule conjI[rotated], fastforce simp: cap_tag_defs)
3958  apply clarsimp
3959  apply (rule context_conjI)
3960   (* get_capMappedASID_CL *)
3961   apply (fastforce elim!: ccap_relationE[where c="ArchObjectCap (PML4Cap _ _)"]
3962                    simp: neq_Nil_conv valid_cap_simps' isCap_simps
3963                          get_capMappedASID_CL_def cap_pml4_cap_lift cap_to_H_simps
3964                    split: if_split_asm)
3965  apply (clarsimp simp: dc_simp neq_Nil_conv[where xs=extraCaps]
3966                        excaps_in_mem_def slotcap_in_mem_def
3967                 dest!: sym[where s="ArchObjectCap cp" for cp])
3968  apply (clarsimp simp: word_less_nat_alt hd_conv_nth dest!: length_ineq_not_Nil)
3969  apply (rule conjI, fastforce simp: mask_def)
3970  apply (rule conjI[rotated])
3971   (* array assertion and page_map_l4_at' *)
3972   apply (clarsimp simp: mask_def typ_heap_simps')
3973   apply (erule array_assertion_page_map_l4_at')
3974   apply (subgoal_tac "s \<turnstile>' fst (extraCaps ! 0)")
3975    prefer 2
3976    apply (fastforce simp: neq_Nil_conv excaps_in_mem_def slotcap_in_mem_def dest: ctes_of_valid')
3977   apply (clarsimp simp: valid_cap_simps' isCap_simps cap_lift_pml4_cap)
3978   apply (erule ccap_relationE[where c="ArchObjectCap (PML4Cap _ _)"])
3979   apply (clarsimp simp: get_capMappedASID_CL_def)
3980   apply (subst cap_lift_PML4Cap_Base[symmetric]; (assumption | rule sym, assumption))
3981  apply (clarsimp simp: rf_sr_ksCurThread)
3982  apply (rule conjI, fastforce simp: ThreadState_Restart_def mask_def)
3983  (* ccap_relation *)
3984  apply (rule conjI)
3985   apply (erule ccap_relationE[where c="ArchObjectCap (PDPointerTableCap _ _)"])
3986   apply (clarsimp simp: ccap_relation_def cap_pdpt_cap_lift isCap_simps cap_to_H_def)
3987   apply (rule conjI)
3988    (* get_capMappedASID_CL *)
3989    apply (fastforce simp: get_capMappedASID_CL_def mask_def
3990                     dest!: cap_lift_pml4_cap)
3991   apply (fastforce simp: word_bw_assocs sign_extend_canonical_address le_def[symmetric] mask_def
3992                          le_user_vtop_canonical_address[simplified user_vtop_def X64.pptrUserTop_def word_le_nat_alt])
3993  (* cmpl4e_relation *)
3994  apply (rule conjI)
3995   apply (clarsimp simp: cpml4e_relation_def superuser_from_vm_rights_def writable_from_vm_rights_def
3996                         x64CacheDisabled_def attribsFromWord_def word_and_1 nth_shiftr
3997                   split: if_split)
3998   apply (clarsimp elim!: ccap_relationE dest!: cap_lift_pdpt_cap simp: cap_to_H_simps cap_pdpt_cap_lift_def)
3999   (* addrFromPPtr *)
4000   apply (rule addrFromPPtr_mask_middle_pml4ShiftBits[simplified pml4ShiftBits_def bit_simps, simplified])
4001    apply (fastforce simp: valid_cap_simps' capAligned_def bit_simps)
4002   apply (fastforce dest!: pd_pointer_table_pdpte_atI'[where x=0, simplified bit_simps, simplified]
4003                    intro!: obj_at_kernel_mappings'
4004                    simp: valid_cap_simps' typ_at_to_obj_at_arches)
4005  apply (fastforce simp: mask_def typ_heap_simps')
4006  done
4007
4008lemma decodeX64ModeMMUInvocation_ccorres:
4009  "\<lbrakk> interpret_excaps extraCaps' = excaps_map extraCaps; \<not> isPageCap cp; \<not> isPageTableCap cp;
4010     \<not> isASIDControlCap cp; \<not> isASIDPoolCap cp \<rbrakk>
4011   \<Longrightarrow>
4012   ccorres (intr_and_se_rel \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
4013       (invs' and (\<lambda>s. ksCurThread s = thread) and ct_active' and sch_act_simple
4014              and (excaps_in_mem extraCaps \<circ> ctes_of)
4015              and cte_wp_at' (diminished' (ArchObjectCap cp) \<circ> cteCap) slot
4016              and (\<lambda>s. \<forall>v \<in> set extraCaps. ex_cte_cap_wp_to' isCNodeCap (snd v) s)
4017              and valid_cap' (ArchObjectCap cp)
4018              and sysargs_rel args buffer and valid_objs')
4019       (UNIV \<inter> {s. label___unsigned_long_' s = label}
4020             \<inter> {s. unat (length___unsigned_long_' s) = length args}
4021             \<inter> {s. cte_' s = cte_Ptr slot}
4022             \<inter> {s. extraCaps___struct_extra_caps_C_' s = extraCaps'}
4023             \<inter> {s. ccap_relation (ArchObjectCap cp) (cap_' s)}
4024             \<inter> {s. buffer_' s = option_to_ptr buffer})
4025       hs
4026       (decodeX64MMUInvocation label args cptr slot cp extraCaps
4027              >>= invocationCatch thread isBlocking isCall InvokeArchObject)
4028       (Call decodeX86ModeMMUInvocation_'proc)"
4029  supply Collect_const[simp del] if_cong[cong]
4030  apply (cinit' lift: label___unsigned_long_' length___unsigned_long_' cte_'
4031                      extraCaps___struct_extra_caps_C_' cap_' buffer_')
4032   apply csymbr
4033   apply (clarsimp simp: cap_get_tag_isCap_ArchObject
4034                   cong: StateSpace.state.fold_congs globals.fold_congs)
4035   apply (rule ccorres_Cond_rhs)
4036    apply (clarsimp simp: decodeX64MMUInvocation_def isCap_simps throwError_bind invocationCatch_def)
4037    apply (rule syscall_error_throwError_ccorres_n)
4038    apply (solves \<open>simp add: syscall_error_to_H_cases\<close>)
4039   apply (rule ccorres_Cond_rhs)
4040    apply (rule ccorres_trim_returnE, simp+)
4041    apply (rule ccorres_call,
4042           rule decodeX64PDPTInvocation_ccorres;
4043           solves simp)
4044   apply (rule ccorres_Cond_rhs)
4045    apply (rule ccorres_trim_returnE, simp+)
4046    apply (rule ccorres_call,
4047           rule decodeX64PageDirectoryInvocation_ccorres;
4048           solves simp)
4049   apply (fastforce intro: ccorres_fail simp: decodeX64MMUInvocation_def)
4050  apply (simp add: o_def)
4051  done
4052
4053lemma decodeX64MMUInvocation_ccorres:
4054  notes Collect_const[simp del] if_cong[cong]
4055  shows
4056  "\<lbrakk> interpret_excaps extraCaps' = excaps_map extraCaps \<rbrakk>
4057   \<Longrightarrow>
4058   ccorres (intr_and_se_rel \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
4059       (invs' and (\<lambda>s. ksCurThread s = thread) and ct_active' and sch_act_simple
4060              and (excaps_in_mem extraCaps \<circ> ctes_of)
4061              and cte_wp_at' (diminished' (ArchObjectCap cp) \<circ> cteCap) slot
4062              and (\<lambda>s. \<forall>v \<in> set extraCaps. ex_cte_cap_wp_to' isCNodeCap (snd v) s)
4063              and sysargs_rel args buffer and valid_objs')
4064       (UNIV \<inter> {s. invLabel_' s = label}
4065             \<inter> {s. unat (length___unsigned_long_' s) = length args}
4066             \<inter> {s. cte_' s = cte_Ptr slot}
4067             \<inter> {s. excaps_' s = extraCaps'}
4068             \<inter> {s. ccap_relation (ArchObjectCap cp) (cap_' s)}
4069             \<inter> {s. buffer_' s = option_to_ptr buffer}) []
4070       (decodeX64MMUInvocation label args cptr slot cp extraCaps
4071              >>= invocationCatch thread isBlocking isCall InvokeArchObject)
4072       (Call decodeX86MMUInvocation_'proc)"
4073  apply (cinit' lift: invLabel_' length___unsigned_long_' cte_'
4074                      excaps_' cap_' buffer_')
4075   apply csymbr
4076   apply (simp add: cap_get_tag_isCap_ArchObject
4077                    X64_H.decodeInvocation_def
4078                    invocation_eq_use_types
4079              cong: StateSpace.state.fold_congs globals.fold_congs)
4080   apply (rule ccorres_Cond_rhs)
4081    apply (rule ccorres_trim_returnE, simp+)
4082    apply (rule ccorres_call,
4083           rule decodeX64FrameInvocation_ccorres,
4084           simp+)[1]
4085   apply (rule ccorres_Cond_rhs)
4086    apply (rule ccorres_trim_returnE, simp+)
4087    apply (rule ccorres_call,
4088           rule decodeX64PageTableInvocation_ccorres, simp+)[1]
4089   (* ASIDControlCap *)
4090   apply (rule ccorres_Cond_rhs)
4091    apply (rule ccorres_trim_returnE, simp+)
4092    apply (rule ccorres_rhs_assoc)+
4093    apply csymbr+
4094    apply (rule ccorres_Cond_rhs_Seq)
4095     apply (rule ccorres_equals_throwError)
4096      apply (fastforce simp: decodeX64MMUInvocation_def decodeX64ASIDControlInvocation_def isCap_simps
4097                             throwError_bind invocationCatch_def
4098                      split: invocation_label.split arch_invocation_label.split)
4099     apply ccorres_rewrite
4100     apply (rule syscall_error_throwError_ccorres_n[simplified id_def o_def])
4101     apply (fastforce simp: syscall_error_to_H_cases)
4102    (* X64ASIDControlMakePool *)
4103    apply (clarsimp simp: decodeX64MMUInvocation_def decodeX64ASIDControlInvocation_def isCap_simps)
4104    apply (simp add: word_less_nat_alt list_case_If2 split_def)
4105    apply csymbr
4106    apply (rule ccorres_Cond_rhs_Seq)
4107     (* args malformed *)
4108     apply (rule ccorres_cond_true_seq | simp)+
4109     apply (simp add: throwError_bind invocationCatch_def)
4110     apply ccorres_rewrite
4111     apply (rule syscall_error_throwError_ccorres_n[simplified id_def o_def])
4112     apply (fastforce simp: syscall_error_to_H_cases)
4113    apply (simp add: interpret_excaps_test_null excaps_map_def)
4114    apply csymbr
4115    apply (rule ccorres_Cond_rhs_Seq)
4116     (* extraCaps malformed *)
4117     apply (rule ccorres_cond_true_seq | simp)+
4118     apply (simp add: throwError_bind invocationCatch_def)
4119     apply ccorres_rewrite
4120     apply (rule syscall_error_throwError_ccorres_n[simplified id_def o_def])
4121     apply (fastforce simp: syscall_error_to_H_cases)
4122    apply csymbr
4123    apply (simp add: interpret_excaps_test_null[OF Suc_leI])
4124    apply (rule ccorres_Cond_rhs_Seq)
4125     apply (simp add: length_ineq_not_Nil throwError_bind invocationCatch_def)
4126     apply ccorres_rewrite
4127     apply (rule syscall_error_throwError_ccorres_n[simplified id_def o_def])
4128     apply (simp add: syscall_error_to_H_cases)
4129    apply (subgoal_tac "1 < length extraCaps")
4130     prefer 2
4131     apply (rule neq_le_trans, simp)
4132     apply (fastforce simp: Suc_leI)
4133    apply (simp add: Let_def split_def liftE_bindE bind_assoc length_ineq_not_Nil)
4134    apply (rule ccorres_add_return)
4135    apply (ctac add: getSyscallArg_ccorres_foo[where args=args and n=0 and buffer=buffer])
4136      apply (rule ccorres_add_return)
4137      apply (ctac add: getSyscallArg_ccorres_foo[where args=args and n=1 and buffer=buffer])
4138        apply csymbr
4139        apply (rule ccorres_add_return,
4140               rule_tac xf'=untyped_' and
4141                        r'="(\<lambda>rv _ un.
4142                              (cap_get_tag un = scast cap_untyped_cap) = isUntypedCap rv \<and>
4143                              (isUntypedCap rv \<longrightarrow> ccap_relation rv un))
4144                            (fst (extraCaps ! 0))"
4145                        in ccorres_split_nothrow)
4146            apply (rule_tac P="excaps_in_mem extraCaps \<circ> ctes_of"
4147                            in ccorres_from_vcg[where P'=UNIV])
4148            apply (rule allI, rule conseqPre, vcg)
4149            apply (frule length_ineq_not_Nil[where xs=extraCaps])
4150            apply (clarsimp simp: return_def neq_Nil_conv excaps_in_mem_def
4151                                  slotcap_in_mem_def)
4152            apply (drule interpret_excaps_eq[rule_format, where n=0], simp)
4153            apply (simp add: mask_def[where n=4])
4154            apply (erule(1) cmap_relationE1[OF cmap_relation_cte])
4155            apply (rule conjI, fastforce intro: typ_heap_simps)
4156            apply (drule ccte_relation_ccap_relation)
4157            apply (simp add: typ_heap_simps cap_get_tag_isCap)
4158           apply ceqv
4159          apply (rename_tac untyped')
4160          apply (rule getSlotCap_ccorres_fudge_n[where vals=extraCaps and n=1])
4161          apply (rule ccorres_move_c_guard_cte)
4162          apply ctac
4163            apply (rule ccorres_assert2)
4164            apply (rule ccorres_pre_gets_x86KSASIDTable_ksArchState)
4165            apply (rename_tac "x86KSASIDTab")
4166            apply (rule ccorres_rhs_assoc2, rule ccorres_rhs_assoc2,
4167                   rule ccorres_rhs_assoc2)
4168            apply (rule ccorres_add_return)
4169            apply (rule_tac r'="\<lambda>rv rv'. rv' = (case [p \<leftarrow> assocs x86KSASIDTab.
4170                                                      fst p < 2 ^ asid_high_bits \<and> snd p = None]
4171                                                of [] \<Rightarrow> 2 ^ asid_high_bits | x # xs \<Rightarrow> fst x)"
4172                    and xf'=i_' in ccorres_split_nothrow)
4173                apply (rule_tac P="\<forall>x \<in> ran x86KSASIDTab. x \<noteq> 0"
4174                                in ccorres_gen_asm)
4175                apply (rule_tac P="\<lambda>s. x86KSASIDTab = x64KSASIDTable (ksArchState s)"
4176                                in ccorres_from_vcg[where P'=UNIV])
4177                apply (clarsimp simp: return_def)
4178                apply (rule HoarePartial.SeqSwap)
4179                 apply (rule_tac I="{t. (\<sigma>, t) \<in> rf_sr \<and> i_' t \<le> 2 ^ asid_high_bits
4180                                        \<and> x86KSASIDTab = x64KSASIDTable (ksArchState \<sigma>)
4181                                        \<and> (\<forall>x < i_' t. x86KSASIDTab x \<noteq> None)
4182                                        \<and> ret__int_' t = from_bool (i_' t < 2 ^ asid_high_bits \<and>
4183                                                                    x86KSASIDTab (i_' t) \<noteq> None)}"
4184                                 in HoarePartial.reannotateWhileNoGuard)
4185                 apply (rule HoarePartial.While[OF order_refl])
4186                  apply (rule conseqPre, vcg)
4187                  apply (clarsimp simp: asidHighBits_handy_convs
4188                                        word_sle_def word_sless_def
4189                                        word_less_nat_alt[symmetric]
4190                                        from_bool_0)
4191                  apply (cut_tac P="\<lambda>y. y < i_' x + 1 = rhs y" for rhs in allI,
4192                         rule less_x_plus_1)
4193                   apply (fastforce simp: max_word_def asid_high_bits_def)
4194                  apply (clarsimp simp: rf_sr_x86KSASIDTable from_bool_def
4195                                        asid_high_bits_word_bits
4196                                        option_to_ptr_def option_to_0_def
4197                                        order_less_imp_le
4198                                        linorder_not_less
4199                                        order_antisym[OF inc_le])
4200                  apply (clarsimp simp: true_def false_def
4201                                 split: option.split if_split)
4202                  apply (auto simp: asid_high_bits_def word_le_nat_alt
4203                                    word_less_nat_alt unat_add_lem[THEN iffD1]
4204                                    Kernel_C_defs)[1]
4205                 apply (clarsimp simp: from_bool_0)
4206                 apply (case_tac "i_' x = 2 ^ asid_high_bits")
4207                  apply (clarsimp split: list.split)
4208                  apply (drule_tac f="\<lambda>xs. (a, b) \<in> set xs" in arg_cong)
4209                  apply (clarsimp simp: in_assocs_is_fun)
4210                  apply fastforce
4211                 apply (frule(1) neq_le_trans)
4212                 apply (subst filter_assocs_Cons)
4213                   apply fastforce
4214                  apply simp
4215                 apply simp
4216                apply (rule conseqPre, vcg)
4217                apply (clarsimp simp: asidHighBits_handy_convs word_sle_def
4218                                      word_sless_def from_bool_0
4219                                      rf_sr_x86KSASIDTable[where n=0, simplified])
4220                apply (simp add: asid_high_bits_def option_to_ptr_def option_to_0_def
4221                                 from_bool_def Kernel_C_defs
4222                          split: option.split if_split)
4223                apply fastforce
4224               apply ceqv
4225              apply (rule ccorres_Guard_Seq)+
4226              apply (simp add: whenE_bindE_throwError_to_if if_to_top_of_bind)
4227              apply (rule_tac Q=\<top> and Q'=\<top> in ccorres_if_cond_throws[rotated -1])
4228                 apply clarsimp
4229                 apply (rule conseqPre, vcg, rule subset_refl)
4230                apply (clarsimp simp: asid_high_bits_word_bits asidHighBits_handy_convs null_def)
4231                apply (clarsimp split: list.split)
4232                apply (fastforce dest!: filter_eq_ConsD)
4233               apply (simp add: throwError_bind invocationCatch_def)
4234               apply (rule syscall_error_throwError_ccorres_n[simplified id_def o_def])
4235               apply (fastforce simp: syscall_error_to_H_cases)
4236              apply (rule ccorres_Guard_Seq)+
4237              apply (simp add: invocationCatch_use_injection_handler
4238                               injection_bindE[OF refl refl] injection_handler_If
4239                               injection_handler_returnOk bindE_assoc
4240                               injection_handler_throwError
4241                         cong: if_cong)
4242              apply csymbr
4243              apply csymbr
4244              apply csymbr
4245              apply (rule ccorres_symb_exec_r)
4246                apply (rule_tac xf'=ret__int_' in ccorres_abstract, ceqv)
4247                  apply (rule_tac P="rv'a = from_bool (\<not> (isUntypedCap (fst (hd extraCaps)) \<and>
4248                                                          capBlockSize (fst (hd extraCaps)) = objBits (makeObject ::asidpool)))"
4249                                  in ccorres_gen_asm2)
4250                apply (rule ccorres_symb_exec_r)
4251                  apply (rule_tac xf'=ret__int_' in ccorres_abstract, ceqv)
4252                  apply (rule_tac P="rv'b = from_bool (\<not> (isUntypedCap (fst (hd extraCaps)) \<and>
4253                                                          capBlockSize (fst (hd extraCaps)) = objBits (makeObject ::asidpool) \<and>
4254                                                          \<not> capIsDevice (fst (hd extraCaps))))"
4255                                  in ccorres_gen_asm2)
4256                  apply (clarsimp simp:  to_bool_if cond_throw_whenE bindE_assoc)
4257                  apply (rule ccorres_split_when_throwError_cond[where Q = \<top> and Q' = \<top>])
4258                     apply fastforce
4259                    apply (rule syscall_error_throwError_ccorres_n[simplified id_def])
4260                    apply (clarsimp simp: syscall_error_rel_def shiftL_nat syscall_error_to_H_cases)
4261                   prefer 2
4262                   apply vcg
4263                  apply clarsimp
4264                  apply (ctac add: ccorres_injection_handler_csum1[OF ensureNoChildren_ccorres])
4265                     apply (clarsimp simp: Collect_False)
4266                     apply csymbr
4267                     apply csymbr
4268                     apply (ctac add: ccorres_injection_handler_csum1
4269                                        [OF lookupTargetSlot_ccorres,
4270                                         unfolded lookupTargetSlot_def])
4271                        apply (simp add: Collect_False split_def)
4272                        apply csymbr
4273                        apply (ctac add: ccorres_injection_handler_csum1
4274                                                [OF ensureEmptySlot_ccorres])
4275                           apply (simp add: ccorres_invocationCatch_Inr
4276                                            performInvocation_def
4277                                            X64_H.performInvocation_def
4278                                            performX64MMUInvocation_def)
4279                           apply (simp add: liftE_bindE)
4280                           apply ccorres_rewrite
4281                           apply (rule_tac P="\<lambda>s. ksCurThread s = thread" in ccorres_cross_over_guard)
4282                           apply (ctac add: setThreadState_ccorres)
4283                             apply (simp only: liftE_bindE[symmetric])
4284                             apply (rule ccorres_seq_skip'[THEN iffD1])
4285                             apply (ctac (no_vcg) add: performASIDControlInvocation_ccorres
4286                                                       [where idx = "capFreeIndex (fst (extraCaps ! 0))"])
4287                               apply (rule ccorres_alternative2)
4288                               apply (rule ccorres_returnOk_skip)
4289                              apply (rule ccorres_inst[where P=\<top> and P'=UNIV], simp)
4290                             apply wp
4291                            apply (wp sts_invs_minor' sts_Restart_ct_active)
4292                           apply simp
4293                           apply (vcg exspec=setThreadState_modifies)
4294                          apply ccorres_rewrite
4295                          apply (rule ccorres_return_C_errorE, simp+)
4296                         apply (wp injection_wp[OF refl])
4297                        apply (simp add: all_ex_eq_helper)
4298                        (* This manual conseqPost prevents the VCG from instantiating False *)
4299                        apply (rule_tac Q'=UNIV and A'="{}" in conseqPost)
4300                          apply (vcg exspec=ensureEmptySlot_modifies)
4301                         apply (frule length_ineq_not_Nil)
4302                         apply (clarsimp simp: null_def ThreadState_Restart_def mask_def hd_conv_nth
4303                                               isCap_simps rf_sr_ksCurThread cap_get_tag_UntypedCap
4304                                               word_le_make_less asid_high_bits_def
4305                                         split: list.split)
4306                         apply (frule interpret_excaps_eq[rule_format, where n=0], fastforce)
4307                         apply (fastforce simp: interpret_excaps_test_null excaps_map_def split_def)
4308                        apply fastforce
4309                       apply ccorres_rewrite
4310                       apply (rule ccorres_return_C_errorE, simp+)
4311                      apply (wp injection_wp[OF refl] hoare_drop_imps)
4312                     apply (simp add: split_def all_ex_eq_helper)
4313                     apply (vcg exspec=lookupTargetSlot_modifies)
4314                    apply simp
4315                    apply ccorres_rewrite
4316                    apply (rule ccorres_return_C_errorE, simp+)
4317                   apply (wp injection_wp[OF refl] ensureNoChildren_wp)
4318                  apply (simp add: all_ex_eq_helper cap_get_tag_isCap)
4319                  apply (vcg exspec=ensureNoChildren_modifies)
4320                 apply clarsimp
4321                 apply vcg
4322                apply clarsimp
4323                apply (rule conseqPre, vcg, clarsimp)
4324               apply clarsimp
4325               apply vcg
4326              apply clarsimp
4327              apply (rule conseqPre, vcg, clarsimp)
4328             apply wp
4329            apply (simp add: cap_get_tag_isCap)
4330            apply (rule HoarePartial.SeqSwap)
4331             apply (rule_tac I="\<lbrace>Prop \<acute>ksCurThread \<acute>root\<rbrace>"
4332                         and Q="\<lbrace>Bonus \<acute>i \<longrightarrow> Prop \<acute>ksCurThread \<acute>root\<rbrace>"
4333                         for Prop Bonus in HoarePartial.reannotateWhileNoGuard)
4334             apply (rule HoarePartial.While[OF order_refl])
4335              apply (rule conseqPre, vcg)
4336              apply clarify
4337              apply (rule conjI)
4338               apply clarify
4339              apply (simp (no_asm))
4340              apply clarify
4341             apply clarsimp
4342            apply vcg
4343           apply simp
4344           apply (rule hoare_drop_imps)
4345           apply wp
4346          apply simp
4347          apply vcg
4348         apply simp
4349         apply wp
4350        apply vcg
4351       apply wp
4352      apply simp
4353      apply (vcg exspec=getSyscallArg_modifies)
4354     apply simp
4355     apply wp
4356    apply simp
4357    apply (vcg exspec=getSyscallArg_modifies)
4358   (* ASIDPoolCap case *)
4359   apply (rule ccorres_Cond_rhs)
4360    apply (simp add: imp_conjR[symmetric] decodeX64MMUInvocation_def)
4361    apply (rule ccorres_rhs_assoc)+
4362    apply csymbr+
4363    apply (rule ccorres_Cond_rhs_Seq)
4364     apply (clarsimp simp: isCap_simps decodeX64ASIDPoolInvocation_def)
4365     apply ccorres_rewrite
4366     apply (rule ccorres_equals_throwError)
4367      apply (fastforce simp: throwError_bind invocationCatch_def
4368                      split: invocation_label.split arch_invocation_label.split)
4369     apply (rule syscall_error_throwError_ccorres_n)
4370     apply (fastforce simp: syscall_error_to_H_cases)
4371    apply (simp add: interpret_excaps_test_null excaps_map_def
4372                     list_case_If2 split_def)
4373    apply (rule ccorres_Cond_rhs_Seq)
4374     apply ccorres_rewrite
4375     apply (clarsimp simp: isCap_simps decodeX64ASIDPoolInvocation_def
4376                           throwError_bind invocationCatch_def)
4377     apply (rule syscall_error_throwError_ccorres_n[simplified dc_def id_def o_def])
4378     apply (fastforce simp: syscall_error_to_H_cases)
4379    apply (clarsimp simp: isCap_simps decodeX64ASIDPoolInvocation_def split: list.split)
4380    apply csymbr
4381    apply (rule ccorres_add_return)
4382    apply (rule ccorres_Guard_Seq)
4383    apply (rule ccorres_rhs_assoc2, rule ccorres_rhs_assoc2,
4384           rule ccorres_rhs_assoc2)
4385    apply (rule_tac R="excaps_in_mem extraCaps \<circ> ctes_of" and
4386                    R'="UNIV" and
4387                    val="from_bool (\<not> (isArchObjectCap (fst (extraCaps ! 0)) \<and>
4388                                       isPML4Cap (capCap (fst (extraCaps ! 0)))) \<or>
4389                                    capPML4MappedASID (capCap (fst (extraCaps ! 0))) \<noteq> None)" and
4390                    xf'=ret__int_' in ccorres_symb_exec_r_known_rv)
4391       apply vcg
4392       apply clarsimp
4393       apply (frule interpret_excaps_eq[rule_format, where n=0], simp)
4394       apply (clarsimp simp: excaps_in_mem_def)
4395       apply (frule (1) slotcap_in_mem_PML4)
4396       apply (clarsimp simp: typ_heap_simps' from_bool_0 split: if_split)
4397       apply (fastforce simp: isCap_simps asidInvalid_def cap_lift_pml4_cap cap_to_H_simps
4398                              get_capMappedASID_CL_def true_def c_valid_cap_def cl_valid_cap_def
4399                        elim!: ccap_relationE split: if_splits)
4400      apply ceqv
4401     apply (rule ccorres_Cond_rhs_Seq)
4402      apply ccorres_rewrite
4403      apply (rule_tac v="Inl (InvalidCapability 1)" in ccorres_equals_throwError)
4404       apply (fastforce simp: isCap_simps throwError_bind invocationCatch_def
4405                        split: capability.split arch_capability.split)
4406      apply (rule syscall_error_throwError_ccorres_n)
4407      apply (fastforce simp: syscall_error_to_H_cases)
4408     apply (clarsimp simp: isCap_simps Kernel_C_defs liftE_bindE bind_assoc)
4409     apply (rule ccorres_pre_gets_x86KSASIDTable_ksArchState)
4410     apply csymbr
4411     apply (rule ccorres_Guard_Seq)+
4412     apply (rule ccorres_add_return)
4413     apply (rule_tac r'="\<lambda>_ rv'. rv' = option_to_ptr (x (ucast (asid_high_bits_of (capASIDBase cp))))
4414                                 \<and> x (ucast (asid_high_bits_of (capASIDBase cp))) \<noteq> Some 0"
4415                 and xf'=pool_' in ccorres_split_nothrow)
4416         apply (rule_tac P="\<lambda>s. x = x64KSASIDTable (ksArchState s)
4417                                \<and> valid_arch_state' s \<and> s \<turnstile>' ArchObjectCap cp"
4418                         in ccorres_from_vcg[where P'=UNIV])
4419         apply (rule allI, rule conseqPre, vcg)
4420         apply (clarsimp simp: return_def valid_arch_state'_def valid_asid_table'_def)
4421         apply (frule cap_get_tag_isCap_ArchObject(2))
4422         apply (clarsimp simp: isCap_simps)
4423         apply (erule_tac v=cap in ccap_relationE)
4424         apply (clarsimp simp: cap_lift_asid_pool_cap cap_to_H_simps valid_cap_simps'
4425                               cap_asid_pool_cap_lift_def)
4426         apply (subst rf_sr_x86KSASIDTable, assumption)
4427          apply (simp add: asid_high_bits_word_bits)
4428          apply (rule shiftr_less_t2n)
4429          apply (fastforce simp: asid_low_bits_def asid_high_bits_def asid_wf_def mask_def
4430                                 asid_bits_def word_le_make_less)
4431         apply (fastforce simp: ucast_asid_high_bits_is_shift asid_wf_def mask_def)
4432        apply ceqv
4433       apply (simp add: whenE_bindE_throwError_to_if if_to_top_of_bind cong: if_cong)
4434       apply (rule_tac Q=\<top> and Q'=\<top> in ccorres_if_cond_throws[rotated -1])
4435          apply vcg
4436         apply (simp add: option_to_0_def option_to_ptr_def split: option.split)
4437         apply fastforce
4438        apply (simp add: throwError_bind invocationCatch_def)
4439        apply (rule ccorres_from_vcg_throws[where P=\<top> and P'=UNIV])
4440        apply (rule allI, rule conseqPre, vcg)
4441        apply (clarsimp simp: throwError_def return_def
4442                              syscall_error_rel_def exception_defs
4443                              syscall_error_to_H_cases false_def)
4444        apply (simp add: lookup_fault_lift_invalid_root)
4445       apply csymbr
4446       apply (simp add: liftME_def bindE_assoc if_to_top_of_bind)
4447       apply (rule_tac Q=\<top> and Q'=\<top> in ccorres_if_cond_throws[rotated -1])
4448          apply vcg
4449         apply (frule cap_get_tag_isCap_ArchObject(2))
4450         apply (clarsimp simp: isCap_simps)
4451         apply (erule_tac v=cap in ccap_relationE)
4452         apply (fastforce simp: cap_lift_asid_pool_cap cap_to_H_simps valid_cap_simps'
4453                                cap_asid_pool_cap_lift_def)
4454        apply (simp add: throwError_bind invocationCatch_def)
4455        apply (rule syscall_error_throwError_ccorres_n)
4456        apply (fastforce simp: syscall_error_to_H_cases)
4457       apply csymbr
4458       apply csymbr
4459       apply (rule ccorres_rhs_assoc2, rule ccorres_rhs_assoc2,
4460              rule ccorres_rhs_assoc2)
4461       apply (simp add: bind_assoc liftE_bindE)
4462       apply (rule_tac xf'=i_' and r'="\<lambda>rv rv'. rv' = (case [(x, y) \<leftarrow> assocs (inv ASIDPool rv).
4463                                                              x \<le> 2 ^ asid_low_bits - 1 \<and> x + capASIDBase cp \<noteq> 0
4464                                                              \<and> y = None] of [] \<Rightarrow> 2 ^ asid_low_bits
4465                                                             | x # xs \<Rightarrow> fst x)"
4466                       in ccorres_split_nothrow)
4467           apply (rule ccorres_add_return2)
4468           apply (rule ccorres_pre_getObject_asidpool)
4469           apply (rule_tac P="\<forall>x \<in> ran (inv ASIDPool xa). x \<noteq> 0"
4470                           in ccorres_gen_asm)
4471           apply (rule_tac P="ko_at' xa (capASIDPool cp)"
4472                           in ccorres_from_vcg[where P'=UNIV])
4473           apply (clarsimp simp: option_to_0_def option_to_ptr_def
4474                                 return_def)
4475           apply (rule HoarePartial.SeqSwap)
4476            apply (rule_tac I="{t. (\<sigma>, t) \<in> rf_sr \<and> i_' t \<le> 2 ^ asid_low_bits
4477                                 \<and> ko_at' xa (capASIDPool cp) \<sigma>
4478                                 \<and> (\<exists>v. cslift t (ap_Ptr (capASIDPool cp))
4479                                         = Some v \<and> (\<forall>x < i_' t. capASIDBase cp + x = 0
4480                                                        \<or> asid_map_get_tag (index (array_C v) (unat x)) \<noteq> scast asid_map_asid_map_none )
4481                                        \<and> ret__int_' t = from_bool (i_' t < 2 ^ asid_low_bits
4482                                             \<and> (capASIDBase cp + (i_' t) = 0
4483                                                 \<or> asid_map_get_tag (index (array_C v) (unat (i_' t))) \<noteq> scast asid_map_asid_map_none)))}"
4484                         in HoarePartial.reannotateWhileNoGuard)
4485            apply (rule HoarePartial.While[OF order_refl])
4486             apply (rule conseqPre, vcg)
4487             apply (clarsimp simp: asidLowBits_handy_convs
4488                                   word_sle_def word_sless_def from_bool_0)
4489             apply (subgoal_tac "capASIDBase_CL (cap_asid_pool_cap_lift cap)
4490                                     = capASIDBase cp")
4491              apply (subgoal_tac "\<And>x. (x < (i_' xb + 1))
4492                                        = (x < i_' xb \<or> x = i_' xb)")
4493               apply (clarsimp simp: inc_le from_bool_def typ_heap_simps
4494                                     asid_low_bits_def not_less field_simps
4495                                     false_def
4496                              split: if_split bool.splits)
4497               apply unat_arith
4498              apply (rule iffI)
4499               apply (rule disjCI)
4500               apply (drule plus_one_helper)
4501               apply simp
4502              apply (subgoal_tac "i_' xb < i_' xb + 1")
4503               apply (erule_tac P="x < y" for x y in disjE, simp_all)[1]
4504              apply (rule plus_one_helper2 [OF order_refl])
4505              apply (rule notI, drule max_word_wrap)
4506              apply (fastforce simp: max_word_def asid_low_bits_def)
4507             apply (simp add: cap_get_tag_isCap_ArchObject[symmetric])
4508             apply (frule cap_get_tag_isCap_unfolded_H_cap)
4509             apply (clarsimp simp: cap_lift_asid_pool_cap cap_to_H_def
4510                                   cap_asid_pool_cap_lift_def
4511                            elim!: ccap_relationE)
4512            apply (clarsimp simp: from_bool_0)
4513            apply (erule cmap_relationE1[OF rf_sr_cpspace_asidpool_relation],
4514                   erule ko_at_projectKO_opt)
4515            apply (clarsimp simp: typ_heap_simps casid_pool_relation_def
4516                                    inv_ASIDPool array_relation_def
4517                             split: asidpool.split_asm asid_pool_C.split_asm)
4518            apply (case_tac "i_' xb = 2 ^ asid_low_bits")
4519             apply (clarsimp split: list.split)
4520             apply (drule_tac f="\<lambda>xs. (a, ba) \<in> set xs" in arg_cong)
4521             apply (clarsimp simp: in_assocs_is_fun)
4522             apply (drule spec, drule(1) mp)
4523             apply (simp add: asid_low_bits_word_bits)
4524             apply (drule spec, drule(1) mp)
4525             apply (simp add: option_to_ptr_def option_to_0_def field_simps)
4526             apply (clarsimp simp: asid_map_relation_def asid_map_lift_def
4527                            split: if_split_asm option.splits asid_map_CL.split_asm )
4528            apply (frule(1) neq_le_trans)
4529            apply (subst filter_assocs_Cons)
4530              apply (simp add: split_def asid_low_bits_word_bits)
4531              apply (rule conjI, assumption)
4532              apply (clarsimp simp add: field_simps)
4533              apply (drule spec, drule(1) mp)
4534              apply (simp add: asid_map_relation_def asid_map_lift_asid_map_none
4535                        split: option.split_asm)
4536             apply (simp add: asid_low_bits_word_bits)
4537             apply (erule allEI, rule impI, drule(1) mp)
4538             apply (clarsimp simp: field_simps)
4539             apply (drule_tac x=xa in spec)
4540             apply (simp add: asid_map_relation_def asid_map_lift_def
4541                       split: if_split_asm option.splits asid_map_CL.split_asm)
4542            apply simp
4543           apply (rule conseqPre, vcg)
4544           apply (clarsimp simp: asidLowBits_handy_convs
4545                                 signed_shift_guard_simpler_64 asid_low_bits_def
4546                                 word_sless_def word_sle_def)
4547           apply (erule cmap_relationE1[OF rf_sr_cpspace_asidpool_relation],
4548                  erule ko_at_projectKO_opt)
4549           apply (clarsimp simp: typ_heap_simps from_bool_def split: if_split)
4550           apply (frule cap_get_tag_isCap_unfolded_H_cap)
4551           apply (clarsimp simp: cap_lift_asid_pool_cap cap_to_H_def
4552                                 cap_asid_pool_cap_lift_def false_def
4553                                 ucast_minus ucast_nat_def
4554                          elim!: ccap_relationE)
4555          apply ceqv
4556         apply (rule ccorres_Guard_Seq)+
4557         apply (simp add: if_to_top_of_bind)
4558         apply (rule ccorres_if_cond_throws[where Q=\<top> and Q'=\<top>, rotated -1])
4559            apply vcg
4560           apply (clarsimp simp: null_def split: list.split
4561                           dest!: filter_eq_ConsD)
4562           apply (simp add: asid_low_bits_def)
4563          apply (simp add: throwError_bind invocationCatch_def)
4564          apply (rule syscall_error_throwError_ccorres_n)
4565          apply (simp add: syscall_error_to_H_cases)
4566         apply (simp add: returnOk_bind ccorres_invocationCatch_Inr
4567                          performInvocation_def
4568                          X64_H.performInvocation_def
4569                          performX64MMUInvocation_def liftE_bindE)
4570         apply csymbr
4571         apply (ctac add: setThreadState_ccorres)
4572           apply (simp only: liftE_bindE[symmetric])
4573           apply (ctac(no_vcg) add: performASIDPoolInvocation_ccorres)
4574             apply (rule ccorres_alternative2)
4575             apply (rule ccorres_return_CE, simp+)[1]
4576            apply (rule ccorres_inst[where P=\<top> and P'=UNIV])
4577            apply simp
4578           apply wp
4579          apply simp
4580          apply (wp sts_invs_minor')
4581         apply simp
4582         apply (vcg exspec=setThreadState_modifies)
4583        apply simp
4584        apply (wp getASID_wp)
4585       apply simp
4586       apply (rule HoarePartial.SeqSwap)
4587        apply (rule_tac I="\<lbrace>\<forall>rv. Prop \<acute>ksCurThread \<acute>vspaceCapSlot rv\<rbrace>"
4588                    and Q="\<lbrace>\<forall>rv. Bonus \<acute>i rv \<longrightarrow> Prop \<acute>ksCurThread \<acute>vspaceCapSlot rv\<rbrace>"
4589                    for Prop Bonus in HoarePartial.reannotateWhileNoGuard)
4590        apply vcg
4591         apply fastforce
4592        apply clarsimp
4593       apply vcg
4594      apply simp
4595      (* HACK: rewrites to fix schematic dependency problems *)
4596      apply (rule_tac t=v0 and s="capASIDPool cp" in subst, fastforce)
4597      apply (rule_tac t=v1 and s="capASIDBase cp" in subst, fastforce)
4598      apply (rule_tac t=b and s="snd (extraCaps ! 0)" in subst, fastforce)
4599      apply (rule return_wp)
4600     apply vcg
4601    apply (rule_tac t=v0 and s="capASIDPool cp" in subst, fastforce)
4602    apply (rule_tac t=v1 and s="capASIDBase cp" in subst, fastforce)
4603    apply (rule_tac t=b and s="snd (extraCaps ! 0)" in subst, fastforce)
4604    apply vcg
4605   (* Mode stuff *)
4606   apply (rule ccorres_trim_returnE; simp)
4607   apply (rule ccorres_call,
4608          rule decodeX64ModeMMUInvocation_ccorres;
4609          simp)
4610  apply (clarsimp simp: o_def)
4611  apply (rule conjI)
4612   apply (clarsimp simp: cte_wp_at_ctes_of ct_in_state'_def
4613                         interpret_excaps_eq excaps_map_def)
4614   apply (frule(1) ctes_of_valid', simp only: diminished_valid'[symmetric])
4615   apply (cases "extraCaps")
4616    apply simp
4617   apply (frule interpret_excaps_eq[rule_format, where n=0], simp)
4618   apply (clarsimp simp: mask_def[where n=4] excaps_in_mem_def
4619                         slotcap_in_mem_def typ_at_to_obj_at_arches
4620                         obj_at'_weakenE[OF _ TrueI] invs_arch_state'
4621                         unat_lt2p[where 'a=machine_word_len, folded word_bits_def])
4622   apply (rule conjI)
4623    apply (frule invs_arch_state')
4624    apply (clarsimp simp: valid_arch_state'_def valid_asid_table'_def neq_Nil_conv)
4625    apply (frule interpret_excaps_eq[rule_format, where n=1], simp)
4626    apply (clarsimp simp: sysargs_rel_to_n word_less_nat_alt)
4627    apply (drule length_le_helper | clarsimp simp: linorder_not_less)+
4628    apply (rule conjI)
4629     apply clarsimp
4630    apply (clarsimp simp: tcb_at_invs' Kernel_C.asidLowBits_def)
4631    apply (clarsimp simp: invs_valid_objs')
4632    apply (rule conjI, fastforce)
4633    apply (clarsimp simp: ctes_of_valid' invs_valid_objs' isCap_simps)
4634    apply (clarsimp simp: ex_cte_cap_wp_to'_def cte_wp_at_ctes_of
4635                          invs_sch_act_wf' dest!: isCapDs(1))
4636    apply (intro conjI)
4637            apply (simp add: Invariants_H.invs_queues)
4638           apply (simp add: valid_tcb_state'_def)
4639          apply (fastforce elim!: pred_tcb'_weakenE dest!: st_tcb_at_idle_thread')
4640         apply (clarsimp simp: st_tcb_at'_def obj_at'_def)
4641         apply (case_tac "tcbState obj", (simp add: runnable'_def)+)[1]
4642        apply simp
4643       apply (simp add: objBits_simps archObjSize_def)
4644      apply fastforce
4645     apply (clarsimp simp: null_def neq_Nil_conv)
4646     apply (drule_tac f="\<lambda>xs. (a, bb) \<in> set xs" in arg_cong)
4647      apply (clarsimp simp: in_assocs_is_fun)
4648     apply (clarsimp simp: le_mask_asid_bits_helper)
4649    apply (simp add: is_aligned_shiftl_self)
4650   apply (intro conjI impI)
4651     apply clarsimp
4652     apply (drule obj_at_valid_objs', clarsimp)
4653     apply (clarsimp simp: projectKOs valid_obj'_def inv_ASIDPool
4654                    split: asidpool.split_asm)
4655    apply clarsimp
4656    apply (drule obj_at_valid_objs', clarsimp)
4657    apply (clarsimp simp: projectKOs valid_obj'_def inv_ASIDPool
4658                    split: asidpool.split_asm)
4659    apply (clarsimp simp: isCap_simps Let_def valid_cap'_def
4660                          maskCapRights_def[where ?x1.0="ArchObjectCap cp" for cp]
4661                          X64_H.maskCapRights_def
4662                   split: arch_capability.split_asm)
4663    apply (clarsimp simp: null_def neq_Nil_conv mask_def field_simps
4664                          asid_low_bits_word_bits asidInvalid_def
4665                   dest!: filter_eq_ConsD)
4666    apply (subst is_aligned_add_less_t2n[rotated], assumption+)
4667       apply (simp add: asid_low_bits_def asid_bits_def)
4668      apply (simp add: asid_wf_def)
4669     apply simp
4670    apply (auto simp: ct_in_state'_def valid_tcb_state'_def
4671               dest!: st_tcb_at_idle_thread'
4672               elim!: pred_tcb'_weakenE)[1]
4673  apply (clarsimp simp: cte_wp_at_ctes_of asidHighBits_handy_convs
4674                        word_sle_def word_sless_def asidLowBits_handy_convs
4675                        rf_sr_ksCurThread "StrictC'_thread_state_defs"
4676                        mask_def[where n=4]
4677                  cong: if_cong)
4678  apply (clarsimp simp: to_bool_def ccap_relation_isDeviceCap2 objBits_simps
4679                        archObjSize_def pageBits_def from_bool_def case_bool_If)
4680  apply (rule conjI)
4681   (* Is Asid Control Cap *)
4682   apply (clarsimp simp: neq_Nil_conv excaps_in_mem_def excaps_map_def)
4683   apply (frule interpret_excaps_eq[rule_format, where n=0], simp)
4684   apply (frule interpret_excaps_eq[rule_format, where n=1], simp)
4685   apply (clarsimp simp: mask_def[where n=4] slotcap_in_mem_def
4686                         ccap_rights_relation_def rightsFromWord_wordFromRights)
4687   apply (clarsimp simp: asid_high_bits_word_bits split: list.split_asm)
4688    apply (clarsimp simp: cap_untyped_cap_lift_def cap_lift_untyped_cap
4689                          cap_to_H_def[split_simps cap_CL.split]
4690                          hd_conv_nth length_ineq_not_Nil Kernel_C_defs
4691                   elim!: ccap_relationE)
4692   apply (clarsimp simp: to_bool_def unat_eq_of_nat
4693                         objBits_simps archObjSize_def pageBits_def from_bool_def case_bool_If
4694                  split: if_splits)
4695  apply (clarsimp simp: asid_low_bits_word_bits isCap_simps neq_Nil_conv
4696                        excaps_map_def excaps_in_mem_def
4697                        p2_gt_0[where 'a=machine_word_len, folded word_bits_def])
4698  apply (frule cap_get_tag_isCap_unfolded_H_cap(13))
4699  apply (frule ctes_of_valid', clarsimp)
4700  apply (simp only: diminished_valid'[symmetric])
4701  apply (frule interpret_excaps_eq[rule_format, where n=0], simp)
4702  apply (rule conjI)
4703   apply (clarsimp simp: cap_lift_asid_pool_cap cap_lift_page_directory_cap
4704                         cap_to_H_def to_bool_def valid_cap'_def
4705                         cap_page_directory_cap_lift_def
4706                         cap_asid_pool_cap_lift_def mask_def
4707                         asid_shiftr_low_bits_less[unfolded mask_def asid_bits_def] word_and_le1
4708                  elim!: ccap_relationE split: if_split_asm)
4709   apply (clarsimp split: list.split)
4710  apply (clarsimp simp: cap_lift_asid_pool_cap cap_lift_page_directory_cap
4711                        cap_to_H_def to_bool_def
4712                        cap_page_directory_cap_lift_def
4713                 elim!: ccap_relationE split: if_split_asm)
4714  apply (erule rf_sr_cte_at_validD[rotated])
4715  apply (fastforce simp: slotcap_in_mem_def2)
4716  done
4717
4718lemma ucast_ucast_first_port_is_ucast:
4719  "cap_lift y = Some (Cap_io_port_cap x) \<Longrightarrow> UCAST(16 \<rightarrow> 32) (UCAST(64 \<rightarrow> 16) (capIOPortFirstPort_CL x))
4720          = UCAST(64 \<rightarrow> 32) (capIOPortFirstPort_CL x)"
4721  apply (clarsimp simp: cap_lift_def Let_def cap_tag_defs mask_def split: if_split_asm)
4722  by word_bitwise
4723
4724lemma ucast_ucast_last_port_is_ucast:
4725  "cap_lift y = Some (Cap_io_port_cap x) \<Longrightarrow> UCAST(16 \<rightarrow> 32) (UCAST(64 \<rightarrow> 16) (capIOPortLastPort_CL x))
4726          = UCAST(64 \<rightarrow> 32) (capIOPortLastPort_CL x)"
4727  apply (clarsimp simp: cap_lift_def Let_def cap_tag_defs mask_def split: if_split_asm)
4728  by word_bitwise
4729
4730lemma ensurePortOperationAllowed_ccorres:
4731  "cap = IOPortCap f l \<Longrightarrow> ccorres (syscall_error_rel \<currency> dc)
4732            (liftxf errstate id (\<lambda>y. ()) ret__unsigned_long_')
4733            (\<top>)
4734            (UNIV \<inter> {s. ccap_relation (ArchObjectCap cap) (cap_' s)}
4735                  \<inter> {s. start_port_' s = ucast start}
4736                  \<inter> {s. size___unsigned_' s = of_nat magnitude} )
4737            hs
4738            (ensurePortOperationAllowed cap start magnitude)
4739            (Call ensurePortOperationAllowed_'proc)"
4740  supply Collect_const[simp del]
4741  apply (cinit lift: cap_' start_port_' size___unsigned_')
4742   apply (simp cong: StateSpace.state.fold_congs globals.fold_congs)
4743   apply csymbr
4744   apply csymbr
4745   apply csymbr
4746   apply csymbr
4747   apply csymbr
4748   apply (rule ccorres_assertE)
4749   apply (rule ccorres_assertE)
4750   apply (frule cap_get_tag_isCap_unfolded_H_cap)
4751   apply (erule ccap_relationE)
4752   apply (clarsimp simp: cap_to_H_def Let_def cap_io_port_cap_lift_def
4753                         ucast_ucast_first_port_is_ucast ucast_ucast_last_port_is_ucast whenE_def
4754                  split: if_split_asm cap_CL.split_asm)
4755   apply (rule ccorres_add_returnOk)
4756   apply (rule ccorres_inst[where P=\<top> and P'=UNIV], rule ccorres_guard_imp)
4757     apply (rule ccorres_Cond_rhs_Seq)
4758      apply clarsimp
4759      apply ccorres_rewrite
4760      apply (rule ccorres_from_vcg_throws[where P'=UNIV and P=\<top>])
4761      apply (rule allI, rule conseqPre, vcg)
4762      apply (clarsimp simp: throwError_def return_def syscall_error_rel_def syscall_error_to_H_cases
4763                            exception_defs)
4764     apply (clarsimp simp: whenE_def)
4765     apply (rule ccorres_return_CE)
4766       apply clarsimp
4767      apply clarsimp
4768     apply clarsimp
4769    apply clarsimp
4770   apply clarsimp
4771  apply clarsimp
4772  by (auto dest!: cap_get_tag_isCap_unfolded_H_cap)
4773
4774lemma setMessageInfo_ksCurThread_ccorres:
4775  "ccorres dc xfdc (tcb_at' thread and (\<lambda>s. ksCurThread s = thread))
4776           (UNIV \<inter> \<lbrace>mi = message_info_to_H mi'\<rbrace>) hs
4777           (setMessageInfo thread mi)
4778           (\<acute>ret__unsigned_long :== CALL wordFromMessageInfo(mi');;
4779            CALL setRegister(\<acute>ksCurThread,
4780                             scast Kernel_C.msgInfoRegister,
4781                             \<acute>ret__unsigned_long))"
4782  unfolding setMessageInfo_def
4783  apply (rule ccorres_guard_imp2)
4784   apply ctac
4785     apply simp
4786     apply (ctac add: setRegister_ccorres)
4787    apply wp
4788   apply vcg
4789  apply (simp add: X64_H.msgInfoRegister_def X64.msgInfoRegister_def
4790                   Kernel_C.msgInfoRegister_def Kernel_C.RSI_def rf_sr_ksCurThread)
4791  done
4792
4793lemma invokeX86PortIn8_ccorres:
4794  notes Collect_const[simp del] dc_simp[simp del]
4795  shows
4796  "ccorres ((intr_and_se_rel \<circ> Inr) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
4797       (valid_objs' and valid_queues and ct_in_state' ((=) Restart) and
4798        (\<lambda>s. ksCurThread s = thread \<and> sch_act_wf (ksSchedulerAction s) s))
4799       (UNIV \<inter> \<lbrace>\<acute>invLabel = scast Kernel_C.X86IOPortIn8\<rbrace>
4800             \<inter> \<lbrace>\<acute>port = port\<rbrace>
4801             \<inter> \<lbrace>\<acute>call = from_bool isCall\<rbrace>)
4802       hs
4803       (doE reply \<leftarrow> performX64PortInvocation (InvokeIOPort (IOPortInvocation port IOPortIn8));
4804           liftE (replyOnRestart thread reply isCall) odE)
4805       (Call invokeX86PortIn_'proc)"
4806  apply (clarsimp simp: performX64PortInvocation_def portIn_def liftE_bindE bind_assoc
4807                        Let_def replyOnRestart_def)
4808  apply (cinit' lift: invLabel_' port_' call_')
4809   apply (clarsimp simp: n_msgRegisters_def)
4810   apply ccorres_rewrite
4811   apply (rule ccorres_rhs_assoc)+
4812   apply (ctac add: in8_ccorres)
4813     apply csymbr
4814     apply (clarsimp simp: liftE_def when_def replyFromKernel_def if_to_top_of_bind bind_assoc setMRs_single)
4815     apply (rule ccorres_symb_exec_l[OF _ _ _ empty_fail_getThreadState])
4816       apply (rule ccorres_if_lhs[OF _ ccorres_False[where P'=UNIV]])
4817       apply (rule ccorres_if_lhs)
4818        apply ccorres_rewrite
4819        apply (rule ccorres_symb_exec_l[OF _ _ _ empty_fail_lookupIPCBuffer])
4820          apply (rule ccorres_rhs_assoc)+
4821          apply (ctac add: setRegister_ccorres)
4822            apply (rule ccorres_rhs_assoc)+
4823            apply (ctac add: setRegister_ccorres)
4824              apply csymbr
4825              apply csymbr
4826              apply (rule ccorres_rhs_assoc2)
4827              apply (ctac (no_vcg) add: setMessageInfo_ksCurThread_ccorres)
4828                apply (rule_tac P="\<lambda>s. ksCurThread s = thread" in ccorres_cross_over_guard)
4829                apply (ctac (no_vcg) add: setThreadState_ccorres)
4830                 apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
4831                 apply (rule allI, rule conseqPre, vcg)
4832                 apply (clarsimp simp: return_def dc_def)
4833                apply wpsimp+
4834              apply (clarsimp simp: ThreadState_Running_def mask_def rf_sr_ksCurThread)
4835             apply wpsimp
4836            apply (vcg exspec=setRegister_modifies)
4837           apply wpsimp
4838          apply (vcg exspec=setRegister_modifies)
4839         apply wpsimp
4840        apply wpsimp
4841       apply ccorres_rewrite
4842       apply (ctac (no_vcg) add: setThreadState_ccorres)
4843        apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
4844        apply (rule allI, rule conseqPre, vcg)
4845        apply (clarsimp simp: return_def dc_def)
4846       apply (wpsimp wp: gts_wp' hoare_vcg_all_lift hoare_drop_imps)
4847      apply wpsimp
4848     apply (wpsimp wp: gts_wp')
4849    apply (wpsimp wp: hoare_vcg_all_lift hoare_vcg_const_imp_lift hoare_vcg_imp_lift')+
4850   apply (vcg exspec=in8_modifies)
4851  by (clarsimp simp: ct_in_state'_def pred_tcb_at'_def obj_at'_def projectKOs
4852                     ThreadState_Running_def mask_def rf_sr_ksCurThread
4853                     X64_H.badgeRegister_def X64.badgeRegister_def "StrictC'_register_defs"
4854                     X64.capRegister_def msgRegisters_unfold message_info_to_H_def
4855                     msgRegisters_ccorres[where n=0, simplified n_msgRegisters_def,
4856                                           simplified, symmetric])
4857
4858lemma invokeX86PortIn16_ccorres:
4859  notes Collect_const[simp del] dc_simp[simp del]
4860  shows
4861  "ccorres ((intr_and_se_rel \<circ> Inr) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
4862       (valid_objs' and valid_queues and ct_in_state' ((=) Restart) and
4863        (\<lambda>s. ksCurThread s = thread \<and> sch_act_wf (ksSchedulerAction s) s))
4864       (UNIV \<inter> \<lbrace>\<acute>invLabel = scast Kernel_C.X86IOPortIn16\<rbrace>
4865             \<inter> \<lbrace>\<acute>port = port\<rbrace>
4866             \<inter> \<lbrace>\<acute>call = from_bool isCall\<rbrace>)
4867       hs
4868       (doE reply \<leftarrow> performX64PortInvocation (InvokeIOPort (IOPortInvocation port IOPortIn16));
4869           liftE (replyOnRestart thread reply isCall) odE)
4870       (Call invokeX86PortIn_'proc)"
4871  apply (clarsimp simp: performX64PortInvocation_def portIn_def liftE_bindE bind_assoc
4872                        Let_def replyOnRestart_def)
4873  apply (cinit' lift: invLabel_' port_' call_')
4874   apply (clarsimp simp: n_msgRegisters_def arch_invocation_label_defs)
4875   apply ccorres_rewrite
4876   apply (rule ccorres_rhs_assoc)+
4877   apply (ctac add: in16_ccorres)
4878     apply csymbr
4879     apply (clarsimp simp: liftE_def when_def replyFromKernel_def if_to_top_of_bind bind_assoc setMRs_single)
4880     apply (rule ccorres_symb_exec_l[OF _ _ _ empty_fail_getThreadState])
4881       apply (rule ccorres_if_lhs[OF _ ccorres_False[where P'=UNIV]])
4882       apply (rule ccorres_if_lhs)
4883        apply ccorres_rewrite
4884        apply (rule ccorres_symb_exec_l[OF _ _ _ empty_fail_lookupIPCBuffer])
4885          apply (rule ccorres_rhs_assoc)+
4886          apply (ctac add: setRegister_ccorres)
4887            apply (rule ccorres_rhs_assoc)+
4888            apply (ctac add: setRegister_ccorres)
4889              apply csymbr
4890              apply csymbr
4891              apply (rule ccorres_rhs_assoc2)
4892              apply (ctac (no_vcg) add: setMessageInfo_ksCurThread_ccorres)
4893                apply (rule_tac P="\<lambda>s. ksCurThread s = thread" in ccorres_cross_over_guard)
4894                apply (ctac (no_vcg) add: setThreadState_ccorres)
4895                 apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
4896                 apply (rule allI, rule conseqPre, vcg)
4897                 apply (clarsimp simp: return_def dc_def)
4898                apply wpsimp+
4899              apply (clarsimp simp: ThreadState_Running_def mask_def rf_sr_ksCurThread)
4900             apply wpsimp
4901            apply (vcg exspec=setRegister_modifies)
4902           apply wpsimp
4903          apply (vcg exspec=setRegister_modifies)
4904         apply wpsimp
4905        apply wpsimp
4906       apply ccorres_rewrite
4907       apply (ctac (no_vcg) add: setThreadState_ccorres)
4908        apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
4909        apply (rule allI, rule conseqPre, vcg)
4910        apply (clarsimp simp: return_def dc_def)
4911       apply (wpsimp wp: gts_wp' hoare_vcg_all_lift hoare_drop_imps)
4912      apply wpsimp
4913     apply (wpsimp wp: gts_wp')
4914    apply (wpsimp wp: hoare_vcg_all_lift hoare_vcg_const_imp_lift hoare_vcg_imp_lift')+
4915   apply (vcg exspec=in16_modifies)
4916  by (clarsimp simp: ct_in_state'_def pred_tcb_at'_def obj_at'_def projectKOs
4917                     ThreadState_Running_def mask_def rf_sr_ksCurThread
4918                     X64_H.badgeRegister_def X64.badgeRegister_def "StrictC'_register_defs"
4919                     X64.capRegister_def msgRegisters_unfold message_info_to_H_def
4920                     msgRegisters_ccorres[where n=0, simplified n_msgRegisters_def,
4921                                           simplified, symmetric])
4922
4923lemma invokeX86PortIn32_ccorres:
4924  notes Collect_const[simp del] dc_simp[simp del]
4925  shows
4926  "ccorres ((intr_and_se_rel \<circ> Inr) \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
4927       (valid_objs' and valid_queues and ct_in_state' ((=) Restart) and
4928        (\<lambda>s. ksCurThread s = thread \<and> sch_act_wf (ksSchedulerAction s) s))
4929       (UNIV \<inter> \<lbrace>\<acute>invLabel = scast Kernel_C.X86IOPortIn32\<rbrace>
4930             \<inter> \<lbrace>\<acute>port = port\<rbrace>
4931             \<inter> \<lbrace>\<acute>call = from_bool isCall\<rbrace>)
4932       hs
4933       (doE reply \<leftarrow> performX64PortInvocation (InvokeIOPort (IOPortInvocation port IOPortIn32));
4934           liftE (replyOnRestart thread reply isCall) odE)
4935       (Call invokeX86PortIn_'proc)"
4936  apply (clarsimp simp: performX64PortInvocation_def portIn_def liftE_bindE bind_assoc
4937                        Let_def replyOnRestart_def)
4938  apply (cinit' lift: invLabel_' port_' call_')
4939   apply (clarsimp simp: n_msgRegisters_def arch_invocation_label_defs)
4940   apply ccorres_rewrite
4941   apply (ctac add: in32_ccorres)
4942     apply (clarsimp simp: liftE_def when_def replyFromKernel_def if_to_top_of_bind bind_assoc setMRs_single)
4943     apply (rule ccorres_symb_exec_l[OF _ _ _ empty_fail_getThreadState])
4944       apply (rule ccorres_if_lhs[OF _ ccorres_False[where P'=UNIV]])
4945       apply (rule ccorres_if_lhs)
4946        apply ccorres_rewrite
4947        apply (rule ccorres_symb_exec_l[OF _ _ _ empty_fail_lookupIPCBuffer])
4948          apply (rule ccorres_rhs_assoc)+
4949          apply (ctac add: setRegister_ccorres)
4950            apply (rule ccorres_rhs_assoc)+
4951            apply (ctac add: setRegister_ccorres)
4952              apply csymbr
4953              apply csymbr
4954              apply (rule ccorres_rhs_assoc2)
4955              apply (ctac (no_vcg) add: setMessageInfo_ksCurThread_ccorres)
4956                apply (rule_tac P="\<lambda>s. ksCurThread s = thread" in ccorres_cross_over_guard)
4957                apply (ctac (no_vcg) add: setThreadState_ccorres)
4958                 apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
4959                 apply (rule allI, rule conseqPre, vcg)
4960                 apply (clarsimp simp: return_def dc_def)
4961                apply wpsimp+
4962              apply (clarsimp simp: ThreadState_Running_def mask_def rf_sr_ksCurThread)
4963             apply wpsimp
4964            apply (vcg exspec=setRegister_modifies)
4965           apply wpsimp
4966          apply (vcg exspec=setRegister_modifies)
4967         apply wpsimp
4968        apply wpsimp
4969       apply ccorres_rewrite
4970       apply (ctac (no_vcg) add: setThreadState_ccorres)
4971        apply (rule_tac P=\<top> and P'=UNIV in ccorres_from_vcg_throws)
4972        apply (rule allI, rule conseqPre, vcg)
4973        apply (clarsimp simp: return_def dc_def)
4974       apply (wpsimp wp: gts_wp' hoare_vcg_all_lift hoare_drop_imps)
4975      apply wpsimp
4976     apply (wpsimp wp: gts_wp')
4977    apply (wpsimp wp: hoare_vcg_all_lift hoare_vcg_const_imp_lift hoare_vcg_imp_lift')+
4978   apply (vcg exspec=in32_modifies)
4979  by (clarsimp simp: ct_in_state'_def pred_tcb_at'_def obj_at'_def projectKOs
4980                     ThreadState_Running_def mask_def rf_sr_ksCurThread
4981                     X64_H.badgeRegister_def X64.badgeRegister_def "StrictC'_register_defs"
4982                     X64.capRegister_def msgRegisters_unfold message_info_to_H_def
4983                     msgRegisters_ccorres[where n=0, simplified n_msgRegisters_def,
4984                                           simplified, symmetric])
4985
4986lemma invokeX86PortOut8_ccorres:
4987  notes Collect_const[simp del] dc_simp[simp del]
4988  shows
4989  "ccorres (cintr \<currency> (\<lambda>rv rv'. rv = [])) (liftxf errstate id (K ()) ret__unsigned_long_')
4990       invs'
4991       (UNIV \<inter> \<lbrace>\<acute>invLabel = scast Kernel_C.X86IOPortOut8\<rbrace>
4992             \<inter> \<lbrace>\<acute>port = port\<rbrace>
4993             \<inter> \<lbrace>\<acute>data___unsigned = ucast data\<rbrace>)
4994       hs
4995       (performX64PortInvocation (InvokeIOPort (IOPortInvocation port (IOPortOut8 data))))
4996       (Call invokeX86PortOut_'proc)"
4997  apply (cinit lift: invLabel_' port_' data___unsigned_' simp: portOut_def liftE_def return_returnOk)
4998   apply (clarsimp, ccorres_rewrite)
4999   apply (ctac (no_vcg) add: out8_ccorres)
5000    apply (rule ccorres_return_CE, simp+)[1]
5001   apply wp
5002  apply clarsimp
5003  done
5004
5005lemma invokeX86PortOut16_ccorres:
5006  notes Collect_const[simp del] dc_simp[simp del]
5007  shows
5008  "ccorres (cintr \<currency> (\<lambda>rv rv'. rv = [])) (liftxf errstate id (K ()) ret__unsigned_long_')
5009       invs'
5010       (UNIV \<inter> \<lbrace>\<acute>invLabel = scast Kernel_C.X86IOPortOut16\<rbrace>
5011             \<inter> \<lbrace>\<acute>port = port\<rbrace>
5012             \<inter> \<lbrace>\<acute>data___unsigned = ucast data\<rbrace>)
5013       hs
5014       (performX64PortInvocation (InvokeIOPort (IOPortInvocation port (IOPortOut16 data))))
5015       (Call invokeX86PortOut_'proc)"
5016  apply (cinit lift: invLabel_' port_' data___unsigned_' simp: portOut_def liftE_def return_returnOk)
5017   apply (clarsimp simp:arch_invocation_label_defs, ccorres_rewrite)
5018   apply (ctac (no_vcg) add: out16_ccorres)
5019    apply (rule ccorres_return_CE, simp+)[1]
5020   apply wp
5021  apply clarsimp
5022  done
5023
5024lemma invokeX86PortOut32_ccorres:
5025  notes Collect_const[simp del] dc_simp[simp del]
5026  shows
5027  "ccorres (cintr \<currency> (\<lambda>rv rv'. rv = [])) (liftxf errstate id (K ()) ret__unsigned_long_')
5028       invs'
5029       (UNIV \<inter> \<lbrace>\<acute>invLabel = scast Kernel_C.X86IOPortOut32\<rbrace>
5030             \<inter> \<lbrace>\<acute>port = port\<rbrace>
5031             \<inter> \<lbrace>\<acute>data___unsigned = ucast data\<rbrace>)
5032       hs
5033       (performX64PortInvocation (InvokeIOPort (IOPortInvocation port (IOPortOut32 data))))
5034       (Call invokeX86PortOut_'proc)"
5035  apply (cinit lift: invLabel_' port_' data___unsigned_' simp: portOut_def liftE_def return_returnOk)
5036   apply (clarsimp simp:arch_invocation_label_defs, ccorres_rewrite)
5037   apply (ctac (no_vcg) add: out32_ccorres)
5038    apply (rule ccorres_return_CE, simp+)[1]
5039   apply wp
5040  apply clarsimp
5041  done
5042
5043lemma setIOPortMask_valid_mdb'[wp]:
5044  "\<lbrace>valid_mdb'\<rbrace> setIOPortMask f l b \<lbrace>\<lambda>_. valid_mdb'\<rbrace>"
5045  by (wpsimp simp: setIOPortMask_def valid_mdb'_def)
5046
5047lemma invokeX86PortControl_ccorres:
5048  notes Collect_const[simp del]
5049  shows
5050  "ccorres (cintr \<currency> (\<lambda>rv rv'. rv = [])) (liftxf errstate id (K ()) ret__unsigned_long_')
5051       (invs' and cte_at' src and K (f \<le> l))
5052       (UNIV \<inter> \<lbrace>\<acute>first_port = f\<rbrace> \<inter> \<lbrace>\<acute>last_port = l\<rbrace>
5053             \<inter> \<lbrace>\<acute>ioportSlot = cte_Ptr dest\<rbrace>
5054             \<inter> \<lbrace>\<acute>controlSlot = cte_Ptr src\<rbrace>)
5055       hs
5056       (performX64PortInvocation (InvokeIOPortControl (IOPortControlIssue f l dest src)))
5057       (Call invokeX86PortControl_'proc)"
5058  apply (clarsimp simp: isCap_simps performX64PortInvocation_def Let_def)
5059  apply (cinit' lift: first_port_' last_port_' ioportSlot_' controlSlot_')
5060   apply (clarsimp cong: StateSpace.state.fold_congs globals.fold_congs)
5061   apply (rule ccorres_Guard_Seq)
5062   apply (clarsimp simp: liftE_def bind_assoc return_returnOk)
5063   apply (ctac add: setIOPortMask_ccorres)
5064     apply csymbr
5065     apply (ctac(no_vcg) add: cteInsert_ccorres)
5066      apply (rule ccorres_return_CE, simp+)[1]
5067     apply wp
5068    apply wp
5069   apply (vcg exspec=setIOPortMask_modifies)
5070  apply (clarsimp simp: is_simple_cap'_def isCap_simps invs_mdb' invs_valid_objs' invs_pspace_aligned'
5071                        invs_pspace_canonical' valid_cap_simps' capAligned_def word_bits_def)
5072  apply (rule conjI)
5073   apply (clarsimp simp: cap_io_port_cap_lift ccap_relation_def cap_to_H_def mask_def)
5074   apply word_bitwise
5075  apply (clarsimp simp: rf_sr_def cstate_relation_def carch_state_relation_def
5076                        global_ioport_bitmap_relation_def Let_def typ_heap_simps)
5077  done
5078
5079lemma unat_length_4_helper:
5080  "\<lbrakk>unat (l::machine_word) = length args; \<not> l < 4\<rbrakk> \<Longrightarrow> \<exists>x xa xb xc xs. args = x#xa#xb#xc#xs"
5081  apply (case_tac args; clarsimp simp: unat_eq_0)
5082  by (rename_tac list, case_tac list, clarsimp, unat_arith)+
5083
5084lemma ucast_drop_big_mask:
5085  "UCAST(64 \<rightarrow> 16) (x && 0xFFFF) = UCAST(64 \<rightarrow> 16) x"
5086  by word_bitwise
5087
5088lemma first_port_last_port_compare:
5089  "UCAST(16 \<rightarrow> 32 signed) (UCAST(64 \<rightarrow> 16) (xa && 0xFFFF))
5090        <s UCAST(16 \<rightarrow> 32 signed) (UCAST(64 \<rightarrow> 16) (x && 0xFFFF))
5091       = (UCAST(64 \<rightarrow> 16) xa < UCAST(64 \<rightarrow> 16) x)"
5092  apply (clarsimp simp: word_sless_alt ucast_drop_big_mask)
5093  apply (subst sint_ucast_eq_uint, clarsimp simp: is_down)+
5094  by (simp add: word_less_alt)
5095
5096(* FIXME X64: move to state rel? *)
5097abbreviation
5098  "x86KSAllocatedIOPorts_ptr == ioport_table_Ptr (symbol_table ''x86KSAllocatedIOPorts'')"
5099
5100(* FIXME X64: use earlier or replace with earlier def? *)
5101definition
5102  port_mask :: "16 word \<Rightarrow> 16 word \<Rightarrow> machine_word"
5103where
5104  "port_mask start end =
5105     mask (unat (end && mask wordRadix)) && ~~ mask (unat (start && mask wordRadix))"
5106
5107lemma shiftr_le_mask:
5108  fixes w :: "'a::len word"
5109  shows "w >> n \<le> mask (LENGTH('a) - n)"
5110  by (metis and_mask_eq_iff_shiftr_0 le_mask_iff shiftr_mask_eq word_size)
5111
5112lemma word_minus_1_shiftr:
5113  notes word_unat.Rep_inject[simp del]
5114  fixes w :: "'a::len word"
5115  assumes low_bits_zero: "w && mask n = 0"
5116  assumes neq_zero: "w \<noteq> 0"
5117  shows "(w - 1) >> n = (w >> n) - 1"
5118  using neq_zero low_bits_zero
5119  apply (subgoal_tac "n < LENGTH('a)")
5120   prefer 2
5121   apply (metis not_less ucast_id ucast_mask_drop)
5122  apply (rule_tac t="w - 1" and s="(w && ~~ mask n) - 1" in subst,
5123         fastforce simp: low_bits_zero mask_eq_x_eq_0)
5124  apply (clarsimp simp: mask_eq_0_eq_x neg_mask_is_div lt1_neq0[symmetric])
5125  apply (subst shiftr_div_2n_w, fastforce simp: word_size)+
5126  apply (rule word_uint.Rep_eqD)
5127  apply (simp only: uint_word_ariths uint_div uint_power_lower)
5128  apply (subst mod_pos_pos_trivial, fastforce, fastforce)+
5129  apply (subst mod_pos_pos_trivial)
5130    apply (fastforce simp: le_diff_eq uint_2p_alt word_le_def)
5131   apply (subst uint_1[symmetric])
5132   apply (fastforce intro: uint_sub_lt2p)
5133  apply (subst int_div_sub_1, fastforce)
5134  apply (clarsimp simp: and_mask_dvd low_bits_zero)
5135  apply (subst mod_pos_pos_trivial)
5136    apply (metis linorder_not_less mult_zero_left shiftr_div_2n shiftr_div_2n_w uint_eq_0
5137                 uint_le_0_iff word_less_1 word_uint.Rep_inject word_size zle_diff1_eq)
5138   apply (metis shiftr_div_2n uint_1 uint_sub_lt2p)
5139  apply fastforce
5140  done
5141
5142lemma ucast_shiftr:
5143  "UCAST('a::len \<rightarrow> 'b::len) w >> n = UCAST('a \<rightarrow> 'b) ((w && mask LENGTH('b)) >> n)"
5144  apply (rule word_eqI[rule_format]; rule iffI; clarsimp simp: nth_ucast nth_shiftr word_size)
5145  apply (drule test_bit_size; simp add: word_size)
5146  done
5147
5148lemma mask_eq_ucast_shiftr:
5149  assumes mask: "w && mask LENGTH('b) = w"
5150  shows "UCAST('a::len \<rightarrow> 'b::len) w >> n = UCAST('a \<rightarrow> 'b) (w >> n)"
5151  by (rule ucast_shiftr[where 'a='a and 'b='b, of w n, simplified mask])
5152
5153lemma mask_eq_ucast_shiftl:
5154  assumes "w && mask (LENGTH('a) - n) = w"
5155  shows "UCAST('a::len \<rightarrow> 'b::len) w << n = UCAST('a \<rightarrow> 'b) (w << n)"
5156  apply (rule subst[where P="\<lambda>c. ucast c << n = ucast (c << n)", OF assms])
5157  apply (rule word_eqI[rule_format]; rule iffI;
5158         clarsimp simp: nth_ucast nth_shiftl word_size;
5159         drule test_bit_size; simp add: word_size)
5160  done
5161
5162lemma mask_le_mono:
5163  "m \<le> n \<Longrightarrow> mask m \<le> mask n"
5164  apply (subst and_mask_eq_iff_le_mask[symmetric])
5165  by (auto intro: word_eqI simp: word_size)
5166
5167lemma word_and_mask_eq_le_mono:
5168  "w && mask m = w \<Longrightarrow> m \<le> n \<Longrightarrow> w && mask n = w"
5169  apply (simp add: and_mask_eq_iff_le_mask)
5170  by (erule order.trans, erule mask_le_mono)
5171
5172lemma first_last_highbits_eq_port_set:
5173  fixes f l :: "16 word"
5174  fixes arr :: "machine_word[1024]"
5175  shows "\<lbrakk>unat f \<le> unat l; unat (f && 0x3F) \<le> unat (l && 0x3F);
5176          unat (f >> 6) = unat (l >> 6);
5177          0 < unat (arr.[unat (l >> 6)] && mask (Suc (unat (l && 0x3F)))
5178                  && ~~ mask (unat (f && 0x3F)))\<rbrakk>
5179         \<Longrightarrow> \<exists>port::16 word.
5180                unat f \<le> unat port \<and> unat port \<le> unat l
5181              \<and> arr.[unat (port >> 6)] !! unat (port && 0x3F)"
5182    apply (frule word_exists_nth[OF word_neq_0_conv[THEN iffD2], OF unat_less_impl_less, simplified],
5183                clarsimp simp: word_size)
5184  apply (rule_tac x="(l && ~~ mask 6) + of_nat i" in exI)
5185  apply (clarsimp simp: neg_mask_bang mask_def[where n=6, simplified, symmetric])
5186  apply (frule test_bit_size, clarsimp simp: word_size less_Suc_eq_le)
5187  apply (frule_tac n=i in word_of_nat_less[where x=64 and 'a=16, simplified])
5188  apply (rule conjI)
5189   subgoal for i
5190     (* f \<le> port *)
5191     apply (frule_tac n=i in word_of_nat_less[where x=64 and 'a=16, simplified])
5192     apply (subst AND_NOT_mask_plus_AND_mask_eq[symmetric, where n=6])
5193     apply (subst unat_plus_simple[THEN iffD1])
5194      apply (clarsimp simp: AND_NOT_mask_plus_AND_mask_eq word_and_le2)
5195     apply (subst unat_plus_simple[THEN iffD1])
5196      apply (rule word_random[where x'="2^6 - 1"], rule is_aligned_no_overflow',
5197                     simp add: is_aligned_neg_mask2)
5198      apply (rule word_less_sub_1, simp)
5199     by (simp add: unat_of_nat word_shiftr_eq_mask)
5200  apply (rule conjI)
5201   subgoal for i
5202     (* port \<le> l *)
5203     apply (subgoal_tac "(((l && ~~ mask 6) + of_nat i) >> 6) = l >> 6")
5204      apply (frule_tac w1="f" in word_le_split_mask[where n=6, THEN iffD1, OF word_le_nat_alt[THEN iffD2]])
5205      apply (erule disjE)
5206       apply clarsimp
5207      apply (rule iffD1[OF word_le_nat_alt])
5208      apply (rule word_le_split_mask[where n=6, THEN iffD2])
5209      apply clarsimp
5210      apply (subst mask_eqs(1)[symmetric])
5211      apply (simp add: mask_AND_NOT_mask)
5212      apply (subst and_mask_eq_iff_le_mask[THEN iffD2])
5213       apply (clarsimp simp add: mask_def dest!: word_less_sub_1)
5214      apply (clarsimp simp: word_le_nat_alt unat_of_nat)
5215     by (simp add: aligned_shift')
5216  subgoal for i
5217    (* arr set *)
5218    apply (simp add: aligned_shift')
5219    apply (subst mask_eqs(1)[symmetric])
5220    apply (simp add: mask_AND_NOT_mask)
5221    apply (subst and_mask_eq_iff_le_mask[THEN iffD2])
5222     apply (clarsimp simp add: mask_def dest!: word_less_sub_1)
5223    by (simp add: unat_of_nat)
5224  done
5225
5226lemma port_set_in_first_word:
5227  fixes f l :: "16 word"
5228  fixes arr :: "machine_word[1024]"
5229  shows "\<lbrakk>unat f \<le> unat l; unat (f >> 6) < unat (l >> 6);
5230           0 < unat (arr.[unat (f >> 6)] && ~~ mask (unat (f && mask 6)))\<rbrakk>
5231       \<Longrightarrow> \<exists>port::16 word. unat f \<le> unat port \<and> unat port \<le> unat l \<and>
5232              arr.[unat (port >> 6)] !! unat (port && mask 6)"
5233  apply (frule word_exists_nth[OF word_neq_0_conv[THEN iffD2], OF unat_less_impl_less, simplified],
5234                clarsimp simp: word_size)
5235  apply (rule_tac x="(f && ~~ mask 6) + of_nat i" in exI)
5236  apply (clarsimp simp: neg_mask_bang mask_def[where n=6, simplified, symmetric])
5237  apply (frule test_bit_size, clarsimp simp: word_size)
5238  apply (frule_tac n=i in word_of_nat_less[where x=64 and 'a=16, simplified])
5239  apply (rule conjI)
5240   subgoal for i
5241     (* f \<le> port *)
5242     apply (frule_tac n=i in word_of_nat_less[where x=64 and 'a=16, simplified])
5243     apply (subst AND_NOT_mask_plus_AND_mask_eq[symmetric, where n=6])
5244     apply (subst unat_plus_simple[THEN iffD1])
5245      apply (clarsimp simp: AND_NOT_mask_plus_AND_mask_eq word_and_le2)
5246     apply (subst unat_plus_simple[THEN iffD1])
5247      apply (rule word_random[where x'="2^6 - 1"], rule is_aligned_no_overflow',
5248                     simp add: is_aligned_neg_mask2)
5249      apply (rule word_less_sub_1, simp)
5250     by (simp add: unat_of_nat)
5251  apply (rule conjI)
5252   subgoal for i
5253     (* port \<le> l *)
5254     apply (subgoal_tac "(((f && ~~ mask 6) + of_nat i) >> 6) = f >> 6")
5255      apply (frule_tac w1="f" in word_le_split_mask[where n=6, THEN iffD1, OF word_le_nat_alt[THEN iffD2]])
5256      apply (erule disjE)
5257       apply (rule iffD1[OF word_le_nat_alt])
5258       apply (rule word_le_split_mask[where n=6, THEN iffD2])
5259       apply clarsimp
5260      apply clarsimp
5261     by (simp add: aligned_shift')
5262  subgoal for i
5263    (* arr set *)
5264    apply (simp add: aligned_shift')
5265    apply (subst mask_eqs(1)[symmetric])
5266    apply (simp add: mask_AND_NOT_mask)
5267    apply (subst and_mask_eq_iff_le_mask[THEN iffD2])
5268     apply (clarsimp simp add: mask_def dest!: word_less_sub_1)
5269    by (simp add: unat_of_nat)
5270  done
5271
5272lemma word_not_exists_nth:
5273  "(w::'a::len word) = 0 \<Longrightarrow> \<forall>i<LENGTH('a). \<not> w !! i"
5274  by (clarsimp simp: nth_0)
5275
5276lemma bitmap_word_zero_no_bits_set1:
5277  fixes f l :: "16 word"
5278  fixes arr :: "machine_word[1024]"
5279  shows "\<lbrakk>unat (f && mask 6) \<le> unat (l && mask 6);
5280        unat (f >> 6) = unat (l >> 6);
5281        arr.[unat (l >> 6)] && mask (Suc (unat (l && mask 6))) &&
5282              ~~ mask (unat (f && mask 6)) = 0\<rbrakk>
5283   \<Longrightarrow> \<forall>port::16 word.
5284        unat f \<le> unat port \<and> unat port \<le> unat l \<longrightarrow>
5285        \<not>arr.[unat (port >> 6)] !! unat (port && mask 6)"
5286  apply clarsimp
5287  apply (drule word_not_exists_nth)
5288  apply (simp only: all_nat_less_eq)
5289  apply (cut_tac w=port in unat_and_mask_less_2p[of 6, simplified mask_def, simplified]; simp)
5290  apply (drule_tac x="unat (port && 0x3F)" in bspec, clarsimp)
5291  apply (frule_tac v1=port in word_le_split_mask[where n=6, THEN iffD1, OF word_le_nat_alt[THEN iffD2]])
5292  apply (frule_tac w1=port in word_le_split_mask[where n=6, THEN iffD1, OF word_le_nat_alt[THEN iffD2]])
5293  apply (subgoal_tac "port >> 6 = l >> 6")
5294   apply (clarsimp simp: word_size neg_mask_bang not_less not_le)
5295   apply (clarsimp simp: word_unat.Rep_inject word_le_nat_alt mask_def)
5296  apply (erule disjE; clarsimp simp:word_less_nat_alt)
5297  done
5298
5299lemma ucast_shiftl_6_absorb:
5300  fixes f l :: "16 word"
5301  assumes "f \<le> l"
5302  assumes "f >> 6 < l >> 6"
5303  shows "UCAST(16\<rightarrow>32 signed) ((f >> 6) + 1) << 6 = UCAST(16 \<rightarrow> 32 signed) (((f >> 6) + 1) << 6)"
5304  using assms
5305  by (word_bitwise, auto)
5306
5307(* FIXME: move *)
5308lemma word_highbits_bounded_highbits_eq:
5309  "\<lbrakk>x \<le> (y :: 'a::len word); y < (x >> n) + 1 << n\<rbrakk> \<Longrightarrow> x >> n  = y >> n"
5310  apply (cases "n < LENGTH('a)")
5311   prefer 2
5312   apply (simp add: shiftr_eq_0)
5313  apply (drule_tac n=n in le_shiftr)
5314  apply (subst (asm) word_shiftl_add_distrib)
5315  apply (drule_tac word_less_sub_1)
5316  apply (simp only: add_diff_eq[symmetric] mask_def[symmetric] and_not_mask[symmetric])
5317  apply (drule_tac u=y and n=n in le_shiftr)
5318  apply (subgoal_tac "(x && ~~ mask n) + mask n >> n \<le> x >> n")
5319   apply fastforce
5320  apply (subst aligned_shift')
5321     apply (fastforce simp: mask_lt_2pn)
5322    apply (fastforce simp: is_aligned_neg_mask)
5323   apply fastforce
5324  apply (fastforce simp: mask_shift)
5325  done
5326
5327lemma bitmap_word_zero_no_bits_set2:
5328  fixes f l :: "16 word"
5329  fixes arr :: "machine_word[1024]"
5330  shows "\<lbrakk>unat f \<le> unat l; unat (f >> 6) < unat (l >> 6);
5331        arr.[unat (f >> 6)] && ~~ mask (unat (f && mask 6)) = 0\<rbrakk>
5332            \<Longrightarrow> \<forall>port::16 word. unat f \<le> unat port \<and>
5333                   unat port < unat (UCAST(16 \<rightarrow> 32 signed) ((f >> 6) + 1) << 6) \<longrightarrow>
5334        \<not>arr.[unat (port >> 6)] !! unat (port && mask 6)"
5335  apply (clarsimp simp: word_le_nat_alt[symmetric] word_less_nat_alt[symmetric] ucast_shiftl_6_absorb)
5336  apply (frule (1) word_highbits_bounded_highbits_eq)
5337  apply simp
5338  apply (frule_tac v1=port in word_le_split_mask[where n=6, THEN iffD1])
5339  apply (erule disjE; clarsimp)
5340  apply (drule word_not_exists_nth)
5341  apply (cut_tac w=port in unat_and_mask_less_2p[of 6]; simp)
5342  apply (drule_tac x="unat (port && mask 6)" in spec, clarsimp simp: neg_mask_bang not_le word_le_nat_alt)
5343  done
5344
5345lemma word_eq_cast_unsigned:
5346  "(x = y) = (UCAST ('a signed \<rightarrow> ('a :: len)) x = ucast y)"
5347  by (simp add: word_eq_iff nth_ucast)
5348
5349lemma isIOPortRangeFree_spec:
5350  notes ucast_mask = ucast_and_mask[where n=6, simplified mask_def, simplified]
5351  notes not_max_word_simps = and_not_max_word shiftr_not_max_word and_mask_not_max_word
5352  notes ucast_cmp_ucast = ucast_le_ucast ucast_less_ucast
5353  notes array_assert = array_assertion_shrink_right[OF array_ptr_valid_array_assertionD]
5354  notes unat_arith_simps' = unat_arith_simps[where 'a=16] unat_arith_simps[where 'a="32 signed"]
5355  notes word_unat.Rep_inject[simp del] int_unat[simp del]
5356  defines "port_array s \<equiv> h_val (hrs_mem (t_hrs_' (globals s))) x86KSAllocatedIOPorts_ptr"
5357  shows
5358  "\<forall>\<sigma>. \<Gamma> \<turnstile>
5359    {s. s = \<sigma> \<and> first_port_' s \<le> last_port_' s \<and> s \<Turnstile>\<^sub>c x86KSAllocatedIOPorts_ptr }
5360      Call isIOPortRangeFree_'proc
5361    {t. globals t = globals \<sigma> \<and>
5362        ret__unsigned_long_' t = from_bool
5363          (\<forall>port. first_port_' \<sigma> \<le> port \<and> port \<le> last_port_' \<sigma>
5364                   \<longrightarrow> \<not> port_array \<sigma>.[unat (port >> wordRadix)] !! unat (port && mask wordRadix))}"
5365  apply (rule allI)
5366  subgoal for \<sigma>
5367  apply (hoare_rule HoarePartial.ProcNoRec1)
5368  apply (simp add: scast_ucast_up_eq_ucast word_upcast_shiftr[symmetric] ucast_mask[symmetric]
5369                   word_upcast_0_sle)
5370  apply (subst whileAnno_subst_invariant
5371                 [where I="{s. globals s = globals \<sigma>
5372                             \<and> first_port_' \<sigma> \<le> last_port_' \<sigma>
5373                             \<and> \<sigma> \<Turnstile>\<^sub>c x86KSAllocatedIOPorts_ptr
5374                             \<and> ucast (first_port_' \<sigma> >> wordRadix) < low_word_' s
5375                             \<and> low_word_' s <= high_word_' s
5376                             \<and> high_word_' s = ucast (last_port_' \<sigma> >> wordRadix)
5377                             \<and> high_index_' s = ucast (last_port_' \<sigma> && mask wordRadix)
5378                             \<and> (\<forall>port. first_port_' \<sigma> \<le> port \<and> ucast port < low_word_' s << wordRadix
5379                                        \<longrightarrow> \<not> port_array \<sigma>.[unat (port >> wordRadix)]
5380                                                !! unat (port && mask wordRadix))}"])
5381  apply (simp add: port_array_def)
5382  apply (rule conseqPre, vcg)
5383    apply (all \<open>clarsimp simp: hrs_simps false_def from_bool_0 wordRadix_def is_up is_down
5384                               unat_ucast_upcast uint_up_ucast sint_ucast_eq_uint up_ucast_inj_eq
5385                               not_max_word_simps[THEN ucast_increment, simplified max_word_def]
5386                               ucast_cmp_ucast ucast_cmp_ucast[where 'a=16 and y="0x40", simplified]\<close>)
5387    subgoal for mem htd first_port last_port low_word
5388      (* Loop invariant is preserved. *)
5389      apply (frule neg_msb_le_mono[OF _ word_upcast_neg_msb], simp)
5390      apply (simp add: word_sless_iff_less word_sle_iff_le word_upcast_neg_msb
5391                       Word_Lemmas.sint_eq_uint uint_nat)
5392      apply (frule less_le_trans[OF _ ucast_le_ucast[THEN iffD2],
5393                                 OF _ _ shiftr_le_mask[unfolded mask_def]]; simp)
5394      apply (intro conjI impI allI; (simp add: unat_arith_simps; fail)?)
5395       apply (drule word_exists_nth; clarsimp)
5396       subgoal for i
5397         (* Early return false. *)
5398         apply (rule_tac x="ucast ((low_word << 6) || of_nat i)" in exI)
5399         apply (match premises in L: \<open>_ < low_word\<close> and U: \<open>low_word < _\<close> (multi) and V: \<open>test_bit _ _\<close>
5400                  \<Rightarrow> \<open>match premises in _[thin]: _ (multi) \<Rightarrow> \<open>insert L U V\<close>\<close>)
5401         apply (frule test_bit_size; simp add: word_size)
5402         apply (rule revcut_rl[OF shiftl_shiftr_id[of 6 low_word]], simp, fastforce elim: less_trans)
5403         apply (rule revcut_rl[OF and_mask_eq_iff_le_mask[of low_word 10, THEN iffD2]],
5404                fastforce simp: mask_def elim: order.trans[where b="0x3FF", OF less_imp_le])
5405         apply (rule revcut_rl[OF le_mask_iff[of "of_nat i :: 32 signed word" 6, THEN iffD1]],
5406                fastforce intro: word_of_nat_le simp: mask_def)
5407         apply (frule and_mask_eq_iff_shiftr_0[where w="of_nat i", THEN iffD2])
5408         apply (subst mask_eq_ucast_shiftr)
5409          apply (simp add: word_ao_dist)
5410          apply (rule arg_cong2[where f=bitOR]; rule and_mask_eq_iff_le_mask[THEN iffD2])
5411           apply (rule le_mask_shiftl_le_mask, simp)
5412           apply (erule order_trans[where y="0x3FF", OF less_imp_le], simp add: mask_def)
5413          apply (erule order_trans[OF less_imp_le[OF of_nat_mono_maybe], rotated]; simp add: mask_def)
5414         apply (simp add: unat_ucast_eq_unat_and_mask ucast_and_mask[symmetric]
5415                          shiftr_over_or_dist word_ao_dist unat_of_nat_eq word_and_mask_eq_le_mono)
5416         apply (thin_tac "test_bit _ _")
5417         apply (rule conjI;
5418                rule word_le_split_mask[where n=6, THEN iffD2, OF disjI1];
5419                rule ucast_less_ucast[where 'b="32 signed", THEN iffD1])
5420         apply (simp_all add: ucast_shiftr ucast_ucast_mask word_ao_dist
5421                              word_and_mask_eq_le_mono[of "of_nat i"]
5422                              word_and_mask_eq_le_mono[of low_word]
5423                              shiftr_over_and_dist shiftr_over_or_dist
5424                              shiftr_mask2)
5425         done
5426      subgoal for port
5427        (* Continue to next loop iteration. *)
5428        apply (case_tac "ucast port < low_word << 6"; (simp add: unat_arith_simps; fail)?)
5429        apply (clarsimp simp: not_less)
5430        apply (subgoal_tac "low_word = ucast (port >> 6)", simp add: unat_arith_simps')
5431        apply (match premises in L: \<open>low_word << 6 \<le> ucast port\<close> and
5432                                 U: \<open>ucast port < low_word + 1 << 6\<close> and
5433                                 T: \<open>low_word < 0x3FF\<close>for port
5434                 \<Rightarrow> \<open>match premises in _[thin]: _ (multi) \<Rightarrow> \<open>insert T L U\<close>\<close>)
5435        apply (drule le_shiftr[where n=6 and v="ucast port"],
5436               drule le_shiftr[where n=6 and u="ucast port", OF word_less_sub_1])
5437        apply (simp add: word_minus_1_shiftr[OF shiftl_mask_is_0,
5438                                             OF word_shift_nonzero[where m=10,
5439                                                                   OF inc_le _ less_is_non_zero_p1,
5440                                                                   OF less_trans]]
5441                         shiftl_shiftr_id[OF _ less_trans]
5442                         shiftl_shiftr_id[OF _ le_less_trans[OF inc_le]]
5443                         word_upcast_shiftr)
5444        done
5445      done
5446   subgoal for mem htd first_port last_port low_word
5447     (* Invariant plus loop termination condition is sufficient to establish VCG postcondition. *)
5448     apply (frule neg_msb_le_mono[OF _ word_upcast_neg_msb], simp)
5449     apply (simp add: word_sless_iff_less word_sle_iff_le word_upcast_neg_msb
5450                      Word_Lemmas.sint_eq_uint uint_nat)
5451     apply (cut_tac word_and_mask_le_2pm1[of last_port 6], simp)
5452     apply (cut_tac shiftr_le_mask[of last_port 6, simplified mask_def], simp)
5453     apply (intro conjI allI impI; (simp add: unat_arith_simps; fail)?)
5454     apply (drule word_exists_nth; clarsimp simp: word_size ucast_less_ucast)
5455      subgoal for i
5456        (* return false. *)
5457        apply (rule exI[of _ "last_port && ~~ mask 6 || of_nat i"])
5458        apply (rule revcut_rl[OF le_mask_iff[of "of_nat i :: 16 word" 6, THEN iffD1]],
5459               fastforce intro: word_of_nat_le simp: mask_def)
5460        apply (frule and_mask_eq_iff_shiftr_0[where w="of_nat i", THEN iffD2])
5461        apply (simp add: shiftr_over_or_dist word_ao_dist mask_AND_NOT_mask unat_of_nat_eq)
5462        apply (rule conjI; rule word_le_split_mask[where n=6, THEN iffD2];
5463               simp add: ucast_less_ucast shiftr_over_or_dist
5464                         word_ao_dist mask_AND_NOT_mask)
5465        apply (rule word_of_nat_le, rule nat_Suc_less_le_imp)
5466        apply (erule rsubst[where P="\<lambda>c. i < c"])
5467        apply (simp add: unat_arith_simps)
5468        done
5469     subgoal for port
5470       (* return true. *)
5471       apply (case_tac "ucast port < UCAST (16 \<rightarrow> 32 signed) (last_port >> 6) << 6")
5472        apply (fastforce simp: unat_arith_simps)
5473       apply (clarsimp simp: not_less ucast_le_ucast)
5474       apply (subgoal_tac "port >> 6 = last_port >> 6")
5475        prefer 2
5476        apply (drule ucast_mono_le[where 'a="32 signed" and 'b="16"])
5477         apply (rule order_less_le_trans, rule ucast_less, simp+)
5478        apply (drule_tac u="ucast _" and v=port and n=6 in le_shiftr)
5479        apply (clarsimp simp: ucast_shiftr shiftl_shiftr1)
5480        apply (simp add: shiftl_shiftr1 word_size shiftr_then_mask_commute[where n=6 and m=10, simplified, symmetric]
5481                         mask_twice ucast_and_mask)
5482        apply (drule_tac u=port and n=6 in le_shiftr)+
5483        apply (clarsimp simp add: shiftr_mask_eq' word_size)
5484       apply (clarsimp simp: word_not_exists_nth)
5485       apply (drule_tac x="unat (port && mask 6)" and v=0 in word_eqD)
5486       apply (simp add: word_size unat_and_mask_less_2p[where m=6, simplified])
5487       apply (simp add: word_less_nat_alt[symmetric])
5488       apply (erule notE, rule plus_one_helper2)
5489        apply (subst (asm) word_le_split_mask[where n=6 and v=last_port])+
5490        apply simp
5491       apply (simp add: unat_arith_simps)
5492       done
5493     done
5494  subgoal
5495    (* VCG precondition is sufficient to establish loop invariant. *)
5496    apply (frule word_le_split_mask[where n=6, THEN iffD1])
5497    apply (simp add: unat_arith_simps)
5498    apply (cut_tac unat_shiftr_less_2p[of 6 10 "first_port_' \<sigma>"]; simp)
5499    apply (cut_tac unat_and_mask_less_2p[of 6 "first_port_' \<sigma>"]; simp)
5500    apply (cut_tac unat_and_mask_less_2p[of 6 "last_port_' \<sigma>"]; simp)
5501    apply (simp add: uint_nat mask_def[where n=6] mask_def[where n=64] less_Suc_eq_le Suc_le_eq)
5502    apply (clarsimp intro!: word_less_imp_sless
5503                      simp: unat_ucast_no_overflow_le word_upcast_neg_msb msb_nth zero_sle_ucast_up
5504                            is_down unat_eq_0
5505          | rule conjI
5506          | erule (2) first_last_highbits_eq_port_set[simplified mask_def[where n=6], simplified]
5507          | erule (2) port_set_in_first_word[simplified mask_def[where n=6], simplified]
5508          | solves \<open>drule (2) bitmap_word_zero_no_bits_set1[simplified mask_def[where n=6], simplified],
5509                     clarsimp\<close>
5510          | solves \<open>drule (2) bitmap_word_zero_no_bits_set2[simplified mask_def[where n=6], simplified],
5511                     clarsimp\<close>
5512          )+
5513    done
5514  done
5515  done
5516
5517lemma foldl_all_False:
5518  "(\<not> foldl (\<lambda>b x. b \<or> f x) False xs) = (\<forall>x \<in> set xs. \<not> f x)"
5519  apply (subst foldl_fun_or_alt)
5520  apply (fold orList_def)
5521  apply (simp add: orList_False image_subset_iff)
5522  done
5523
5524lemma isIOPortRangeFree_ccorres:
5525  notes Collect_const[simp del]
5526  shows
5527  "ccorres (\<lambda>rv rv'. rv' = from_bool rv) ret__unsigned_long_'
5528     (\<lambda>_. f \<le> l)
5529     (UNIV \<inter> \<lbrace>\<acute>first_port = f\<rbrace> \<inter> \<lbrace>\<acute>last_port = l\<rbrace>)
5530     hs
5531     (isIOPortRangeFree f l)
5532     (Call isIOPortRangeFree_'proc)"
5533  apply (intro ccorres_from_vcg hoarep_conseq_spec_state[OF isIOPortRangeFree_spec] allI)
5534  apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def carch_state_relation_def
5535                        global_ioport_bitmap_relation_def2
5536                        isIOPortRangeFree_def gets_def get_def monad_simps in_monad
5537                intro!: arg_cong[where f=from_bool])
5538  apply (match premises in H: \<open>cioport_bitmap_to_H _ = _\<close> and O: \<open>first_port_' s \<le> last_port_' s\<close> for s
5539           \<Rightarrow> \<open>match premises in _[thin]: _(multi) \<Rightarrow> \<open>insert O H\<close>\<close>)
5540  apply (clarsimp simp: cioport_bitmap_to_H_def wordRadix_def port_mask_def foldl_all_False)
5541  apply (rule all_cong, drule_tac x=port in fun_cong)
5542  by simp
5543
5544lemma decodeIOPortControlInvocation_ccorres:
5545  notes if_cong[cong] Collect_const[simp del]
5546  assumes "interpret_excaps extraCaps' = excaps_map extraCaps"
5547  assumes "isIOPortControlCap cp"
5548  shows
5549  "ccorres (intr_and_se_rel \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
5550       (invs' and (\<lambda>s. ksCurThread s = thread) and ct_active' and sch_act_simple
5551              and (excaps_in_mem extraCaps \<circ> ctes_of)
5552              and cte_wp_at' (diminished' (ArchObjectCap cp) \<circ> cteCap) slot
5553              and (\<lambda>s. \<forall>v \<in> set extraCaps. ex_cte_cap_wp_to' isCNodeCap (snd v) s)
5554              and (\<lambda>s. \<forall>v \<in> set extraCaps. s \<turnstile>' fst v)
5555              and sysargs_rel args buffer and valid_objs')
5556       (UNIV \<inter> {s. invLabel_' s = label}
5557             \<inter> {s. unat (length___unsigned_long_' s) = length args}
5558             \<inter> {s. slot_' s = cte_Ptr slot}
5559             \<inter> {s. excaps_' s = extraCaps'}
5560             \<inter> {s. ccap_relation (ArchObjectCap cp) (cap_' s)}
5561             \<inter> {s. buffer_' s = option_to_ptr buffer}) []
5562       (decodeX64PortInvocation label args slot cp (map fst extraCaps)
5563              >>= invocationCatch thread isBlocking isCall InvokeArchObject)
5564       (Call decodeX86PortControlInvocation_'proc)"
5565   (is "ccorres _ _ ?apre ?cpre _ _ _")
5566proof -
5567  from assms show ?thesis
5568    apply (clarsimp simp: isCap_simps decodeX64PortInvocation_def Let_def)
5569    apply (cinit' lift: invLabel_' length___unsigned_long_' slot_' excaps_' cap_' buffer_')
5570     apply (clarsimp cong: StateSpace.state.fold_congs globals.fold_congs)
5571     apply (rule ccorres_Cond_rhs_Seq, clarsimp simp: invocation_eq_use_types, ccorres_rewrite)
5572      apply (rule ccorres_equals_throwError)
5573       apply (fastforce simp: throwError_bind invocationCatch_def
5574                       split: invocation_label.splits arch_invocation_label.splits)
5575      apply (rule syscall_error_throwError_ccorres_n)
5576      apply (clarsimp simp: syscall_error_to_H_cases)
5577     apply (clarsimp simp: invocation_eq_use_types)
5578     apply csymbr
5579     apply (rule ccorres_Cond_rhs_Seq)
5580      apply (simp add: if_1_0_0 word_less_nat_alt throwError_bind invocationCatch_def)
5581      apply (rule ccorres_cond_true_seq)
5582      apply (rule ccorres_equals_throwError)
5583       apply (simp add: throwError_bind split: list.split)
5584       apply fastforce
5585      apply (rule syscall_error_throwError_ccorres_n)
5586      apply (simp add: syscall_error_to_H_cases)
5587     apply csymbr
5588     apply (simp add: if_1_0_0 interpret_excaps_test_null excaps_map_def)
5589     apply (rule ccorres_Cond_rhs_Seq)
5590      apply (simp add: throwError_bind invocationCatch_def)
5591      apply (rule ccorres_equals_throwError)
5592       apply (fastforce simp: throwError_bind split: list.split)
5593      apply (rule syscall_error_throwError_ccorres_n)
5594      apply (simp add: syscall_error_to_H_cases)
5595     apply ccorres_rewrite
5596     apply (frule (1) unat_length_4_helper)
5597     apply (clarsimp simp: neq_Nil_conv)
5598     apply (rule ccorres_add_return,
5599                  ctac add: getSyscallArg_ccorres_foo[where args=args and buffer=buffer and n=0])
5600       apply csymbr
5601       apply (rule ccorres_add_return,
5602                    ctac add: getSyscallArg_ccorres_foo[where args=args and buffer=buffer and n=1])
5603         apply csymbr
5604         apply (rule ccorres_add_return,
5605                      ctac add: getSyscallArg_ccorres_foo[where args=args and buffer=buffer and n=2])
5606           apply (rule ccorres_add_return,
5607                        ctac add: getSyscallArg_ccorres_foo[where args=args and buffer=buffer and n=3])
5608             apply (rule getSlotCap_ccorres_fudge_n[where vals=extraCaps and n=0])
5609             apply (rule ccorres_move_c_guard_cte)
5610             apply ctac
5611               apply (rule ccorres_assert2)
5612               apply (clarsimp simp: first_port_last_port_compare)
5613               apply (rule ccorres_Cond_rhs_Seq)
5614                apply (rule ccorres_equals_throwError)
5615                 apply (fastforce simp: whenE_def throwError_bind invocationCatch_def)
5616                apply ccorres_rewrite
5617                apply (rule syscall_error_throwError_ccorres_n[simplified dc_def id_def o_def])
5618                apply (clarsimp simp: syscall_error_to_H_cases)
5619               apply (clarsimp simp: ucast_drop_big_mask)
5620               apply (clarsimp simp: invocationCatch_use_injection_handler injection_bindE[OF refl refl]
5621                                     injection_handler_If injection_handler_throwError
5622                                     injection_liftE[OF refl] injection_handler_returnOk bindE_assoc)
5623               apply (clarsimp simp: liftE_bindE)
5624               apply (ctac(no_vcg) add: isIOPortRangeFree_ccorres)
5625                apply clarsimp
5626                apply (rule ccorres_Cond_rhs_Seq)
5627                 apply (clarsimp simp: from_bool_0 injection_handler_throwError)
5628                 apply ccorres_rewrite
5629                 apply (rule syscall_error_throwError_ccorres_n[simplified dc_def id_def o_def])
5630                 apply (clarsimp simp: syscall_error_to_H_cases)
5631                apply (clarsimp simp: from_bool_neq_0 injection_handler_returnOk)
5632                apply (ctac add: ccorres_injection_handler_csum1
5633                                   [OF lookupTargetSlot_ccorres, unfolded lookupTargetSlot_def])
5634                   apply (simp add: Collect_False split_def)
5635                   apply csymbr
5636                   apply (ctac add: ccorres_injection_handler_csum1[OF ensureEmptySlot_ccorres])
5637                      apply (simp add: injection_handler_returnOk ccorres_invocationCatch_Inr
5638                                       performInvocation_def bindE_assoc X64_H.performInvocation_def)
5639                      apply ccorres_rewrite
5640                      apply (rule_tac P="\<lambda>s. thread = ksCurThread s" in ccorres_cross_over_guard)
5641                      apply (ctac add: setThreadState_ccorres)
5642                        apply (ctac(no_vcg) add: invokeX86PortControl_ccorres
5643                                                   [simplified dc_def o_def id_def])
5644                        apply clarsimp
5645                        apply (rule ccorres_alternative2)
5646                        apply (rule ccorres_return_CE, simp+)[1]
5647                        apply (rule ccorres_return_C_errorE, simp+)[1]
5648                        apply wp
5649                       apply (wp sts_invs_minor')
5650                      apply (simp add: Collect_const_mem)
5651                      apply (vcg exspec=setThreadState_modifies)
5652                     apply simp
5653                     apply (rule ccorres_split_throws, ccorres_rewrite)
5654                      apply (rule ccorres_return_C_errorE, simp+)[1]
5655                     apply vcg
5656                    apply (wpsimp wp: injection_wp_E[OF refl])
5657                   apply (simp add: Collect_const_mem all_ex_eq_helper)
5658                   apply (vcg exspec=ensureEmptySlot_modifies)
5659                  apply simp
5660                  apply (rule ccorres_split_throws, ccorres_rewrite)
5661                   apply (rule ccorres_return_C_errorE, simp+)[1]
5662                  apply vcg
5663                 apply simp
5664                 apply (wpsimp wp: injection_wp_E[OF refl] hoare_drop_imps)
5665                apply (simp add: all_ex_eq_helper)
5666                apply (vcg exspec=lookupTargetSlot_modifies)
5667               apply (wpsimp wp: isIOPortRangeFree_wp)
5668              apply (rule_tac Q="\<lambda>rv. invs' and valid_cap' a and st_tcb_at' runnable' thread
5669                                      and sch_act_simple and cte_wp_at' \<top> slot
5670                                      and (\<lambda>s. thread = ksCurThread s)" in hoare_strengthen_post)
5671               apply (wpsimp wp: getSlotCap_wp)
5672              apply (clarsimp simp: unat_less_2p_word_bits invs_queues invs_valid_objs'
5673                                    valid_tcb_state'_def
5674                                    invs_sch_act_wf' st_tcb_strg'[rule_format] st_tcb_at'_def obj_at'_def
5675                                    projectKOs word_le_not_less
5676                             split: thread_state.splits)
5677              apply (case_tac "tcbState obj"; clarsimp)
5678               apply (drule invs_valid_idle',
5679                           clarsimp simp: valid_idle'_def pred_tcb_at'_def obj_at'_def projectKOs)+
5680             apply (simp add: all_ex_eq_helper)
5681             apply vcg
5682            apply wp
5683           apply (simp add: all_ex_eq_helper, vcg exspec=getSyscallArg_modifies)
5684          apply wp
5685         apply (simp add: all_ex_eq_helper, vcg exspec=getSyscallArg_modifies)
5686        apply wp
5687       apply (simp add: all_ex_eq_helper, vcg exspec=getSyscallArg_modifies)
5688      apply (rule_tac Q="\<lambda>rv. ?apre" in hoare_strengthen_post)
5689       apply wp
5690      apply (clarsimp simp: sysargs_rel_to_n excaps_in_mem_def slotcap_in_mem_def cte_wp_at_ctes_of
5691                      interpret_excaps_eq
5692               dest!: ct_active_runnable')
5693      apply (clarsimp simp: ct_in_state'_def)
5694     apply (rule_tac P="UNIV" in conseqPre)
5695      apply (simp add: all_ex_eq_helper, vcg exspec=getSyscallArg_modifies)
5696     apply (clarsimp simp: interpret_excaps_eq rf_sr_ksCurThread ThreadState_Restart_def mask_def)
5697     apply (clarsimp simp: rf_sr_def cstate_relation_def Let_def)
5698    apply clarsimp
5699    apply (rule conjI, clarsimp simp: sysargs_rel_to_n o_def dest!: unat_length_4_helper)
5700    apply (clarsimp simp: o_def)
5701    done
5702qed
5703
5704lemma unat_length_2_helper:
5705  "\<lbrakk>unat (l::machine_word) = length args; \<not> l < 2\<rbrakk> \<Longrightarrow> \<exists>x xa xs. args = x#xa#xs"
5706  apply (case_tac args; clarsimp simp: unat_eq_0)
5707  by (case_tac list; clarsimp simp: unat_eq_1)
5708
5709lemma ct_active_st_tcb_at_minor':
5710  assumes "ct_active' s"
5711  shows "st_tcb_at' (\<lambda>st'. tcb_st_refs_of' st' = {} \<and> st' \<noteq> Inactive \<and> st' \<noteq> IdleThreadState) (ksCurThread s) s"
5712        "st_tcb_at' runnable' (ksCurThread s) s"
5713   using assms
5714   by (clarsimp simp: st_tcb_at'_def ct_in_state'_def obj_at'_def projectKOs,
5715              case_tac "tcbState obj"; clarsimp)+
5716
5717(* FIXME: move to Lib *)
5718lemma length_Suc_0_conv:
5719  "length x = Suc 0 = (\<exists>y. x = [y])"
5720  by (induct x; clarsimp)
5721
5722lemma decodeIOPortInvocation_ccorres:
5723  notes if_cong[cong]
5724  assumes "interpret_excaps extraCaps' = excaps_map extraCaps"
5725  assumes "isIOPortCap cp"
5726  shows
5727  "ccorres (intr_and_se_rel \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
5728       (invs' and (\<lambda>s. ksCurThread s = thread) and ct_active' and sch_act_simple
5729              and (excaps_in_mem extraCaps \<circ> ctes_of)
5730              and cte_wp_at' (diminished' (ArchObjectCap cp) \<circ> cteCap) slot
5731              and (\<lambda>s. \<forall>v \<in> set extraCaps. ex_cte_cap_wp_to' isCNodeCap (snd v) s)
5732              and sysargs_rel args buffer and valid_objs')
5733       (UNIV \<inter> {s. invLabel_' s = label}
5734             \<inter> {s. unat (length___unsigned_long_' s) = length args}
5735             \<inter> {s. slot_' s = cte_Ptr slot}
5736             \<inter> {s. excaps_' s = extraCaps'}
5737             \<inter> {s. ccap_relation (ArchObjectCap cp) (cap_' s)}
5738             \<inter> {s. buffer_' s = option_to_ptr buffer}
5739             \<inter> {s. call_' s = from_bool isCall}) []
5740       (decodeX64PortInvocation label args slot cp (map fst extraCaps)
5741              >>= invocationCatch thread isBlocking isCall InvokeArchObject)
5742       (Call decodeX86PortInvocation_'proc)"
5743proof -
5744  from assms show ?thesis
5745  supply Collect_const[simp del]
5746  apply (clarsimp simp: isCap_simps decodeX64PortInvocation_def Let_def)
5747  apply (cinit' lift: invLabel_' length___unsigned_long_' slot_' excaps_' cap_' buffer_' call_')
5748   apply (clarsimp cong: StateSpace.state.fold_congs globals.fold_congs)
5749   apply (rule ccorres_Cond_rhs) (* IN invocations *)
5750    apply (erule ccorres_disj_division)
5751     \<comment> \<open>In8\<close>
5752       apply (clarsimp simp: invocation_eq_use_types cong: list.case_cong)
5753       apply (rule ccorres_rhs_assoc)+
5754       apply (rule ccorres_Cond_rhs_Seq) (* length error *)
5755        apply (clarsimp simp: throwError_bind invocationCatch_def)
5756        apply ccorres_rewrite
5757        apply (rule syscall_error_throwError_ccorres_n)
5758        apply (clarsimp simp: syscall_error_to_H_cases)
5759       apply (case_tac args; clarsimp simp: unat_gt_0[symmetric])
5760       apply (rule ccorres_add_return,
5761                 ctac add: getSyscallArg_ccorres_foo[where args=args and buffer=buffer and n=0])
5762         apply csymbr
5763         apply ccorres_rewrite
5764         apply (clarsimp simp: invocationCatch_use_injection_handler injection_bindE[OF refl refl]
5765                               injection_handler_If injection_handler_throwError
5766                               injection_liftE[OF refl] injection_handler_returnOk bindE_assoc)
5767         apply (ctac add: ccorres_injection_handler_csum1 [OF ensurePortOperationAllowed_ccorres])
5768            apply (clarsimp, ccorres_rewrite)
5769            apply (clarsimp simp: injection_handler_returnOk ccorres_invocationCatch_Inr bindE_assoc returnOk_bind liftE_bindE)
5770            apply (ctac add: setThreadState_ccorres)
5771              apply (simp add: X64_H.performInvocation_def performInvocation_def returnOk_bind bindE_assoc)
5772              apply (rule ccorres_nondet_refinement)
5773               apply (rule is_nondet_refinement_bindE)
5774                apply (rule is_nondet_refinement_refl)
5775               apply (simp split: if_split)
5776               apply (rule conjI[rotated], rule impI, rule is_nondet_refinement_refl)
5777               apply (rule impI, rule is_nondet_refinement_alternative1)
5778              apply (simp add: ucast_mask_drop[where n=16, simplified mask_def, simplified])
5779              apply (rule ccorres_add_returnOk)
5780              apply (ctac (no_vcg) add:invokeX86PortIn8_ccorres)
5781                apply (rule ccorres_return_CE, simp+)[1]
5782               apply (rule ccorres_return_C_errorE, simp+)[1]
5783              apply wp
5784             apply (wpsimp wp: ct_in_state'_set sts_running_valid_queues)
5785            apply (simp add: Collect_const_mem intr_and_se_rel_def cintr_def exception_defs)
5786            apply (vcg exspec=setThreadState_modifies)
5787           apply clarsimp
5788           apply ccorres_rewrite
5789           apply (rule ccorres_return_C_errorE, simp+)[1]
5790          apply (wpsimp wp: injection_wp_E[where f=Inl])
5791         apply (vcg exspec=ensurePortOperationAllowed_modifies)
5792        apply wpsimp
5793       apply (vcg exspec=getSyscallArg_modifies)
5794    apply (erule ccorres_disj_division)
5795    \<comment> \<open>In16\<close>
5796       apply (clarsimp simp: invocation_eq_use_types cong: list.case_cong)
5797       apply (rule ccorres_rhs_assoc)+
5798       apply (rule ccorres_Cond_rhs_Seq) (* length error *)
5799        apply (clarsimp simp: throwError_bind invocationCatch_def)
5800        apply ccorres_rewrite
5801        apply (rule syscall_error_throwError_ccorres_n)
5802        apply (clarsimp simp: syscall_error_to_H_cases)
5803       apply (case_tac args; clarsimp simp: unat_gt_0[symmetric])
5804       apply (rule ccorres_add_return,
5805                 ctac add: getSyscallArg_ccorres_foo[where args=args and buffer=buffer and n=0])
5806         apply csymbr
5807         apply ccorres_rewrite
5808         apply (clarsimp simp: invocationCatch_use_injection_handler injection_bindE[OF refl refl]
5809                               injection_handler_If injection_handler_throwError
5810                               injection_liftE[OF refl] injection_handler_returnOk bindE_assoc)
5811        apply (ctac add: ccorres_injection_handler_csum1 [OF ensurePortOperationAllowed_ccorres])
5812           apply (clarsimp, ccorres_rewrite)
5813           apply (clarsimp simp: injection_handler_returnOk ccorres_invocationCatch_Inr bindE_assoc returnOk_bind liftE_bindE)
5814           apply (ctac add: setThreadState_ccorres)
5815             apply (simp add: X64_H.performInvocation_def performInvocation_def returnOk_bind bindE_assoc)
5816             apply (rule ccorres_nondet_refinement)
5817              apply (rule is_nondet_refinement_bindE)
5818               apply (rule is_nondet_refinement_refl)
5819              apply (simp split: if_split)
5820              apply (rule conjI[rotated], rule impI, rule is_nondet_refinement_refl)
5821             apply (rule impI, rule is_nondet_refinement_alternative1)
5822             apply (simp add: ucast_mask_drop[where n=16, simplified mask_def, simplified])
5823             apply (rule ccorres_add_returnOk)
5824             apply (ctac (no_vcg) add:invokeX86PortIn16_ccorres)
5825               apply (rule ccorres_return_CE, simp+)[1]
5826              apply (rule ccorres_return_C_errorE, simp+)[1]
5827             apply wp
5828            apply (wpsimp wp: ct_in_state'_set sts_running_valid_queues)
5829           apply (simp add: Collect_const_mem intr_and_se_rel_def cintr_def exception_defs)
5830           apply (vcg exspec=setThreadState_modifies)
5831          apply clarsimp
5832          apply ccorres_rewrite
5833          apply (rule ccorres_return_C_errorE, simp+)[1]
5834         apply (wpsimp wp: injection_wp_E[where f=Inl])
5835        apply (vcg exspec=ensurePortOperationAllowed_modifies)
5836       apply wpsimp
5837      apply (vcg exspec=getSyscallArg_modifies)
5838    \<comment> \<open>In32\<close>
5839       apply (clarsimp simp: invocation_eq_use_types cong: list.case_cong)
5840       apply (rule ccorres_rhs_assoc)+
5841       apply (rule ccorres_Cond_rhs_Seq) (* length error *)
5842        apply (clarsimp simp: throwError_bind invocationCatch_def)
5843        apply ccorres_rewrite
5844        apply (rule syscall_error_throwError_ccorres_n)
5845        apply (clarsimp simp: syscall_error_to_H_cases)
5846       apply (case_tac args; clarsimp simp: unat_gt_0[symmetric])
5847       apply (rule ccorres_add_return,
5848                 ctac add: getSyscallArg_ccorres_foo[where args=args and buffer=buffer and n=0])
5849         apply csymbr
5850         apply ccorres_rewrite
5851         apply (clarsimp simp: invocationCatch_use_injection_handler injection_bindE[OF refl refl]
5852                               injection_handler_If injection_handler_throwError
5853                               injection_liftE[OF refl] injection_handler_returnOk bindE_assoc)
5854        apply (ctac add: ccorres_injection_handler_csum1 [OF ensurePortOperationAllowed_ccorres])
5855           apply (clarsimp, ccorres_rewrite)
5856           apply (clarsimp simp: injection_handler_returnOk ccorres_invocationCatch_Inr bindE_assoc returnOk_bind liftE_bindE)
5857           apply (ctac add: setThreadState_ccorres)
5858             apply (simp add: X64_H.performInvocation_def performInvocation_def returnOk_bind bindE_assoc)
5859             apply (rule ccorres_nondet_refinement)
5860              apply (rule is_nondet_refinement_bindE)
5861               apply (rule is_nondet_refinement_refl)
5862              apply (simp split: if_split)
5863              apply (rule conjI[rotated], rule impI, rule is_nondet_refinement_refl)
5864             apply (rule impI, rule is_nondet_refinement_alternative1)
5865             apply (simp add: ucast_mask_drop[where n=16, simplified mask_def, simplified])
5866             apply (rule ccorres_add_returnOk)
5867             apply (ctac (no_vcg) add:invokeX86PortIn32_ccorres)
5868               apply (rule ccorres_return_CE, simp+)[1]
5869              apply (rule ccorres_return_C_errorE, simp+)[1]
5870             apply wp
5871            apply (wpsimp wp: ct_in_state'_set sts_running_valid_queues)
5872           apply (simp add: Collect_const_mem intr_and_se_rel_def cintr_def exception_defs)
5873           apply (vcg exspec=setThreadState_modifies)
5874          apply clarsimp
5875          apply ccorres_rewrite
5876          apply (rule ccorres_return_C_errorE, simp+)[1]
5877         apply (wpsimp wp: injection_wp_E[where f=Inl])
5878        apply (vcg exspec=ensurePortOperationAllowed_modifies)
5879       apply wpsimp
5880      apply (vcg exspec=getSyscallArg_modifies)
5881   apply clarsimp
5882   apply (rule ccorres_Cond_rhs) (* OUT invocations *)
5883    apply (erule ccorres_disj_division)
5884    \<comment> \<open>Out8\<close>
5885     apply (clarsimp simp: invocation_eq_use_types cong: list.case_cong)
5886     apply (rule ccorres_rhs_assoc)+
5887     apply (rule ccorres_Cond_rhs_Seq) (* length error *)
5888      apply (rule ccorres_equals_throwError)
5889       apply (drule x_less_2_0_1, erule disjE;
5890                    clarsimp simp: throwError_bind invocationCatch_def length_Suc_0_conv
5891                            dest!: sym[where s="Suc 0"];
5892                    fastforce)
5893      apply ccorres_rewrite
5894      apply (rule syscall_error_throwError_ccorres_n)
5895      apply (clarsimp simp: syscall_error_to_H_cases)
5896     apply (frule (1) unat_length_2_helper)
5897    apply clarsimp
5898    apply (rule ccorres_add_return,
5899              ctac add: getSyscallArg_ccorres_foo[where args=args and buffer=buffer and n=0])
5900      apply csymbr
5901      apply (rule ccorres_add_return,
5902              ctac add: getSyscallArg_ccorres_foo[where args=args and buffer=buffer and n=1])
5903        apply csymbr
5904        apply ccorres_rewrite
5905        apply (clarsimp simp: invocationCatch_use_injection_handler injection_bindE[OF refl refl]
5906                            injection_handler_If injection_handler_throwError
5907                            injection_liftE[OF refl] injection_handler_returnOk bindE_assoc)
5908        apply (rule ccorres_rhs_assoc)+
5909        apply (ctac add: ccorres_injection_handler_csum1 [OF ensurePortOperationAllowed_ccorres])
5910           apply (clarsimp, ccorres_rewrite, csymbr)
5911           apply (clarsimp simp: injection_handler_returnOk ccorres_invocationCatch_Inr bindE_assoc returnOk_bind liftE_bindE)
5912           apply (ctac add: setThreadState_ccorres)
5913             apply (simp add: X64_H.performInvocation_def performInvocation_def returnOk_bind bindE_assoc)
5914             apply (simp add: ucast_mask_drop[where n=16, simplified mask_def, simplified]
5915                           ucast_mask_drop[where n=32, simplified mask_def, simplified])
5916             apply (ctac add: invokeX86PortOut8_ccorres)
5917                apply clarsimp
5918                 apply (rule ccorres_alternative2)
5919                 apply (rule ccorres_return_CE, simp+)[1]
5920                apply (rule ccorres_return_C_errorE, simp+)[1]
5921               apply wp
5922              apply (vcg exspec=invokeX86PortOut_modifies)
5923             apply (wp sts_invs_minor')
5924            apply (vcg exspec=setThreadState_modifies)
5925           apply (clarsimp, ccorres_rewrite, csymbr)
5926           apply (rule ccorres_return_C_errorE, simp+)[1]
5927          apply (wpsimp wp: injection_wp_E[where f=Inl])
5928         apply (vcg exspec=ensurePortOperationAllowed_modifies)
5929        apply wpsimp
5930       apply (vcg exspec=getSyscallArg_modifies)
5931      apply wpsimp
5932     apply (vcg exspec=getSyscallArg_modifies)
5933    apply (erule ccorres_disj_division)
5934    \<comment> \<open>Out16\<close>
5935     apply (clarsimp simp: invocation_eq_use_types cong: list.case_cong)
5936     apply (rule ccorres_rhs_assoc)+
5937     apply (rule ccorres_Cond_rhs_Seq) (* length error *)
5938      apply (rule ccorres_equals_throwError)
5939       apply (drule x_less_2_0_1, erule disjE;
5940                    clarsimp simp: throwError_bind invocationCatch_def length_Suc_0_conv
5941                            dest!: sym[where s="Suc 0"];
5942                    fastforce)
5943      apply ccorres_rewrite
5944      apply (rule syscall_error_throwError_ccorres_n)
5945      apply (clarsimp simp: syscall_error_to_H_cases)
5946     apply (frule (1) unat_length_2_helper)
5947    apply clarsimp
5948    apply (rule ccorres_add_return,
5949              ctac add: getSyscallArg_ccorres_foo[where args=args and buffer=buffer and n=0])
5950      apply csymbr
5951      apply (rule ccorres_add_return,
5952              ctac add: getSyscallArg_ccorres_foo[where args=args and buffer=buffer and n=1])
5953        apply csymbr
5954        apply ccorres_rewrite
5955        apply (clarsimp simp: invocationCatch_use_injection_handler injection_bindE[OF refl refl]
5956                            injection_handler_If injection_handler_throwError
5957                            injection_liftE[OF refl] injection_handler_returnOk bindE_assoc)
5958        apply (rule ccorres_rhs_assoc)+
5959        apply (ctac add: ccorres_injection_handler_csum1 [OF ensurePortOperationAllowed_ccorres])
5960           apply (clarsimp, ccorres_rewrite, csymbr)
5961           apply (clarsimp simp: injection_handler_returnOk ccorres_invocationCatch_Inr bindE_assoc returnOk_bind liftE_bindE)
5962           apply (ctac add: setThreadState_ccorres)
5963             apply (simp add: X64_H.performInvocation_def performInvocation_def returnOk_bind bindE_assoc)
5964             apply (simp add: ucast_mask_drop[where n=16, simplified mask_def, simplified]
5965                           ucast_mask_drop[where n=32, simplified mask_def, simplified])
5966             apply (ctac add: invokeX86PortOut16_ccorres)
5967                apply clarsimp
5968                apply (rule ccorres_alternative2)
5969                apply (rule ccorres_return_CE, simp+)[1]
5970               apply (rule ccorres_return_C_errorE, simp+)[1]
5971              apply wp
5972             apply (vcg exspec=invokeX86PortOut_modifies)
5973            apply (wp sts_invs_minor')
5974           apply (vcg exspec=setThreadState_modifies)
5975          apply (clarsimp, ccorres_rewrite, csymbr)
5976          apply (rule ccorres_return_C_errorE, simp+)[1]
5977         apply (wpsimp wp: injection_wp_E[where f=Inl])
5978        apply (vcg exspec=ensurePortOperationAllowed_modifies)
5979       apply wpsimp
5980      apply (vcg exspec=getSyscallArg_modifies)
5981     apply wpsimp
5982    apply (vcg exspec=getSyscallArg_modifies)
5983    \<comment> \<open>Out32\<close>
5984     apply (clarsimp simp: invocation_eq_use_types cong: list.case_cong)
5985     apply (rule ccorres_rhs_assoc)+
5986     apply (rule ccorres_Cond_rhs_Seq) (* length error *)
5987      apply (rule ccorres_equals_throwError)
5988       apply (drule x_less_2_0_1, erule disjE;
5989                    clarsimp simp: throwError_bind invocationCatch_def length_Suc_0_conv
5990                            dest!: sym[where s="Suc 0"];
5991                    fastforce)
5992      apply ccorres_rewrite
5993      apply (rule syscall_error_throwError_ccorres_n)
5994      apply (clarsimp simp: syscall_error_to_H_cases)
5995     apply (frule (1) unat_length_2_helper)
5996    apply clarsimp
5997    apply (rule ccorres_add_return,
5998              ctac add: getSyscallArg_ccorres_foo[where args=args and buffer=buffer and n=0])
5999      apply csymbr
6000      apply (rule ccorres_add_return,
6001              ctac add: getSyscallArg_ccorres_foo[where args=args and buffer=buffer and n=1])
6002        apply csymbr
6003        apply ccorres_rewrite
6004        apply (clarsimp simp: invocationCatch_use_injection_handler injection_bindE[OF refl refl]
6005                            injection_handler_If injection_handler_throwError
6006                            injection_liftE[OF refl] injection_handler_returnOk bindE_assoc)
6007        apply (rule ccorres_rhs_assoc)+
6008        apply (ctac add: ccorres_injection_handler_csum1 [OF ensurePortOperationAllowed_ccorres])
6009           apply (clarsimp, ccorres_rewrite, csymbr)
6010           apply (clarsimp simp: injection_handler_returnOk ccorres_invocationCatch_Inr bindE_assoc returnOk_bind liftE_bindE)
6011           apply (ctac add: setThreadState_ccorres)
6012             apply (simp add: X64_H.performInvocation_def performInvocation_def returnOk_bind bindE_assoc)
6013             apply (simp add: ucast_mask_drop[where n=16, simplified mask_def, simplified]
6014                           ucast_mask_drop[where n=32, simplified mask_def, simplified])
6015             apply (ctac add: invokeX86PortOut32_ccorres)
6016                apply clarsimp
6017                apply (rule ccorres_alternative2)
6018                apply (rule ccorres_return_CE, simp+)[1]
6019               apply (rule ccorres_return_C_errorE, simp+)[1]
6020              apply wp
6021             apply (vcg exspec=invokeX86PortOut_modifies)
6022            apply (wp sts_invs_minor')
6023           apply (vcg exspec=setThreadState_modifies)
6024          apply (clarsimp, ccorres_rewrite, csymbr)
6025          apply (rule ccorres_return_C_errorE, simp+)[1]
6026         apply (wpsimp wp: injection_wp_E[where f=Inl])
6027        apply (vcg exspec=ensurePortOperationAllowed_modifies)
6028       apply wpsimp
6029      apply (vcg exspec=getSyscallArg_modifies)
6030     apply wpsimp
6031    apply (vcg exspec=getSyscallArg_modifies)
6032   apply (clarsimp simp: invocation_eq_use_types)
6033   apply (rule ccorres_equals_throwError)
6034    apply (fastforce simp: throwError_bind invocationCatch_def
6035                    split: invocation_label.splits arch_invocation_label.splits)
6036   apply (rule syscall_error_throwError_ccorres_n)
6037   apply (clarsimp simp: syscall_error_to_H_cases)
6038  apply (clarsimp simp: arch_invocation_label_defs sysargs_rel_to_n valid_tcb_state'_def tcb_at_invs'
6039                        invs_queues invs_sch_act_wf' ct_active_st_tcb_at_minor' rf_sr_ksCurThread
6040                        ThreadState_Restart_def mask_def
6041                        ucast_mask_drop[where n=16, simplified mask_def, simplified])
6042  apply (safe, simp_all add: unat_eq_0 unat_eq_1)
6043         apply (clarsimp dest!: unat_length_2_helper simp: ThreadState_Restart_def mask_def syscall_error_rel_def
6044                                | (thin_tac "P" for P)+, word_bitwise)+
6045  done
6046qed
6047
6048lemma Mode_decodeInvocation_ccorres:
6049  assumes "interpret_excaps extraCaps' = excaps_map extraCaps"
6050  assumes "isPageCap cp \<or> isPageTableCap cp \<or> isPageDirectoryCap cp \<or> isPDPointerTableCap cp \<or> isPML4Cap cp"
6051  shows
6052  "ccorres (intr_and_se_rel \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
6053       (invs' and (\<lambda>s. ksCurThread s = thread) and ct_active' and sch_act_simple
6054              and (excaps_in_mem extraCaps \<circ> ctes_of)
6055              and cte_wp_at' (diminished' (ArchObjectCap cp) \<circ> cteCap) slot
6056              and (\<lambda>s. \<forall>v \<in> set extraCaps. ex_cte_cap_wp_to' isCNodeCap (snd v) s)
6057              and sysargs_rel args buffer and valid_objs')
6058       (UNIV \<inter> {s. label___unsigned_long_' s = label}
6059             \<inter> {s. unat (length___unsigned_long_' s) = length args}
6060             \<inter> {s. slot_' s = cte_Ptr slot}
6061             \<inter> {s. extraCaps___struct_extra_caps_C_' s = extraCaps'}
6062             \<inter> {s. ccap_relation (ArchObjectCap cp) (cap_' s)}
6063             \<inter> {s. buffer_' s = option_to_ptr buffer}) []
6064       (decodeX64MMUInvocation label args cptr slot cp extraCaps
6065          >>= invocationCatch thread isBlocking isCall Invocations_H.invocation.InvokeArchObject)
6066       (Call Mode_decodeInvocation_'proc)"
6067  using assms
6068  apply (cinit' lift: label___unsigned_long_' length___unsigned_long_' slot_'
6069                      extraCaps___struct_extra_caps_C_' cap_' buffer_')
6070   apply csymbr
6071   apply (simp only: cap_get_tag_isCap_ArchObject X64_H.decodeInvocation_def)
6072   apply (rule ccorres_cond_true)
6073   apply (rule ccorres_trim_returnE[rotated 2, OF ccorres_call, OF decodeX64MMUInvocation_ccorres], simp+)[1]
6074  apply (clarsimp simp: o_def)
6075  done
6076
6077lemma Arch_decodeInvocation_ccorres:
6078  notes if_cong[cong]
6079  assumes "interpret_excaps extraCaps' = excaps_map extraCaps"
6080  shows
6081  "ccorres (intr_and_se_rel \<currency> dc) (liftxf errstate id (K ()) ret__unsigned_long_')
6082       (invs' and (\<lambda>s. ksCurThread s = thread) and ct_active' and sch_act_simple
6083              and (excaps_in_mem extraCaps \<circ> ctes_of)
6084              and cte_wp_at' (diminished' (ArchObjectCap cp) \<circ> cteCap) slot
6085              and (\<lambda>s. \<forall>v \<in> set extraCaps. ex_cte_cap_wp_to' isCNodeCap (snd v) s)
6086              and sysargs_rel args buffer and valid_objs')
6087       (UNIV \<inter> {s. invLabel_' s = label}
6088             \<inter> {s. unat (length___unsigned_long_' s) = length args}
6089             \<inter> {s. slot_' s = cte_Ptr slot}
6090             \<inter> {s. excaps_' s = extraCaps'}
6091             \<inter> {s. ccap_relation (ArchObjectCap cp) (cap_' s)}
6092             \<inter> {s. buffer_' s = option_to_ptr buffer}
6093             \<inter> {s. call_' s = from_bool isCall }) []
6094       (Arch.decodeInvocation label args cptr slot cp extraCaps
6095              >>= invocationCatch thread isBlocking isCall InvokeArchObject)
6096       (Call Arch_decodeInvocation_'proc)"
6097  (is "ccorres ?r ?xf ?P (?P' slot_') [] ?a ?c")
6098proof -
6099  note trim_call = ccorres_trim_returnE[rotated 2, OF ccorres_call]
6100  from assms show ?thesis
6101    apply (cinit' lift: invLabel_' length___unsigned_long_' slot_' excaps_' cap_' buffer_' call_')
6102     apply csymbr
6103     apply (simp only: cap_get_tag_isCap_ArchObject X64_H.decodeInvocation_def)
6104     apply (rule ccorres_Cond_rhs)
6105      apply (subgoal_tac "\<not> isIOCap cp", simp)
6106       apply (rule trim_call[OF decodeX64MMUInvocation_ccorres], simp+)[1]
6107      apply (fastforce simp: isCap_simps isIOCap_def)
6108     apply (rule ccorres_Cond_rhs)
6109      apply (subgoal_tac "isIOCap cp", simp)
6110       apply (rule trim_call[OF decodeIOPortControlInvocation_ccorres], simp+)[1]
6111      apply (fastforce simp: isCap_simps isIOCap_def)
6112     apply (rule ccorres_Cond_rhs)
6113      apply (subgoal_tac "isIOCap cp", simp)
6114       apply (rule trim_call[OF decodeIOPortInvocation_ccorres], simp+)[1]
6115      apply (fastforce simp: isCap_simps isIOCap_def)
6116     apply (subgoal_tac "\<not> isIOCap cp", simp)
6117      apply (rule trim_call[OF Mode_decodeInvocation_ccorres])
6118            apply assumption
6119           apply (cases cp; simp add: isCap_simps)
6120          apply simp+
6121     apply (cases cp; simp add: isCap_simps isIOCap_def)
6122    apply (clarsimp simp: o_def excaps_in_mem_def slotcap_in_mem_def)
6123    apply (drule (1) bspec)
6124    apply (clarsimp simp: ctes_of_valid')
6125    done
6126qed
6127
6128end
6129
6130end
6131