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