| author | wenzelm | 
| Wed, 14 Mar 2012 15:37:51 +0100 | |
| changeset 46920 | 5f44c8bea84e | 
| parent 39758 | b8a53e3a0ee2 | 
| child 60304 | 3f429b7d8eb5 | 
| permissions | -rw-r--r-- | 
| 13673 | 1  | 
(* Title: HOL/MicroJava/Comp/TranslCompTp.thy  | 
2  | 
Author: Martin Strecker  | 
|
3  | 
*)  | 
|
4  | 
||
| 15481 | 5  | 
theory TranslCompTp  | 
6  | 
imports Index "../BV/JVMType"  | 
|
7  | 
begin  | 
|
8  | 
||
| 13673 | 9  | 
(**********************************************************************)  | 
10  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
11  | 
definition comb :: "['a \<Rightarrow> 'b list \<times> 'c, 'c \<Rightarrow> 'b list \<times> 'd, 'a] \<Rightarrow> 'b list \<times> 'd" where  | 
| 13673 | 12  | 
"comb == (\<lambda> f1 f2 x0. let (xs1, x1) = f1 x0;  | 
13  | 
(xs2, x2) = f2 x1  | 
|
14  | 
in (xs1 @ xs2, x2))"  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
15  | 
|
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
16  | 
definition comb_nil :: "'a \<Rightarrow> 'b list \<times> 'a" where  | 
| 13673 | 17  | 
"comb_nil a == ([], a)"  | 
18  | 
||
| 
35355
 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 
wenzelm 
parents: 
35102 
diff
changeset
 | 
19  | 
notation (xsymbols)  | 
| 
 
613e133966ea
modernized syntax declarations, and make them actually work with authentic syntax;
 
wenzelm 
parents: 
35102 
diff
changeset
 | 
