-
Notifications
You must be signed in to change notification settings - Fork 0
/
DetWP.thy
152 lines (135 loc) · 4.62 KB
/
DetWP.thy
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
(*
* Copyright 2014, General Dynamics C4 Systems
*
* This software may be distributed and modified according to the terms of
* the GNU General Public License version 2. Note that NO WARRANTY is provided.
* See "LICENSE_GPLv2.txt" for details.
*
* @TAG(GD_GPL)
*)
theory DetWP
imports "../../lib/clib/DetWPLib" Include_C
begin
lemma det_wp_doMachineOp [wp]:
"det_wp (\<lambda>_. P) f \<Longrightarrow> det_wp (\<lambda>_. P) (doMachineOp f)"
apply (simp add: doMachineOp_def split_def)
apply (rule det_wp_pre, wp)
apply (erule det_wp_select_f)
apply wp
apply simp
done
lemma det_wp_loadWordUser [wp]:
"det_wp (pointerInUserData x and K (is_aligned x 2)) (loadWordUser x)"
apply (simp add: loadWordUser_def loadWord_def)
apply (rule det_wp_pre, wp)
apply (rule det_wp_pre, wp)
apply clarsimp
apply assumption
apply wp
apply (clarsimp simp: is_aligned_mask)
done
declare det_wp_liftM[wp]
declare det_wp_assert_opt[wp]
declare det_wp_when[wp]
declare det_wp_unless[wp]
declare word_neq_0_conv [simp del]
lemma det_wp_loadObject_default [wp]:
"det_wp (\<lambda>s. \<exists>obj. projectKO_opt ko = Some (obj::'a) \<and>
is_aligned p (objBits obj) \<and> q = p
\<and> case_option True (\<lambda>x. 2 ^ (objBits obj) \<le> x - p) n)
(loadObject_default p q n ko :: ('a::pre_storable) kernel)"
apply (simp add: loadObject_default_def split_def projectKO_def
alignCheck_def alignError_def magnitudeCheck_def
unless_def)
apply (rule det_wp_pre)
apply (wp case_option_wp)
apply (clarsimp simp: is_aligned_mask[symmetric])
apply simp
done
lemma det_wp_getTCB [wp]:
"det_wp (tcb_at' t) (getObject t :: tcb kernel)"
apply (simp add: getObject_def split_def)
apply (rule det_wp_pre)
apply (wp|wpc)+
apply (clarsimp simp add: obj_at'_def projectKOs objBits_simps
cong: conj_cong)
apply (simp add: lookupAround2_known1)
apply (rule ps_clear_lookupAround2, assumption+)
apply simp
apply (erule is_aligned_no_overflow)
apply (simp add: word_bits_def)
done
lemma det_wp_setObject_other [wp]:
fixes ob :: "'a :: pspace_storable"
assumes x: "updateObject ob = updateObject_default ob"
shows "det_wp (obj_at' (\<lambda>k::'a. objBits k = objBits ob) ptr)
(setObject ptr ob)"
apply (simp add: setObject_def x split_def updateObject_default_def
magnitudeCheck_def
projectKO_def2 alignCheck_def alignError_def)
apply (rule det_wp_pre)
apply (wp )
apply (clarsimp simp: is_aligned_mask[symmetric] obj_at'_def
objBits_def[symmetric] projectKOs
project_inject lookupAround2_known1)
apply (erule(1) ps_clear_lookupAround2)
apply simp
apply (erule is_aligned_get_word_bits)
apply (subst add_diff_eq[symmetric])
apply (erule is_aligned_no_wrap')
apply simp
apply simp
apply fastforce
done
lemma det_wp_setTCB [wp]:
"det_wp (tcb_at' t) (setObject t (v::tcb))"
apply (rule det_wp_pre)
apply (wp|wpc|simp)+
apply (clarsimp simp: objBits_simps)
done
lemma det_wp_threadGet [wp]:
"det_wp (tcb_at' t) (threadGet f t)"
apply (simp add: threadGet_def)
apply (rule det_wp_pre, wp)
apply simp
done
lemma det_wp_threadSet [wp]:
"det_wp (tcb_at' t) (threadSet f t)"
apply (simp add: threadSet_def)
apply (rule det_wp_pre, wp)
apply simp
done
lemma det_wp_asUser [wp]:
"det f \<Longrightarrow> det_wp (tcb_at' t) (asUser t f)"
apply (simp add: asUser_def split_def)
apply (rule det_wp_pre)
apply wp
apply (drule det_wp_det)
apply (erule det_wp_select_f)
apply wp
apply (rule_tac Q="\<lambda>_. tcb_at' t" in hoare_post_imp)
apply simp
apply wp
apply simp
done
lemma det_wp_getMRs:
"det_wp (tcb_at' thread and case_option \<top> valid_ipc_buffer_ptr' buffer) (getMRs thread buffer mi)"
apply (clarsimp simp: getMRs_def)
apply (rule det_wp_pre)
apply (wp det_mapM det_getRegister order_refl det_wp_mapM)
apply (simp add: word_size)
apply (wp asUser_inv mapM_wp' getRegister_inv)
apply clarsimp
apply (rule conjI)
apply (simp add: pointerInUserData_def)
apply (erule valid_ipc_buffer_ptr'D2)
apply (rule word_mult_less_mono1)
apply (erule order_le_less_trans)
apply (simp add: msgMaxLength_def max_ipc_words)
apply simp
apply (simp add: max_ipc_words)
apply (simp add: is_aligned_mult_triv2 [where n = 2, simplified] word_bits_conv)
apply (erule valid_ipc_buffer_ptr_aligned_2)
apply (simp add: is_aligned_mult_triv2 [where n = 2, simplified] word_bits_conv)
done
end