author | wenzelm |
Wed, 11 Dec 2024 10:40:57 +0100 | |
changeset 81577 | a712bf5ccab0 |
parent 80914 | d97fdabd9e2b |
child 82695 | d93ead9ac6df |
permissions | -rw-r--r-- |
12857 | 1 |
(* Title: HOL/Bali/AxSem.thy |
12854 | 2 |
Author: David von Oheimb |
3 |
*) |
|
4 |
||
62042 | 5 |
subsection \<open>Axiomatic semantics of Java expressions and statements |
12854 | 6 |
(see also Eval.thy) |
62042 | 7 |
\<close> |
16417 | 8 |
theory AxSem imports Evaln TypeSafe begin |
12854 | 9 |
|
62042 | 10 |
text \<open> |
12854 | 11 |
design issues: |
12 |
\begin{itemize} |
|
13 |
\item a strong version of validity for triples with premises, namely one that |
|
14 |
takes the recursive depth needed to complete execution, enables |
|
15 |
correctness proof |
|
16 |
\item auxiliary variables are handled first-class (-> Thomas Kleymann) |
|
17 |
\item expressions not flattened to elementary assignments (as usual for |
|
18 |
axiomatic semantics) but treated first-class => explicit result value |
|
19 |
handling |
|
20 |
\item intermediate values not on triple, but on assertion level |
|
21 |
(with result entry) |
|
22 |
\item multiple results with semantical substitution mechnism not requiring a |
|
23 |
stack |
|
24 |
\item because of dynamic method binding, terms need to be dependent on state. |
|
25 |
this is also useful for conditional expressions and statements |
|
26 |
\item result values in triples exactly as in eval relation (also for xcpt |
|
27 |
states) |
|
28 |
\item validity: additional assumption of state conformance and well-typedness, |
|
29 |
which is required for soundness and thus rule hazard required of completeness |
|
30 |
\end{itemize} |
|
31 |
||
32 |
restrictions: |
|
33 |
\begin{itemize} |
|
34 |
\item all triples in a derivation are of the same type (due to weak |
|
35 |
polymorphism) |
|
36 |
\end{itemize} |
|
62042 | 37 |
\<close> |
12854 | 38 |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67399
diff
changeset
|
39 |
type_synonym res = vals \<comment> \<open>result entry\<close> |
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
40 |
|
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
41 |
abbreviation (input) |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
42 |
Val where "Val x == In1 x" |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
43 |
|
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
44 |
abbreviation (input) |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
45 |
Var where "Var x == In2 x" |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
46 |
|
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
47 |
abbreviation (input) |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
48 |
Vals where "Vals x == In3 x" |
12854 | 49 |
|
50 |
syntax |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
51 |
"_Val" :: "[pttrn] => pttrn" (\<open>Val:_\<close> [951] 950) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
52 |
"_Var" :: "[pttrn] => pttrn" (\<open>Var:_\<close> [951] 950) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
53 |
"_Vals" :: "[pttrn] => pttrn" (\<open>Vals:_\<close> [951] 950) |
12854 | 54 |
|
55 |
translations |
|
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
56 |
"\<lambda>Val:v . b" == "(\<lambda>v. b) \<circ> CONST the_In1" |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
57 |
"\<lambda>Var:v . b" == "(\<lambda>v. b) \<circ> CONST the_In2" |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
58 |
"\<lambda>Vals:v. b" == "(\<lambda>v. b) \<circ> CONST the_In3" |
12854 | 59 |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67399
diff
changeset
|
60 |
\<comment> \<open>relation on result values, state and auxiliary variables\<close> |
41778 | 61 |
type_synonym 'a assn = "res \<Rightarrow> state \<Rightarrow> 'a \<Rightarrow> bool" |
12854 | 62 |
translations |
35431 | 63 |
(type) "'a assn" <= (type) "vals \<Rightarrow> state \<Rightarrow> 'a \<Rightarrow> bool" |
12854 | 64 |
|
37956 | 65 |
definition |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
66 |
assn_imp :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> bool" (infixr \<open>\<Rightarrow>\<close> 25) |
37956 | 67 |
where "(P \<Rightarrow> Q) = (\<forall>Y s Z. P Y s Z \<longrightarrow> Q Y s Z)" |
12854 | 68 |
|
69 |
lemma assn_imp_def2 [iff]: "(P \<Rightarrow> Q) = (\<forall>Y s Z. P Y s Z \<longrightarrow> Q Y s Z)" |
|
70 |
apply (unfold assn_imp_def) |
|
71 |
apply (rule HOL.refl) |
|
72 |
done |
|
73 |
||
74 |
||
58887 | 75 |
subsubsection "assertion transformers" |
12854 | 76 |
|
77 |
subsection "peek-and" |
|
78 |
||
37956 | 79 |
definition |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
80 |
peek_and :: "'a assn \<Rightarrow> (state \<Rightarrow> bool) \<Rightarrow> 'a assn" (infixl \<open>\<and>.\<close> 13) |
37956 | 81 |
where "(P \<and>. p) = (\<lambda>Y s Z. P Y s Z \<and> p s)" |
12854 | 82 |
|
83 |
lemma peek_and_def2 [simp]: "peek_and P p Y s = (\<lambda>Z. (P Y s Z \<and> p s))" |
|
84 |
apply (unfold peek_and_def) |
|
85 |
apply (simp (no_asm)) |
|
86 |
done |
|
87 |
||
88 |
lemma peek_and_Not [simp]: "(P \<and>. (\<lambda>s. \<not> f s)) = (P \<and>. Not \<circ> f)" |
|
89 |
apply (rule ext) |
|
90 |
apply (rule ext) |
|
91 |
apply (simp (no_asm)) |
|
92 |
done |
|
93 |
||
94 |
lemma peek_and_and [simp]: "peek_and (peek_and P p) p = peek_and P p" |
|
95 |
apply (unfold peek_and_def) |
|
96 |
apply (simp (no_asm)) |
|
97 |
done |
|
98 |
||
99 |
lemma peek_and_commut: "(P \<and>. p \<and>. q) = (P \<and>. q \<and>. p)" |
|
100 |
apply (rule ext) |
|
101 |
apply (rule ext) |
|
102 |
apply (rule ext) |
|
103 |
apply auto |
|
104 |
done |
|
105 |
||
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
106 |
abbreviation |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
107 |
Normal :: "'a assn \<Rightarrow> 'a assn" |
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
108 |
where "Normal P == P \<and>. normal" |
12854 | 109 |
|
110 |
lemma peek_and_Normal [simp]: "peek_and (Normal P) p = Normal (peek_and P p)" |
|
111 |
apply (rule ext) |
|
112 |
apply (rule ext) |
|
113 |
apply (rule ext) |
|
114 |
apply auto |
|
115 |
done |
|
116 |
||
117 |
subsection "assn-supd" |
|
118 |
||
37956 | 119 |
definition |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
120 |
assn_supd :: "'a assn \<Rightarrow> (state \<Rightarrow> state) \<Rightarrow> 'a assn" (infixl \<open>;.\<close> 13) |
37956 | 121 |
where "(P ;. f) = (\<lambda>Y s' Z. \<exists>s. P Y s Z \<and> s' = f s)" |
12854 | 122 |
|
123 |
lemma assn_supd_def2 [simp]: "assn_supd P f Y s' Z = (\<exists>s. P Y s Z \<and> s' = f s)" |
|
124 |
apply (unfold assn_supd_def) |
|
125 |
apply (simp (no_asm)) |
|
126 |
done |
|
127 |
||
128 |
subsection "supd-assn" |
|
129 |
||
37956 | 130 |
definition |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
131 |
supd_assn :: "(state \<Rightarrow> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (infixr \<open>.;\<close> 13) |
37956 | 132 |
where "(f .; P) = (\<lambda>Y s. P Y (f s))" |
12854 | 133 |
|
134 |
||
135 |
lemma supd_assn_def2 [simp]: "(f .; P) Y s = P Y (f s)" |
|
136 |
apply (unfold supd_assn_def) |
|
137 |
apply (simp (no_asm)) |
|
138 |
done |
|
139 |
||
140 |
lemma supd_assn_supdD [elim]: "((f .; Q) ;. f) Y s Z \<Longrightarrow> Q Y s Z" |
|
141 |
apply auto |
|
142 |
done |
|
143 |
||
144 |
lemma supd_assn_supdI [elim]: "Q Y s Z \<Longrightarrow> (f .; (Q ;. f)) Y s Z" |
|
145 |
apply (auto simp del: split_paired_Ex) |
|
146 |
done |
|
147 |
||
148 |
subsection "subst-res" |
|
149 |
||
37956 | 150 |
definition |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
151 |
subst_res :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn" (\<open>_\<leftarrow>_\<close> [60,61] 60) |
37956 | 152 |
where "P\<leftarrow>w = (\<lambda>Y. P w)" |
12854 | 153 |
|
154 |
lemma subst_res_def2 [simp]: "(P\<leftarrow>w) Y = P w" |
|
155 |
apply (unfold subst_res_def) |
|
156 |
apply (simp (no_asm)) |
|
157 |
done |
|
158 |
||
159 |
lemma subst_subst_res [simp]: "P\<leftarrow>w\<leftarrow>v = P\<leftarrow>w" |
|
160 |
apply (rule ext) |
|
161 |
apply (simp (no_asm)) |
|
162 |
done |
|
163 |
||
164 |
lemma peek_and_subst_res [simp]: "(P \<and>. p)\<leftarrow>w = (P\<leftarrow>w \<and>. p)" |
|
165 |
apply (rule ext) |
|
166 |
apply (rule ext) |
|
167 |
apply (simp (no_asm)) |
|
168 |
done |
|
169 |
||
170 |
(*###Do not work for some strange (unification?) reason |
|
171 |
lemma subst_res_Val_beta [simp]: "(\<lambda>Y. P (the_In1 Y))\<leftarrow>Val v = (\<lambda>Y. P v)" |
|
172 |
apply (rule ext) |
|
173 |
by simp |
|
174 |
||
175 |
lemma subst_res_Var_beta [simp]: "(\<lambda>Y. P (the_In2 Y))\<leftarrow>Var vf = (\<lambda>Y. P vf)"; |
|
176 |
apply (rule ext) |
|
177 |
by simp |
|
178 |
||
179 |
lemma subst_res_Vals_beta [simp]: "(\<lambda>Y. P (the_In3 Y))\<leftarrow>Vals vs = (\<lambda>Y. P vs)"; |
|
180 |
apply (rule ext) |
|
181 |
by simp |
|
182 |
*) |
|
183 |
||
184 |
subsection "subst-Bool" |
|
185 |
||
37956 | 186 |
definition |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
187 |
subst_Bool :: "'a assn \<Rightarrow> bool \<Rightarrow> 'a assn" (\<open>_\<leftarrow>=_\<close> [60,61] 60) |
37956 | 188 |
where "P\<leftarrow>=b = (\<lambda>Y s Z. \<exists>v. P (Val v) s Z \<and> (normal s \<longrightarrow> the_Bool v=b))" |
12854 | 189 |
|
190 |
lemma subst_Bool_def2 [simp]: |
|
191 |
"(P\<leftarrow>=b) Y s Z = (\<exists>v. P (Val v) s Z \<and> (normal s \<longrightarrow> the_Bool v=b))" |
|
192 |
apply (unfold subst_Bool_def) |
|
193 |
apply (simp (no_asm)) |
|
194 |
done |
|
195 |
||
196 |
lemma subst_Bool_the_BoolI: "P (Val b) s Z \<Longrightarrow> (P\<leftarrow>=the_Bool b) Y s Z" |
|
197 |
apply auto |
|
198 |
done |
|
199 |
||
200 |
subsection "peek-res" |
|
201 |
||
37956 | 202 |
definition |
203 |
peek_res :: "(res \<Rightarrow> 'a assn) \<Rightarrow> 'a assn" |
|
204 |
where "peek_res Pf = (\<lambda>Y. Pf Y Y)" |
|
12854 | 205 |
|
206 |
syntax |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
207 |
"_peek_res" :: "pttrn \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (\<open>\<lambda>_:. _\<close> [0,3] 3) |
80768 | 208 |
syntax_consts |
209 |
"_peek_res" == peek_res |
|
12854 | 210 |
translations |
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
211 |
"\<lambda>w:. P" == "CONST peek_res (\<lambda>w. P)" |
12854 | 212 |
|
213 |
lemma peek_res_def2 [simp]: "peek_res P Y = P Y Y" |
|
214 |
apply (unfold peek_res_def) |
|
215 |
apply (simp (no_asm)) |
|
216 |
done |
|
217 |
||
218 |
lemma peek_res_subst_res [simp]: "peek_res P\<leftarrow>w = P w\<leftarrow>w" |
|
219 |
apply (rule ext) |
|
220 |
apply (simp (no_asm)) |
|
221 |
done |
|
222 |
||
223 |
(* unused *) |
|
224 |
lemma peek_subst_res_allI: |
|
225 |
"(\<And>a. T a (P (f a)\<leftarrow>f a)) \<Longrightarrow> \<forall>a. T a (peek_res P\<leftarrow>f a)" |
|
226 |
apply (rule allI) |
|
227 |
apply (simp (no_asm)) |
|
228 |
apply fast |
|
229 |
done |
|
230 |
||
231 |
subsection "ign-res" |
|
232 |
||
37956 | 233 |
definition |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
234 |
ign_res :: "'a assn \<Rightarrow> 'a assn" (\<open>_\<down>\<close> [1000] 1000) |
37956 | 235 |
where "P\<down> = (\<lambda>Y s Z. \<exists>Y. P Y s Z)" |
12854 | 236 |
|
237 |
lemma ign_res_def2 [simp]: "P\<down> Y s Z = (\<exists>Y. P Y s Z)" |
|
238 |
apply (unfold ign_res_def) |
|
239 |
apply (simp (no_asm)) |
|
240 |
done |
|
241 |
||
242 |
lemma ign_ign_res [simp]: "P\<down>\<down> = P\<down>" |
|
243 |
apply (rule ext) |
|
244 |
apply (rule ext) |
|
245 |
apply (rule ext) |
|
246 |
apply (simp (no_asm)) |
|
247 |
done |
|
248 |
||
249 |
lemma ign_subst_res [simp]: "P\<down>\<leftarrow>w = P\<down>" |
|
250 |
apply (rule ext) |
|
251 |
apply (rule ext) |
|
252 |
apply (rule ext) |
|
253 |
apply (simp (no_asm)) |
|
254 |
done |
|
255 |
||
256 |
lemma peek_and_ign_res [simp]: "(P \<and>. p)\<down> = (P\<down> \<and>. p)" |
|
257 |
apply (rule ext) |
|
258 |
apply (rule ext) |
|
259 |
apply (rule ext) |
|
260 |
apply (simp (no_asm)) |
|
261 |
done |
|
262 |
||
263 |
subsection "peek-st" |
|
264 |
||
37956 | 265 |
definition |
266 |
peek_st :: "(st \<Rightarrow> 'a assn) \<Rightarrow> 'a assn" |
|
267 |
where "peek_st P = (\<lambda>Y s. P (store s) Y s)" |
|
12854 | 268 |
|
269 |
syntax |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
270 |
"_peek_st" :: "pttrn \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (\<open>\<lambda>_.. _\<close> [0,3] 3) |
80768 | 271 |
syntax_consts |
272 |
"_peek_st" == peek_st |
|
12854 | 273 |
translations |
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
274 |
"\<lambda>s.. P" == "CONST peek_st (\<lambda>s. P)" |
12854 | 275 |
|
276 |
lemma peek_st_def2 [simp]: "(\<lambda>s.. Pf s) Y s = Pf (store s) Y s" |
|
277 |
apply (unfold peek_st_def) |
|
278 |
apply (simp (no_asm)) |
|
279 |
done |
|
280 |
||
281 |
lemma peek_st_triv [simp]: "(\<lambda>s.. P) = P" |
|
282 |
apply (rule ext) |
|
283 |
apply (rule ext) |
|
284 |
apply (simp (no_asm)) |
|
285 |
done |
|
286 |
||
287 |
lemma peek_st_st [simp]: "(\<lambda>s.. \<lambda>s'.. P s s') = (\<lambda>s.. P s s)" |
|
288 |
apply (rule ext) |
|
289 |
apply (rule ext) |
|
290 |
apply (simp (no_asm)) |
|
291 |
done |
|
292 |
||
293 |
lemma peek_st_split [simp]: "(\<lambda>s.. \<lambda>Y s'. P s Y s') = (\<lambda>Y s. P (store s) Y s)" |
|
294 |
apply (rule ext) |
|
295 |
apply (rule ext) |
|
296 |
apply (simp (no_asm)) |
|
297 |
done |
|
298 |
||
299 |
lemma peek_st_subst_res [simp]: "(\<lambda>s.. P s)\<leftarrow>w = (\<lambda>s.. P s\<leftarrow>w)" |
|
300 |
apply (rule ext) |
|
301 |
apply (simp (no_asm)) |
|
302 |
done |
|
303 |
||
304 |
lemma peek_st_Normal [simp]: "(\<lambda>s..(Normal (P s))) = Normal (\<lambda>s.. P s)" |
|
305 |
apply (rule ext) |
|
306 |
apply (rule ext) |
|
307 |
apply (simp (no_asm)) |
|
308 |
done |
|
309 |
||
310 |
subsection "ign-res-eq" |
|
311 |
||
37956 | 312 |
definition |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
313 |
ign_res_eq :: "'a assn \<Rightarrow> res \<Rightarrow> 'a assn" (\<open>_\<down>=_\<close> [60,61] 60) |
37956 | 314 |
where "P\<down>=w \<equiv> (\<lambda>Y:. P\<down> \<and>. (\<lambda>s. Y=w))" |
12854 | 315 |
|
316 |
lemma ign_res_eq_def2 [simp]: "(P\<down>=w) Y s Z = ((\<exists>Y. P Y s Z) \<and> Y=w)" |
|
317 |
apply (unfold ign_res_eq_def) |
|
318 |
apply auto |
|
319 |
done |
|
320 |
||
321 |
lemma ign_ign_res_eq [simp]: "(P\<down>=w)\<down> = P\<down>" |
|
322 |
apply (rule ext) |
|
323 |
apply (rule ext) |
|
324 |
apply (rule ext) |
|
325 |
apply (simp (no_asm)) |
|
326 |
done |
|
327 |
||
328 |
(* unused *) |
|
329 |
lemma ign_res_eq_subst_res: "P\<down>=w\<leftarrow>w = P\<down>" |
|
330 |
apply (rule ext) |
|
331 |
apply (rule ext) |
|
332 |
apply (rule ext) |
|
333 |
apply (simp (no_asm)) |
|
334 |
done |
|
335 |
||
336 |
(* unused *) |
|
337 |
lemma subst_Bool_ign_res_eq: "((P\<leftarrow>=b)\<down>=x) Y s Z = ((P\<leftarrow>=b) Y s Z \<and> Y=x)" |
|
338 |
apply (simp (no_asm)) |
|
339 |
done |
|
340 |
||
341 |
subsection "RefVar" |
|
342 |
||
37956 | 343 |
definition |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
344 |
RefVar :: "(state \<Rightarrow> vvar \<times> state) \<Rightarrow> 'a assn \<Rightarrow> 'a assn" (infixr \<open>..;\<close> 13) |
37956 | 345 |
where "(vf ..; P) = (\<lambda>Y s. let (v,s') = vf s in P (Var v) s')" |
12854 | 346 |
|
347 |
lemma RefVar_def2 [simp]: "(vf ..; P) Y s = |
|
348 |
P (Var (fst (vf s))) (snd (vf s))" |
|
349 |
apply (unfold RefVar_def Let_def) |
|
350 |
apply (simp (no_asm) add: split_beta) |
|
351 |
done |
|
352 |
||
353 |
subsection "allocation" |
|
354 |
||
37956 | 355 |
definition |
356 |
Alloc :: "prog \<Rightarrow> obj_tag \<Rightarrow> 'a assn \<Rightarrow> 'a assn" |
|
357 |
where "Alloc G otag P = (\<lambda>Y s Z. \<forall>s' a. G\<turnstile>s \<midarrow>halloc otag\<succ>a\<rightarrow> s'\<longrightarrow> P (Val (Addr a)) s' Z)" |
|
12854 | 358 |
|
37956 | 359 |
definition |
360 |
SXAlloc :: "prog \<Rightarrow> 'a assn \<Rightarrow> 'a assn" |
|
361 |
where "SXAlloc G P = (\<lambda>Y s Z. \<forall>s'. G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s' \<longrightarrow> P Y s' Z)" |
|
12854 | 362 |
|
363 |
||
364 |
lemma Alloc_def2 [simp]: "Alloc G otag P Y s Z = |
|
365 |
(\<forall>s' a. G\<turnstile>s \<midarrow>halloc otag\<succ>a\<rightarrow> s'\<longrightarrow> P (Val (Addr a)) s' Z)" |
|
366 |
apply (unfold Alloc_def) |
|
367 |
apply (simp (no_asm)) |
|
368 |
done |
|
369 |
||
370 |
lemma SXAlloc_def2 [simp]: |
|
371 |
"SXAlloc G P Y s Z = (\<forall>s'. G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s' \<longrightarrow> P Y s' Z)" |
|
372 |
apply (unfold SXAlloc_def) |
|
373 |
apply (simp (no_asm)) |
|
374 |
done |
|
375 |
||
58887 | 376 |
subsubsection "validity" |
12854 | 377 |
|
37956 | 378 |
definition |
379 |
type_ok :: "prog \<Rightarrow> term \<Rightarrow> state \<Rightarrow> bool" where |
|
380 |
"type_ok G t s = |
|
381 |
(\<exists>L T C A. (normal s \<longrightarrow> \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T \<and> |
|
382 |
\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>t\<guillemotright>A ) |
|
383 |
\<and> s\<Colon>\<preceq>(G,L))" |
|
12854 | 384 |
|
58310 | 385 |
datatype 'a triple = triple "('a assn)" "term" "('a assn)" (** should be |
12854 | 386 |
something like triple = \<forall>'a. triple ('a assn) term ('a assn) **) |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
387 |
(\<open>{(1_)}/ _\<succ>/ {(1_)}\<close> [3,65,3] 75) |
41778 | 388 |
type_synonym 'a triples = "'a triple set" |
12854 | 389 |
|
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
390 |
abbreviation |
12854 | 391 |
var_triple :: "['a assn, var ,'a assn] \<Rightarrow> 'a triple" |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
392 |
(\<open>{(1_)}/ _=\<succ>/ {(1_)}\<close> [3,80,3] 75) |
61989 | 393 |
where "{P} e=\<succ> {Q} == {P} In2 e\<succ> {Q}" |
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
394 |
|
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
395 |
abbreviation |
12854 | 396 |
expr_triple :: "['a assn, expr ,'a assn] \<Rightarrow> 'a triple" |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
397 |
(\<open>{(1_)}/ _-\<succ>/ {(1_)}\<close> [3,80,3] 75) |
61989 | 398 |
where "{P} e-\<succ> {Q} == {P} In1l e\<succ> {Q}" |
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
399 |
|
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
400 |
abbreviation |
12854 | 401 |
exprs_triple :: "['a assn, expr list ,'a assn] \<Rightarrow> 'a triple" |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
402 |
(\<open>{(1_)}/ _\<doteq>\<succ>/ {(1_)}\<close> [3,65,3] 75) |
61989 | 403 |
where "{P} e\<doteq>\<succ> {Q} == {P} In3 e\<succ> {Q}" |
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
404 |
|
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
405 |
abbreviation |
12854 | 406 |
stmt_triple :: "['a assn, stmt, 'a assn] \<Rightarrow> 'a triple" |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
407 |
(\<open>{(1_)}/ ._./ {(1_)}\<close> [3,65,3] 75) |
61989 | 408 |
where "{P} .c. {Q} == {P} In1r c\<succ> {Q}" |
12854 | 409 |
|
61989 | 410 |
notation (ASCII) |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
411 |
triple (\<open>{(1_)}/ _>/ {(1_)}\<close> [3,65,3]75) and |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
412 |
var_triple (\<open>{(1_)}/ _=>/ {(1_)}\<close> [3,80,3] 75) and |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
413 |
expr_triple (\<open>{(1_)}/ _->/ {(1_)}\<close> [3,80,3] 75) and |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
414 |
exprs_triple (\<open>{(1_)}/ _#>/ {(1_)}\<close> [3,65,3] 75) |
12854 | 415 |
|
416 |
lemma inj_triple: "inj (\<lambda>(P,t,Q). {P} t\<succ> {Q})" |
|
13585 | 417 |
apply (rule inj_onI) |
12854 | 418 |
apply auto |
419 |
done |
|
420 |
||
421 |
lemma triple_inj_eq: "({P} t\<succ> {Q} = {P'} t'\<succ> {Q'} ) = (P=P' \<and> t=t' \<and> Q=Q')" |
|
422 |
apply auto |
|
423 |
done |
|
424 |
||
35416
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents:
35069
diff
changeset
|
425 |
definition mtriples :: "('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<Rightarrow> 'sig \<Rightarrow> expr) \<Rightarrow> |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
426 |
('c \<Rightarrow> 'sig \<Rightarrow> 'a assn) \<Rightarrow> ('c \<times> 'sig) set \<Rightarrow> 'a triples" (\<open>{{(1_)}/ _-\<succ>/ {(1_)} | _}\<close>[3,65,3,65]75) where |
37956 | 427 |
"{{P} tf-\<succ> {Q} | ms} = (\<lambda>(C,sig). {Normal(P C sig)} tf C sig-\<succ> {Q C sig})`ms" |
12854 | 428 |
|
37956 | 429 |
definition |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
430 |
triple_valid :: "prog \<Rightarrow> nat \<Rightarrow> 'a triple \<Rightarrow> bool" (\<open>_\<Turnstile>_:_\<close> [61,0, 58] 57) |
37956 | 431 |
where |
432 |
"G\<Turnstile>n:t = |
|
433 |
(case t of {P} t\<succ> {Q} \<Rightarrow> |
|
434 |
\<forall>Y s Z. P Y s Z \<longrightarrow> type_ok G t s \<longrightarrow> |
|
435 |
(\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s') \<longrightarrow> Q Y' s' Z))" |
|
12854 | 436 |
|
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
437 |
abbreviation |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
438 |
triples_valid:: "prog \<Rightarrow> nat \<Rightarrow> 'a triples \<Rightarrow> bool" (\<open>_|\<Turnstile>_:_\<close> [61,0, 58] 57) |
61989 | 439 |
where "G|\<Turnstile>n:ts == Ball ts (triple_valid G n)" |
37956 | 440 |
|
61989 | 441 |
notation (ASCII) |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
442 |
triples_valid ( \<open>_||=_:_\<close> [61,0, 58] 57) |
37956 | 443 |
|
444 |
||
445 |
definition |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
446 |
ax_valids :: "prog \<Rightarrow> 'b triples \<Rightarrow> 'a triples \<Rightarrow> bool" (\<open>_,_|\<Turnstile>_\<close> [61,58,58] 57) |
37956 | 447 |
where "(G,A|\<Turnstile>ts) = (\<forall>n. G|\<Turnstile>n:A \<longrightarrow> G|\<Turnstile>n:ts)" |
12854 | 448 |
|
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
449 |
abbreviation |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
450 |
ax_valid :: "prog \<Rightarrow> 'b triples \<Rightarrow> 'a triple \<Rightarrow> bool" (\<open>_,_\<Turnstile>_\<close> [61,58,58] 57) |
61989 | 451 |
where "G,A \<Turnstile>t == G,A|\<Turnstile>{t}" |
35067
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents:
32960
diff
changeset
|
452 |
|
61989 | 453 |
notation (ASCII) |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
454 |
ax_valid ( \<open>_,_|=_\<close> [61,58,58] 57) |
12854 | 455 |
|
456 |
||
457 |
lemma triple_valid_def2: "G\<Turnstile>n:{P} t\<succ> {Q} = |
|
458 |
(\<forall>Y s Z. P Y s Z |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
459 |
\<longrightarrow> (\<exists>L. (normal s \<longrightarrow> (\<exists> C T A. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T \<and> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
460 |
\<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>dom (locals (store s))\<guillemotright>t\<guillemotright>A)) \<and> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
461 |
s\<Colon>\<preceq>(G,L)) |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
462 |
\<longrightarrow> (\<forall>Y' s'. G\<turnstile>s \<midarrow>t\<succ>\<midarrow>n\<rightarrow> (Y',s')\<longrightarrow> Q Y' s' Z))" |
12854 | 463 |
apply (unfold triple_valid_def type_ok_def) |
464 |
apply (simp (no_asm)) |
|
465 |
done |
|
466 |
||
467 |
||
468 |
declare split_paired_All [simp del] split_paired_Ex [simp del] |
|
62390 | 469 |
declare if_split [split del] if_split_asm [split del] |
12854 | 470 |
option.split [split del] option.split_asm [split del] |
62042 | 471 |
setup \<open>map_theory_simpset (fn ctxt => ctxt delloop "split_all_tac")\<close> |
472 |
setup \<open>map_theory_claset (fn ctxt => ctxt delSWrapper "split_all_tac")\<close> |
|
12854 | 473 |
|
23747 | 474 |
inductive |
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
475 |
ax_derivs :: "prog \<Rightarrow> 'a triples \<Rightarrow> 'a triples \<Rightarrow> bool" (\<open>_,_|\<turnstile>_\<close> [61,58,58] 57) |
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
80768
diff
changeset
|
476 |
and ax_deriv :: "prog \<Rightarrow> 'a triples \<Rightarrow> 'a triple \<Rightarrow> bool" (\<open>_,_\<turnstile>_\<close> [61,58,58] 57) |
21765 | 477 |
for G :: prog |
478 |
where |
|
12854 | 479 |
|
21765 | 480 |
"G,A \<turnstile>t \<equiv> G,A|\<turnstile>{t}" |
481 |
||
482 |
| empty: " G,A|\<turnstile>{}" |
|
483 |
| insert:"\<lbrakk>G,A\<turnstile>t; G,A|\<turnstile>ts\<rbrakk> \<Longrightarrow> |
|
12854 | 484 |
G,A|\<turnstile>insert t ts" |
485 |
||
21765 | 486 |
| asm: "ts\<subseteq>A \<Longrightarrow> G,A|\<turnstile>ts" |
12854 | 487 |
|
488 |
(* could be added for convenience and efficiency, but is not necessary |
|
489 |
cut: "\<lbrakk>G,A'|\<turnstile>ts; G,A|\<turnstile>A'\<rbrakk> \<Longrightarrow> |
|
490 |
G,A |\<turnstile>ts" |
|
491 |
*) |
|
21765 | 492 |
| weaken:"\<lbrakk>G,A|\<turnstile>ts'; ts \<subseteq> ts'\<rbrakk> \<Longrightarrow> G,A|\<turnstile>ts" |
12854 | 493 |
|
21765 | 494 |
| conseq:"\<forall>Y s Z . P Y s Z \<longrightarrow> (\<exists>P' Q'. G,A\<turnstile>{P'} t\<succ> {Q'} \<and> (\<forall>Y' s'. |
12854 | 495 |
(\<forall>Y Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow> |
496 |
Q Y' s' Z )) |
|
497 |
\<Longrightarrow> G,A\<turnstile>{P } t\<succ> {Q }" |
|
498 |
||
21765 | 499 |
| hazard:"G,A\<turnstile>{P \<and>. Not \<circ> type_ok G t} t\<succ> {Q}" |
12854 | 500 |
|
28524 | 501 |
| Abrupt: "G,A\<turnstile>{P\<leftarrow>(undefined3 t) \<and>. Not \<circ> normal} t\<succ> {P}" |
12854 | 502 |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67399
diff
changeset
|
503 |
\<comment> \<open>variables\<close> |
21765 | 504 |
| LVar: " G,A\<turnstile>{Normal (\<lambda>s.. P\<leftarrow>Var (lvar vn s))} LVar vn=\<succ> {P}" |
12854 | 505 |
|
21765 | 506 |
| FVar: "\<lbrakk>G,A\<turnstile>{Normal P} .Init C. {Q}; |
12854 | 507 |
G,A\<turnstile>{Q} e-\<succ> {\<lambda>Val:a:. fvar C stat fn a ..; R}\<rbrakk> \<Longrightarrow> |
12925
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12859
diff
changeset
|
508 |
G,A\<turnstile>{Normal P} {accC,C,stat}e..fn=\<succ> {R}" |
12854 | 509 |
|
21765 | 510 |
| AVar: "\<lbrakk>G,A\<turnstile>{Normal P} e1-\<succ> {Q}; |
12854 | 511 |
\<forall>a. G,A\<turnstile>{Q\<leftarrow>Val a} e2-\<succ> {\<lambda>Val:i:. avar G i a ..; R}\<rbrakk> \<Longrightarrow> |
512 |
G,A\<turnstile>{Normal P} e1.[e2]=\<succ> {R}" |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67399
diff
changeset
|
513 |
\<comment> \<open>expressions\<close> |
12854 | 514 |
|
21765 | 515 |
| NewC: "\<lbrakk>G,A\<turnstile>{Normal P} .Init C. {Alloc G (CInst C) Q}\<rbrakk> \<Longrightarrow> |
12854 | 516 |
G,A\<turnstile>{Normal P} NewC C-\<succ> {Q}" |
517 |
||
21765 | 518 |
| NewA: "\<lbrakk>G,A\<turnstile>{Normal P} .init_comp_ty T. {Q}; G,A\<turnstile>{Q} e-\<succ> |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
28524
diff
changeset
|
519 |
{\<lambda>Val:i:. abupd (check_neg i) .; Alloc G (Arr T (the_Intg i)) R}\<rbrakk> \<Longrightarrow> |
12854 | 520 |
G,A\<turnstile>{Normal P} New T[e]-\<succ> {R}" |
521 |
||
21765 | 522 |
| Cast: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:v:. \<lambda>s.. |
12854 | 523 |
abupd (raise_if (\<not>G,s\<turnstile>v fits T) ClassCast) .; Q\<leftarrow>Val v}\<rbrakk> \<Longrightarrow> |
524 |
G,A\<turnstile>{Normal P} Cast T e-\<succ> {Q}" |
|
525 |
||
21765 | 526 |
| Inst: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:v:. \<lambda>s.. |
12854 | 527 |
Q\<leftarrow>Val (Bool (v\<noteq>Null \<and> G,s\<turnstile>v fits RefT T))}\<rbrakk> \<Longrightarrow> |
528 |
G,A\<turnstile>{Normal P} e InstOf T-\<succ> {Q}" |
|
529 |
||
21765 | 530 |
| Lit: "G,A\<turnstile>{Normal (P\<leftarrow>Val v)} Lit v-\<succ> {P}" |
12854 | 531 |
|
21765 | 532 |
| UnOp: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:v:. Q\<leftarrow>Val (eval_unop unop v)}\<rbrakk> |
13337
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents:
12925
diff
changeset
|
533 |
\<Longrightarrow> |
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents:
12925
diff
changeset
|
534 |
G,A\<turnstile>{Normal P} UnOp unop e-\<succ> {Q}" |
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents:
12925
diff
changeset
|
535 |
|
21765 | 536 |
| BinOp: |
13337
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents:
12925
diff
changeset
|
537 |
"\<lbrakk>G,A\<turnstile>{Normal P} e1-\<succ> {Q}; |
13384 | 538 |
\<forall>v1. G,A\<turnstile>{Q\<leftarrow>Val v1} |
539 |
(if need_second_arg binop v1 then (In1l e2) else (In1r Skip))\<succ> |
|
540 |
{\<lambda>Val:v2:. R\<leftarrow>Val (eval_binop binop v1 v2)}\<rbrakk> |
|
13337
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents:
12925
diff
changeset
|
541 |
\<Longrightarrow> |
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents:
12925
diff
changeset
|
542 |
G,A\<turnstile>{Normal P} BinOp binop e1 e2-\<succ> {R}" |
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents:
12925
diff
changeset
|
543 |
|
21765 | 544 |
| Super:" G,A\<turnstile>{Normal (\<lambda>s.. P\<leftarrow>Val (val_this s))} Super-\<succ> {P}" |
12854 | 545 |
|
21765 | 546 |
| Acc: "\<lbrakk>G,A\<turnstile>{Normal P} va=\<succ> {\<lambda>Var:(v,f):. Q\<leftarrow>Val v}\<rbrakk> \<Longrightarrow> |
12854 | 547 |
G,A\<turnstile>{Normal P} Acc va-\<succ> {Q}" |
548 |
||
21765 | 549 |
| Ass: "\<lbrakk>G,A\<turnstile>{Normal P} va=\<succ> {Q}; |
12854 | 550 |
\<forall>vf. G,A\<turnstile>{Q\<leftarrow>Var vf} e-\<succ> {\<lambda>Val:v:. assign (snd vf) v .; R}\<rbrakk> \<Longrightarrow> |
551 |
G,A\<turnstile>{Normal P} va:=e-\<succ> {R}" |
|
552 |
||
21765 | 553 |
| Cond: "\<lbrakk>G,A \<turnstile>{Normal P} e0-\<succ> {P'}; |
12854 | 554 |
\<forall>b. G,A\<turnstile>{P'\<leftarrow>=b} (if b then e1 else e2)-\<succ> {Q}\<rbrakk> \<Longrightarrow> |
555 |
G,A\<turnstile>{Normal P} e0 ? e1 : e2-\<succ> {Q}" |
|
556 |
||
21765 | 557 |
| Call: |
12854 | 558 |
"\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {Q}; \<forall>a. G,A\<turnstile>{Q\<leftarrow>Val a} args\<doteq>\<succ> {R a}; |
559 |
\<forall>a vs invC declC l. G,A\<turnstile>{(R a\<leftarrow>Vals vs \<and>. |
|
560 |
(\<lambda>s. declC=invocation_declclass G mode (store s) a statT \<lparr>name=mn,parTs=pTs\<rparr> \<and> |
|
561 |
invC = invocation_class mode (store s) a statT \<and> |
|
562 |
l = locals (store s)) ;. |
|
563 |
init_lvars G declC \<lparr>name=mn,parTs=pTs\<rparr> mode a vs) \<and>. |
|
564 |
(\<lambda>s. normal s \<longrightarrow> G\<turnstile>mode\<rightarrow>invC\<preceq>statT)} |
|
565 |
Methd declC \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> {set_lvars l .; S}\<rbrakk> \<Longrightarrow> |
|
12925
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12859
diff
changeset
|
566 |
G,A\<turnstile>{Normal P} {accC,statT,mode}e\<cdot>mn({pTs}args)-\<succ> {S}" |
12854 | 567 |
|
21765 | 568 |
| Methd:"\<lbrakk>G,A\<union> {{P} Methd-\<succ> {Q} | ms} |\<turnstile> {{P} body G-\<succ> {Q} | ms}\<rbrakk> \<Longrightarrow> |
12854 | 569 |
G,A|\<turnstile>{{P} Methd-\<succ> {Q} | ms}" |
570 |
||
21765 | 571 |
| Body: "\<lbrakk>G,A\<turnstile>{Normal P} .Init D. {Q}; |
12854 | 572 |
G,A\<turnstile>{Q} .c. {\<lambda>s.. abupd (absorb Ret) .; R\<leftarrow>(In1 (the (locals s Result)))}\<rbrakk> |
573 |
\<Longrightarrow> |
|
574 |
G,A\<turnstile>{Normal P} Body D c-\<succ> {R}" |
|
575 |
||
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67399
diff
changeset
|
576 |
\<comment> \<open>expression lists\<close> |
12854 | 577 |
|
21765 | 578 |
| Nil: "G,A\<turnstile>{Normal (P\<leftarrow>Vals [])} []\<doteq>\<succ> {P}" |
12854 | 579 |
|
21765 | 580 |
| Cons: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {Q}; |
12854 | 581 |
\<forall>v. G,A\<turnstile>{Q\<leftarrow>Val v} es\<doteq>\<succ> {\<lambda>Vals:vs:. R\<leftarrow>Vals (v#vs)}\<rbrakk> \<Longrightarrow> |
582 |
G,A\<turnstile>{Normal P} e#es\<doteq>\<succ> {R}" |
|
583 |
||
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67399
diff
changeset
|
584 |
\<comment> \<open>statements\<close> |
12854 | 585 |
|
21765 | 586 |
| Skip: "G,A\<turnstile>{Normal (P\<leftarrow>\<diamondsuit>)} .Skip. {P}" |
12854 | 587 |
|
21765 | 588 |
| Expr: "\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {Q\<leftarrow>\<diamondsuit>}\<rbrakk> \<Longrightarrow> |
12854 | 589 |
G,A\<turnstile>{Normal P} .Expr e. {Q}" |
590 |
||
21765 | 591 |
| Lab: "\<lbrakk>G,A\<turnstile>{Normal P} .c. {abupd (absorb l) .; Q}\<rbrakk> \<Longrightarrow> |
12854 | 592 |
G,A\<turnstile>{Normal P} .l\<bullet> c. {Q}" |
593 |
||
21765 | 594 |
| Comp: "\<lbrakk>G,A\<turnstile>{Normal P} .c1. {Q}; |
12854 | 595 |
G,A\<turnstile>{Q} .c2. {R}\<rbrakk> \<Longrightarrow> |
596 |
G,A\<turnstile>{Normal P} .c1;;c2. {R}" |
|
597 |
||
21765 | 598 |
| If: "\<lbrakk>G,A \<turnstile>{Normal P} e-\<succ> {P'}; |
12854 | 599 |
\<forall>b. G,A\<turnstile>{P'\<leftarrow>=b} .(if b then c1 else c2). {Q}\<rbrakk> \<Longrightarrow> |
600 |
G,A\<turnstile>{Normal P} .If(e) c1 Else c2. {Q}" |
|
601 |
(* unfolding variant of Loop, not needed here |
|
602 |
LoopU:"\<lbrakk>G,A \<turnstile>{Normal P} e-\<succ> {P'}; |
|
603 |
\<forall>b. G,A\<turnstile>{P'\<leftarrow>=b} .(if b then c;;While(e) c else Skip).{Q}\<rbrakk> |
|
604 |
\<Longrightarrow> G,A\<turnstile>{Normal P} .While(e) c. {Q}" |
|
605 |
*) |
|
21765 | 606 |
| Loop: "\<lbrakk>G,A\<turnstile>{P} e-\<succ> {P'}; |
12854 | 607 |
G,A\<turnstile>{Normal (P'\<leftarrow>=True)} .c. {abupd (absorb (Cont l)) .; P}\<rbrakk> \<Longrightarrow> |
608 |
G,A\<turnstile>{P} .l\<bullet> While(e) c. {(P'\<leftarrow>=False)\<down>=\<diamondsuit>}" |
|
609 |
||
21765 | 610 |
| Jmp: "G,A\<turnstile>{Normal (abupd (\<lambda>a. (Some (Jump j))) .; P\<leftarrow>\<diamondsuit>)} .Jmp j. {P}" |
12854 | 611 |
|
21765 | 612 |
| Throw:"\<lbrakk>G,A\<turnstile>{Normal P} e-\<succ> {\<lambda>Val:a:. abupd (throw a) .; Q\<leftarrow>\<diamondsuit>}\<rbrakk> \<Longrightarrow> |
12854 | 613 |
G,A\<turnstile>{Normal P} .Throw e. {Q}" |
614 |
||
21765 | 615 |
| Try: "\<lbrakk>G,A\<turnstile>{Normal P} .c1. {SXAlloc G Q}; |
12854 | 616 |
G,A\<turnstile>{Q \<and>. (\<lambda>s. G,s\<turnstile>catch C) ;. new_xcpt_var vn} .c2. {R}; |
617 |
(Q \<and>. (\<lambda>s. \<not>G,s\<turnstile>catch C)) \<Rightarrow> R\<rbrakk> \<Longrightarrow> |
|
618 |
G,A\<turnstile>{Normal P} .Try c1 Catch(C vn) c2. {R}" |
|
619 |
||
21765 | 620 |
| Fin: "\<lbrakk>G,A\<turnstile>{Normal P} .c1. {Q}; |
12854 | 621 |
\<forall>x. G,A\<turnstile>{Q \<and>. (\<lambda>s. x = fst s) ;. abupd (\<lambda>x. None)} |
622 |
.c2. {abupd (abrupt_if (x\<noteq>None) x) .; R}\<rbrakk> \<Longrightarrow> |
|
623 |
G,A\<turnstile>{Normal P} .c1 Finally c2. {R}" |
|
624 |
||
21765 | 625 |
| Done: "G,A\<turnstile>{Normal (P\<leftarrow>\<diamondsuit> \<and>. initd C)} .Init C. {P}" |
12854 | 626 |
|
21765 | 627 |
| Init: "\<lbrakk>the (class G C) = c; |
12854 | 628 |
G,A\<turnstile>{Normal ((P \<and>. Not \<circ> initd C) ;. supd (init_class_obj G C))} |
629 |
.(if C = Object then Skip else Init (super c)). {Q}; |
|
68451 | 630 |
\<forall>l. G,A\<turnstile>{Q \<and>. (\<lambda>s. l = locals (store s)) ;. set_lvars Map.empty} |
12854 | 631 |
.init c. {set_lvars l .; R}\<rbrakk> \<Longrightarrow> |
632 |
G,A\<turnstile>{Normal (P \<and>. Not \<circ> initd C)} .Init C. {R}" |
|
633 |
||
62042 | 634 |
\<comment> \<open>Some dummy rules for the intermediate terms \<open>Callee\<close>, |
635 |
\<open>InsInitE\<close>, \<open>InsInitV\<close>, \<open>FinA\<close> only used by the smallstep |
|
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
67399
diff
changeset
|
636 |
semantics.\<close> |
21765 | 637 |
| InsInitV: " G,A\<turnstile>{Normal P} InsInitV c v=\<succ> {Q}" |
638 |
| InsInitE: " G,A\<turnstile>{Normal P} InsInitE c e-\<succ> {Q}" |
|
639 |
| Callee: " G,A\<turnstile>{Normal P} Callee l e-\<succ> {Q}" |
|
640 |
| FinA: " G,A\<turnstile>{Normal P} .FinA a c. {Q}" |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
641 |
(* |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
642 |
axioms |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
643 |
*) |
12854 | 644 |
|
37956 | 645 |
definition |
646 |
adapt_pre :: "'a assn \<Rightarrow> 'a assn \<Rightarrow> 'a assn \<Rightarrow> 'a assn" |
|
647 |
where "adapt_pre P Q Q' = (\<lambda>Y s Z. \<forall>Y' s'. \<exists>Z'. P Y s Z' \<and> (Q Y' s' Z' \<longrightarrow> Q' Y' s' Z))" |
|
12854 | 648 |
|
649 |
||
58887 | 650 |
subsubsection "rules derived by induction" |
12854 | 651 |
|
652 |
lemma cut_valid: "\<lbrakk>G,A'|\<Turnstile>ts; G,A|\<Turnstile>A'\<rbrakk> \<Longrightarrow> G,A|\<Turnstile>ts" |
|
653 |
apply (unfold ax_valids_def) |
|
654 |
apply fast |
|
655 |
done |
|
656 |
||
657 |
(*if cut is available |
|
658 |
Goal "\<lbrakk>G,A'|\<turnstile>ts; A' \<subseteq> A; \<forall>P Q t. {P} t\<succ> {Q} \<in> A' \<longrightarrow> (\<exists>T. (G,L)\<turnstile>t\<Colon>T) \<rbrakk> \<Longrightarrow> |
|
659 |
G,A|\<turnstile>ts" |
|
660 |
b y etac ax_derivs.cut 1; |
|
661 |
b y eatac ax_derivs.asm 1 1; |
|
662 |
qed "ax_thin"; |
|
663 |
*) |
|
664 |
lemma ax_thin [rule_format (no_asm)]: |
|
665 |
"G,(A'::'a triple set)|\<turnstile>(ts::'a triple set) \<Longrightarrow> \<forall>A. A' \<subseteq> A \<longrightarrow> G,A|\<turnstile>ts" |
|
666 |
apply (erule ax_derivs.induct) |
|
69597 | 667 |
apply (tactic "ALLGOALS (EVERY'[clarify_tac \<^context>, REPEAT o smp_tac \<^context> 1])") |
12854 | 668 |
apply (rule ax_derivs.empty) |
669 |
apply (erule (1) ax_derivs.insert) |
|
670 |
apply (fast intro: ax_derivs.asm) |
|
671 |
(*apply (fast intro: ax_derivs.cut) *) |
|
672 |
apply (fast intro: ax_derivs.weaken) |
|
69597 | 673 |
apply (rule ax_derivs.conseq, intro strip, tactic "smp_tac \<^context> 3 1",clarify, |
674 |
tactic "smp_tac \<^context> 1 1",rule exI, rule exI, erule (1) conjI) |
|
13337
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents:
12925
diff
changeset
|
675 |
(* 37 subgoals *) |
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents:
12925
diff
changeset
|
676 |
prefer 18 (* Methd *) |
23563 | 677 |
apply (rule ax_derivs.Methd, drule spec, erule mp, fast) |
69597 | 678 |
apply (tactic \<open>TRYALL (resolve_tac \<^context> ((funpow 5 tl) @{thms ax_derivs.intros}))\<close>) |
23563 | 679 |
apply auto |
12854 | 680 |
done |
681 |
||
682 |
lemma ax_thin_insert: "G,(A::'a triple set)\<turnstile>(t::'a triple) \<Longrightarrow> G,insert x A\<turnstile>t" |
|
683 |
apply (erule ax_thin) |
|
684 |
apply fast |
|
685 |
done |
|
686 |
||
687 |
lemma subset_mtriples_iff: |
|
688 |
"ts \<subseteq> {{P} mb-\<succ> {Q} | ms} = (\<exists>ms'. ms'\<subseteq>ms \<and> ts = {{P} mb-\<succ> {Q} | ms'})" |
|
689 |
apply (unfold mtriples_def) |
|
690 |
apply (rule subset_image_iff) |
|
691 |
done |
|
692 |
||
693 |
lemma weaken: |
|
67613 | 694 |
"G,(A::'a triple set)|\<turnstile>(ts'::'a triple set) \<Longrightarrow> \<forall>ts. ts \<subseteq> ts' \<longrightarrow> G,A|\<turnstile>ts" |
12854 | 695 |
apply (erule ax_derivs.induct) |
13337
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents:
12925
diff
changeset
|
696 |
(*42 subgoals*) |
69597 | 697 |
apply (tactic "ALLGOALS (strip_tac \<^context>)") |
698 |
apply (tactic \<open>ALLGOALS(REPEAT o (EVERY'[dresolve_tac \<^context> @{thms subset_singletonD}, |
|
699 |
eresolve_tac \<^context> [disjE], |
|
700 |
fast_tac (\<^context> addSIs @{thms ax_derivs.empty})]))\<close>) |
|
701 |
apply (tactic "TRYALL (hyp_subst_tac \<^context>)") |
|
12854 | 702 |
apply (simp, rule ax_derivs.empty) |
703 |
apply (drule subset_insertD) |
|
704 |
apply (blast intro: ax_derivs.insert) |
|
705 |
apply (fast intro: ax_derivs.asm) |
|
706 |
(*apply (blast intro: ax_derivs.cut) *) |
|
707 |
apply (fast intro: ax_derivs.weaken) |
|
69597 | 708 |
apply (rule ax_derivs.conseq, clarify, tactic "smp_tac \<^context> 3 1", blast(* unused *)) |
13337
f75dfc606ac7
Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents:
12925
diff
changeset
|
709 |
(*37 subgoals*) |
69597 | 710 |
apply (tactic \<open>TRYALL (resolve_tac \<^context> ((funpow 5 tl) @{thms ax_derivs.intros}) |
711 |
THEN_ALL_NEW fast_tac \<^context>)\<close>) |
|
12854 | 712 |
(*1 subgoal*) |
713 |
apply (clarsimp simp add: subset_mtriples_iff) |
|
714 |
apply (rule ax_derivs.Methd) |
|
715 |
apply (drule spec) |
|
716 |
apply (erule impE) |
|
717 |
apply (rule exI) |
|
718 |
apply (erule conjI) |
|
719 |
apply (rule HOL.refl) |
|
720 |
oops (* dead end, Methd is to blame *) |
|
721 |
||
722 |
||
58887 | 723 |
subsubsection "rules derived from conseq" |
12854 | 724 |
|
62042 | 725 |
text \<open>In the following rules we often have to give some type annotations like: |
69597 | 726 |
\<^term>\<open>G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q}\<close>. |
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
727 |
Given only the term above without annotations, Isabelle would infer a more |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
728 |
general type were we could have |
69597 | 729 |
different types of auxiliary variables in the assumption set (\<^term>\<open>A\<close>) and |
730 |
in the triple itself (\<^term>\<open>P\<close> and \<^term>\<open>Q\<close>). But |
|
62042 | 731 |
\<open>ax_derivs.Methd\<close> enforces the same type in the inductive definition of |
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
732 |
the derivation. So we have to restrict the types to be able to apply the |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
733 |
rules. |
62042 | 734 |
\<close> |
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
735 |
lemma conseq12: "\<lbrakk>G,(A::'a triple set)\<turnstile>{P'::'a assn} t\<succ> {Q'}; |
12854 | 736 |
\<forall>Y s Z. P Y s Z \<longrightarrow> (\<forall>Y' s'. (\<forall>Y Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow> |
737 |
Q Y' s' Z)\<rbrakk> |
|
738 |
\<Longrightarrow> G,A\<turnstile>{P ::'a assn} t\<succ> {Q }" |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
739 |
apply (rule ax_derivs.conseq) |
12854 | 740 |
apply clarsimp |
741 |
apply blast |
|
742 |
done |
|
743 |
||
62042 | 744 |
\<comment> \<open>Nice variant, since it is so symmetric we might be able to memorise it.\<close> |
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
745 |
lemma conseq12': "\<lbrakk>G,(A::'a triple set)\<turnstile>{P'::'a assn} t\<succ> {Q'}; \<forall>s Y' s'. |
12854 | 746 |
(\<forall>Y Z. P' Y s Z \<longrightarrow> Q' Y' s' Z) \<longrightarrow> |
747 |
(\<forall>Y Z. P Y s Z \<longrightarrow> Q Y' s' Z)\<rbrakk> |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
748 |
\<Longrightarrow> G,A\<turnstile>{P::'a assn } t\<succ> {Q }" |
12854 | 749 |
apply (erule conseq12) |
750 |
apply fast |
|
751 |
done |
|
752 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
753 |
lemma conseq12_from_conseq12': "\<lbrakk>G,(A::'a triple set)\<turnstile>{P'::'a assn} t\<succ> {Q'}; |
12854 | 754 |
\<forall>Y s Z. P Y s Z \<longrightarrow> (\<forall>Y' s'. (\<forall>Y Z'. P' Y s Z' \<longrightarrow> Q' Y' s' Z') \<longrightarrow> |
755 |
Q Y' s' Z)\<rbrakk> |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
756 |
\<Longrightarrow> G,A\<turnstile>{P::'a assn} t\<succ> {Q }" |
12854 | 757 |
apply (erule conseq12') |
758 |
apply blast |
|
759 |
done |
|
760 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
761 |
lemma conseq1: "\<lbrakk>G,(A::'a triple set)\<turnstile>{P'::'a assn} t\<succ> {Q}; P \<Rightarrow> P'\<rbrakk> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
762 |
\<Longrightarrow> G,A\<turnstile>{P::'a assn} t\<succ> {Q}" |
12854 | 763 |
apply (erule conseq12) |
764 |
apply blast |
|
765 |
done |
|
766 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
767 |
lemma conseq2: "\<lbrakk>G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q'}; Q' \<Rightarrow> Q\<rbrakk> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
768 |
\<Longrightarrow> G,A\<turnstile>{P::'a assn} t\<succ> {Q}" |
12854 | 769 |
apply (erule conseq12) |
770 |
apply blast |
|
771 |
done |
|
772 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
773 |
lemma ax_escape: |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
774 |
"\<lbrakk>\<forall>Y s Z. P Y s Z |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
775 |
\<longrightarrow> G,(A::'a triple set)\<turnstile>{\<lambda>Y' s' (Z'::'a). (Y',s') = (Y,s)} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
776 |
t\<succ> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
777 |
{\<lambda>Y s Z'. Q Y s Z} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
778 |
\<rbrakk> \<Longrightarrow> G,A\<turnstile>{P::'a assn} t\<succ> {Q::'a assn}" |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
779 |
apply (rule ax_derivs.conseq) |
12854 | 780 |
apply force |
781 |
done |
|
782 |
||
783 |
(* unused *) |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
784 |
lemma ax_constant: "\<lbrakk> C \<Longrightarrow> G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q}\<rbrakk> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
785 |
\<Longrightarrow> G,A\<turnstile>{\<lambda>Y s Z. C \<and> P Y s Z} t\<succ> {Q}" |
12854 | 786 |
apply (rule ax_escape (* unused *)) |
787 |
apply clarify |
|
788 |
apply (rule conseq12) |
|
789 |
apply fast |
|
790 |
apply auto |
|
791 |
done |
|
792 |
(*alternative (more direct) proof: |
|
793 |
apply (rule ax_derivs.conseq) *)(* unused *)(* |
|
794 |
apply (fast) |
|
795 |
*) |
|
796 |
||
797 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
798 |
lemma ax_impossible [intro]: |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
799 |
"G,(A::'a triple set)\<turnstile>{\<lambda>Y s Z. False} t\<succ> {Q::'a assn}" |
12854 | 800 |
apply (rule ax_escape) |
801 |
apply clarify |
|
802 |
done |
|
803 |
||
804 |
(* unused *) |
|
67399 | 805 |
lemma ax_nochange_lemma: "\<lbrakk>P Y s; All ((=) w)\<rbrakk> \<Longrightarrow> P w s" |
12854 | 806 |
apply auto |
807 |
done |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
808 |
|
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
809 |
lemma ax_nochange: |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
810 |
"G,(A::(res \<times> state) triple set)\<turnstile>{\<lambda>Y s Z. (Y,s)=Z} t\<succ> {\<lambda>Y s Z. (Y,s)=Z} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
811 |
\<Longrightarrow> G,A\<turnstile>{P::(res \<times> state) assn} t\<succ> {P}" |
12854 | 812 |
apply (erule conseq12) |
813 |
apply auto |
|
814 |
apply (erule (1) ax_nochange_lemma) |
|
815 |
done |
|
816 |
||
817 |
(* unused *) |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
818 |
lemma ax_trivial: "G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {\<lambda>Y s Z. True}" |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
819 |
apply (rule ax_derivs.conseq(* unused *)) |
12854 | 820 |
apply auto |
821 |
done |
|
822 |
||
823 |
(* unused *) |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
824 |
lemma ax_disj: |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
825 |
"\<lbrakk>G,(A::'a triple set)\<turnstile>{P1::'a assn} t\<succ> {Q1}; G,A\<turnstile>{P2::'a assn} t\<succ> {Q2}\<rbrakk> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
826 |
\<Longrightarrow> G,A\<turnstile>{\<lambda>Y s Z. P1 Y s Z \<or> P2 Y s Z} t\<succ> {\<lambda>Y s Z. Q1 Y s Z \<or> Q2 Y s Z}" |
12854 | 827 |
apply (rule ax_escape (* unused *)) |
828 |
apply safe |
|
829 |
apply (erule conseq12, fast)+ |
|
830 |
done |
|
831 |
||
832 |
(* unused *) |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
833 |
lemma ax_supd_shuffle: |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
834 |
"(\<exists>Q. G,(A::'a triple set)\<turnstile>{P::'a assn} .c1. {Q} \<and> G,A\<turnstile>{Q ;. f} .c2. {R}) = |
12854 | 835 |
(\<exists>Q'. G,A\<turnstile>{P} .c1. {f .; Q'} \<and> G,A\<turnstile>{Q'} .c2. {R})" |
836 |
apply (best elim!: conseq1 conseq2) |
|
837 |
done |
|
838 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
839 |
lemma ax_cases: " |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
840 |
\<lbrakk>G,(A::'a triple set)\<turnstile>{P \<and>. C} t\<succ> {Q::'a assn}; |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
841 |
G,A\<turnstile>{P \<and>. Not \<circ> C} t\<succ> {Q}\<rbrakk> \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}" |
12854 | 842 |
apply (unfold peek_and_def) |
843 |
apply (rule ax_escape) |
|
844 |
apply clarify |
|
845 |
apply (case_tac "C s") |
|
846 |
apply (erule conseq12, force)+ |
|
847 |
done |
|
848 |
(*alternative (more direct) proof: |
|
849 |
apply (rule rtac ax_derivs.conseq) *)(* unused *)(* |
|
850 |
apply clarify |
|
851 |
apply (case_tac "C s") |
|
852 |
apply force+ |
|
853 |
*) |
|
854 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
855 |
lemma ax_adapt: "G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
856 |
\<Longrightarrow> G,A\<turnstile>{adapt_pre P Q Q'} t\<succ> {Q'}" |
12854 | 857 |
apply (unfold adapt_pre_def) |
858 |
apply (erule conseq12) |
|
859 |
apply fast |
|
860 |
done |
|
861 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
862 |
lemma adapt_pre_adapts: "G,(A::'a triple set)\<Turnstile>{P::'a assn} t\<succ> {Q} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
863 |
\<longrightarrow> G,A\<Turnstile>{adapt_pre P Q Q'} t\<succ> {Q'}" |
12854 | 864 |
apply (unfold adapt_pre_def) |
865 |
apply (simp add: ax_valids_def triple_valid_def2) |
|
866 |
apply fast |
|
867 |
done |
|
868 |
||
869 |
||
870 |
lemma adapt_pre_weakest: |
|
871 |
"\<forall>G (A::'a triple set) t. G,A\<Turnstile>{P} t\<succ> {Q} \<longrightarrow> G,A\<Turnstile>{P'} t\<succ> {Q'} \<Longrightarrow> |
|
872 |
P' \<Rightarrow> adapt_pre P Q (Q'::'a assn)" |
|
873 |
apply (unfold adapt_pre_def) |
|
874 |
apply (drule spec) |
|
875 |
apply (drule_tac x = "{}" in spec) |
|
876 |
apply (drule_tac x = "In1r Skip" in spec) |
|
877 |
apply (simp add: ax_valids_def triple_valid_def2) |
|
878 |
oops |
|
879 |
||
880 |
lemma peek_and_forget1_Normal: |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
881 |
"G,(A::'a triple set)\<turnstile>{Normal P} t\<succ> {Q::'a assn} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
882 |
\<Longrightarrow> G,A\<turnstile>{Normal (P \<and>. p)} t\<succ> {Q}" |
12854 | 883 |
apply (erule conseq1) |
884 |
apply (simp (no_asm)) |
|
885 |
done |
|
886 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
887 |
lemma peek_and_forget1: |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
888 |
"G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
889 |
\<Longrightarrow> G,A\<turnstile>{P \<and>. p} t\<succ> {Q}" |
12854 | 890 |
apply (erule conseq1) |
891 |
apply (simp (no_asm)) |
|
892 |
done |
|
893 |
||
894 |
lemmas ax_NormalD = peek_and_forget1 [of _ _ _ _ _ normal] |
|
895 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
896 |
lemma peek_and_forget2: |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
897 |
"G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> {Q \<and>. p} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
898 |
\<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}" |
12854 | 899 |
apply (erule conseq2) |
900 |
apply (simp (no_asm)) |
|
901 |
done |
|
902 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
903 |
lemma ax_subst_Val_allI: |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
904 |
"\<forall>v. G,(A::'a triple set)\<turnstile>{(P' v )\<leftarrow>Val v} t\<succ> {(Q v)::'a assn} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
905 |
\<Longrightarrow> \<forall>v. G,A\<turnstile>{(\<lambda>w:. P' (the_In1 w))\<leftarrow>Val v} t\<succ> {Q v}" |
12854 | 906 |
apply (force elim!: conseq1) |
907 |
done |
|
908 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
909 |
lemma ax_subst_Var_allI: |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
910 |
"\<forall>v. G,(A::'a triple set)\<turnstile>{(P' v )\<leftarrow>Var v} t\<succ> {(Q v)::'a assn} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
911 |
\<Longrightarrow> \<forall>v. G,A\<turnstile>{(\<lambda>w:. P' (the_In2 w))\<leftarrow>Var v} t\<succ> {Q v}" |
12854 | 912 |
apply (force elim!: conseq1) |
913 |
done |
|
914 |
||
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
915 |
lemma ax_subst_Vals_allI: |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
916 |
"(\<forall>v. G,(A::'a triple set)\<turnstile>{( P' v )\<leftarrow>Vals v} t\<succ> {(Q v)::'a assn}) |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
917 |
\<Longrightarrow> \<forall>v. G,A\<turnstile>{(\<lambda>w:. P' (the_In3 w))\<leftarrow>Vals v} t\<succ> {Q v}" |
12854 | 918 |
apply (force elim!: conseq1) |
919 |
done |
|
920 |
||
921 |
||
58887 | 922 |
subsubsection "alternative axioms" |
12854 | 923 |
|
924 |
lemma ax_Lit2: |
|
925 |
"G,(A::'a triple set)\<turnstile>{Normal P::'a assn} Lit v-\<succ> {Normal (P\<down>=Val v)}" |
|
926 |
apply (rule ax_derivs.Lit [THEN conseq1]) |
|
927 |
apply force |
|
928 |
done |
|
929 |
lemma ax_Lit2_test_complete: |
|
930 |
"G,(A::'a triple set)\<turnstile>{Normal (P\<leftarrow>Val v)::'a assn} Lit v-\<succ> {P}" |
|
931 |
apply (rule ax_Lit2 [THEN conseq2]) |
|
932 |
apply force |
|
933 |
done |
|
934 |
||
935 |
lemma ax_LVar2: "G,(A::'a triple set)\<turnstile>{Normal P::'a assn} LVar vn=\<succ> {Normal (\<lambda>s.. P\<down>=Var (lvar vn s))}" |
|
936 |
apply (rule ax_derivs.LVar [THEN conseq1]) |
|
937 |
apply force |
|
938 |
done |
|
939 |
||
940 |
lemma ax_Super2: "G,(A::'a triple set)\<turnstile> |
|
941 |
{Normal P::'a assn} Super-\<succ> {Normal (\<lambda>s.. P\<down>=Val (val_this s))}" |
|
942 |
apply (rule ax_derivs.Super [THEN conseq1]) |
|
943 |
apply force |
|
944 |
done |
|
945 |
||
946 |
lemma ax_Nil2: |
|
947 |
"G,(A::'a triple set)\<turnstile>{Normal P::'a assn} []\<doteq>\<succ> {Normal (P\<down>=Vals [])}" |
|
948 |
apply (rule ax_derivs.Nil [THEN conseq1]) |
|
949 |
apply force |
|
950 |
done |
|
951 |
||
952 |
||
58887 | 953 |
subsubsection "misc derived structural rules" |
12854 | 954 |
|
955 |
(* unused *) |
|
956 |
lemma ax_finite_mtriples_lemma: "\<lbrakk>F \<subseteq> ms; finite ms; \<forall>(C,sig)\<in>ms. |
|
957 |
G,(A::'a triple set)\<turnstile>{Normal (P C sig)::'a assn} mb C sig-\<succ> {Q C sig}\<rbrakk> \<Longrightarrow> |
|
958 |
G,A|\<turnstile>{{P} mb-\<succ> {Q} | F}" |
|
959 |
apply (frule (1) finite_subset) |
|
24038 | 960 |
apply (erule rev_mp) |
12854 | 961 |
apply (erule thin_rl) |
962 |
apply (erule finite_induct) |
|
963 |
apply (unfold mtriples_def) |
|
964 |
apply (clarsimp intro!: ax_derivs.empty ax_derivs.insert)+ |
|
965 |
apply force |
|
966 |
done |
|
967 |
lemmas ax_finite_mtriples = ax_finite_mtriples_lemma [OF subset_refl] |
|
968 |
||
969 |
lemma ax_derivs_insertD: |
|
970 |
"G,(A::'a triple set)|\<turnstile>insert (t::'a triple) ts \<Longrightarrow> G,A\<turnstile>t \<and> G,A|\<turnstile>ts" |
|
971 |
apply (fast intro: ax_derivs.weaken) |
|
972 |
done |
|
973 |
||
974 |
lemma ax_methods_spec: |
|
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
60754
diff
changeset
|
975 |
"\<lbrakk>G,(A::'a triple set)|\<turnstile>case_prod f ` ms; (C,sig) \<in> ms\<rbrakk>\<Longrightarrow> G,A\<turnstile>((f C sig)::'a triple)" |
12854 | 976 |
apply (erule ax_derivs.weaken) |
977 |
apply (force del: image_eqI intro: rev_image_eqI) |
|
978 |
done |
|
979 |
||
980 |
(* this version is used to avoid using the cut rule *) |
|
981 |
lemma ax_finite_pointwise_lemma [rule_format]: "\<lbrakk>F \<subseteq> ms; finite ms\<rbrakk> \<Longrightarrow> |
|
982 |
((\<forall>(C,sig)\<in>F. G,(A::'a triple set)\<turnstile>(f C sig::'a triple)) \<longrightarrow> (\<forall>(C,sig)\<in>ms. G,A\<turnstile>(g C sig::'a triple))) \<longrightarrow> |
|
61424
c3658c18b7bc
prod_case as canonical name for product type eliminator
haftmann
parents:
60754
diff
changeset
|
983 |
G,A|\<turnstile>case_prod f ` F \<longrightarrow> G,A|\<turnstile>case_prod g ` F" |
12854 | 984 |
apply (frule (1) finite_subset) |
24038 | 985 |
apply (erule rev_mp) |
12854 | 986 |
apply (erule thin_rl) |
987 |
apply (erule finite_induct) |
|
988 |
apply clarsimp+ |
|
989 |
apply (drule ax_derivs_insertD) |
|
990 |
apply (rule ax_derivs.insert) |
|
991 |
apply (simp (no_asm_simp) only: split_tupled_all) |
|
992 |
apply (auto elim: ax_methods_spec) |
|
993 |
done |
|
994 |
lemmas ax_finite_pointwise = ax_finite_pointwise_lemma [OF subset_refl] |
|
995 |
||
996 |
lemma ax_no_hazard: |
|
997 |
"G,(A::'a triple set)\<turnstile>{P \<and>. type_ok G t} t\<succ> {Q::'a assn} \<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Q}" |
|
998 |
apply (erule ax_cases) |
|
999 |
apply (rule ax_derivs.hazard [THEN conseq1]) |
|
1000 |
apply force |
|
1001 |
done |
|
1002 |
||
1003 |
lemma ax_free_wt: |
|
1004 |
"(\<exists>T L C. \<lparr>prg=G,cls=C,lcl=L\<rparr>\<turnstile>t\<Colon>T) |
|
1005 |
\<longrightarrow> G,(A::'a triple set)\<turnstile>{Normal P} t\<succ> {Q::'a assn} \<Longrightarrow> |
|
1006 |
G,A\<turnstile>{Normal P} t\<succ> {Q}" |
|
1007 |
apply (rule ax_no_hazard) |
|
1008 |
apply (rule ax_escape) |
|
1009 |
apply clarify |
|
1010 |
apply (erule mp [THEN conseq12]) |
|
1011 |
apply (auto simp add: type_ok_def) |
|
1012 |
done |
|
1013 |
||
69597 | 1014 |
ML \<open>ML_Thms.bind_thms ("ax_Abrupts", sum3_instantiate \<^context> @{thm ax_derivs.Abrupt})\<close> |
12854 | 1015 |
declare ax_Abrupts [intro!] |
1016 |
||
21765 | 1017 |
lemmas ax_Normal_cases = ax_cases [of _ _ _ normal] |
12854 | 1018 |
|
1019 |
lemma ax_Skip [intro!]: "G,(A::'a triple set)\<turnstile>{P\<leftarrow>\<diamondsuit>} .Skip. {P::'a assn}" |
|
1020 |
apply (rule ax_Normal_cases) |
|
1021 |
apply (rule ax_derivs.Skip) |
|
1022 |
apply fast |
|
1023 |
done |
|
45605 | 1024 |
lemmas ax_SkipI = ax_Skip [THEN conseq1] |
12854 | 1025 |
|
1026 |
||
58887 | 1027 |
subsubsection "derived rules for methd call" |
12854 | 1028 |
|
1029 |
lemma ax_Call_known_DynT: |
|
1030 |
"\<lbrakk>G\<turnstile>IntVir\<rightarrow>C\<preceq>statT; |
|
1031 |
\<forall>a vs l. G,A\<turnstile>{(R a\<leftarrow>Vals vs \<and>. (\<lambda>s. l = locals (store s)) ;. |
|
1032 |
init_lvars G C \<lparr>name=mn,parTs=pTs\<rparr> IntVir a vs)} |
|
1033 |
Methd C \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> {set_lvars l .; S}; |
|
1034 |
\<forall>a. G,A\<turnstile>{Q\<leftarrow>Val a} args\<doteq>\<succ> |
|
1035 |
{R a \<and>. (\<lambda>s. C = obj_class (the (heap (store s) (the_Addr a))) \<and> |
|
1036 |
C = invocation_declclass |
|
1037 |
G IntVir (store s) a statT \<lparr>name=mn,parTs=pTs\<rparr> )}; |
|
1038 |
G,(A::'a triple set)\<turnstile>{Normal P} e-\<succ> {Q::'a assn}\<rbrakk> |
|
12925
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12859
diff
changeset
|
1039 |
\<Longrightarrow> G,A\<turnstile>{Normal P} {accC,statT,IntVir}e\<cdot>mn({pTs}args)-\<succ> {S}" |
12854 | 1040 |
apply (erule ax_derivs.Call) |
1041 |
apply safe |
|
1042 |
apply (erule spec) |
|
1043 |
apply (rule ax_escape, clarsimp) |
|
1044 |
apply (drule spec, drule spec, drule spec,erule conseq12) |
|
1045 |
apply force |
|
1046 |
done |
|
1047 |
||
1048 |
||
1049 |
lemma ax_Call_Static: |
|
1050 |
"\<lbrakk>\<forall>a vs l. G,A\<turnstile>{R a\<leftarrow>Vals vs \<and>. (\<lambda>s. l = locals (store s)) ;. |
|
1051 |
init_lvars G C \<lparr>name=mn,parTs=pTs\<rparr> Static any_Addr vs} |
|
1052 |
Methd C \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> {set_lvars l .; S}; |
|
1053 |
G,A\<turnstile>{Normal P} e-\<succ> {Q}; |
|
1054 |
\<forall> a. G,(A::'a triple set)\<turnstile>{Q\<leftarrow>Val a} args\<doteq>\<succ> {(R::val \<Rightarrow> 'a assn) a |
|
1055 |
\<and>. (\<lambda> s. C=invocation_declclass |
|
1056 |
G Static (store s) a statT \<lparr>name=mn,parTs=pTs\<rparr>)} |
|
12925
99131847fb93
Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents:
12859
diff
changeset
|
1057 |
\<rbrakk> \<Longrightarrow> G,A\<turnstile>{Normal P} {accC,statT,Static}e\<cdot>mn({pTs}args)-\<succ> {S}" |
12854 | 1058 |
apply (erule ax_derivs.Call) |
1059 |
apply safe |
|
1060 |
apply (erule spec) |
|
1061 |
apply (rule ax_escape, clarsimp) |
|
59807 | 1062 |
apply (erule_tac V = "P \<longrightarrow> Q" for P Q in thin_rl) |
12854 | 1063 |
apply (drule spec,drule spec,drule spec, erule conseq12) |
20014 | 1064 |
apply (force simp add: init_lvars_def Let_def) |
12854 | 1065 |
done |
1066 |
||
1067 |
lemma ax_Methd1: |
|
1068 |
"\<lbrakk>G,A\<union>{{P} Methd-\<succ> {Q} | ms}|\<turnstile> {{P} body G-\<succ> {Q} | ms}; (C,sig)\<in> ms\<rbrakk> \<Longrightarrow> |
|
1069 |
G,A\<turnstile>{Normal (P C sig)} Methd C sig-\<succ> {Q C sig}" |
|
1070 |
apply (drule ax_derivs.Methd) |
|
1071 |
apply (unfold mtriples_def) |
|
1072 |
apply (erule (1) ax_methods_spec) |
|
1073 |
done |
|
1074 |
||
1075 |
lemma ax_MethdN: |
|
1076 |
"G,insert({Normal P} Methd C sig-\<succ> {Q}) A\<turnstile> |
|
1077 |
{Normal P} body G C sig-\<succ> {Q} \<Longrightarrow> |
|
1078 |
G,A\<turnstile>{Normal P} Methd C sig-\<succ> {Q}" |
|
1079 |
apply (rule ax_Methd1) |
|
1080 |
apply (rule_tac [2] singletonI) |
|
1081 |
apply (unfold mtriples_def) |
|
1082 |
apply clarsimp |
|
1083 |
done |
|
1084 |
||
1085 |
lemma ax_StatRef: |
|
1086 |
"G,(A::'a triple set)\<turnstile>{Normal (P\<leftarrow>Val Null)} StatRef rt-\<succ> {P::'a assn}" |
|
1087 |
apply (rule ax_derivs.Cast) |
|
1088 |
apply (rule ax_Lit2 [THEN conseq2]) |
|
1089 |
apply clarsimp |
|
1090 |
done |
|
1091 |
||
58887 | 1092 |
subsubsection "rules derived from Init and Done" |
12854 | 1093 |
|
1094 |
lemma ax_InitS: "\<lbrakk>the (class G C) = c; C \<noteq> Object; |
|
68451 | 1095 |
\<forall>l. G,A\<turnstile>{Q \<and>. (\<lambda>s. l = locals (store s)) ;. set_lvars Map.empty} |
12854 | 1096 |
.init c. {set_lvars l .; R}; |
1097 |
G,A\<turnstile>{Normal ((P \<and>. Not \<circ> initd C) ;. supd (init_class_obj G C))} |
|
1098 |
.Init (super c). {Q}\<rbrakk> \<Longrightarrow> |
|
1099 |
G,(A::'a triple set)\<turnstile>{Normal (P \<and>. Not \<circ> initd C)} .Init C. {R::'a assn}" |
|
1100 |
apply (erule ax_derivs.Init) |
|
1101 |
apply (simp (no_asm_simp)) |
|
1102 |
apply assumption |
|
1103 |
done |
|
1104 |
||
1105 |
lemma ax_Init_Skip_lemma: |
|
1106 |
"\<forall>l. G,(A::'a triple set)\<turnstile>{P\<leftarrow>\<diamondsuit> \<and>. (\<lambda>s. l = locals (store s)) ;. set_lvars l'} |
|
1107 |
.Skip. {(set_lvars l .; P)::'a assn}" |
|
1108 |
apply (rule allI) |
|
1109 |
apply (rule ax_SkipI) |
|
1110 |
apply clarsimp |
|
1111 |
done |
|
1112 |
||
1113 |
lemma ax_triv_InitS: "\<lbrakk>the (class G C) = c;init c = Skip; C \<noteq> Object; |
|
1114 |
P\<leftarrow>\<diamondsuit> \<Rightarrow> (supd (init_class_obj G C) .; P); |
|
1115 |
G,A\<turnstile>{Normal (P \<and>. initd C)} .Init (super c). {(P \<and>. initd C)\<leftarrow>\<diamondsuit>}\<rbrakk> \<Longrightarrow> |
|
1116 |
G,(A::'a triple set)\<turnstile>{Normal P\<leftarrow>\<diamondsuit>} .Init C. {(P \<and>. initd C)::'a assn}" |
|
1117 |
apply (rule_tac C = "initd C" in ax_cases) |
|
1118 |
apply (rule conseq1, rule ax_derivs.Done, clarsimp) |
|
1119 |
apply (simp (no_asm)) |
|
1120 |
apply (erule (1) ax_InitS) |
|
1121 |
apply simp |
|
1122 |
apply (rule ax_Init_Skip_lemma) |
|
1123 |
apply (erule conseq1) |
|
1124 |
apply force |
|
1125 |
done |
|
1126 |
||
1127 |
lemma ax_Init_Object: "wf_prog G \<Longrightarrow> G,(A::'a triple set)\<turnstile> |
|
1128 |
{Normal ((supd (init_class_obj G Object) .; P\<leftarrow>\<diamondsuit>) \<and>. Not \<circ> initd Object)} |
|
1129 |
.Init Object. {(P \<and>. initd Object)::'a assn}" |
|
1130 |
apply (rule ax_derivs.Init) |
|
1131 |
apply (drule class_Object, force) |
|
1132 |
apply (simp_all (no_asm)) |
|
1133 |
apply (rule_tac [2] ax_Init_Skip_lemma) |
|
1134 |
apply (rule ax_SkipI, force) |
|
1135 |
done |
|
1136 |
||
1137 |
lemma ax_triv_Init_Object: "\<lbrakk>wf_prog G; |
|
1138 |
(P::'a assn) \<Rightarrow> (supd (init_class_obj G Object) .; P)\<rbrakk> \<Longrightarrow> |
|
1139 |
G,(A::'a triple set)\<turnstile>{Normal P\<leftarrow>\<diamondsuit>} .Init Object. {P \<and>. initd Object}" |
|
1140 |
apply (rule_tac C = "initd Object" in ax_cases) |
|
1141 |
apply (rule conseq1, rule ax_derivs.Done, clarsimp) |
|
1142 |
apply (erule ax_Init_Object [THEN conseq1]) |
|
1143 |
apply force |
|
1144 |
done |
|
1145 |
||
1146 |
||
58887 | 1147 |
subsubsection "introduction rules for Alloc and SXAlloc" |
12854 | 1148 |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1149 |
lemma ax_SXAlloc_Normal: |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1150 |
"G,(A::'a triple set)\<turnstile>{P::'a assn} .c. {Normal Q} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1151 |
\<Longrightarrow> G,A\<turnstile>{P} .c. {SXAlloc G Q}" |
12854 | 1152 |
apply (erule conseq2) |
1153 |
apply (clarsimp elim!: sxalloc_elim_cases simp add: split_tupled_all) |
|
1154 |
done |
|
1155 |
||
1156 |
lemma ax_Alloc: |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1157 |
"G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1158 |
{Normal (\<lambda>Y (x,s) Z. (\<forall>a. new_Addr (heap s) = Some a \<longrightarrow> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1159 |
Q (Val (Addr a)) (Norm(init_obj G (CInst C) (Heap a) s)) Z)) \<and>. |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1160 |
heap_free (Suc (Suc 0))} |
12854 | 1161 |
\<Longrightarrow> G,A\<turnstile>{P} t\<succ> {Alloc G (CInst C) Q}" |
1162 |
apply (erule conseq2) |
|
1163 |
apply (auto elim!: halloc_elim_cases) |
|
1164 |
done |
|
1165 |
||
1166 |
lemma ax_Alloc_Arr: |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1167 |
"G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1168 |
{\<lambda>Val:i:. Normal (\<lambda>Y (x,s) Z. \<not>the_Intg i<0 \<and> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1169 |
(\<forall>a. new_Addr (heap s) = Some a \<longrightarrow> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1170 |
Q (Val (Addr a)) (Norm (init_obj G (Arr T (the_Intg i)) (Heap a) s)) Z)) \<and>. |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1171 |
heap_free (Suc (Suc 0))} |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1172 |
\<Longrightarrow> |
12854 | 1173 |
G,A\<turnstile>{P} t\<succ> {\<lambda>Val:i:. abupd (check_neg i) .; Alloc G (Arr T(the_Intg i)) Q}" |
1174 |
apply (erule conseq2) |
|
1175 |
apply (auto elim!: halloc_elim_cases) |
|
1176 |
done |
|
1177 |
||
1178 |
lemma ax_SXAlloc_catch_SXcpt: |
|
13688
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1179 |
"\<lbrakk>G,(A::'a triple set)\<turnstile>{P::'a assn} t\<succ> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1180 |
{(\<lambda>Y (x,s) Z. x=Some (Xcpt (Std xn)) \<and> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1181 |
(\<forall>a. new_Addr (heap s) = Some a \<longrightarrow> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1182 |
Q Y (Some (Xcpt (Loc a)),init_obj G (CInst (SXcpt xn)) (Heap a) s) Z)) |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1183 |
\<and>. heap_free (Suc (Suc 0))}\<rbrakk> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1184 |
\<Longrightarrow> |
a0b16d42d489
"Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents:
13585
diff
changeset
|
1185 |
G,A\<turnstile>{P} t\<succ> {SXAlloc G (\<lambda>Y s Z. Q Y s Z \<and> G,s\<turnstile>catch SXcpt xn)}" |
12854 | 1186 |
apply (erule conseq2) |
1187 |
apply (auto elim!: sxalloc_elim_cases halloc_elim_cases) |
|
1188 |
done |
|
1189 |
||
1190 |
end |