| author | berghofe | 
| Tue, 24 Apr 2007 15:17:22 +0200 | |
| changeset 22781 | 18fbba942a80 | 
| parent 16417 | 9bc16273c2d4 | 
| child 23755 | 1c4672d130b1 | 
| permissions | -rw-r--r-- | 
| 11376 | 1  | 
(* Title: HOL/NanoJava/OpSem.thy  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: David von Oheimb  | 
|
4  | 
Copyright 2001 Technische Universitaet Muenchen  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
header "Operational Evaluation Semantics"  | 
|
8  | 
||
| 16417 | 9  | 
theory OpSem imports State begin  | 
| 11376 | 10  | 
|
11  | 
consts  | 
|
| 11476 | 12  | 
exec :: "(state \<times> stmt \<times> nat \<times> state) set"  | 
13  | 
eval :: "(state \<times> expr \<times> val \<times> nat \<times> state) set"  | 
|
| 11376 | 14  | 
syntax (xsymbols)  | 
| 11508 | 15  | 
 exec :: "[state,stmt,    nat,state] => bool" ("_ -_-_\<rightarrow> _"  [98,90,   65,98] 89)
 | 
16  | 
 eval :: "[state,expr,val,nat,state] => bool" ("_ -_\<succ>_-_\<rightarrow> _"[98,95,99,65,98] 89)
 | 
|
| 11376 | 17  | 
syntax  | 
| 11508 | 18  | 
 exec :: "[state,stmt,    nat,state] => bool" ("_ -_-_-> _"  [98,90,   65,98]89)
 | 
19  | 
 eval :: "[state,expr,val,nat,state] => bool" ("_ -_>_-_-> _"[98,95,99,65,98]89)
 | 
|
| 11376 | 20  | 
translations  | 
| 11476 | 21  | 
"s -c -n-> s'" == "(s, c, n, s') \<in> exec"  | 
22  | 
"s -e>v-n-> s'" == "(s, e, v, n, s') \<in> eval"  | 
|
| 11376 | 23  | 
|
| 11476 | 24  | 
inductive exec eval intros  | 
| 11376 | 25  | 
Skip: " s -Skip-n-> s"  | 
26  | 
||
27  | 
Comp: "[| s0 -c1-n-> s1; s1 -c2-n-> s2 |] ==>  | 
|
28  | 
s0 -c1;; c2-n-> s2"  | 
|
29  | 
||
| 11476 | 30  | 
Cond: "[| s0 -e>v-n-> s1; s1 -(if v\<noteq>Null then c1 else c2)-n-> s2 |] ==>  | 
31  | 
s0 -If(e) c1 Else c2-n-> s2"  | 
|
32  | 
||
33  | 
LoopF:" s0<x> = Null ==>  | 
|
34  | 
s0 -While(x) c-n-> s0"  | 
|
35  | 
LoopT:"[| s0<x> \<noteq> Null; s0 -c-n-> s1; s1 -While(x) c-n-> s2 |] ==>  | 
|
36  | 
s0 -While(x) c-n-> s2"  | 
|
| 11376 | 37  | 
|
| 11476 | 38  | 
LAcc: " s -LAcc x>s<x>-n-> s"  | 
39  | 
||
40  | 
LAss: " s -e>v-n-> s' ==>  | 
|
41  | 
s -x:==e-n-> lupd(x\<mapsto>v) s'"  | 
|
42  | 
||
43  | 
FAcc: " s -e>Addr a-n-> s' ==>  | 
|
44  | 
s -e..f>get_field s' a f-n-> s'"  | 
|
45  | 
||
46  | 
FAss: "[| s0 -e1>Addr a-n-> s1; s1 -e2>v-n-> s2 |] ==>  | 
|
47  | 
s0 -e1..f:==e2-n-> upd_obj a f v s2"  | 
|
| 11376 | 48  | 
|
49  | 
NewC: " new_Addr s = Addr a ==>  | 
|
| 11476 | 50  | 
s -new C>Addr a-n-> new_obj a C s"  | 
| 11376 | 51  | 
|
| 11476 | 52  | 
Cast: "[| s -e>v-n-> s';  | 
53  | 
case v of Null => True | Addr a => obj_class s' a \<preceq>C C |] ==>  | 
|
54  | 
s -Cast C e>v-n-> s'"  | 
|
| 11376 | 55  | 
|
| 11476 | 56  | 
Call: "[| s0 -e1>a-n-> s1; s1 -e2>p-n-> s2;  | 
| 11772 | 57  | 
lupd(This\<mapsto>a)(lupd(Par\<mapsto>p)(del_locs s2)) -Meth (C,m)-n-> s3  | 
| 11476 | 58  | 
     |] ==> s0 -{C}e1..m(e2)>s3<Res>-n-> set_locs s2 s3"
 | 
| 11376 | 59  | 
|
| 
11497
 
0e66e0114d9a
corrected initialization of locals, streamlined Impl
 
oheimb 
parents: 
11486 
diff
changeset
 | 
60  | 
Meth: "[| s<This> = Addr a; D = obj_class s a; D\<preceq>C C;  | 
| 
 
0e66e0114d9a
corrected initialization of locals, streamlined Impl
 
oheimb 
parents: 
11486 
diff
changeset
 | 
