src/HOL/MicroJava/BV/BVExample.thy
author noschinl
Fri, 17 Apr 2015 10:49:57 +0200
changeset 60108 d7fe3e0aca85
parent 58886 8a6cac7c7247
child 61361 8b5f00202e1a
permissions -rw-r--r--
rewrite: add default pattern "in concl" for more cases
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/BV/BVExample.thy
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
     2
    Author:     Gerwin Klein
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
     3
*)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
     4
58886
8a6cac7c7247 modernized header;
wenzelm
parents: 52866
diff changeset
     5
section {* Example Welltypings \label{sec:BVExample} *}
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
     6
23022
9872ef956276 added qualification for ambiguous definition names
haftmann
parents: 22271
diff changeset
     7
theory BVExample
41413
64cd30d6b0b8 explicit file specifications -- avoid secondary load path;
wenzelm
parents: 40077
diff changeset
     8
imports
64cd30d6b0b8 explicit file specifications -- avoid secondary load path;
wenzelm
parents: 40077
diff changeset
     9
  "../JVM/JVMListExample"
64cd30d6b0b8 explicit file specifications -- avoid secondary load path;
wenzelm
parents: 40077
diff changeset
    10
  BVSpecTypeSafe
64cd30d6b0b8 explicit file specifications -- avoid secondary load path;
wenzelm
parents: 40077
diff changeset
    11
  JVM
23022
9872ef956276 added qualification for ambiguous definition names
haftmann
parents: 22271
diff changeset
    12
begin
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    13
12972
da7345ff18e1 fixed document
kleing
parents: 12951
diff changeset
    14
text {*
da7345ff18e1 fixed document
kleing
parents: 12951
diff changeset
    15
  This theory shows type correctness of the example program in section 
da7345ff18e1 fixed document
kleing
parents: 12951
diff changeset
    16
  \ref{sec:JVMListExample} (p. \pageref{sec:JVMListExample}) by
da7345ff18e1 fixed document
kleing
parents: 12951
diff changeset
    17
  explicitly providing a welltyping. It also shows that the start
da7345ff18e1 fixed document
kleing
parents: 12951
diff changeset
    18
  state of the program conforms to the welltyping; hence type safe
da7345ff18e1 fixed document
kleing
parents: 12951
diff changeset
    19
  execution is guaranteed.
da7345ff18e1 fixed document
kleing
parents: 12951
diff changeset
    20
*}
da7345ff18e1 fixed document
kleing
parents: 12951
diff changeset
    21
58886
8a6cac7c7247 modernized header;
wenzelm
parents: 52866
diff changeset
    22
subsection "Setup"
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    23
13101
kleing
parents: 13092
diff changeset
    24
text {* Abbreviations for definitions we will have to use often in the
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    25
proofs below: *}
13101
kleing
parents: 13092
diff changeset
    26
