src/HOL/MicroJava/Comp/TranslCompTp.thy
author wenzelm
Thu, 11 Feb 2010 00:45:02 +0100
changeset 35102 cc7a0b9f938c
parent 32960 69916a850301
child 35355 613e133966ea
child 35416 d8d7d1b785af
permissions -rwxr-xr-x
modernized translations;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/Comp/TranslCompTp.thy
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     2
    Author:     Martin Strecker
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     3
*)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     4
15481
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 14981
diff changeset
     5
theory TranslCompTp
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 14981
diff changeset
     6
imports Index "../BV/JVMType"
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 14981
diff changeset
     7
begin
fc075ae929e4 the new subst tactic, by Lucas Dixon
paulson
parents: 14981
diff changeset
     8
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
     9
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    10
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    11
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    12
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    13
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    14
constdefs
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    15
  comb :: "['a \<Rightarrow> 'b list \<times> 'c, 'c \<Rightarrow> 'b list \<times> 'd, 'a] \<Rightarrow> 'b list \<times> 'd" 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    16
  "comb == (\<lambda> f1 f2 x0. let (xs1, x1) = f1 x0; 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    17
                            (xs2, x2) = f2 x1 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    18
                        in  (xs1 @ xs2, x2))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    19
  comb_nil :: "'a \<Rightarrow> 'b list \<times> 'a"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    20
  "comb_nil a == ([], a)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    21
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    22
syntax (xsymbols)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    23
  "comb" :: "['a \<Rightarrow> 'b list \<times> 'c, 'c \<Rightarrow> 'b list \<times> 'd, 'a] \<Rightarrow> 'b list \<times> 'd"    
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    24
            (infixr "\<box>" 55)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    25
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    26
lemma comb_nil_left [simp]: "comb_nil \<box> f = f"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    27
by (simp add: comb_def comb_nil_def split_beta)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    28
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    29
lemma comb_nil_right [simp]: "f \<box> comb_nil = f"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    30
by (simp add: comb_def comb_nil_def split_beta)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    31
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    32
lemma comb_assoc [simp]: "(fa \<box> fb) \<box> fc = fa \<box> (fb \<box> fc)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    33
by (simp add: comb_def split_beta)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    34
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    35
lemma comb_inv: "(xs', x') = (f1 \<box> f2) x0 \<Longrightarrow>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    36
  \<exists> xs1 x1 xs2 x2. (xs1, x1) = (f1 x0) \<and> (xs2, x2) = f2 x1 \<and> xs'= xs1 @ xs2 \<and> x'=x2"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    37
apply (case_tac "f1 x0")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    38
apply (case_tac "f2 x1")
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    39
apply (simp add: comb_def split_beta)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    40
done
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    41
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    42
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    43
35102
cc7a0b9f938c modernized translations;
wenzelm
parents: 32960
diff changeset
    44
abbreviation (input)
cc7a0b9f938c modernized translations;
wenzelm
parents: 32960
diff changeset
    45
  mt_of :: "method_type \<times> state_type \<Rightarrow> method_type"
cc7a0b9f938c modernized translations;
wenzelm
parents: 32960
diff changeset
    46
  where "mt_of == fst"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    47
35102
cc7a0b9f938c modernized translations;
wenzelm
parents: 32960
diff changeset
    48
abbreviation (input)
cc7a0b9f938c modernized translations;
wenzelm
parents: 32960
diff changeset
    49
  sttp_of :: "method_type \<times> state_type \<Rightarrow> state_type"
cc7a0b9f938c modernized translations;
wenzelm
parents: 32960
diff changeset
    50
  where "sttp_of == snd"