61  | 
init_locs D m s -Impl (D,m)-n-> s' |] ==>  | 
| 11507 | 62  | 
s -Meth (C,m)-n-> s'"  | 
| 11376 | 63  | 
|
| 11508 | 64  | 
Impl: " s -body Cm- n-> s' ==>  | 
65  | 
s -Impl Cm-Suc n-> s'"  | 
|
| 11376 | 66  | 
|
| 11476 | 67  | 
|
| 11376 | 68  | 
inductive_cases exec_elim_cases':  | 
| 11486 | 69  | 
"s -Skip -n\<rightarrow> t"  | 
70  | 
"s -c1;; c2 -n\<rightarrow> t"  | 
|
71  | 
"s -If(e) c1 Else c2-n\<rightarrow> t"  | 
|
72  | 
"s -While(x) c -n\<rightarrow> t"  | 
|
73  | 
"s -x:==e -n\<rightarrow> t"  | 
|
74  | 
"s -e1..f:==e2 -n\<rightarrow> t"  | 
|
| 11508 | 75  | 
inductive_cases Meth_elim_cases: "s -Meth Cm -n\<rightarrow> t"  | 
76  | 
inductive_cases Impl_elim_cases: "s -Impl Cm -n\<rightarrow> t"  | 
|
| 11376 | 77  | 
lemmas exec_elim_cases = exec_elim_cases' Meth_elim_cases Impl_elim_cases  | 
| 11476 | 78  | 
inductive_cases eval_elim_cases:  | 
| 11486 | 79  | 
"s -new C \<succ>v-n\<rightarrow> t"  | 
80  | 
"s -Cast C e \<succ>v-n\<rightarrow> t"  | 
|
81  | 
"s -LAcc x \<succ>v-n\<rightarrow> t"  | 
|
82  | 
"s -e..f \<succ>v-n\<rightarrow> t"  | 
|
83  | 
				  "s -{C}e1..m(e2)  \<succ>v-n\<rightarrow> t"
 | 
|
| 11376 | 84  | 
|
| 11476 | 85  | 
lemma exec_eval_mono [rule_format]:  | 
86  | 
"(s -c -n\<rightarrow> t \<longrightarrow> (\<forall>m. n \<le> m \<longrightarrow> s -c -m\<rightarrow> t)) \<and>  | 
|
87  | 
(s -e\<succ>v-n\<rightarrow> t \<longrightarrow> (\<forall>m. n \<le> m \<longrightarrow> s -e\<succ>v-m\<rightarrow> t))"  | 
|
88  | 
apply (rule exec_eval.induct)  | 
|
89  | 
prefer 14 (* Impl *)  | 
|
| 11376 | 90  | 
apply clarify  | 
91  | 
apply (rename_tac n)  | 
|
92  | 
apply (case_tac n)  | 
|
| 11508 | 93  | 
apply (blast intro:exec_eval.intros)+  | 
| 11376 | 94  | 
done  | 
| 11476 | 95  | 
lemmas exec_mono = exec_eval_mono [THEN conjunct1, rule_format]  | 
96  | 
lemmas eval_mono = exec_eval_mono [THEN conjunct2, rule_format]  | 
|
97  | 
||
98  | 
lemma exec_exec_max: "\<lbrakk>s1 -c1- n1 \<rightarrow> t1 ; s2 -c2- n2\<rightarrow> t2\<rbrakk> \<Longrightarrow>  | 
|
99  | 
s1 -c1-max n1 n2\<rightarrow> t1 \<and> s2 -c2-max n1 n2\<rightarrow> t2"  | 
|
100  | 
by (fast intro: exec_mono le_maxI1 le_maxI2)  | 
|
| 11376 | 101  | 
|
| 11476 | 102  | 
lemma eval_exec_max: "\<lbrakk>s1 -c- n1 \<rightarrow> t1 ; s2 -e\<succ>v- n2\<rightarrow> t2\<rbrakk> \<Longrightarrow>  | 
103  | 
s1 -c-max n1 n2\<rightarrow> t1 \<and> s2 -e\<succ>v-max n1 n2\<rightarrow> t2"  | 
|
104  | 
by (fast intro: eval_mono exec_mono le_maxI1 le_maxI2)  | 
|
105  | 
||
106  | 
lemma eval_eval_max: "\<lbrakk>s1 -e1\<succ>v1- n1 \<rightarrow> t1 ; s2 -e2\<succ>v2- n2\<rightarrow> t2\<rbrakk> \<Longrightarrow>  | 
|
107  | 
s1 -e1\<succ>v1-max n1 n2\<rightarrow> t1 \<and> s2 -e2\<succ>v2-max n1 n2\<rightarrow> t2"  | 
|
108  | 
by (fast intro: eval_mono le_maxI1 le_maxI2)  | 
|
109  | 
||
110  | 
lemma eval_eval_exec_max:  | 
|
111  | 
"\<lbrakk>s1 -e1\<succ>v1-n1\<rightarrow> t1; s2 -e2\<succ>v2-n2\<rightarrow> t2; s3 -c-n3\<rightarrow> t3\<rbrakk> \<Longrightarrow>  | 
|
112  | 
s1 -e1\<succ>v1-max (max n1 n2) n3\<rightarrow> t1 \<and>  | 
|
113  | 
s2 -e2\<succ>v2-max (max n1 n2) n3\<rightarrow> t2 \<and>  | 
|
114  | 
s3 -c -max (max n1 n2) n3\<rightarrow> t3"  | 
|
115  | 
apply (drule (1) eval_eval_max, erule thin_rl)  | 
|
116  | 
by (fast intro: exec_mono eval_mono le_maxI1 le_maxI2)  | 
|
| 11376 | 117  | 
|
| 11565 | 118  | 
lemma Impl_body_eq: "(\<lambda>t. \<exists>n. Z -Impl M-n\<rightarrow> t) = (\<lambda>t. \<exists>n. Z -body M-n\<rightarrow> t)"  | 
| 11376 | 119  | 
apply (rule ext)  | 
| 11476 | 120  | 
apply (fast elim: exec_elim_cases intro: exec_eval.Impl)  | 
| 11376 | 121  | 
done  | 
122  | 
||
123  | 
end  |