20  | 
comb (infixr "\<box>" 55)  | 
| 13673 | 21  | 
|
22  | 
lemma comb_nil_left [simp]: "comb_nil \<box> f = f"  | 
|
23  | 
by (simp add: comb_def comb_nil_def split_beta)  | 
|
24  | 
||
25  | 
lemma comb_nil_right [simp]: "f \<box> comb_nil = f"  | 
|
26  | 
by (simp add: comb_def comb_nil_def split_beta)  | 
|
27  | 
||
28  | 
lemma comb_assoc [simp]: "(fa \<box> fb) \<box> fc = fa \<box> (fb \<box> fc)"  | 
|
29  | 
by (simp add: comb_def split_beta)  | 
|
30  | 
||
31  | 
lemma comb_inv: "(xs', x') = (f1 \<box> f2) x0 \<Longrightarrow>  | 
|
32  | 
\<exists> xs1 x1 xs2 x2. (xs1, x1) = (f1 x0) \<and> (xs2, x2) = f2 x1 \<and> xs'= xs1 @ xs2 \<and> x'=x2"  | 
|
33  | 
apply (case_tac "f1 x0")  | 
|
34  | 
apply (case_tac "f2 x1")  | 
|
35  | 
apply (simp add: comb_def split_beta)  | 
|
36  | 
done  | 
|
37  | 
||
38  | 
(**********************************************************************)  | 
|
39  | 
||
| 35102 | 40  | 
abbreviation (input)  | 
41  | 
mt_of :: "method_type \<times> state_type \<Rightarrow> method_type"  | 
|
42  | 
where "mt_of == fst"  | 
|
| 13673 | 43  | 
|
| 35102 | 44  | 
abbreviation (input)  | 
45  | 
sttp_of :: "method_type \<times> state_type \<Rightarrow> state_type"  | 
|
46  | 
where "sttp_of == snd"  | 
|
| 13673 | 47  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
48  | 
definition nochangeST :: "state_type \<Rightarrow> method_type \<times> state_type" where  | 
| 13673 | 49  | 
"nochangeST sttp == ([Some sttp], sttp)"  | 
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
50  | 
|
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
51  | 
definition pushST :: "[ty list, state_type] \<Rightarrow> method_type \<times> state_type" where  | 
| 13673 | 52  | 
"pushST tps == (\<lambda> (ST, LT). ([Some (ST, LT)], (tps @ ST, LT)))"  | 
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
53  | 
|
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
54  | 
definition dupST :: "state_type \<Rightarrow> method_type \<times> state_type" where  | 
| 13673 | 55  | 
"dupST == (\<lambda> (ST, LT). ([Some (ST, LT)], (hd ST # ST, LT)))"  | 
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
56  | 
|
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
57  | 
definition dup_x1ST :: "state_type \<Rightarrow> method_type \<times> state_type" where  | 
| 13673 | 58  | 
"dup_x1ST == (\<lambda> (ST, LT). ([Some (ST, LT)],  | 
59  | 
(hd ST # hd (tl ST) # hd ST # (tl (tl ST)), LT)))"  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
60  | 
|
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
61  | 
definition popST :: "[nat, state_type] \<Rightarrow> method_type \<times> state_type" where  | 
| 13673 | 62  | 
"popST n == (\<lambda> (ST, LT). ([Some (ST, LT)], (drop n ST, LT)))"  | 
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
63  | 
|
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
64  | 
definition replST :: "[nat, ty, state_type] \<Rightarrow> method_type \<times> state_type" where  | 
| 13673 | 65  | 
"replST n tp == (\<lambda> (ST, LT). ([Some (ST, LT)], (tp # (drop n ST), LT)))"  | 
66  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
67  | 
definition storeST :: "[nat, ty, state_type] \<Rightarrow> method_type \<times> state_type" where  | 
| 13673 | 68  | 
"storeST i tp == (\<lambda> (ST, LT). ([Some (ST, LT)], (tl ST, LT [i:= OK tp])))"  | 
69  | 
||
70  | 
||
71  | 
(* Expressions *)  | 
|
72  | 
||
| 39758 | 73  | 
primrec compTpExpr :: "java_mb \<Rightarrow> java_mb prog \<Rightarrow> expr \<Rightarrow>  | 
74  | 
state_type \<Rightarrow> method_type \<times> state_type"  | 
|
75  | 
and compTpExprs :: "java_mb \<Rightarrow> java_mb prog \<Rightarrow> expr list \<Rightarrow>  | 
|
76  | 
state_type \<Rightarrow> method_type \<times> state_type"  | 
|
77  | 
where  | 
|
| 13673 | 78  | 
"compTpExpr jmb G (NewC c) = pushST [Class c]"  | 
| 39758 | 79  | 
| "compTpExpr jmb G (Cast c e) = (compTpExpr jmb G e) \<box> (replST 1 (Class c))"  | 
80  | 
| "compTpExpr jmb G (Lit val) = pushST [the (typeof (\<lambda>v. None) val)]"  | 
|
81  | 
| "compTpExpr jmb G (BinOp bo e1 e2) =  | 
|
| 13673 | 82  | 
(compTpExpr jmb G e1) \<box> (compTpExpr jmb G e2) \<box>  | 
83  | 
(case bo of  | 
|
84  | 
Eq => popST 2 \<box> pushST [PrimT Boolean] \<box> popST 1 \<box> pushST [PrimT Boolean]  | 
|
85  | 
| Add => replST 2 (PrimT Integer))"  | 
|
| 39758 | 86  | 
| "compTpExpr jmb G (LAcc vn) = (\<lambda> (ST, LT).  | 
| 13673 | 87  | 
pushST [ok_val (LT ! (index jmb vn))] (ST, LT))"  | 
| 39758 | 88  | 
| "compTpExpr jmb G (vn::=e) =  | 
| 13673 | 89  | 
(compTpExpr jmb G e) \<box> dupST \<box> (popST 1)"  | 
| 39758 | 90  | 
| "compTpExpr jmb G ( {cn}e..fn ) = 
 | 
| 13673 | 91  | 
(compTpExpr jmb G e) \<box> replST 1 (snd (the (field (G,cn) fn)))"  | 
| 39758 | 92  | 
| "compTpExpr jmb G (FAss cn e1 fn e2 ) =  | 
| 13673 | 93  | 
(compTpExpr jmb G e1) \<box> (compTpExpr jmb G e2) \<box> dup_x1ST \<box> (popST 2)"  | 
| 39758 | 94  | 
| "compTpExpr jmb G ({C}a..mn({fpTs}ps)) =
 | 
