src/HOL/Bali/Eval.thy
author kleing
Tue, 13 May 2003 08:59:21 +0200
changeset 14024 213dcc39358f
parent 13688 a0b16d42d489
child 14981 e73f8140af78
permissions -rw-r--r--
HOL-Real -> HOL-Complex
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12857
a4386cc9b1c3 tuned header;
wenzelm
parents: 12854
diff changeset
     1
(*  Title:      HOL/Bali/Eval.thy
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     2
    ID:         $Id$
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     3
    Author:     David von Oheimb
12858
wenzelm
parents: 12857
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     5
*)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     6
header {* Operational evaluation (big-step) semantics of Java expressions and 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     7
          statements
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     8
*}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     9
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    10
theory Eval = State + DeclConcepts:
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    11
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    12
text {*
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    13
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    14
improvements over Java Specification 1.0:
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    15
\begin{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    16
\item dynamic method lookup does not need to consider the return type 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    17
      (cf.15.11.4.4)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    18
\item throw raises a NullPointer exception if a null reference is given, and 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    19
      each throw of a standard exception yield a fresh exception object 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    20
      (was not specified)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    21
\item if there is not enough memory even to allocate an OutOfMemory exception,
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    22
  evaluation/execution fails, i.e. simply stops (was not specified)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    23
\item array assignment checks lhs (and may throw exceptions) before evaluating 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    24
      rhs
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    25
\item fixed exact positions of class initializations 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    26
      (immediate at first active use)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    27
\end{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    28
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    29
design issues:
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    30
\begin{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    31
\item evaluation vs. (single-step) transition semantics
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    32
  evaluation semantics chosen, because:
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    33
  \begin{itemize} 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    34
  \item[++] less verbose and therefore easier to read (and to handle in proofs)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    35
  \item[+]  more abstract
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    36
  \item[+]  intermediate values (appearing in recursive rules) need not be 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    37
     stored explicitly, e.g. no call body construct or stack of invocation 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    38
     frames containing local variables and return addresses for method calls 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    39
     needed
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    40
  \item[+]  convenient rule induction for subject reduction theorem
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    41
  \item[-]  no interleaving (for parallelism) can be described
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    42
  \item[-]  stating a property of infinite executions requires the meta-level 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    43
     argument that this property holds for any finite prefixes of it 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    44
     (e.g. stopped using a counter that is decremented to zero and then 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    45
     throwing an exception)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    46
  \end{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    47
\item unified evaluation for variables, expressions, expression lists, 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    48
      statements
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    49
\item the value entry in statement rules is redundant 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    50
\item the value entry in rules is irrelevant in case of exceptions, but its full
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    51
  inclusion helps to make the rule structure independent of exception occurence.
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    52
\item as irrelevant value entries are ignored, it does not matter if they are 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    53
      unique.
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    54
  For simplicity, (fixed) arbitrary values are preferred over "free" values.
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    55
\item the rule format is such that the start state may contain an exception.
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    56
  \begin{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    57
  \item[++] faciliates exception handling
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    58
  \item[+]  symmetry
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    59
  \end{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    60
\item the rules are defined carefully in order to be applicable even in not
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    61
  type-correct situations (yielding undefined values),
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    62
  e.g. @{text "the_Addr (Val (Bool b)) = arbitrary"}.
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    63
  \begin{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    64
  \item[++] fewer rules 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    65
  \item[-]  less readable because of auxiliary functions like @{text the_Addr}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    66
  \end{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    67
  Alternative: "defensive" evaluation throwing some InternalError exception
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    68
               in case of (impossible, for correct programs) type mismatches
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    69
\item there is exactly one rule per syntactic construct
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    70
  \begin{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    71
  \item[+] no redundancy in case distinctions
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    72
  \end{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    73
\item halloc fails iff there is no free heap address. When there is
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    74
  only one free heap address left, it returns an OutOfMemory exception.
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    75
  In this way it is guaranteed that when an OutOfMemory exception is thrown for
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    76
  the first time, there is a free location on the heap to allocate it.
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    77
\item the allocation of objects that represent standard exceptions is deferred 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    78
      until execution of any enclosing catch clause, which is transparent to 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    79
      the program.
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    80
  \begin{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    81
  \item[-]  requires an auxiliary execution relation
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    82
  \item[++] avoids copies of allocation code and awkward case distinctions 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    83
           (whether there is enough memory to allocate the exception) in 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    84
            evaluation rules
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    85
  \end{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    86
\item unfortunately @{text new_Addr} is not directly executable because of 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    87
      Hilbert operator.
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    88
\end{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    89
simplifications:
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    90
\begin{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    91
\item local variables are initialized with default values 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    92
      (no definite assignment)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    93
\item garbage collection not considered, therefore also no finalizers
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    94
\item stack overflow and memory overflow during class initialization not 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    95
      modelled
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    96
\item exceptions in initializations not replaced by ExceptionInInitializerError
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    97
\end{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    98
*}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    99
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   100
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   101
types vvar  =         "val \<times> (val \<Rightarrow> state \<Rightarrow> state)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   102
      vals  =        "(val, vvar, val list) sum3"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   103
translations
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   104
     "vvar" <= (type) "val \<times> (val \<Rightarrow> state \<Rightarrow> state)"
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   105
     "vals" <= (type)"(val, vvar, val list) sum3" 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   106
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   107
text {* To avoid redundancy and to reduce the number of rules, there is only 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   108
 one evaluation rule for each syntactic term. This is also true for variables
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   109
 (e.g. see the rules below for @{text LVar}, @{text FVar} and @{text AVar}). 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   110
 So evaluation of a variable must capture both possible further uses: 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   111
 read (rule @{text Acc}) or write (rule @{text Ass}) to the variable. 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   112
 Therefor a variable evaluates to a special value @{term vvar}, which is 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   113
 a pair, consisting of the current value (for later read access) and an update 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   114
 function (for later write access). Because
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   115
 during assignment to an array variable an exception may occur if the types
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   116
 don't match, the update function is very generic: it transforms the 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   117
 full state. This generic update function causes some technical trouble during
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   118
 some proofs (e.g. type safety, correctness of definite assignment). There we
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   119
 need to prove some additional invariant on this update function to prove the
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   120
 assignment correct, since the update function could potentially alter the whole
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   121
 state in an arbitrary manner. This invariant must be carried around through
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   122
 the whole induction.
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   123
 So for future approaches it may be better not to take
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   124
 such a generic update function, but only to store the address and the kind
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   125
 of variable (array (+ element type), local variable or field) for later 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   126
 assignment. 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   127
*}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   128
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   129
syntax (xsymbols)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   130
  dummy_res :: "vals" ("\<diamondsuit>")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   131
translations
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   132
  "\<diamondsuit>" == "In1 Unit"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   133
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   134
syntax 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   135
  val_inj_vals:: "expr \<Rightarrow> term" ("\<lfloor>_\<rfloor>\<^sub>e" 1000)
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   136
  var_inj_vals::  "var \<Rightarrow> term"  ("\<lfloor>_\<rfloor>\<^sub>v" 1000)
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   137
  lst_inj_vals:: "expr list \<Rightarrow> term" ("\<lfloor>_\<rfloor>\<^sub>l" 1000)
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   138
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   139
translations 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   140
  "\<lfloor>e\<rfloor>\<^sub>e" \<rightharpoonup> "In1 e"
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   141
  "\<lfloor>v\<rfloor>\<^sub>v" \<rightharpoonup> "In2 v"
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   142
  "\<lfloor>es\<rfloor>\<^sub>l" \<rightharpoonup> "In3 es"
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   143
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   144
constdefs
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   145
  arbitrary3 :: "('al + 'ar, 'b, 'c) sum3 \<Rightarrow> vals"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   146
 "arbitrary3 \<equiv> sum3_case (In1 \<circ> sum_case (\<lambda>x. arbitrary) (\<lambda>x. Unit))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   147
                     (\<lambda>x. In2 arbitrary) (\<lambda>x. In3 arbitrary)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   148
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   149
lemma [simp]: "arbitrary3 (In1l x) = In1 arbitrary"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   150
by (simp add: arbitrary3_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   151
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   152
lemma [simp]: "arbitrary3 (In1r x) = \<diamondsuit>"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   153
by (simp add: arbitrary3_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   154
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   155
lemma [simp]: "arbitrary3 (In2  x) = In2 arbitrary"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   156
by (simp add: arbitrary3_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   157
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   158
lemma [simp]: "arbitrary3 (In3  x) = In3 arbitrary"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   159
by (simp add: arbitrary3_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   160
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   161
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   162
section "exception throwing and catching"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   163
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   164
constdefs
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   165
  throw :: "val \<Rightarrow> abopt \<Rightarrow> abopt"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   166
 "throw a' x \<equiv> abrupt_if True (Some (Xcpt (Loc (the_Addr a')))) (np a' x)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   167
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   168
lemma throw_def2: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   169
 "throw a' x = abrupt_if True (Some (Xcpt (Loc (the_Addr a')))) (np a' x)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   170
apply (unfold throw_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   171
apply (simp (no_asm))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   172
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   173
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   174
constdefs
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   175
  fits    :: "prog \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ty \<Rightarrow> bool" ("_,_\<turnstile>_ fits _"[61,61,61,61]60)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   176
 "G,s\<turnstile>a' fits T  \<equiv> (\<exists>rt. T=RefT rt) \<longrightarrow> a'=Null \<or> G\<turnstile>obj_ty(lookup_obj s a')\<preceq>T"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   177
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   178
lemma fits_Null [simp]: "G,s\<turnstile>Null fits T"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   179
by (simp add: fits_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   180
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   181
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   182
lemma fits_Addr_RefT [simp]:
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   183
  "G,s\<turnstile>Addr a fits RefT t = G\<turnstile>obj_ty (the (heap s a))\<preceq>RefT t"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   184
by (simp add: fits_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   185
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   186
lemma fitsD: "\<And>X. G,s\<turnstile>a' fits T \<Longrightarrow> (\<exists>pt. T = PrimT pt) \<or>  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   187
  (\<exists>t. T = RefT t) \<and> a' = Null \<or>  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   188
  (\<exists>t. T = RefT t) \<and> a' \<noteq> Null \<and>  G\<turnstile>obj_ty (lookup_obj s a')\<preceq>T"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   189
apply (unfold fits_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   190
apply (case_tac "\<exists>pt. T = PrimT pt")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   191
apply  simp_all
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   192
apply (case_tac "T")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   193
defer 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   194
apply (case_tac "a' = Null")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   195
apply  simp_all
13601
fd3e3d6b37b2 Adapted to new simplifier.
berghofe
parents: 13462
diff changeset
   196
apply rules
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   197
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   198
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   199
constdefs
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   200
  catch ::"prog \<Rightarrow> state \<Rightarrow> qtname \<Rightarrow> bool"      ("_,_\<turnstile>catch _"[61,61,61]60)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   201
 "G,s\<turnstile>catch C\<equiv>\<exists>xc. abrupt s=Some (Xcpt xc) \<and> 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   202
                    G,store s\<turnstile>Addr (the_Loc xc) fits Class C"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   203
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   204
lemma catch_Norm [simp]: "\<not>G,Norm s\<turnstile>catch tn"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   205
apply (unfold catch_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   206
apply (simp (no_asm))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   207
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   208
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   209
lemma catch_XcptLoc [simp]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   210
  "G,(Some (Xcpt (Loc a)),s)\<turnstile>catch C = G,s\<turnstile>Addr a fits Class C"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   211
apply (unfold catch_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   212
apply (simp (no_asm))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   213
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   214
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   215
lemma catch_Jump [simp]: "\<not>G,(Some (Jump j),s)\<turnstile>catch tn"
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   216
apply (unfold catch_def)
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   217
apply (simp (no_asm))
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   218
done
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   219
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   220
lemma catch_Error [simp]: "\<not>G,(Some (Error e),s)\<turnstile>catch tn"
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   221
apply (unfold catch_def)
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   222
apply (simp (no_asm))
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   223
done
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   224
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   225
constdefs
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   226
  new_xcpt_var :: "vname \<Rightarrow> state \<Rightarrow> state"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   227
 "new_xcpt_var vn \<equiv> 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   228
     \<lambda>(x,s). Norm (lupd(VName vn\<mapsto>Addr (the_Loc (the_Xcpt (the x)))) s)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   229
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   230
lemma new_xcpt_var_def2 [simp]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   231
 "new_xcpt_var vn (x,s) = 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   232
    Norm (lupd(VName vn\<mapsto>Addr (the_Loc (the_Xcpt (the x)))) s)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   233
apply (unfold new_xcpt_var_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   234
apply (simp (no_asm))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   235
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   236
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   237
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   238
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   239
section "misc"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   240
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   241
constdefs
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   242
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   243
  assign     :: "('a \<Rightarrow> state \<Rightarrow> state) \<Rightarrow> 'a \<Rightarrow> state \<Rightarrow> state"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   244
 "assign f v \<equiv> \<lambda>(x,s). let (x',s') = (if x = None then f v else id) (x,s)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   245
		   in  (x',if x' = None then s' else s)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   246
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   247
(*
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   248
lemma assign_Norm_Norm [simp]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   249
"f v \<lparr>abrupt=None,store=s\<rparr> = \<lparr>abrupt=None,store=s'\<rparr> 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   250
 \<Longrightarrow> assign f v \<lparr>abrupt=None,store=s\<rparr> = \<lparr>abrupt=None,store=s'\<rparr>"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   251
by (simp add: assign_def Let_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   252
*)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   253
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   254
lemma assign_Norm_Norm [simp]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   255
"f v (Norm s) = Norm s' \<Longrightarrow> assign f v (Norm s) = Norm s'"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   256
by (simp add: assign_def Let_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   257
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   258
(*
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   259
lemma assign_Norm_Some [simp]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   260
  "\<lbrakk>abrupt (f v \<lparr>abrupt=None,store=s\<rparr>) = Some y\<rbrakk> 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   261
   \<Longrightarrow> assign f v \<lparr>abrupt=None,store=s\<rparr> = \<lparr>abrupt=Some y,store =s\<rparr>"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   262
by (simp add: assign_def Let_def split_beta)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   263
*)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   264
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   265
lemma assign_Norm_Some [simp]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   266
  "\<lbrakk>abrupt (f v (Norm s)) = Some y\<rbrakk> 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   267
   \<Longrightarrow> assign f v (Norm s) = (Some y,s)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   268
by (simp add: assign_def Let_def split_beta)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   269
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   270
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   271
lemma assign_Some [simp]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   272
"assign f v (Some x,s) = (Some x,s)" 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   273
by (simp add: assign_def Let_def split_beta)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   274
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   275
lemma assign_Some1 [simp]:  "\<not> normal s \<Longrightarrow> assign f v s = s" 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   276
by (auto simp add: assign_def Let_def split_beta)
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   277
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   278
lemma assign_supd [simp]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   279
"assign (\<lambda>v. supd (f v)) v (x,s)  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   280
  = (x, if x = None then f v s else s)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   281
apply auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   282
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   283
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   284
lemma assign_raise_if [simp]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   285
  "assign (\<lambda>v (x,s). ((raise_if (b s v) xcpt) x, f v s)) v (x, s) =  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   286
  (raise_if (b s v) xcpt x, if x=None \<and> \<not>b s v then f v s else s)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   287
apply (case_tac "x = None")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   288
apply auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   289
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   290
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   291
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   292
(*
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   293
lemma assign_raise_if [simp]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   294
  "assign (\<lambda>v s. \<lparr>abrupt=(raise_if (b (store s) v) xcpt) (abrupt s),
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   295
                  store = f v (store s)\<rparr>) v s =  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   296
  \<lparr>abrupt=raise_if (b (store s) v) xcpt (abrupt s),
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   297
   store= if (abrupt s)=None \<and> \<not>b (store s) v 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   298
             then f v (store s) else (store s)\<rparr>"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   299
apply (case_tac "abrupt s = None")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   300
apply auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   301
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   302
*)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   303
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   304
constdefs
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   305
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   306
  init_comp_ty :: "ty \<Rightarrow> stmt"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   307
 "init_comp_ty T \<equiv> if (\<exists>C. T = Class C) then Init (the_Class T) else Skip"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   308
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   309
lemma init_comp_ty_PrimT [simp]: "init_comp_ty (PrimT pt) = Skip"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   310
apply (unfold init_comp_ty_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   311
apply (simp (no_asm))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   312
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   313
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   314
constdefs
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   315
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   316
 invocation_class  :: "inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> qtname"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   317
 "invocation_class m s a' statT 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   318
    \<equiv> (case m of
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   319
         Static \<Rightarrow> if (\<exists> statC. statT = ClassT statC) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   320
                      then the_Class (RefT statT) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   321
                      else Object
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   322
       | SuperM \<Rightarrow> the_Class (RefT statT)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   323
       | IntVir \<Rightarrow> obj_class (lookup_obj s a'))"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   324
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   325
invocation_declclass::"prog \<Rightarrow> inv_mode \<Rightarrow> st \<Rightarrow> val \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> qtname"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   326
"invocation_declclass G m s a' statT sig 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   327
   \<equiv> declclass (the (dynlookup G statT 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   328
                                (invocation_class m s a' statT)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   329
                                sig))" 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   330
  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   331
lemma invocation_class_IntVir [simp]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   332
"invocation_class IntVir s a' statT = obj_class (lookup_obj s a')"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   333
by (simp add: invocation_class_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   334
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   335
lemma dynclass_SuperM [simp]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   336
 "invocation_class SuperM s a' statT = the_Class (RefT statT)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   337
by (simp add: invocation_class_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   338
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   339
lemma invocation_class_Static [simp]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   340
  "invocation_class Static s a' statT = (if (\<exists> statC. statT = ClassT statC) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   341
                                            then the_Class (RefT statT) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   342
                                            else Object)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   343
by (simp add: invocation_class_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   344
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   345
constdefs
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   346
  init_lvars :: "prog \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> inv_mode \<Rightarrow> val \<Rightarrow> val list \<Rightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   347
		   state \<Rightarrow> state"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   348
 "init_lvars G C sig mode a' pvs 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   349
   \<equiv> \<lambda> (x,s). 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   350
      let m = mthd (the (methd G C sig));
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   351
          l = \<lambda> k. 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   352
              (case k of
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   353
                 EName e 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   354
                   \<Rightarrow> (case e of 
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   355
                         VNam v \<Rightarrow> (empty ((pars m)[\<mapsto>]pvs)) v
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   356
                       | Res    \<Rightarrow> None)
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   357
               | This 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   358
                   \<Rightarrow> (if mode=Static then None else Some a'))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   359
      in set_lvars l (if mode = Static then x else np a' x,s)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   360
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   361
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   362
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   363
lemma init_lvars_def2: --{* better suited for simplification *} 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   364
"init_lvars G C sig mode a' pvs (x,s) =  
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   365
  set_lvars 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   366
    (\<lambda> k. 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   367
       (case k of
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   368
          EName e 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   369
            \<Rightarrow> (case e of 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   370
                  VNam v 
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   371
                  \<Rightarrow> (empty ((pars (mthd (the (methd G C sig))))[\<mapsto>]pvs)) v
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   372
               | Res \<Rightarrow> None)
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   373
        | This 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   374
            \<Rightarrow> (if mode=Static then None else Some a')))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   375
    (if mode = Static then x else np a' x,s)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   376
apply (unfold init_lvars_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   377
apply (simp (no_asm) add: Let_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   378
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   379
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   380
constdefs
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   381
  body :: "prog \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> expr"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   382
 "body G C sig \<equiv> let m = the (methd G C sig) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   383
                 in Body (declclass m) (stmt (mbody (mthd m)))"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   384
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   385
lemma body_def2: --{* better suited for simplification *} 
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   386
"body G C sig = Body  (declclass (the (methd G C sig))) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   387
                      (stmt (mbody (mthd (the (methd G C sig)))))"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   388
apply (unfold body_def Let_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   389
apply auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   390
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   391
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   392
section "variables"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   393
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   394
constdefs
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   395
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   396
  lvar :: "lname \<Rightarrow> st \<Rightarrow> vvar"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   397
 "lvar vn s \<equiv> (the (locals s vn), \<lambda>v. supd (lupd(vn\<mapsto>v)))"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   398
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   399
  fvar :: "qtname \<Rightarrow> bool \<Rightarrow> vname \<Rightarrow> val \<Rightarrow> state \<Rightarrow> vvar \<times> state"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   400
 "fvar C stat fn a' s 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   401
    \<equiv> let (oref,xf) = if stat then (Stat C,id)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   402
                              else (Heap (the_Addr a'),np a');
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   403
	          n = Inl (fn,C); 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   404
                  f = (\<lambda>v. supd (upd_gobj oref n v)) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   405
      in ((the (values (the (globs (store s) oref)) n),f),abupd xf s)"
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   406
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   407
  avar :: "prog \<Rightarrow> val \<Rightarrow> val \<Rightarrow> state \<Rightarrow> vvar \<times> state"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   408
 "avar G i' a' s 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   409
    \<equiv> let   oref = Heap (the_Addr a'); 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   410
               i = the_Intg i'; 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   411
               n = Inr i;
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   412
        (T,k,cs) = the_Arr (globs (store s) oref); 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   413
               f = (\<lambda>v (x,s). (raise_if (\<not>G,s\<turnstile>v fits T) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   414
                                           ArrStore x
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   415
                              ,upd_gobj oref n v s)) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   416
      in ((the (cs n),f)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   417
         ,abupd (raise_if (\<not>i in_bounds k) IndOutBound \<circ> np a') s)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   418
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   419
lemma fvar_def2: --{* better suited for simplification *} 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   420
"fvar C stat fn a' s =  
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   421
  ((the 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   422
     (values 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   423
      (the (globs (store s) (if stat then Stat C else Heap (the_Addr a')))) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   424
      (Inl (fn,C)))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   425
   ,(\<lambda>v. supd (upd_gobj (if stat then Stat C else Heap (the_Addr a')) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   426
                        (Inl (fn,C)) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   427
                        v)))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   428
  ,abupd (if stat then id else np a') s)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   429
  "
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   430
apply (unfold fvar_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   431
apply (simp (no_asm) add: Let_def split_beta)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   432
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   433
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   434
lemma avar_def2: --{* better suited for simplification *} 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   435
"avar G i' a' s =  
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   436
  ((the ((snd(snd(the_Arr (globs (store s) (Heap (the_Addr a')))))) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   437
           (Inr (the_Intg i')))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   438
   ,(\<lambda>v (x,s').  (raise_if (\<not>G,s'\<turnstile>v fits (fst(the_Arr (globs (store s)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   439
                                                   (Heap (the_Addr a')))))) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   440
                            ArrStore x
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   441
                 ,upd_gobj  (Heap (the_Addr a')) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   442
                               (Inr (the_Intg i')) v s')))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   443
  ,abupd (raise_if (\<not>(the_Intg i') in_bounds (fst(snd(the_Arr (globs (store s) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   444
                   (Heap (the_Addr a'))))))) IndOutBound \<circ> np a')
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   445
          s)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   446
apply (unfold avar_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   447
apply (simp (no_asm) add: Let_def split_beta)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   448
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   449
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   450
constdefs
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   451
check_field_access::
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   452
"prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> vname \<Rightarrow> bool \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   453
"check_field_access G accC statDeclC fn stat a' s
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   454
 \<equiv> let oref = if stat then Stat statDeclC
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   455
                      else Heap (the_Addr a');
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   456
       dynC = case oref of
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   457
                   Heap a \<Rightarrow> obj_class (the (globs (store s) oref))
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   458
                 | Stat C \<Rightarrow> C;
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   459
       f    = (the (table_of (DeclConcepts.fields G dynC) (fn,statDeclC)))
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   460
   in abupd 
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   461
        (error_if (\<not> G\<turnstile>Field fn (statDeclC,f) in dynC dyn_accessible_from accC)
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   462
                  AccessViolation)
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   463
        s"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   464
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   465
constdefs
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   466
check_method_access:: 
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   467
  "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> inv_mode \<Rightarrow>  sig \<Rightarrow> val \<Rightarrow> state \<Rightarrow> state"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   468
"check_method_access G accC statT mode sig  a' s
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   469
 \<equiv> let invC = invocation_class mode (store s) a' statT;
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   470
       dynM = the (dynlookup G statT invC sig)
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   471
   in abupd 
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   472
        (error_if (\<not> G\<turnstile>Methd sig dynM in invC dyn_accessible_from accC)
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   473
                  AccessViolation)
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   474
        s"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   475
       
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   476
section "evaluation judgments"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   477
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   478
consts
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   479
  eval   :: "prog \<Rightarrow> (state \<times> term    \<times> vals \<times> state) set"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   480
  halloc::  "prog \<Rightarrow> (state \<times> obj_tag \<times> loc  \<times> state) set"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   481
  sxalloc:: "prog \<Rightarrow> (state                  \<times> state) set"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   482
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   483
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   484
syntax
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   485
eval ::"[prog,state,term,vals*state]=>bool"("_|-_ -_>-> _"  [61,61,80,   61]60)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   486
exec ::"[prog,state,stmt      ,state]=>bool"("_|-_ -_-> _"   [61,61,65,   61]60)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   487
evar ::"[prog,state,var  ,vvar,state]=>bool"("_|-_ -_=>_-> _"[61,61,90,61,61]60)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   488
eval_::"[prog,state,expr ,val, state]=>bool"("_|-_ -_->_-> _"[61,61,80,61,61]60)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   489
evals::"[prog,state,expr list ,
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   490
		    val  list ,state]=>bool"("_|-_ -_#>_-> _"[61,61,61,61,61]60)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   491
hallo::"[prog,state,obj_tag,
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   492
	             loc,state]=>bool"("_|-_ -halloc _>_-> _"[61,61,61,61,61]60)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   493
sallo::"[prog,state        ,state]=>bool"("_|-_ -sxalloc-> _"[61,61,      61]60)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   494
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   495
syntax (xsymbols)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   496
eval ::"[prog,state,term,vals\<times>state]\<Rightarrow>bool" ("_\<turnstile>_ \<midarrow>_\<succ>\<rightarrow> _"  [61,61,80,   61]60)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   497
exec ::"[prog,state,stmt      ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_\<rightarrow> _"   [61,61,65,   61]60)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   498
evar ::"[prog,state,var  ,vvar,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_=\<succ>_\<rightarrow> _"[61,61,90,61,61]60)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   499
eval_::"[prog,state,expr ,val ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_-\<succ>_\<rightarrow> _"[61,61,80,61,61]60)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   500
evals::"[prog,state,expr list ,
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   501
		    val  list ,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>_\<doteq>\<succ>_\<rightarrow> _"[61,61,61,61,61]60)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   502
hallo::"[prog,state,obj_tag,
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   503
	             loc,state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>halloc _\<succ>_\<rightarrow> _"[61,61,61,61,61]60)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   504
sallo::"[prog,state,        state]\<Rightarrow>bool"("_\<turnstile>_ \<midarrow>sxalloc\<rightarrow> _"[61,61,      61]60)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   505
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   506
translations
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   507
  "G\<turnstile>s \<midarrow>t   \<succ>\<rightarrow>  w___s' " == "(s,t,w___s') \<in> eval G"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   508
  "G\<turnstile>s \<midarrow>t   \<succ>\<rightarrow> (w,  s')" <= "(s,t,w,  s') \<in> eval G"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   509
  "G\<turnstile>s \<midarrow>t   \<succ>\<rightarrow> (w,x,s')" <= "(s,t,w,x,s') \<in> eval G"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   510
  "G\<turnstile>s \<midarrow>c    \<rightarrow>  (x,s')" <= "G\<turnstile>s \<midarrow>In1r c\<succ>\<rightarrow> (\<diamondsuit>,x,s')"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   511
  "G\<turnstile>s \<midarrow>c    \<rightarrow>     s' " == "G\<turnstile>s \<midarrow>In1r c\<succ>\<rightarrow> (\<diamondsuit>  ,  s')"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   512
  "G\<turnstile>s \<midarrow>e-\<succ>v \<rightarrow>  (x,s')" <= "G\<turnstile>s \<midarrow>In1l e\<succ>\<rightarrow> (In1 v ,x,s')"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   513
  "G\<turnstile>s \<midarrow>e-\<succ>v \<rightarrow>     s' " == "G\<turnstile>s \<midarrow>In1l e\<succ>\<rightarrow> (In1 v ,  s')"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   514
  "G\<turnstile>s \<midarrow>e=\<succ>vf\<rightarrow>  (x,s')" <= "G\<turnstile>s \<midarrow>In2  e\<succ>\<rightarrow> (In2 vf,x,s')"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   515
  "G\<turnstile>s \<midarrow>e=\<succ>vf\<rightarrow>     s' " == "G\<turnstile>s \<midarrow>In2  e\<succ>\<rightarrow> (In2 vf,  s')"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   516
  "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v \<rightarrow>  (x,s')" <= "G\<turnstile>s \<midarrow>In3  e\<succ>\<rightarrow> (In3 v ,x,s')"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   517
  "G\<turnstile>s \<midarrow>e\<doteq>\<succ>v \<rightarrow>     s' " == "G\<turnstile>s \<midarrow>In3  e\<succ>\<rightarrow> (In3 v ,  s')"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   518
  "G\<turnstile>s \<midarrow>halloc oi\<succ>a\<rightarrow> (x,s')" <= "(s,oi,a,x,s') \<in> halloc G"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   519
  "G\<turnstile>s \<midarrow>halloc oi\<succ>a\<rightarrow>    s' " == "(s,oi,a,  s') \<in> halloc G"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   520
  "G\<turnstile>s \<midarrow>sxalloc\<rightarrow>     (x,s')" <= "(s     ,x,s') \<in> sxalloc G"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   521
  "G\<turnstile>s \<midarrow>sxalloc\<rightarrow>        s' " == "(s     ,  s') \<in> sxalloc G"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   522
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   523
inductive "halloc G" intros --{* allocating objects on the heap, cf. 12.5 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   524
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   525
  Abrupt: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   526
  "G\<turnstile>(Some x,s) \<midarrow>halloc oi\<succ>arbitrary\<rightarrow> (Some x,s)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   527
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   528
  New:	  "\<lbrakk>new_Addr (heap s) = Some a; 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   529
	    (x,oi') = (if atleast_free (heap s) (Suc (Suc 0)) then (None,oi)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   530
		       else (Some (Xcpt (Loc a)),CInst (SXcpt OutOfMemory)))\<rbrakk>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   531
            \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   532
	    G\<turnstile>Norm s \<midarrow>halloc oi\<succ>a\<rightarrow> (x,init_obj G oi' (Heap a) s)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   533
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   534
inductive "sxalloc G" intros --{* allocating exception objects for
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   535
	 	 	      standard exceptions (other than OutOfMemory) *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   536
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   537
  Norm:	 "G\<turnstile> Norm              s   \<midarrow>sxalloc\<rightarrow>  Norm             s"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   538
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   539
  Jmp:   "G\<turnstile>(Some (Jump j), s)  \<midarrow>sxalloc\<rightarrow> (Some (Jump j), s)"
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   540
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   541
  Error: "G\<turnstile>(Some (Error e), s)  \<midarrow>sxalloc\<rightarrow> (Some (Error e), s)"
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   542
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   543
  XcptL: "G\<turnstile>(Some (Xcpt (Loc a) ),s)  \<midarrow>sxalloc\<rightarrow> (Some (Xcpt (Loc a)),s)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   544
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   545
  SXcpt: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>halloc (CInst (SXcpt xn))\<succ>a\<rightarrow> (x,s1)\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   546
	  G\<turnstile>(Some (Xcpt (Std xn)),s0) \<midarrow>sxalloc\<rightarrow> (Some (Xcpt (Loc a)),s1)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   547
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   548
13384
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   549
text {* 
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   550
\par
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   551
*} (* dummy text command to break paragraph for latex;
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   552
              large paragraphs exhaust memory of debian pdflatex *)
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   553
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   554
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   555
inductive "eval G" intros
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   556
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   557
--{* propagation of abrupt completion *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   558
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   559
  --{* cf. 14.1, 15.5 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   560
  Abrupt: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   561
   "G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<rightarrow> (arbitrary3 t,(Some xc,s))"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   562
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   563
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   564
--{* execution of statements *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   565
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   566
  --{* cf. 14.5 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   567
  Skip:	 			    "G\<turnstile>Norm s \<midarrow>Skip\<rightarrow> Norm s"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   568
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   569
  --{* cf. 14.7 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   570
  Expr:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   571
				  G\<turnstile>Norm s0 \<midarrow>Expr e\<rightarrow> s1"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   572
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   573
  Lab:  "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c \<rightarrow> s1\<rbrakk> \<Longrightarrow>
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   574
                                G\<turnstile>Norm s0 \<midarrow>l\<bullet> c\<rightarrow> abupd (absorb l) s1"
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   575
  --{* cf. 14.2 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   576
  Comp:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1 \<rightarrow> s1;
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   577
	  G\<turnstile>     s1 \<midarrow>c2 \<rightarrow> s2\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   578
				 G\<turnstile>Norm s0 \<midarrow>c1;; c2\<rightarrow> s2"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   579
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   580
  --{* cf. 14.8.2 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   581
  If:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1;
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   582
	  G\<turnstile>     s1\<midarrow>(if the_Bool b then c1 else c2)\<rightarrow> s2\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   583
		       G\<turnstile>Norm s0 \<midarrow>If(e) c1 Else c2 \<rightarrow> s2"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   584
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   585
  --{* cf. 14.10, 14.10.1 *}
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   586
  
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   587
  --{* A continue jump from the while body @{term c} is handled by 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   588
     this rule. If a continue jump with the proper label was invoked inside 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   589
     @{term c} this label (Cont l) is deleted out of the abrupt component of 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   590
     the state before the iterative evaluation of the while statement.
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   591
     A break jump is handled by the Lab Statement @{text "Lab l (while\<dots>)"}.
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   592
  *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   593
  Loop:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>b\<rightarrow> s1;
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   594
	  if the_Bool b 
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   595
             then (G\<turnstile>s1 \<midarrow>c\<rightarrow> s2 \<and> 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   596
                   G\<turnstile>(abupd (absorb (Cont l)) s2) \<midarrow>l\<bullet> While(e) c\<rightarrow> s3)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   597
	     else s3 = s1\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   598
			      G\<turnstile>Norm s0 \<midarrow>l\<bullet> While(e) c\<rightarrow> s3"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   599
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   600
  Jmp: "G\<turnstile>Norm s \<midarrow>Jmp j\<rightarrow> (Some (Jump j), s)"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   601
   
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   602
  --{* cf. 14.16 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   603
  Throw: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   604
				 G\<turnstile>Norm s0 \<midarrow>Throw e\<rightarrow> abupd (throw a') s1"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   605
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   606
  --{* cf. 14.18.1 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   607
  Try:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> s1; G\<turnstile>s1 \<midarrow>sxalloc\<rightarrow> s2; 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   608
	  if G,s2\<turnstile>catch C then G\<turnstile>new_xcpt_var vn s2 \<midarrow>c2\<rightarrow> s3 else s3 = s2\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   609
		  G\<turnstile>Norm s0 \<midarrow>Try c1 Catch(C vn) c2\<rightarrow> s3"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   610
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   611
  --{* cf. 14.18.2 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   612
  Fin:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>c1\<rightarrow> (x1,s1);
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   613
	  G\<turnstile>Norm s1 \<midarrow>c2\<rightarrow> s2;
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   614
          s3=(if (\<exists> err. x1=Some (Error err)) 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   615
              then (x1,s1) 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   616
              else abupd (abrupt_if (x1\<noteq>None) x1) s2) \<rbrakk> 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   617
          \<Longrightarrow>
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   618
          G\<turnstile>Norm s0 \<midarrow>c1 Finally c2\<rightarrow> s3"
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   619
  --{* cf. 12.4.2, 8.5 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   620
  Init:	"\<lbrakk>the (class G C) = c;
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   621
	  if inited C (globs s0) then s3 = Norm s0
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   622
	  else (G\<turnstile>Norm (init_class_obj G C s0) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   623
		  \<midarrow>(if C = Object then Skip else Init (super c))\<rightarrow> s1 \<and>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   624
	       G\<turnstile>set_lvars empty s1 \<midarrow>init c\<rightarrow> s2 \<and> s3 = restore_lvars s1 s2)\<rbrakk> 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   625
              \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   626
		 G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s3"
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   627
   --{* This class initialisation rule is a little bit inaccurate. Look at the
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   628
      exact sequence:
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   629
      (1) The current class object (the static fields) are initialised
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   630
           (@{text init_class_obj}),
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   631
      (2) the superclasses are initialised,
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   632
      (3) the static initialiser of the current class is invoked.
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   633
      More precisely we should expect another ordering, namely 2 1 3.
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   634
      But we can't just naively toggle 1 and 2. By calling 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   635
      @{text init_class_obj} 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   636
      before initialising the superclasses, we also implicitly record that
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   637
      we have started to initialise the current class (by setting an 
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   638
      value for the class object). This becomes 
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   639
      crucial for the completeness proof of the axiomatic semantics 
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   640
      @{text "AxCompl.thy"}. Static initialisation requires an induction on 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   641
      the number of classes not yet initialised (or to be more precise, 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   642
      classes were the initialisation has not yet begun). 
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   643
      So we could first assign a dummy value to the class before
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   644
      superclass initialisation and afterwards set the correct values.
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   645
      But as long as we don't take memory overflow into account 
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   646
      when allocating class objects, we can leave things as they are for 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   647
      convenience. 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   648
   *}
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   649
--{* evaluation of expressions *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   650
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   651
  --{* cf. 15.8.1, 12.4.1 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   652
  NewC:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s1;
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   653
	  G\<turnstile>     s1 \<midarrow>halloc (CInst C)\<succ>a\<rightarrow> s2\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   654
	                          G\<turnstile>Norm s0 \<midarrow>NewC C-\<succ>Addr a\<rightarrow> s2"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   655
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   656
  --{* cf. 15.9.1, 12.4.1 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   657
  NewA:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>init_comp_ty T\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>i'\<rightarrow> s2; 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   658
	  G\<turnstile>abupd (check_neg i') s2 \<midarrow>halloc (Arr T (the_Intg i'))\<succ>a\<rightarrow> s3\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   659
	                        G\<turnstile>Norm s0 \<midarrow>New T[e]-\<succ>Addr a\<rightarrow> s3"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   660
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   661
  --{* cf. 15.15 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   662
  Cast:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   663
	  s2 = abupd (raise_if (\<not>G,store s1\<turnstile>v fits T) ClassCast) s1\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   664
			        G\<turnstile>Norm s0 \<midarrow>Cast T e-\<succ>v\<rightarrow> s2"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   665
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   666
  --{* cf. 15.19.2 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   667
  Inst:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1;
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   668
	  b = (v\<noteq>Null \<and> G,store s1\<turnstile>v fits RefT T)\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   669
			      G\<turnstile>Norm s0 \<midarrow>e InstOf T-\<succ>Bool b\<rightarrow> s1"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   670
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   671
  --{* cf. 15.7.1 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   672
  Lit:	"G\<turnstile>Norm s \<midarrow>Lit v-\<succ>v\<rightarrow> Norm s"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   673
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   674
  UnOp: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>v\<rightarrow> s1\<rbrakk> 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   675
         \<Longrightarrow> G\<turnstile>Norm s0 \<midarrow>UnOp unop e-\<succ>(eval_unop unop v)\<rightarrow> s1"
13384
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   676
 
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   677
  BinOp: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1; 
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   678
          G\<turnstile>s1 \<midarrow>(if need_second_arg binop v1 then (In1l e2) else (In1r Skip))
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   679
                \<succ>\<rightarrow> (In1 v2,s2)
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   680
          \<rbrakk> 
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   681
         \<Longrightarrow> G\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>(eval_binop binop v1 v2)\<rightarrow> s2"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   682
   
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   683
  --{* cf. 15.10.2 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   684
  Super: "G\<turnstile>Norm s \<midarrow>Super-\<succ>val_this s\<rightarrow> Norm s"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   685
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   686
  --{* cf. 15.2 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   687
  Acc:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(v,f)\<rightarrow> s1\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   688
	                          G\<turnstile>Norm s0 \<midarrow>Acc va-\<succ>v\<rightarrow> s1"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   689
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   690
  --{* cf. 15.25.1 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   691
  Ass:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>va=\<succ>(w,f)\<rightarrow> s1;
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   692
          G\<turnstile>     s1 \<midarrow>e-\<succ>v  \<rightarrow> s2\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   693
				   G\<turnstile>Norm s0 \<midarrow>va:=e-\<succ>v\<rightarrow> assign f v s2"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   694
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   695
  --{* cf. 15.24 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   696
  Cond:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e0-\<succ>b\<rightarrow> s1;
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   697
          G\<turnstile>     s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<rightarrow> s2\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   698
			    G\<turnstile>Norm s0 \<midarrow>e0 ? e1 : e2-\<succ>v\<rightarrow> s2"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   699
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   700
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   701
-- {* The interplay of  @{term Call}, @{term Methd} and @{term Body}:
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   702
      Method invocation is split up into these three rules:
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   703
      \begin{itemize}
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   704
      \item [@{term Call}] Calculates the target address and evaluates the
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   705
                           arguments of the method, and then performs dynamic
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   706
                           or static lookup of the method, corresponding to the
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   707
                           call mode. Then the @{term Methd} rule is evaluated
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   708
                           on the calculated declaration class of the method
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   709
                           invocation.
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   710
      \item [@{term Methd}] A syntactic bridge for the folded method body.
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   711
                            It is used by the axiomatic semantics to add the
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   712
                            proper hypothesis for recursive calls of the method.
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   713
      \item [@{term Body}] An extra syntactic entity for the unfolded method
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   714
                           body was introduced to properly trigger class 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   715
                           initialisation. Without class initialisation we 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   716
                           could just evaluate the body statement. 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   717
      \end{itemize}
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   718
   *}
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   719
  --{* cf. 15.11.4.1, 15.11.4.2, 15.11.4.4, 15.11.4.5 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   720
  Call:	
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   721
  "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1; G\<turnstile>s1 \<midarrow>args\<doteq>\<succ>vs\<rightarrow> s2;
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   722
    D = invocation_declclass G mode (store s2) a' statT \<lparr>name=mn,parTs=pTs\<rparr>;
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   723
    s3=init_lvars G D \<lparr>name=mn,parTs=pTs\<rparr> mode a' vs s2;
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   724
    s3' = check_method_access G accC statT mode \<lparr>name=mn,parTs=pTs\<rparr> a' s3;
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   725
    G\<turnstile>s3' \<midarrow>Methd D \<lparr>name=mn,parTs=pTs\<rparr>-\<succ>v\<rightarrow> s4\<rbrakk>
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   726
   \<Longrightarrow>
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   727
       G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn({pTs}args)-\<succ>v\<rightarrow> (restore_lvars s2 s4)"
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   728
--{* The accessibility check is after @{term init_lvars}, to keep it simple. 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   729
   @{term init_lvars} already tests for the absence of a null-pointer 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   730
   reference in case of an instance method invocation.
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   731
*}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   732
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   733
  Methd:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>body G D sig-\<succ>v\<rightarrow> s1\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   734
				G\<turnstile>Norm s0 \<midarrow>Methd D sig-\<succ>v\<rightarrow> s1"
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   735
  
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   736
  Body:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init D\<rightarrow> s1; G\<turnstile>s1 \<midarrow>c\<rightarrow> s2; 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   737
          s3 = (if (\<exists> l. abrupt s2 = Some (Jump (Break l)) \<or>  
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   738
                         abrupt s2 = Some (Jump (Cont l)))
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   739
                  then abupd (\<lambda> x. Some (Error CrossMethodJump)) s2 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   740
                  else s2)\<rbrakk> \<Longrightarrow>
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   741
           G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>the (locals (store s2) Result)
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   742
              \<rightarrow>abupd (absorb Ret) s3"
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   743
  --{* cf. 14.15, 12.4.1 *}
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   744
  --{* We filter out a break/continue in @{term s2}, so that we can proof 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   745
     definite assignment
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   746
     correct, without the need of conformance of the state. By this the
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   747
     different parts of the typesafety proof can be disentangled a little. *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   748
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   749
--{* evaluation of variables *}
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   750
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   751
  --{* cf. 15.13.1, 15.7.2 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   752
  LVar:	"G\<turnstile>Norm s \<midarrow>LVar vn=\<succ>lvar vn s\<rightarrow> Norm s"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   753
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   754
  --{* cf. 15.10.1, 12.4.1 *}
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   755
  FVar:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init statDeclC\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e-\<succ>a\<rightarrow> s2;
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   756
	  (v,s2') = fvar statDeclC stat fn a s2;
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   757
          s3 = check_field_access G accC statDeclC fn stat a s2' \<rbrakk> \<Longrightarrow>
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   758
	  G\<turnstile>Norm s0 \<midarrow>{accC,statDeclC,stat}e..fn=\<succ>v\<rightarrow> s3"
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   759
 --{* The accessibility check is after @{term fvar}, to keep it simple. 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   760
    @{term fvar} already tests for the absence of a null-pointer reference 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   761
    in case of an instance field
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   762
  *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   763
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   764
  --{* cf. 15.12.1, 15.25.1 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   765
  AVar:	"\<lbrakk>G\<turnstile> Norm s0 \<midarrow>e1-\<succ>a\<rightarrow> s1; G\<turnstile>s1 \<midarrow>e2-\<succ>i\<rightarrow> s2;
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   766
	  (v,s2') = avar G i a s2\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   767
	              G\<turnstile>Norm s0 \<midarrow>e1.[e2]=\<succ>v\<rightarrow> s2'"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   768
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   769
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   770
--{* evaluation of expression lists *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   771
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   772
  --{* cf. 15.11.4.2 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   773
  Nil:
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   774
				    "G\<turnstile>Norm s0 \<midarrow>[]\<doteq>\<succ>[]\<rightarrow> Norm s0"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   775
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   776
  --{* cf. 15.6.4 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   777
  Cons:	"\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e -\<succ> v \<rightarrow> s1;
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   778
          G\<turnstile>     s1 \<midarrow>es\<doteq>\<succ>vs\<rightarrow> s2\<rbrakk> \<Longrightarrow>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   779
				   G\<turnstile>Norm s0 \<midarrow>e#es\<doteq>\<succ>v#vs\<rightarrow> s2"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   780
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   781
(* Rearrangement of premisses:
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   782
[0,1(Abrupt),2(Skip),8(Jmp),4(Lab),30(Nil),31(Cons),27(LVar),17(Cast),18(Inst),
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   783
 17(Lit),18(UnOp),19(BinOp),20(Super),21(Acc),3(Expr),5(Comp),25(Methd),26(Body),23(Cond),6(If),
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   784
 7(Loop),11(Fin),9(Throw),13(NewC),14(NewA),12(Init),22(Ass),10(Try),28(FVar),
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   785
 29(AVar),24(Call)]
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   786
*)
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   787
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   788
ML {*
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   789
bind_thm ("eval_induct_", rearrange_prems 
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   790
[0,1,2,8,4,30,31,27,15,16,
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   791
 17,18,19,20,21,3,5,25,26,23,6,
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   792
 7,11,9,13,14,12,22,10,28,
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   793
 29,24] (thm "eval.induct"))
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   794
*}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   795
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   796
13384
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   797
text {* 
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   798
\par
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   799
*} (* dummy text command to break paragraph for latex;
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   800
              large paragraphs exhaust memory of debian pdflatex *)
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   801
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   802
lemmas eval_induct = eval_induct_ [split_format and and and and and and and and
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   803
   and and and and and and s1 (* Acc *) and and s2 (* Comp *) and and and and 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   804
   and and 
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   805
   s2 (* Fin *) and and s2 (* NewC *)] 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   806
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   807
declare split_if     [split del] split_if_asm     [split del] 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   808
        option.split [split del] option.split_asm [split del]
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   809
inductive_cases halloc_elim_cases: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   810
  "G\<turnstile>(Some xc,s) \<midarrow>halloc oi\<succ>a\<rightarrow> s'"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   811
  "G\<turnstile>(Norm    s) \<midarrow>halloc oi\<succ>a\<rightarrow> s'"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   812
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   813
inductive_cases sxalloc_elim_cases:
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   814
 	"G\<turnstile> Norm                 s  \<midarrow>sxalloc\<rightarrow> s'"
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   815
        "G\<turnstile>(Some (Jump j),s) \<midarrow>sxalloc\<rightarrow> s'"
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   816
 	"G\<turnstile>(Some (Error e),s) \<midarrow>sxalloc\<rightarrow> s'"
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   817
        "G\<turnstile>(Some (Xcpt (Loc a )),s) \<midarrow>sxalloc\<rightarrow> s'"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   818
 	"G\<turnstile>(Some (Xcpt (Std xn)),s) \<midarrow>sxalloc\<rightarrow> s'"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   819
inductive_cases sxalloc_cases: "G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s'"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   820
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   821
lemma sxalloc_elim_cases2: "\<lbrakk>G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s';  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   822
       \<And>s.   \<lbrakk>s' = Norm s\<rbrakk> \<Longrightarrow> P;  
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   823
       \<And>j s. \<lbrakk>s' = (Some (Jump j),s)\<rbrakk> \<Longrightarrow> P;
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   824
       \<And>e s. \<lbrakk>s' = (Some (Error e),s)\<rbrakk> \<Longrightarrow> P;
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   825
       \<And>a s. \<lbrakk>s' = (Some (Xcpt (Loc a)),s)\<rbrakk> \<Longrightarrow> P  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   826
      \<rbrakk> \<Longrightarrow> P"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   827
apply cut_tac 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   828
apply (erule sxalloc_cases)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   829
apply blast+
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   830
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   831
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   832
declare not_None_eq [simp del] (* IntDef.Zero_def [simp del] *)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   833
declare split_paired_All [simp del] split_paired_Ex [simp del]
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   834
ML_setup {*
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   835
simpset_ref() := simpset() delloop "split_all_tac"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   836
*}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   837
inductive_cases eval_cases: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> vs'"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   838
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   839
inductive_cases eval_elim_cases [cases set]:
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   840
        "G\<turnstile>(Some xc,s) \<midarrow>t                              \<succ>\<rightarrow> vs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   841
	"G\<turnstile>Norm s \<midarrow>In1r Skip                           \<succ>\<rightarrow> xs'"
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   842
        "G\<turnstile>Norm s \<midarrow>In1r (Jmp j)                        \<succ>\<rightarrow> xs'"
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   843
        "G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> c)                         \<succ>\<rightarrow> xs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   844
	"G\<turnstile>Norm s \<midarrow>In3  ([])                           \<succ>\<rightarrow> vs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   845
	"G\<turnstile>Norm s \<midarrow>In3  (e#es)                         \<succ>\<rightarrow> vs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   846
	"G\<turnstile>Norm s \<midarrow>In1l (Lit w)                        \<succ>\<rightarrow> vs'"
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   847
        "G\<turnstile>Norm s \<midarrow>In1l (UnOp unop e)                  \<succ>\<rightarrow> vs'"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   848
        "G\<turnstile>Norm s \<midarrow>In1l (BinOp binop e1 e2)            \<succ>\<rightarrow> vs'"
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   849
	"G\<turnstile>Norm s \<midarrow>In2  (LVar vn)                      \<succ>\<rightarrow> vs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   850
	"G\<turnstile>Norm s \<midarrow>In1l (Cast T e)                     \<succ>\<rightarrow> vs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   851
	"G\<turnstile>Norm s \<midarrow>In1l (e InstOf T)                   \<succ>\<rightarrow> vs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   852
	"G\<turnstile>Norm s \<midarrow>In1l (Super)                        \<succ>\<rightarrow> vs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   853
	"G\<turnstile>Norm s \<midarrow>In1l (Acc va)                       \<succ>\<rightarrow> vs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   854
	"G\<turnstile>Norm s \<midarrow>In1r (Expr e)                       \<succ>\<rightarrow> xs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   855
	"G\<turnstile>Norm s \<midarrow>In1r (c1;; c2)                      \<succ>\<rightarrow> xs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   856
	"G\<turnstile>Norm s \<midarrow>In1l (Methd C sig)                  \<succ>\<rightarrow> xs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   857
	"G\<turnstile>Norm s \<midarrow>In1l (Body D c)                     \<succ>\<rightarrow> xs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   858
	"G\<turnstile>Norm s \<midarrow>In1l (e0 ? e1 : e2)                 \<succ>\<rightarrow> vs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   859
	"G\<turnstile>Norm s \<midarrow>In1r (If(e) c1 Else c2)             \<succ>\<rightarrow> xs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   860
	"G\<turnstile>Norm s \<midarrow>In1r (l\<bullet> While(e) c)                \<succ>\<rightarrow> xs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   861
	"G\<turnstile>Norm s \<midarrow>In1r (c1 Finally c2)                \<succ>\<rightarrow> xs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   862
	"G\<turnstile>Norm s \<midarrow>In1r (Throw e)                      \<succ>\<rightarrow> xs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   863
	"G\<turnstile>Norm s \<midarrow>In1l (NewC C)                       \<succ>\<rightarrow> vs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   864
	"G\<turnstile>Norm s \<midarrow>In1l (New T[e])                     \<succ>\<rightarrow> vs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   865
	"G\<turnstile>Norm s \<midarrow>In1l (Ass va e)                     \<succ>\<rightarrow> vs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   866
	"G\<turnstile>Norm s \<midarrow>In1r (Try c1 Catch(tn vn) c2)       \<succ>\<rightarrow> xs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   867
	"G\<turnstile>Norm s \<midarrow>In2  ({accC,statDeclC,stat}e..fn)   \<succ>\<rightarrow> vs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   868
	"G\<turnstile>Norm s \<midarrow>In2  (e1.[e2])                      \<succ>\<rightarrow> vs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   869
	"G\<turnstile>Norm s \<midarrow>In1l ({accC,statT,mode}e\<cdot>mn({pT}p)) \<succ>\<rightarrow> vs'"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
   870
	"G\<turnstile>Norm s \<midarrow>In1r (Init C)                       \<succ>\<rightarrow> xs'"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   871
declare not_None_eq [simp]  (* IntDef.Zero_def [simp] *)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   872
declare split_paired_All [simp] split_paired_Ex [simp]
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   873
ML_setup {*
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   874
simpset_ref() := simpset() addloop ("split_all_tac", split_all_tac)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   875
*}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   876
declare split_if     [split] split_if_asm     [split] 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   877
        option.split [split] option.split_asm [split]
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   878
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   879
lemma eval_Inj_elim: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   880
 "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w,s') 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   881
 \<Longrightarrow> case t of 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   882
       In1 ec \<Rightarrow> (case ec of 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   883
                    Inl e \<Rightarrow> (\<exists>v. w = In1 v) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   884
                  | Inr c \<Rightarrow> w = \<diamondsuit>)  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   885
     | In2 e \<Rightarrow> (\<exists>v. w = In2 v) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   886
     | In3 e \<Rightarrow> (\<exists>v. w = In3 v)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   887
apply (erule eval_cases)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   888
apply auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   889
apply (induct_tac "t")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   890
apply (induct_tac "a")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   891
apply auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   892
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   893
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   894
text {* The following simplification procedures set up the proper injections of
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   895
 terms and their corresponding values in the evaluation relation:
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   896
 E.g. an expression 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   897
 (injection @{term In1l} into terms) always evaluates to ordinary values 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   898
 (injection @{term In1} into generalised values @{term vals}). 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
   899
*}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   900
ML_setup {*
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   901
fun eval_fun nam inj rhs =
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   902
let
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   903
  val name = "eval_" ^ nam ^ "_eq"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   904
  val lhs = "G\<turnstile>s \<midarrow>" ^ inj ^ " t\<succ>\<rightarrow> (w, s')"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   905
  val () = qed_goal name (the_context()) (lhs ^ " = (" ^ rhs ^ ")") 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   906
	(K [Auto_tac, ALLGOALS (ftac (thm "eval_Inj_elim")) THEN Auto_tac])
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   907
  fun is_Inj (Const (inj,_) $ _) = true
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   908
    | is_Inj _                   = false
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   909
  fun pred (_ $ (Const ("Pair",_) $ _ $ 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   910
      (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ x $ _ ))) $ _ ) = is_Inj x
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   911
in
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13384
diff changeset
   912
  cond_simproc name lhs pred (thm name)
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   913
end
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   914
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   915
val eval_expr_proc =eval_fun "expr" "In1l" "\<exists>v.  w=In1 v   \<and> G\<turnstile>s \<midarrow>t-\<succ>v \<rightarrow> s'"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   916
val eval_var_proc  =eval_fun "var"  "In2"  "\<exists>vf. w=In2 vf  \<and> G\<turnstile>s \<midarrow>t=\<succ>vf\<rightarrow> s'"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   917
val eval_exprs_proc=eval_fun "exprs""In3"  "\<exists>vs. w=In3 vs  \<and> G\<turnstile>s \<midarrow>t\<doteq>\<succ>vs\<rightarrow> s'"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   918
val eval_stmt_proc =eval_fun "stmt" "In1r" "     w=\<diamondsuit> \<and> G\<turnstile>s \<midarrow>t    \<rightarrow> s'";
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   919
Addsimprocs [eval_expr_proc,eval_var_proc,eval_exprs_proc,eval_stmt_proc];
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   920
bind_thms ("AbruptIs", sum3_instantiate (thm "eval.Abrupt"))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   921
*}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   922
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   923
declare halloc.Abrupt [intro!] eval.Abrupt [intro!]  AbruptIs [intro!] 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   924
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   925
text{* @{text Callee},@{text InsInitE}, @{text InsInitV}, @{text FinA} are only
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   926
used in smallstep semantics, not in the bigstep semantics. So their is no
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   927
valid evaluation of these terms 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   928
*}
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   929
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   930
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   931
lemma eval_Callee: "G\<turnstile>Norm s\<midarrow>Callee l e-\<succ>v\<rightarrow> s' = False"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   932
proof -
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   933
  { fix s t v s'
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   934
    assume eval: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (v,s')" and
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   935
         normal: "normal s" and
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   936
         callee: "t=In1l (Callee l e)"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   937
    then have "False"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   938
    proof (induct)
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   939
    qed (auto)
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   940
  }  
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   941
  then show ?thesis
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   942
    by (cases s') fastsimp
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   943
qed
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   944
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   945
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   946
lemma eval_InsInitE: "G\<turnstile>Norm s\<midarrow>InsInitE c e-\<succ>v\<rightarrow> s' = False"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   947
proof -
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   948
  { fix s t v s'
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   949
    assume eval: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (v,s')" and
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   950
         normal: "normal s" and
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   951
         callee: "t=In1l (InsInitE c e)"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   952
    then have "False"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   953
    proof (induct)
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   954
    qed (auto)
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   955
  }
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   956
  then show ?thesis
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   957
    by (cases s') fastsimp
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   958
qed
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   959
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   960
lemma eval_InsInitV: "G\<turnstile>Norm s\<midarrow>InsInitV c w=\<succ>v\<rightarrow> s' = False"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   961
proof -
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   962
  { fix s t v s'
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   963
    assume eval: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (v,s')" and
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   964
         normal: "normal s" and
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   965
         callee: "t=In2 (InsInitV c w)"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   966
    then have "False"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   967
    proof (induct)
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   968
    qed (auto)
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   969
  }  
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   970
  then show ?thesis
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   971
    by (cases s') fastsimp
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   972
qed
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   973
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   974
lemma eval_FinA: "G\<turnstile>Norm s\<midarrow>FinA a c\<rightarrow> s' = False"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   975
proof -
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   976
  { fix s t v s'
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   977
    assume eval: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (v,s')" and
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   978
         normal: "normal s" and
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   979
         callee: "t=In1r (FinA a c)"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   980
    then have "False"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   981
    proof (induct)
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   982
    qed (auto)
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   983
  }  
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   984
  then show ?thesis
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   985
    by (cases s') fastsimp 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   986
qed
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   987
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   988
lemma eval_no_abrupt_lemma: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   989
   "\<And>s s'. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w,s') \<Longrightarrow> normal s' \<longrightarrow> normal s"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   990
by (erule eval_cases, auto)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   991
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   992
lemma eval_no_abrupt: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   993
  "G\<turnstile>(x,s) \<midarrow>t\<succ>\<rightarrow> (w,Norm s') = 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   994
        (x = None \<and> G\<turnstile>Norm s \<midarrow>t\<succ>\<rightarrow> (w,Norm s'))"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   995
apply auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   996
apply (frule eval_no_abrupt_lemma, auto)+
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   997
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   998
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   999
ML {*
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1000
local
12919
d6a0d168291e removed theory Option;
wenzelm
parents: 12858
diff changeset
  1001
  fun is_None (Const ("Datatype.option.None",_)) = true
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1002
    | is_None _ = false
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1003
  fun pred (t as (_ $ (Const ("Pair",_) $
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1004
     (Const ("Pair", _) $ x $ _) $ _ ) $ _)) = is_None x
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1005
in
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1006
  val eval_no_abrupt_proc = 
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13384
diff changeset
  1007
  cond_simproc "eval_no_abrupt" "G\<turnstile>(x,s) \<midarrow>e\<succ>\<rightarrow> (w,Norm s')" pred 
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1008
          (thm "eval_no_abrupt")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1009
end;
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1010
Addsimprocs [eval_no_abrupt_proc]
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1011
*}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1012
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1013
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1014
lemma eval_abrupt_lemma: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1015
  "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (v,s') \<Longrightarrow> abrupt s=Some xc \<longrightarrow> s'= s \<and> v = arbitrary3 t"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1016
by (erule eval_cases, auto)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1017
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1018
lemma eval_abrupt: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1019
 " G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<rightarrow> (w,s') =  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1020
     (s'=(Some xc,s) \<and> w=arbitrary3 t \<and> 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1021
     G\<turnstile>(Some xc,s) \<midarrow>t\<succ>\<rightarrow> (arbitrary3 t,(Some xc,s)))"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1022
apply auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1023
apply (frule eval_abrupt_lemma, auto)+
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1024
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1025
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1026
ML {*
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1027
local
12919
d6a0d168291e removed theory Option;
wenzelm
parents: 12858
diff changeset
  1028
  fun is_Some (Const ("Pair",_) $ (Const ("Datatype.option.Some",_) $ _)$ _) =true
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1029
    | is_Some _ = false
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1030
  fun pred (_ $ (Const ("Pair",_) $
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1031
     _ $ (Const ("Pair", _) $ _ $ (Const ("Pair", _) $ _ $
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1032
       x))) $ _ ) = is_Some x
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1033
in
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1034
  val eval_abrupt_proc = 
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13384
diff changeset
  1035
  cond_simproc "eval_abrupt" 
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1036
               "G\<turnstile>(Some xc,s) \<midarrow>e\<succ>\<rightarrow> (w,s')" pred (thm "eval_abrupt")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1037
end;
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1038
Addsimprocs [eval_abrupt_proc]
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1039
*}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1040
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1041
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1042
lemma LitI: "G\<turnstile>s \<midarrow>Lit v-\<succ>(if normal s then v else arbitrary)\<rightarrow> s"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1043
apply (case_tac "s", case_tac "a = None")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1044
by (auto intro!: eval.Lit)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1045
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1046
lemma SkipI [intro!]: "G\<turnstile>s \<midarrow>Skip\<rightarrow> s"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1047
apply (case_tac "s", case_tac "a = None")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1048
by (auto intro!: eval.Skip)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1049
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1050
lemma ExprI: "G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s' \<Longrightarrow> G\<turnstile>s \<midarrow>Expr e\<rightarrow> s'"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1051
apply (case_tac "s", case_tac "a = None")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1052
by (auto intro!: eval.Expr)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1053
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1054
lemma CompI: "\<lbrakk>G\<turnstile>s \<midarrow>c1\<rightarrow> s1; G\<turnstile>s1 \<midarrow>c2\<rightarrow> s2\<rbrakk> \<Longrightarrow> G\<turnstile>s \<midarrow>c1;; c2\<rightarrow> s2"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1055
apply (case_tac "s", case_tac "a = None")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1056
by (auto intro!: eval.Comp)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1057
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1058
lemma CondI: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1059
  "\<And>s1. \<lbrakk>G\<turnstile>s \<midarrow>e-\<succ>b\<rightarrow> s1; G\<turnstile>s1 \<midarrow>(if the_Bool b then e1 else e2)-\<succ>v\<rightarrow> s2\<rbrakk> \<Longrightarrow> 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1060
         G\<turnstile>s \<midarrow>e ? e1 : e2-\<succ>(if normal s1 then v else arbitrary)\<rightarrow> s2"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1061
apply (case_tac "s", case_tac "a = None")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1062
by (auto intro!: eval.Cond)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1063
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1064
lemma IfI: "\<lbrakk>G\<turnstile>s \<midarrow>e-\<succ>v\<rightarrow> s1; G\<turnstile>s1 \<midarrow>(if the_Bool v then c1 else c2)\<rightarrow> s2\<rbrakk>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1065
                 \<Longrightarrow> G\<turnstile>s \<midarrow>If(e) c1 Else c2\<rightarrow> s2"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1066
apply (case_tac "s", case_tac "a = None")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1067
by (auto intro!: eval.If)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1068
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
  1069
lemma MethdI: "G\<turnstile>s \<midarrow>body G C sig-\<succ>v\<rightarrow> s' 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
  1070
                \<Longrightarrow> G\<turnstile>s \<midarrow>Methd C sig-\<succ>v\<rightarrow> s'"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1071
apply (case_tac "s", case_tac "a = None")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1072
by (auto intro!: eval.Methd)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1073
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
  1074
lemma eval_Call: 
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
  1075
   "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>e-\<succ>a'\<rightarrow> s1; G\<turnstile>s1 \<midarrow>ps\<doteq>\<succ>pvs\<rightarrow> s2;  
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
  1076
     D = invocation_declclass G mode (store s2) a' statT \<lparr>name=mn,parTs=pTs\<rparr>;
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
  1077
     s3 = init_lvars G D \<lparr>name=mn,parTs=pTs\<rparr> mode a' pvs s2;
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
  1078
     s3' = check_method_access G accC statT mode \<lparr>name=mn,parTs=pTs\<rparr> a' s3;
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
  1079
     G\<turnstile>s3'\<midarrow>Methd D \<lparr>name=mn,parTs=pTs\<rparr>-\<succ> v\<rightarrow> s4; 
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
  1080
       s4' = restore_lvars s2 s4\<rbrakk> \<Longrightarrow>  
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12919
diff changeset
  1081
       G\<turnstile>Norm s0 \<midarrow>{accC,statT,mode}e\<cdot>mn({pTs}ps)-\<succ>v\<rightarrow> s4'"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1082
apply (drule eval.Call, assumption)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1083
apply (rule HOL.refl)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1084
apply simp+
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1085
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1086
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1087
lemma eval_Init: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1088
"\<lbrakk>if inited C (globs s0) then s3 = Norm s0 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1089
  else G\<turnstile>Norm (init_class_obj G C s0)  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1090
         \<midarrow>(if C = Object then Skip else Init (super (the (class G C))))\<rightarrow> s1 \<and>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1091
       G\<turnstile>set_lvars empty s1 \<midarrow>(init (the (class G C)))\<rightarrow> s2 \<and> 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1092
      s3 = restore_lvars s1 s2\<rbrakk> \<Longrightarrow>  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1093
  G\<turnstile>Norm s0 \<midarrow>Init C\<rightarrow> s3"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1094
apply (rule eval.Init)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1095
apply auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1096
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1097
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1098
lemma init_done: "initd C s \<Longrightarrow> G\<turnstile>s \<midarrow>Init C\<rightarrow> s"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1099
apply (case_tac "s", simp)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1100
apply (case_tac "a")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1101
apply  safe
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1102
apply (rule eval_Init)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1103
apply   auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1104
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1105
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1106
lemma eval_StatRef: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1107
"G\<turnstile>s \<midarrow>StatRef rt-\<succ>(if abrupt s=None then Null else arbitrary)\<rightarrow> s"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1108
apply (case_tac "s", simp)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1109
apply (case_tac "a = None")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1110
apply (auto del: eval.Abrupt intro!: eval.intros)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1111
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1112
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1113
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1114
lemma SkipD [dest!]: "G\<turnstile>s \<midarrow>Skip\<rightarrow> s' \<Longrightarrow> s' = s" 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1115
apply (erule eval_cases)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1116
by auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1117
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1118
lemma Skip_eq [simp]: "G\<turnstile>s \<midarrow>Skip\<rightarrow> s' = (s = s')"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1119
by auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1120
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1121
(*unused*)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1122
lemma init_retains_locals [rule_format (no_asm)]: "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> (w,s') \<Longrightarrow>  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1123
  (\<forall>C. t=In1r (Init C) \<longrightarrow> locals (store s) = locals (store s'))"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1124
apply (erule eval.induct)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1125
apply (simp (no_asm_use) split del: split_if_asm option.split_asm)+
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1126
apply auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1127
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1128
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1129
lemma halloc_xcpt [dest!]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1130
  "\<And>s'. G\<turnstile>(Some xc,s) \<midarrow>halloc oi\<succ>a\<rightarrow> s' \<Longrightarrow> s'=(Some xc,s)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1131
apply (erule_tac halloc_elim_cases)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1132
by auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1133
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1134
(*
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1135
G\<turnstile>(x,(h,l)) \<midarrow>e\<succ>v\<rightarrow> (x',(h',l'))) \<Longrightarrow> l This = l' This"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1136
G\<turnstile>(x,(h,l)) \<midarrow>s  \<rightarrow> (x',(h',l'))) \<Longrightarrow> l This = l' This"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1137
*)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1138
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1139
lemma eval_Methd: 
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
  1140
  "G\<turnstile>s \<midarrow>In1l(body G C sig)\<succ>\<rightarrow> (w,s') 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
  1141
   \<Longrightarrow> G\<turnstile>s \<midarrow>In1l(Methd C sig)\<succ>\<rightarrow> (w,s')"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1142
apply (case_tac "s")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1143
apply (case_tac "a")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1144
apply clarsimp+
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1145
apply (erule eval.Methd)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1146
apply (drule eval_abrupt_lemma)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1147
apply force
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1148
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1149
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
  1150
lemma eval_Body: "\<lbrakk>G\<turnstile>Norm s0 \<midarrow>Init D\<rightarrow> s1; G\<turnstile>s1 \<midarrow>c\<rightarrow> s2;
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
  1151
                   res=the (locals (store s2) Result);
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
  1152
                   s3 = (if (\<exists> l. abrupt s2 = Some (Jump (Break l)) \<or>
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
  1153
                                  abrupt s2 = Some (Jump (Cont l))) 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
  1154
                   then abupd (\<lambda> x. Some (Error CrossMethodJump)) s2 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
  1155
                   else s2);
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
  1156
                   s4=abupd (absorb Ret) s3\<rbrakk> \<Longrightarrow>
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
  1157
 G\<turnstile>Norm s0 \<midarrow>Body D c-\<succ>res\<rightarrow>s4"
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
  1158
by (auto elim: eval.Body)
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
  1159
13384
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1160
lemma eval_binop_arg2_indep:
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1161
"\<not> need_second_arg binop v1 \<Longrightarrow> eval_binop binop v1 x = eval_binop binop v1 y"
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1162
by (cases binop)
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1163
   (simp_all add: need_second_arg_def)
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1164
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1165
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1166
lemma eval_BinOp_arg2_indepI:
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1167
  assumes eval_e1: "G\<turnstile>Norm s0 \<midarrow>e1-\<succ>v1\<rightarrow> s1" and
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1168
          no_need: "\<not> need_second_arg binop v1" 
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1169
  shows "G\<turnstile>Norm s0 \<midarrow>BinOp binop e1 e2-\<succ>(eval_binop binop v1 v2)\<rightarrow> s1"
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1170
         (is "?EvalBinOp v2")
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1171
proof -
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1172
  from eval_e1 
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1173
  have "?EvalBinOp Unit" 
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1174
    by (rule eval.BinOp)
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1175
       (simp add: no_need)
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1176
  moreover
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1177
  from no_need 
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1178
  have "eval_binop binop v1 Unit = eval_binop binop v1 v2"
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1179
    by (simp add: eval_binop_arg2_indep)
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1180
  ultimately
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1181
  show ?thesis
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1182
    by simp
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1183
qed
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
  1184
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1185
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1186
section "single valued"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1187
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1188
lemma unique_halloc [rule_format (no_asm)]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1189
  "\<And>s as as'. (s,oi,as)\<in>halloc G \<Longrightarrow> (s,oi,as')\<in>halloc G \<longrightarrow> as'=as"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1190
apply (simp (no_asm_simp) only: split_tupled_all)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1191
apply (erule halloc.induct)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1192
apply  (auto elim!: halloc_elim_cases split del: split_if split_if_asm)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1193
apply (drule trans [THEN sym], erule sym) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1194
defer
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1195
apply (drule trans [THEN sym], erule sym)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1196
apply auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1197
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1198
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1199
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1200
lemma single_valued_halloc: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1201
  "single_valued {((s,oi),(a,s')). G\<turnstile>s \<midarrow>halloc oi\<succ>a \<rightarrow> s'}"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1202
apply (unfold single_valued_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1203
by (clarsimp, drule (1) unique_halloc, auto)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1204
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1205
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1206
lemma unique_sxalloc [rule_format (no_asm)]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1207
  "\<And>s s'. G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s' \<Longrightarrow> G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s'' \<longrightarrow> s'' = s'"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1208
apply (simp (no_asm_simp) only: split_tupled_all)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1209
apply (erule sxalloc.induct)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1210
apply   (auto dest: unique_halloc elim!: sxalloc_elim_cases 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1211
              split del: split_if split_if_asm)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1212
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1213
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1214
lemma single_valued_sxalloc: "single_valued {(s,s'). G\<turnstile>s \<midarrow>sxalloc\<rightarrow> s'}"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1215
apply (unfold single_valued_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1216
apply (blast dest: unique_sxalloc)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1217
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1218
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1219
lemma split_pairD: "(x,y) = p \<Longrightarrow> x = fst p & y = snd p"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1220
by auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1221
13688
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13601
diff changeset
  1222
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
  1223
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1224
lemma unique_eval [rule_format (no_asm)]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1225
  "G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> ws \<Longrightarrow> (\<forall>ws'. G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> ws' \<longrightarrow> ws' = ws)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1226
apply (case_tac "ws")
13601
fd3e3d6b37b2 Adapted to new simplifier.
berghofe
parents: 13462
diff changeset
  1227
apply hypsubst
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1228
apply (erule eval_induct)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1229
apply (tactic {* ALLGOALS (EVERY'
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1230
      [strip_tac, rotate_tac ~1, eresolve_tac (thms "eval_elim_cases")]) *})
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
  1231
(* 31 subgoals *)
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
  1232
prefer 28 (* Try *) 
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1233
apply (simp (no_asm_use) only: split add: split_if_asm)
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
  1234
(* 34 subgoals *)
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
  1235
prefer 30 (* Init *)
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1236
apply (case_tac "inited C (globs s0)", (simp only: if_True if_False)+)
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
  1237
prefer 26 (* While *)
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1238
apply (simp (no_asm_use) only: split add: split_if_asm, blast)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1239
apply (drule_tac x="(In1 bb, s1a)" in spec, drule (1) mp, simp)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1240
apply (drule_tac x="(In1 bb, s1a)" in spec, drule (1) mp, simp)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1241
apply blast
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
  1242
(* 33 subgoals *)
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1243
apply (blast dest: unique_sxalloc unique_halloc split_pairD)+
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1244
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1245
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1246
(* unused *)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1247
lemma single_valued_eval: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1248
 "single_valued {((s,t),vs'). G\<turnstile>s \<midarrow>t\<succ>\<rightarrow> vs'}"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1249
apply (unfold single_valued_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1250
by (clarify, drule (1) unique_eval, auto)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1251
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
  1252
end