13673
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    51
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    52
consts
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    53
  compTpExpr  :: "java_mb \<Rightarrow> java_mb prog \<Rightarrow> expr
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    54
                   \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    55
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    56
  compTpExprs :: "java_mb \<Rightarrow> java_mb prog \<Rightarrow> expr list
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    57
                   \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    58
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    59
  compTpStmt  :: "java_mb \<Rightarrow> java_mb prog \<Rightarrow> stmt
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    60
                   \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    61
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    62
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    63
constdefs
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    64
  nochangeST :: "state_type \<Rightarrow> method_type \<times> state_type"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    65
  "nochangeST sttp == ([Some sttp], sttp)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    66
  pushST :: "[ty list, state_type] \<Rightarrow> method_type \<times> state_type"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    67
  "pushST tps == (\<lambda> (ST, LT). ([Some (ST, LT)], (tps @ ST, LT)))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    68
  dupST :: "state_type \<Rightarrow> method_type \<times> state_type"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    69
  "dupST == (\<lambda> (ST, LT). ([Some (ST, LT)], (hd ST # ST, LT)))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    70
  dup_x1ST :: "state_type \<Rightarrow> method_type \<times> state_type"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    71
  "dup_x1ST == (\<lambda> (ST, LT). ([Some (ST, LT)], 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    72
                             (hd ST # hd (tl ST) # hd ST # (tl (tl ST)), LT)))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    73
  popST :: "[nat, state_type] \<Rightarrow> method_type \<times> state_type"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    74
  "popST n == (\<lambda> (ST, LT). ([Some (ST, LT)], (drop n ST, LT)))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    75
  replST :: "[nat, ty, state_type] \<Rightarrow> method_type \<times> state_type"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    76
  "replST n tp == (\<lambda> (ST, LT). ([Some (ST, LT)], (tp # (drop n ST), LT)))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    77
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    78
  storeST :: "[nat, ty, state_type] \<Rightarrow> method_type \<times> state_type"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    79
  "storeST i tp == (\<lambda> (ST, LT). ([Some (ST, LT)], (tl ST, LT [i:= OK tp])))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    80
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    81
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    82
(* Expressions *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    83
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    84
primrec
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    85
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    86
  "compTpExpr jmb G (NewC c) = pushST [Class c]"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    87
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    88
  "compTpExpr jmb G (Cast c e) = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    89
  (compTpExpr jmb G e) \<box> (replST 1 (Class c))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    90
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    91
  "compTpExpr jmb G (Lit val) = pushST [the (typeof (\<lambda>v. None) val)]"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    92
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    93
  "compTpExpr jmb G (BinOp bo e1 e2) = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    94
     (compTpExpr jmb G e1) \<box> (compTpExpr jmb G e2) \<box> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    95
     (case bo of 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    96
       Eq => popST 2 \<box> pushST [PrimT Boolean] \<box> popST 1 \<box> pushST [PrimT Boolean]
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    97
     | Add => replST 2 (PrimT Integer))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    98
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
    99
  "compTpExpr jmb G (LAcc vn) = (\<lambda> (ST, LT). 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   100
      pushST [ok_val (LT ! (index jmb vn))] (ST, LT))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   101
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   102
  "compTpExpr jmb G (vn::=e) = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   103
      (compTpExpr jmb G e) \<box> dupST \<box> (popST 1)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   104
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   105
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   106
  "compTpExpr jmb G ( {cn}e..fn ) = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   107
      (compTpExpr jmb G e) \<box> replST 1 (snd (the (field (G,cn) fn)))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   108
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   109
  "compTpExpr jmb G (FAss cn e1 fn e2 ) = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   110
      (compTpExpr jmb G e1) \<box> (compTpExpr jmb G e2) \<box> dup_x1ST \<box> (popST 2)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   111
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   112
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   113
  "compTpExpr jmb G ({C}a..mn({fpTs}ps)) =
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   114
       (compTpExpr jmb G a) \<box> (compTpExprs jmb G ps) \<box> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   115
       (replST ((length ps) + 1) (method_rT (the (method (G,C) (mn,fpTs)))))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   116
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   117
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   118
(* Expression lists *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   119
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   120
  "compTpExprs jmb G [] = comb_nil"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   121
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   122
  "compTpExprs jmb G (e#es) = (compTpExpr jmb G e) \<box> (compTpExprs jmb G es)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   123
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   124
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   125
(* Statements *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   126
primrec
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   127
   "compTpStmt jmb G Skip = comb_nil"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   128
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   129
   "compTpStmt jmb G (Expr e) =  (compTpExpr jmb G e) \<box> popST 1"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   130
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   131
   "compTpStmt jmb G (c1;; c2) = (compTpStmt jmb G c1) \<box> (compTpStmt jmb G c2)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   132
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   133
   "compTpStmt jmb G (If(e) c1 Else c2) = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   134
      (pushST [PrimT Boolean]) \<box> (compTpExpr jmb G e) \<box> popST 2 \<box>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   135
      (compTpStmt jmb G c1) \<box> nochangeST \<box> (compTpStmt jmb G c2)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   136
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   137
   "compTpStmt jmb G (While(e) c) = 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   138
      (pushST [PrimT Boolean]) \<box> (compTpExpr jmb G e) \<box> popST 2 \<box>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   139
      (compTpStmt jmb G c) \<box> nochangeST"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   140
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   141
constdefs
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   142
  compTpInit  :: "java_mb \<Rightarrow> (vname * ty)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   143
                   \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   144
  "compTpInit jmb == (\<lambda> (vn,ty). (pushST [ty]) \<box>  (storeST (index jmb vn) ty))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   145
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   146
consts
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   147
  compTpInitLvars :: "[java_mb, (vname \<times> ty) list] 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   148
                   \<Rightarrow> state_type \<Rightarrow> method_type \<times> state_type"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   149
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   150
primrec 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   151
  "compTpInitLvars jmb [] = comb_nil"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   152
  "compTpInitLvars jmb (lv#lvars) = (compTpInit jmb lv) \<box> (compTpInitLvars jmb lvars)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   153
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   154
constdefs
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   155
   start_ST :: "opstack_type"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   156
  "start_ST == []"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   157
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   158
   start_LT :: "cname \<Rightarrow> ty list \<Rightarrow> nat \<Rightarrow> locvars_type"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   159
  "start_LT C pTs n ==  (OK (Class C))#((map OK pTs))@(replicate n Err)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   160
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   161
  compTpMethod  :: "[java_mb prog, cname, java_mb mdecl] \<Rightarrow> method_type"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   162
  "compTpMethod G C == \<lambda> ((mn,pTs),rT, jmb). 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   163
                         let (pns,lvars,blk,res) = jmb
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   164
                         in (mt_of
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   165
                            ((compTpInitLvars jmb lvars \<box> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   166
                              compTpStmt jmb G blk \<box> 
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   167
                              compTpExpr jmb G res \<box>
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   168
                              nochangeST)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   169
                                (start_ST, start_LT C pTs (length lvars))))"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   170
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   171
  compTp :: "java_mb prog => prog_type"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   172
  "compTp G C sig == let (D, rT, jmb) = (the (method (G, C) sig))
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   173
                      in compTpMethod G C (sig, rT, jmb)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   174
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   175
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   176
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   177
(**********************************************************************)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   178
  (* Computing the maximum stack size from the method_type *)
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   179
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   180
constdefs
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   181
  ssize_sto :: "(state_type option) \<Rightarrow> nat"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   182
  "ssize_sto sto ==  case sto of None \<Rightarrow> 0 | (Some (ST, LT)) \<Rightarrow> length ST"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   183
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   184
  max_of_list :: "nat list \<Rightarrow> nat"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   185
  "max_of_list xs == foldr max xs 0"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   186
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   187
  max_ssize :: "method_type \<Rightarrow> nat"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   188
  "max_ssize mt == max_of_list (map ssize_sto mt)"
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   189
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   190
2950128b8206 First checkin of compiler
streckem
parents:
diff changeset
   191
end