src/HOL/MicroJava/JVM/JVMListExample.thy
author wenzelm
Sun, 30 Sep 2007 21:55:15 +0200
changeset 24783 5a3e336a2e37
parent 24340 811f78424efc
child 28408 e26aac53723d
permissions -rw-r--r--
avoid internal names;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/JVM/JVMListExample.thy
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
     2
    ID:         $Id$
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
     3
    Author:     Stefan Berghofer
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
     4
*)
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
     5
12973
8040cce614e5 fixed missing label
kleing
parents: 12951
diff changeset
     6
header {* \isaheader{Example for generating executable code from JVM semantics}\label{sec:JVMListExample} *}
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
     7
24340
811f78424efc theory header: more precise imports;
wenzelm
parents: 24225
diff changeset
     8
theory JVMListExample imports "../J/SystemClasses" JVMExec begin
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
     9
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    10
consts
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    11
  list_nam :: cnam
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    12
  test_nam :: cnam
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    13
  append_name :: mname
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    14
  makelist_name :: mname
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    15
  val_nam :: vnam
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    16
  next_nam :: vnam
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    17
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    18
constdefs
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    19
  list_name :: cname
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    20
  "list_name == Cname list_nam"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    21
  
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    22
  test_name :: cname
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    23
  "test_name == Cname test_nam"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    24
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    25
  val_name :: vname
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    26
  "val_name == VName val_nam"
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    27
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    28
  next_name :: vname
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    29
  "next_name == VName next_nam"
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    30
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    31
  append_ins :: bytecode
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    32
  "append_ins == 
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    33
       [Load 0,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    34
        Getfield next_name list_name,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    35
        Dup,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    36
        LitPush Null,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    37
        Ifcmpeq 4,
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    38
        Load 1,       
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    39
        Invoke list_name append_name [Class list_name],
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    40
        Return,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    41
        Pop,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    42
        Load 0,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    43
        Load 1,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    44
        Putfield next_name list_name,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    45
        LitPush Unit,
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    46
        Return]"
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    47
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    48
  list_class :: "jvm_method class"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    49
  "list_class ==
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    50
    (Object,
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    51
     [(val_name, PrimT Integer), (next_name, Class list_name)],
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    52
     [((append_name, [Class list_name]), PrimT Void,
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    53
        (3, 0, append_ins,[(1,2,8,Xcpt NullPointer)]))])"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    54
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    55
  make_list_ins :: bytecode
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    56
  "make_list_ins ==
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    57
       [New list_name,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    58
        Dup,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    59
        Store 0,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    60
        LitPush (Intg 1),
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    61
        Putfield val_name list_name,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    62
        New list_name,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    63
        Dup,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    64
        Store 1,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    65
        LitPush (Intg 2),
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    66
        Putfield val_name list_name,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    67
        New list_name,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    68
        Dup,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    69
        Store 2,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    70
        LitPush (Intg 3),
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    71
        Putfield val_name list_name,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    72
        Load 0,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    73
        Load 1,
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    74
        Invoke list_name append_name [Class list_name],
13101
kleing
parents: 13091
diff changeset
    75
        Pop,
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    76
        Load 0,
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    77
        Load 2,
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    78
        Invoke list_name append_name [Class list_name],
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    79
        Return]"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    80
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    81
  test_class :: "jvm_method class"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    82
  "test_class ==
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    83
    (Object, [],
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    84
     [((makelist_name, []), PrimT Void, (3, 2, make_list_ins,[]))])"
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    85
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    86
  E :: jvm_prog
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    87
  "E == SystemClasses @ [(list_name, list_class), (test_name, test_class)]"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    88
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    89
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    90
types_code
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
    91
  cnam ("string")
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    92
  vnam ("string")
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    93
  mname ("string")
24783
5a3e336a2e37 avoid internal names;
wenzelm
parents: 24340
diff changeset
    94
  loc' ("int")
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    95
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
    96
consts_code
16770
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16644
diff changeset
    97
  "new_Addr" ("\<module>new'_addr {* %x. case x of None => True | Some y => False *}/ {* None *}/ {* Loc *}")
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16644
diff changeset
    98
attach {*
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16644
diff changeset
    99
fun new_addr p none loc hp =
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16644
diff changeset
   100
  let fun nr i = if p (hp (loc i)) then (loc i, none) else nr (i+1);
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16644
diff changeset
   101
  in nr 0 end;
1f1b1fae30e4 Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents: 16644
diff changeset
   102
*}
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
   103
18679
cf9f1584431a generated code: raise Match instead of ERROR;
wenzelm
parents: 17145
diff changeset
   104
  "arbitrary" ("(raise Match)")
22921
475ff421a6a3 consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents: 21404
diff changeset
   105
  "arbitrary :: val" ("{* Unit *}")
475ff421a6a3 consts in consts_code Isar commands are now referred to by usual term syntax
haftmann
parents: 21404
diff changeset
   106
  "arbitrary :: cname" ("{* Object *}")
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
   107
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
   108
  "list_nam" ("\"list\"")
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents: 12911
diff changeset
   109
  "test_nam" ("\"test\"")
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
   110
  "append_name" ("\"append\"")
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
   111
  "makelist_name" ("\"makelist\"")
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
   112
  "val_nam" ("\"val\"")
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
   113
  "next_nam" ("\"next\"")
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
   114
20971
1e7df197b8b8 added code generator setup
haftmann
parents: 18679
diff changeset
   115
definition
1e7df197b8b8 added code generator setup
haftmann
parents: 18679
diff changeset
   116
  "test = exec (E, start_state E test_name makelist_name)"
1e7df197b8b8 added code generator setup
haftmann
parents: 18679
diff changeset
   117
1e7df197b8b8 added code generator setup
haftmann
parents: 18679
diff changeset
   118
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
   119
subsection {* Single step execution *}
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
   120
17145
e623e57b0f44 Adapted to new code generator syntax.
berghofe
parents: 16770
diff changeset
   121
code_module JVM
e623e57b0f44 Adapted to new code generator syntax.
berghofe
parents: 16770
diff changeset
   122
contains
24225
348e982c694b dropped code generator setup garbage
haftmann
parents: 24074
diff changeset
   123
  exec = exec
348e982c694b dropped code generator setup garbage
haftmann
parents: 24074
diff changeset
   124
  test = test
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
   125
24074
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   126
ML {*
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   127
JVM.test;
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   128
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   129
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   130
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   131
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   132
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   133
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   134
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   135
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   136
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   137
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   138
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   139
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   140
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   141
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   142
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   143
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   144
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   145
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   146
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   147
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   148
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   149
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   150
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   151
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   152
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   153
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   154
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   155
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   156
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   157
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   158
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   159
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   160
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   161
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   162
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   163
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   164
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   165
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   166
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   167
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   168
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   169
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   170
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   171
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   172
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   173
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   174
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   175
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   176
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   177
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   178
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   179
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   180
JVM.exec (JVM.E, JVM.the it);
40f414b87655 tuned ML declarations;
wenzelm
parents: 22921
diff changeset
   181
*}
12442
0ecba8660de7 Example for code generator.
berghofe
parents:
diff changeset
   182
12973
8040cce614e5 fixed missing label
kleing
parents: 12951
diff changeset
   183
end