lemmas name_defs   = list_name_def test_name_def val_name_def next_name_def 
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    27
lemmas system_defs = SystemClasses_def ObjectC_def NullPointerC_def 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    28
                     OutOfMemoryC_def ClassCastC_def
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    29
lemmas class_defs  = list_class_def test_class_def
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    30
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    31
text {* These auxiliary proofs are for efficiency: class lookup,
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    32
subclass relation, method and field lookup are computed only once:
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    33
*}
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    34
lemma class_Object [simp]:
28520
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
    35
  "class E Object = Some (undefined, [],[])"
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    36
  by (simp add: class_def system_defs E_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    37
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    38
lemma class_NullPointer [simp]:
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    39
  "class E (Xcpt NullPointer) = Some (Object, [], [])"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    40
  by (simp add: class_def system_defs E_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    41
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    42
lemma class_OutOfMemory [simp]:
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    43
  "class E (Xcpt OutOfMemory) = Some (Object, [], [])"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    44
  by (simp add: class_def system_defs E_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    45
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    46
lemma class_ClassCast [simp]:
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    47
  "class E (Xcpt ClassCast) = Some (Object, [], [])"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    48
  by (simp add: class_def system_defs E_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    49
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    50
lemma class_list [simp]:
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    51
  "class E list_name = Some list_class"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    52
  by (simp add: class_def system_defs E_def name_defs distinct_classes [symmetric])
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    53
 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    54
lemma class_test [simp]:
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    55
  "class E test_name = Some test_class"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    56
  by (simp add: class_def system_defs E_def name_defs distinct_classes [symmetric])
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    57
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    58
lemma E_classes [simp]:
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    59
  "{C. is_class E C} = {list_name, test_name, Xcpt NullPointer, 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    60
                        Xcpt ClassCast, Xcpt OutOfMemory, Object}"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    61
  by (auto simp add: is_class_def class_def system_defs E_def name_defs class_defs)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    62
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    63
text {* The subclass releation spelled out: *}
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    64
lemma subcls1:
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 31998
diff changeset
    65
  "subcls1 E = {(list_name,Object), (test_name,Object), (Xcpt NullPointer, Object),
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 31998
diff changeset
    66
                (Xcpt ClassCast, Object), (Xcpt OutOfMemory, Object)}"
31197
c1c163ec6c44 fine-tuned elimination of comprehensions involving x=t.
nipkow
parents: 31168
diff changeset
    67
apply (simp add: subcls1_def2)
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 28520
diff changeset
    68
apply (simp add: name_defs class_defs system_defs E_def class_def)
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 31998
diff changeset
    69
apply (simp add: Sigma_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 31998
diff changeset
    70
apply auto
31166
a90fe83f58ea "{x. P x & x=t & Q x}" is now rewritten to "if P t & Q t then {t} else {}"
nipkow
parents: 28520
diff changeset
    71
done
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    72
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    73
text {* The subclass relation is acyclic; hence its converse is well founded: *}
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    74
lemma notin_rtrancl:
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 31998
diff changeset
    75
  "(a, b) \<in> r\<^sup>* \<Longrightarrow> a \<noteq> b \<Longrightarrow> (\<And>y. (a, y) \<notin> r) \<Longrightarrow> False"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 31998
diff changeset
    76
  by (auto elim: converse_rtranclE)
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    77
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 31998
diff changeset
    78
lemma acyclic_subcls1_E: "acyclic (subcls1 E)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 31998
diff changeset
    79
  apply (rule acyclicI)
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    80
  apply (simp add: subcls1)
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 31998
diff changeset
    81
  apply (auto dest!: tranclD)
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    82
  apply (auto elim!: notin_rtrancl simp add: name_defs distinct_classes)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    83
  done
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    84
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 31998
diff changeset
    85
lemma wf_subcls1_E: "wf ((subcls1 E)\<inverse>)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents: 31998
diff changeset
    86
  apply (rule finite_acyclic_wf_converse)
23757
087b0a241557 - Renamed inductive2 to inductive
berghofe
parents: 23022
diff changeset
    87
  apply (simp add: subcls1 del: insert_iff)
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    88
  apply (rule acyclic_subcls1_E)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    89
  done  
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    90
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    91
text {* Method and field lookup: *}
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    92
lemma method_Object [simp]:
31852
a16bb835e97d explicit Set constructor for code generated for sets
haftmann
parents: 31197
diff changeset
    93
  "method (E, Object) = Map.empty"
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    94
  by (simp add: method_rec_lemma [OF class_Object wf_subcls1_E])
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    95
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    96
lemma method_append [simp]:
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    97
  "method (E, list_name) (append_name, [Class list_name]) =
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    98
  Some (list_name, PrimT Void, 3, 0, append_ins, [(1, 2, 8, Xcpt NullPointer)])"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
    99
  apply (insert class_list)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   100
  apply (unfold list_class_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   101
  apply (drule method_rec_lemma [OF _ wf_subcls1_E])
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   102
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   103
  done
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   104
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   105
lemma method_makelist [simp]:
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   106
  "method (E, test_name) (makelist_name, []) = 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   107
  Some (test_name, PrimT Void, 3, 2, make_list_ins, [])"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   108
  apply (insert class_test)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   109
  apply (unfold test_class_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   110
  apply (drule method_rec_lemma [OF _ wf_subcls1_E])
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   111
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   112
  done
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   113
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   114
lemma field_val [simp]:
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   115
  "field (E, list_name) val_name = Some (list_name, PrimT Integer)"
23022
9872ef956276 added qualification for ambiguous definition names
haftmann
parents: 22271
diff changeset
   116
  apply (unfold TypeRel.field_def)
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   117
  apply (insert class_list)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   118
  apply (unfold list_class_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   119
  apply (drule fields_rec_lemma [OF _ wf_subcls1_E])
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   120
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   121
  done
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   122
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   123
lemma field_next [simp]:
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   124
  "field (E, list_name) next_name = Some (list_name, Class list_name)"
23022
9872ef956276 added qualification for ambiguous definition names
haftmann
parents: 22271
diff changeset
   125
  apply (unfold TypeRel.field_def)
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   126
  apply (insert class_list)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   127
  apply (unfold list_class_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   128
  apply (drule fields_rec_lemma [OF _ wf_subcls1_E])
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   129
  apply (simp add: name_defs distinct_fields [symmetric])
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   130
  done
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   131
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   132
lemma [simp]: "fields (E, Object) = []"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   133
   by (simp add: fields_rec_lemma [OF class_Object wf_subcls1_E])
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   134
 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   135
lemma [simp]: "fields (E, Xcpt NullPointer) = []"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   136
  by (simp add: fields_rec_lemma [OF class_NullPointer wf_subcls1_E])
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   137
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   138
lemma [simp]: "fields (E, Xcpt ClassCast) = []"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   139
  by (simp add: fields_rec_lemma [OF class_ClassCast wf_subcls1_E])
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   140
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   141
lemma [simp]: "fields (E, Xcpt OutOfMemory) = []"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   142
  by (simp add: fields_rec_lemma [OF class_OutOfMemory wf_subcls1_E])
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   143
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   144
lemma [simp]: "fields (E, test_name) = []"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   145
  apply (insert class_test)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   146
  apply (unfold test_class_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   147
  apply (drule fields_rec_lemma [OF _ wf_subcls1_E])
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   148
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   149
  done
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   150
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   151
lemmas [simp] = is_class_def
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   152
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   153
text {*
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   154
  The next definition and three proof rules implement an algorithm to
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   155
  enumarate natural numbers. The command @{text "apply (elim pc_end pc_next pc_0"} 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   156
  transforms a goal of the form
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   157
  @{prop [display] "pc < n \<Longrightarrow> P pc"} 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   158
  into a series of goals
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   159
  @{prop [display] "P 0"} 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   160
  @{prop [display] "P (Suc 0)"} 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   161
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   162
  @{text "\<dots>"}
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   163
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   164
  @{prop [display] "P n"} 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   165
*}
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
   166
definition intervall :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> bool" ("_ \<in> [_, _')") where
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   167
  "x \<in> [a, b) \<equiv> a \<le> x \<and> x < b"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   168
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   169
lemma pc_0: "x < n \<Longrightarrow> (x \<in> [0, n) \<Longrightarrow> P x) \<Longrightarrow> P x"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   170
  by (simp add: intervall_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   171
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   172
lemma pc_next: "x \<in> [n0, n) \<Longrightarrow> P n0 \<Longrightarrow> (x \<in> [Suc n0, n) \<Longrightarrow> P x) \<Longrightarrow> P x"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   173
  apply (cases "x=n0")
13187
e5434b822a96 Modifications due to enhanced linear arithmetic.
nipkow
parents: 13148
diff changeset
   174
  apply (auto simp add: intervall_def)
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   175
  done
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   176
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   177
lemma pc_end: "x \<in> [n,n) \<Longrightarrow> P x" 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   178
  by (unfold intervall_def) arith
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   179
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   180
58886
8a6cac7c7247 modernized header;
wenzelm
parents: 52866
diff changeset
   181
subsection "Program structure"
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   182
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   183
text {*
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   184
  The program is structurally wellformed:
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   185
*}
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13727
diff changeset
   186
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   187
lemma wf_struct:
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   188
  "wf_prog (\<lambda>G C mb. True) E" (is "wf_prog ?mb E")
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   189
proof -
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   190
  have "unique E" 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   191
    by (simp add: system_defs E_def class_defs name_defs distinct_classes)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   192
  moreover
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   193
  have "set SystemClasses \<subseteq> set E" by (simp add: system_defs E_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   194
  hence "wf_syscls E" by (rule wf_syscls)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   195
  moreover
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   196
  have "wf_cdecl ?mb E ObjectC" by (simp add: wf_cdecl_def ObjectC_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   197
  moreover
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   198
  have "wf_cdecl ?mb E NullPointerC" 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   199
    by (auto elim: notin_rtrancl 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   200
            simp add: wf_cdecl_def name_defs NullPointerC_def subcls1)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   201
  moreover
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   202
  have "wf_cdecl ?mb E ClassCastC" 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   203
    by (auto elim: notin_rtrancl 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   204
            simp add: wf_cdecl_def name_defs ClassCastC_def subcls1)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   205
  moreover
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   206
  have "wf_cdecl ?mb E OutOfMemoryC" 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   207
    by (auto elim: notin_rtrancl 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   208
            simp add: wf_cdecl_def name_defs OutOfMemoryC_def subcls1)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   209
  moreover
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   210
  have "wf_cdecl ?mb E (list_name, list_class)"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   211
    apply (auto elim!: notin_rtrancl 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   212
            simp add: wf_cdecl_def wf_fdecl_def list_class_def 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   213
                      wf_mdecl_def wf_mhead_def subcls1)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   214
    apply (auto simp add: name_defs distinct_classes distinct_fields)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   215
    done    
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   216
  moreover
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   217
  have "wf_cdecl ?mb E (test_name, test_class)" 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   218
    apply (auto elim!: notin_rtrancl 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   219
            simp add: wf_cdecl_def wf_fdecl_def test_class_def 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   220
                      wf_mdecl_def wf_mhead_def subcls1)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   221
    apply (auto simp add: name_defs distinct_classes distinct_fields)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   222
    done       
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   223
  ultimately
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13727
diff changeset
   224
  show ?thesis 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13727
diff changeset
   225
    by (simp add: wf_prog_def ws_prog_def wf_cdecl_mrT_cdecl_mdecl E_def SystemClasses_def)
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   226
qed
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   227
58886
8a6cac7c7247 modernized header;
wenzelm
parents: 52866
diff changeset
   228
subsection "Welltypings"
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   229
text {*
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   230
  We show welltypings of the methods @{term append_name} in class @{term list_name}, 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   231
  and @{term makelist_name} in class @{term test_name}:
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   232
*}
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   233
lemmas eff_simps [simp] = eff_def norm_eff_def xcpt_eff_def
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   234
declare appInvoke [simp del]
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   235
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
   236
definition phi_append :: method_type ("\<phi>\<^sub>a") where
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   237
  "\<phi>\<^sub>a \<equiv> map (\<lambda>(x,y). Some (x, map OK y)) [ 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   238
   (                                    [], [Class list_name, Class list_name]),
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   239
   (                     [Class list_name], [Class list_name, Class list_name]),
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   240
   (                     [Class list_name], [Class list_name, Class list_name]),
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   241
   (    [Class list_name, Class list_name], [Class list_name, Class list_name]),
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   242
   ([NT, Class list_name, Class list_name], [Class list_name, Class list_name]),
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   243
   (                     [Class list_name], [Class list_name, Class list_name]),
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   244
   (    [Class list_name, Class list_name], [Class list_name, Class list_name]),
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   245
   (                          [PrimT Void], [Class list_name, Class list_name]),
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   246
   (                        [Class Object], [Class list_name, Class list_name]),
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   247
   (                                    [], [Class list_name, Class list_name]),
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   248
   (                     [Class list_name], [Class list_name, Class list_name]),
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   249
   (    [Class list_name, Class list_name], [Class list_name, Class list_name]),
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   250
   (                                    [], [Class list_name, Class list_name]),
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   251
   (                          [PrimT Void], [Class list_name, Class list_name])]"
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   252
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   253
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   254
lemma bounded_append [simp]:
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   255
  "check_bounded append_ins [(Suc 0, 2, 8, Xcpt NullPointer)]"
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   256
  apply (simp add: check_bounded_def)
40077
c8a9eaaa2f59 nat_number -> eval_nat_numeral
nipkow
parents: 37024
diff changeset
   257
  apply (simp add: eval_nat_numeral append_ins_def)
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   258
  apply (rule allI, rule impI)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   259
  apply (elim pc_end pc_next pc_0)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   260
  apply auto
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   261
  done
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   262
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   263
lemma types_append [simp]: "check_types E 3 (Suc (Suc 0)) (map OK \<phi>\<^sub>a)"
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   264
  apply (auto simp add: check_types_def phi_append_def JVM_states_unfold)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   265
  apply (unfold list_def)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   266
  apply auto
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   267
  done
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   268
  
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   269
lemma wt_append [simp]:
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   270
  "wt_method E list_name [Class list_name] (PrimT Void) 3 0 append_ins
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   271
             [(Suc 0, 2, 8, Xcpt NullPointer)] \<phi>\<^sub>a"
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   272
  apply (simp add: wt_method_def wt_start_def wt_instr_def)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   273
  apply (simp add: phi_append_def append_ins_def)
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   274
  apply clarify
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   275
  apply (elim pc_end pc_next pc_0)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   276
  apply simp
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44035
diff changeset
   277
  apply (fastforce simp add: match_exception_entry_def sup_state_conv subcls1)
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   278
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   279
  apply simp
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 44035
diff changeset
   280
  apply (fastforce simp add: sup_state_conv subcls1)
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   281
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   282
  apply (simp add: app_def xcpt_app_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   283
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   284
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   285
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   286
  apply (simp add: match_exception_entry_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   287
  apply (simp add: match_exception_entry_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   288
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   289
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   290
  done
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   291
13006
51c5f3f11d16 symbolized
kleing
parents: 12972
diff changeset
   292
text {* Some abbreviations for readability *} 
35102
cc7a0b9f938c modernized translations;
wenzelm
parents: 33954
diff changeset
   293
abbreviation Clist :: ty 
cc7a0b9f938c modernized translations;
wenzelm
parents: 33954
diff changeset
   294
  where "Clist == Class list_name"
cc7a0b9f938c modernized translations;
wenzelm
parents: 33954
diff changeset
   295
abbreviation Ctest :: ty
cc7a0b9f938c modernized translations;
wenzelm
parents: 33954
diff changeset
   296
  where "Ctest == Class test_name"
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   297
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
   298
definition phi_makelist :: method_type ("\<phi>\<^sub>m") where
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   299
  "\<phi>\<^sub>m \<equiv> map (\<lambda>(x,y). Some (x, y)) [ 
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   300
    (                                   [], [OK Ctest, Err     , Err     ]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   301
    (                              [Clist], [OK Ctest, Err     , Err     ]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   302
    (                       [Clist, Clist], [OK Ctest, Err     , Err     ]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   303
    (                              [Clist], [OK Clist, Err     , Err     ]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   304
    (               [PrimT Integer, Clist], [OK Clist, Err     , Err     ]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   305
    (                                   [], [OK Clist, Err     , Err     ]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   306
    (                              [Clist], [OK Clist, Err     , Err     ]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   307
    (                       [Clist, Clist], [OK Clist, Err     , Err     ]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   308
    (                              [Clist], [OK Clist, OK Clist, Err     ]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   309
    (               [PrimT Integer, Clist], [OK Clist, OK Clist, Err     ]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   310
    (                                   [], [OK Clist, OK Clist, Err     ]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   311
    (                              [Clist], [OK Clist, OK Clist, Err     ]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   312
    (                       [Clist, Clist], [OK Clist, OK Clist, Err     ]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   313
    (                              [Clist], [OK Clist, OK Clist, OK Clist]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   314
    (               [PrimT Integer, Clist], [OK Clist, OK Clist, OK Clist]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   315
    (                                   [], [OK Clist, OK Clist, OK Clist]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   316
    (                              [Clist], [OK Clist, OK Clist, OK Clist]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   317
    (                       [Clist, Clist], [OK Clist, OK Clist, OK Clist]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   318
    (                         [PrimT Void], [OK Clist, OK Clist, OK Clist]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   319
    (                                   [], [OK Clist, OK Clist, OK Clist]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   320
    (                              [Clist], [OK Clist, OK Clist, OK Clist]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   321
    (                       [Clist, Clist], [OK Clist, OK Clist, OK Clist]),
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   322
    (                         [PrimT Void], [OK Clist, OK Clist, OK Clist])]"
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   323
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   324
lemma bounded_makelist [simp]: "check_bounded make_list_ins []"
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   325
  apply (simp add: check_bounded_def)
40077
c8a9eaaa2f59 nat_number -> eval_nat_numeral
nipkow
parents: 37024
diff changeset
   326
  apply (simp add: eval_nat_numeral make_list_ins_def)
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   327
  apply (rule allI, rule impI)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   328
  apply (elim pc_end pc_next pc_0)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   329
  apply auto
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   330
  done
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   331
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   332
lemma types_makelist [simp]: "check_types E 3 (Suc (Suc (Suc 0))) (map OK \<phi>\<^sub>m)"
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   333
  apply (auto simp add: check_types_def phi_makelist_def JVM_states_unfold)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   334
  apply (unfold list_def)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   335
  apply auto
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   336
  done
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   337
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   338
lemma wt_makelist [simp]:
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   339
  "wt_method E test_name [] (PrimT Void) 3 2 make_list_ins [] \<phi>\<^sub>m"
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   340
  apply (simp add: wt_method_def)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13187
diff changeset
   341
  apply (simp add: make_list_ins_def phi_makelist_def)
40077
c8a9eaaa2f59 nat_number -> eval_nat_numeral
nipkow
parents: 37024
diff changeset
   342
  apply (simp add: wt_start_def eval_nat_numeral)
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   343
  apply (simp add: wt_instr_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   344
  apply clarify
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   345
  apply (elim pc_end pc_next pc_0)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   346
  apply (simp add: match_exception_entry_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   347
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   348
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   349
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   350
  apply (simp add: match_exception_entry_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   351
  apply (simp add: match_exception_entry_def) 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   352
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   353
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   354
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   355
  apply (simp add: match_exception_entry_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   356
  apply (simp add: match_exception_entry_def) 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   357
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   358
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   359
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   360
  apply (simp add: match_exception_entry_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   361
  apply (simp add: match_exception_entry_def) 
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   362
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   363
  apply (simp add: app_def xcpt_app_def)
13101
kleing
parents: 13092
diff changeset
   364
  apply simp 
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   365
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   366
  apply simp
13101
kleing
parents: 13092
diff changeset
   367
  apply (simp add: app_def xcpt_app_def) 
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   368
  apply simp
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   369
  done
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   370
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   371
text {* The whole program is welltyped: *}
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35102
diff changeset
   372
definition Phi :: prog_type ("\<Phi>") where
13101
kleing
parents: 13092
diff changeset
   373
  "\<Phi> C sg \<equiv> if C = test_name \<and> sg = (makelist_name, []) then \<phi>\<^sub>m else          
kleing
parents: 13092
diff changeset
   374
             if C = list_name \<and> sg = (append_name, [Class list_name]) then \<phi>\<^sub>a else []"
13139
94648e0e4506 fix for change in nat number simplification
kleing
parents: 13101
diff changeset
   375
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   376
lemma wf_prog:
13101
kleing
parents: 13092
diff changeset
   377
  "wt_jvm_prog E \<Phi>" 
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   378
  apply (unfold wt_jvm_prog_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   379
  apply (rule wf_mb'E [OF wf_struct])
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   380
  apply (simp add: E_def)
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   381
  apply clarify
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   382
  apply (fold E_def)
13101
kleing
parents: 13092
diff changeset
   383
  apply (simp add: system_defs class_defs Phi_def) 
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   384
  apply auto
13101
kleing
parents: 13092
diff changeset
   385
  done 
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   386
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   387
58886
8a6cac7c7247 modernized header;
wenzelm
parents: 52866
diff changeset
   388
subsection "Conformance"
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   389
text {* Execution of the program will be typesafe, because its
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   390
  start state conforms to the welltyping: *}
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   391
13052
3bf41c474a88 canonical start state
kleing
parents: 13043
diff changeset
   392
lemma "E,\<Phi> \<turnstile>JVM start_state E test_name makelist_name \<surd>"
3bf41c474a88 canonical start state
kleing
parents: 13043
diff changeset
   393
  apply (rule BV_correct_initial)
3bf41c474a88 canonical start state
kleing
parents: 13043
diff changeset
   394
    apply (rule wf_prog)
3bf41c474a88 canonical start state
kleing
parents: 13043
diff changeset
   395
   apply simp
3bf41c474a88 canonical start state
kleing
parents: 13043
diff changeset
   396
  apply simp
12951
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   397
  done
a9fdcb71d252 introduces SystemClasses and BVExample
kleing
parents:
diff changeset
   398
13092
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   399
58886
8a6cac7c7247 modernized header;
wenzelm
parents: 52866
diff changeset
   400
subsection "Example for code generation: inferring method types"
13092
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   401
28520
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   402
definition test_kil :: "jvm_prog \<Rightarrow> cname \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   403
             exception_table \<Rightarrow> instr list \<Rightarrow> JVMType.state list" where
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   404
  "test_kil G C pTs rT mxs mxl et instr =
13092
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   405
   (let first  = Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err));
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   406
        start  = OK first#(replicate (size instr - 1) (OK None))
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   407
    in  kiljvm G mxs (1+size pTs+mxl) rT et instr start)"
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   408
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   409
lemma [code]:
44035
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   410
  "unstables r step ss = 
47399
b72fa7bf9a10 abandoned almost redundant *_foldr lemmas
haftmann
parents: 45990
diff changeset
   411
   fold (\<lambda>p A. if \<not>stable r step ss p then insert p A else A) [0..<size ss] {}"
44035
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   412
proof -
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   413
  have "unstables r step ss = (UN p:{..<size ss}. if \<not>stable r step ss p then {p} else {})"
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   414
    apply (unfold unstables_def)
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   415
    apply (rule equalityI)
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   416
    apply (rule subsetI)
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   417
    apply (erule CollectE)
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   418
    apply (erule conjE)
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   419
    apply (rule UN_I)
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   420
    apply simp
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   421
    apply simp
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   422
    apply (rule subsetI)
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   423
    apply (erule UN_E)
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   424
    apply (case_tac "\<not> stable r step ss p")
52866
438f578ef1d9 tuned proofs;
wenzelm
parents: 52435
diff changeset
   425
    apply simp_all
44035
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   426
    done
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   427
  also have "\<And>f. (UN p:{..<size ss}. f p) = Union (set (map f [0..<size ss]))" by auto
47399
b72fa7bf9a10 abandoned almost redundant *_foldr lemmas
haftmann
parents: 45990
diff changeset
   428
  also note Sup_set_fold also note fold_map
44035
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   429
  also have "op \<union> \<circ> (\<lambda>p. if \<not> stable r step ss p then {p} else {}) = 
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   430
            (\<lambda>p A. if \<not>stable r step ss p then insert p A else A)"
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   431
    by(auto simp add: fun_eq_iff)
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   432
  finally show ?thesis .
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   433
qed
13092
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   434
44035
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   435
definition some_elem :: "'a set \<Rightarrow> 'a" where [code del]:
28520
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   436
  "some_elem = (%S. SOME x. x : S)"
52435
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51272
diff changeset
   437
code_printing
6646bb548c6b migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
haftmann
parents: 51272
diff changeset
   438
  constant some_elem \<rightharpoonup> (SML) "(case/ _ of/ Set/ xs/ =>/ hd/ xs)"
13092
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   439
28520
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   440
text {* This code setup is just a demonstration and \emph{not} sound! *}
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   441
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   442
lemma False
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   443
proof -
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   444
  have "some_elem (set [False, True]) = False"
44035
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   445
    by eval
28520
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   446
  moreover have "some_elem (set [True, False]) = True"
44035
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   447
    by eval
28520
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   448
  ultimately show False
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   449
    by (simp add: some_elem_def)
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   450
qed
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   451
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   452
lemma [code]:
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45985
diff changeset
   453
  "iter f step ss w = while (\<lambda>(ss, w). \<not> Set.is_empty w)
28520
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   454
    (\<lambda>(ss, w).
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   455
        let p = some_elem w in propa f (step p (ss ! p)) ss (w - {p}))
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   456
    (ss, w)"
45986
c9e50153e5ae moved various set operations to theory Set (resp. Product_Type)
haftmann
parents: 45985
diff changeset
   457
  unfolding iter_def Set.is_empty_def some_elem_def ..
20593
5af400cc64d5 added some stuff for code generation 2
haftmann
parents: 17636
diff changeset
   458
13092
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   459
lemma JVM_sup_unfold [code]:
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   460
 "JVMType.sup S m n = lift2 (Opt.sup
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   461
       (Product.sup (Listn.sup (JType.sup S))
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   462
         (\<lambda>x y. OK (map2 (lift2 (JType.sup S)) x y))))" 
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   463
  apply (unfold JVMType.sup_def JVMType.sl_def Opt.esl_def Err.sl_def
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   464
         stk_esl_def reg_sl_def Product.esl_def  
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   465
         Listn.sl_def upto_esl_def JType.esl_def Err.esl_def) 
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   466
  by simp
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   467
28520
376b9c083b04 tuned code setup
haftmann
parents: 24351
diff changeset
   468
lemmas [code] = JType.sup_def [unfolded exec_lub_def] JVM_le_unfold
44035
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   469
lemmas [code] = lesub_def plussub_def
13092
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   470
45970
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 44890
diff changeset
   471
lemmas [code] =
b6d0cff57d96 adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents: 44890
diff changeset
   472
  JType.sup_def [unfolded exec_lub_def]
44035
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   473
  wf_class_code
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   474
  widen.equation
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   475
  match_exception_entry_def
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   476
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   477
definition test1 where
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   478
  "test1 = test_kil E list_name [Class list_name] (PrimT Void) 3 0
13092
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   479
    [(Suc 0, 2, 8, Xcpt NullPointer)] append_ins"
44035
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   480
definition test2 where
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   481
  "test2 = test_kil E test_name [] (PrimT Void) 3 2 [] make_list_ins"
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   482
51272
9c8d63b4b6be prefer stateless 'ML_val' for tests;
wenzelm
parents: 47399
diff changeset
   483
ML_val {* 
44035
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   484
  @{code test1}; 
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   485
  @{code test2};
322d1657c40c replace old SML code generator by new code generator in MicroJava/JVM and /BV
Andreas Lochbihler
parents: 41413
diff changeset
   486
*}
13092
eae72c47d07f Added example for code generation.
berghofe
parents: 13052
diff changeset
   487
13006
51c5f3f11d16 symbolized
kleing
parents: 12972
diff changeset
   488
end
47399
b72fa7bf9a10 abandoned almost redundant *_foldr lemmas
haftmann
parents: 45990
diff changeset
   489