| 13673 | 95  | 
(compTpExpr jmb G a) \<box> (compTpExprs jmb G ps) \<box>  | 
96  | 
(replST ((length ps) + 1) (method_rT (the (method (G,C) (mn,fpTs)))))"  | 
|
| 39758 | 97  | 
| "compTpExprs jmb G [] = comb_nil"  | 
98  | 
| "compTpExprs jmb G (e#es) = (compTpExpr jmb G e) \<box> (compTpExprs jmb G es)"  | 
|
| 13673 | 99  | 
|
100  | 
||
101  | 
(* Statements *)  | 
|
| 39758 | 102  | 
primrec compTpStmt :: "java_mb \<Rightarrow> java_mb prog \<Rightarrow> stmt \<Rightarrow>  | 
103  | 
state_type \<Rightarrow> method_type \<times> state_type"  | 
|
104  | 
where  | 
|
105  | 
"compTpStmt jmb G Skip = comb_nil"  | 
|
106  | 
| "compTpStmt jmb G (Expr e) = (compTpExpr jmb G e) \<box> popST 1"  | 
|
107  | 
| "compTpStmt jmb G (c1;; c2) = (compTpStmt jmb G c1) \<box> (compTpStmt jmb G c2)"  | 
|
108  | 
| "compTpStmt jmb G (If(e) c1 Else c2) =  | 
|
| 13673 | 109  | 
(pushST [PrimT Boolean]) \<box> (compTpExpr jmb G e) \<box> popST 2 \<box>  | 
110  | 
(compTpStmt jmb G c1) \<box> nochangeST \<box> (compTpStmt jmb G c2)"  | 
|
| 39758 | 111  | 
| "compTpStmt jmb G (While(e) c) =  | 
| 13673 | 112  | 
(pushST [PrimT Boolean]) \<box> (compTpExpr jmb G e) \<box> popST 2 \<box>  | 
113  | 
(compTpStmt jmb G c) \<box> nochangeST"  | 
|
114  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
115  | 
definition compTpInit :: "java_mb \<Rightarrow> (vname * ty)  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
116  | 
\<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type" where  | 
| 13673 | 117  | 
"compTpInit jmb == (\<lambda> (vn,ty). (pushST [ty]) \<box> (storeST (index jmb vn) ty))"  | 
118  | 
||
| 39758 | 119  | 
primrec compTpInitLvars :: "[java_mb, (vname \<times> ty) list] \<Rightarrow>  | 
120  | 
state_type \<Rightarrow> method_type \<times> state_type"  | 
|
121  | 
where  | 
|
| 13673 | 122  | 
"compTpInitLvars jmb [] = comb_nil"  | 
| 39758 | 123  | 
| "compTpInitLvars jmb (lv#lvars) = (compTpInit jmb lv) \<box> (compTpInitLvars jmb lvars)"  | 
| 13673 | 124  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
125  | 
definition start_ST :: "opstack_type" where  | 
| 13673 | 126  | 
"start_ST == []"  | 
127  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
128  | 
definition start_LT :: "cname \<Rightarrow> ty list \<Rightarrow> nat \<Rightarrow> locvars_type" where  | 
| 13673 | 129  | 
"start_LT C pTs n == (OK (Class C))#((map OK pTs))@(replicate n Err)"  | 
130  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
131  | 
definition compTpMethod :: "[java_mb prog, cname, java_mb mdecl] \<Rightarrow> method_type" where  | 
| 13673 | 132  | 
"compTpMethod G C == \<lambda> ((mn,pTs),rT, jmb).  | 
133  | 
let (pns,lvars,blk,res) = jmb  | 
|
134  | 
in (mt_of  | 
|
135  | 
((compTpInitLvars jmb lvars \<box>  | 
|
136  | 
compTpStmt jmb G blk \<box>  | 
|
137  | 
compTpExpr jmb G res \<box>  | 
|
138  | 
nochangeST)  | 
|
139  | 
(start_ST, start_LT C pTs (length lvars))))"  | 
|
140  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
141  | 
definition compTp :: "java_mb prog => prog_type" where  | 
| 13673 | 142  | 
"compTp G C sig == let (D, rT, jmb) = (the (method (G, C) sig))  | 
143  | 
in compTpMethod G C (sig, rT, jmb)"  | 
|
144  | 
||
145  | 
||
146  | 
||
147  | 
(**********************************************************************)  | 
|
148  | 
(* Computing the maximum stack size from the method_type *)  | 
|
149  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
150  | 
definition ssize_sto :: "(state_type option) \<Rightarrow> nat" where  | 
| 13673 | 151  | 
"ssize_sto sto == case sto of None \<Rightarrow> 0 | (Some (ST, LT)) \<Rightarrow> length ST"  | 
152  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
153  | 
definition max_of_list :: "nat list \<Rightarrow> nat" where  | 
| 13673 | 154  | 
"max_of_list xs == foldr max xs 0"  | 
155  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35102 
diff
changeset
 | 
156  | 
definition max_ssize :: "method_type \<Rightarrow> nat" where  | 
| 13673 | 157  | 
"max_ssize mt == max_of_list (map ssize_sto mt)"  | 
158  | 
||
159  | 
end  |