src/HOL/Bali/Term.thy
author wenzelm
Sun, 23 Apr 2017 14:15:09 +0200
changeset 65553 006a274cdbc2
parent 62390 842917225d56
child 67443 3abf6a722518
permissions -rw-r--r--
added missing file (amending f533820e7248);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12857
a4386cc9b1c3 tuned header;
wenzelm
parents: 12854
diff changeset
     1
(*  Title:      HOL/Bali/Term.thy
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     2
    Author:     David von Oheimb
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     3
*)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     4
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
     5
subsection \<open>Java expressions and statements\<close>
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     6
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14981
diff changeset
     7
theory Term imports Value Table begin
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     8
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
     9
text \<open>
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    10
design issues:
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    11
\begin{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    12
\item invocation frames for local variables could be reduced to special static
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    13
  objects (one per method). This would reduce redundancy, but yield a rather
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    14
  non-standard execution model more difficult to understand.
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    15
\item method bodies separated from calls to handle assumptions in axiomat. 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    16
  semantics
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    17
  NB: Body is intended to be in the environment of the called method.
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    18
\item class initialization is regarded as (auxiliary) statement 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    19
      (required for AxSem)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    20
\item result expression of method return is handled by a special result variable
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    21
  result variable is treated uniformly with local variables
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    22
  \begin{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    23
  \item[+] welltypedness and existence of the result/return expression is 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    24
           ensured without extra efford
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    25
  \end{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    26
\end{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    27
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    28
simplifications:
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    29
\begin{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    30
\item expression statement allowed for any expression
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    31
\item This  is modeled as a special non-assignable local variable
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    32
\item Super is modeled as a general expression with the same value as This
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    33
\item access to field x in current class via This.x
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    34
\item NewA creates only one-dimensional arrays;
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    35
  initialization of further subarrays may be simulated with nested NewAs
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    36
\item The 'Lit' constructor is allowed to contain a reference value.
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    37
  But this is assumed to be prohibited in the input language, which is enforced
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    38
  by the type-checking rules.
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    39
\item a call of a static method via a type name may be simulated by a dummy 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    40
      variable
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    41
\item no nested blocks with inner local variables
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    42
\item no synchronized statements
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    43
\item no secondary forms of if, while (e.g. no for) (may be easily simulated)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    44
\item no switch (may be simulated with if)
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    45
\item the \<open>try_catch_finally\<close> statement is divided into the 
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    46
      \<open>try_catch\<close> statement 
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    47
      and a finally statement, which may be considered as try..finally with 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    48
      empty catch
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    49
\item the \<open>try_catch\<close> statement has exactly one catch clause; 
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    50
      multiple ones can be
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    51
  simulated with instanceof
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    52
\item the compiler is supposed to add the annotations {\<open>_\<close>} during 
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    53
      type-checking. This
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    54
  transformation is left out as its result is checked by the type rules anyway
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    55
\end{itemize}
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    56
\<close>
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    57
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
    58
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
    59
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    60
type_synonym locals = "(lname, val) table"  \<comment>\<open>local variables\<close>
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
    61
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
    62
58310
91ea607a34d8 updated news
blanchet
parents: 58251
diff changeset
    63
datatype jump
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    64
        = Break label \<comment>\<open>break\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    65
        | Cont label  \<comment>\<open>continue\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    66
        | Ret         \<comment>\<open>return from method\<close>
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
    67
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    68
datatype xcpt        \<comment>\<open>exception\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    69
        = Loc loc    \<comment>\<open>location of allocated execption object\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    70
        | Std xname  \<comment>\<open>intermediate standard exception, see Eval.thy\<close>
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
    71
58310
91ea607a34d8 updated news
blanchet
parents: 58251
diff changeset
    72
datatype error
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    73
       =  AccessViolation  \<comment>\<open>Access to a member that isn't permitted\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    74
        | CrossMethodJump  \<comment>\<open>Method exits with a break or continue\<close>
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
    75
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    76
datatype abrupt       \<comment>\<open>abrupt completion\<close> 
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    77
        = Xcpt xcpt   \<comment>\<open>exception\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    78
        | Jump jump   \<comment>\<open>break, continue, return\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    79
        | Error error \<comment> \<open>runtime errors, we wan't to detect and proof absent
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    80
                            in welltyped programms\<close>
41778
5f79a9e42507 modernized specifications;
wenzelm
parents: 37956
diff changeset
    81
type_synonym
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
    82
  abopt  = "abrupt option"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
    83
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    84
text \<open>Local variable store and exception. 
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
    85
Anticipation of State.thy used by smallstep semantics. For a method call, 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
    86
we save the local variables of the caller in the term Callee to restore them 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
    87
after method return. Also an exception must be restored after the finally
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    88
statement\<close>
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
    89
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
    90
translations
35431
8758fe1fc9f8 cleanup type translations;
wenzelm
parents: 35416
diff changeset
    91
 (type) "locals" <= (type) "(lname, val) table"
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
    92
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    93
datatype inv_mode                  \<comment>\<open>invocation mode for method calls\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    94
        = Static                   \<comment>\<open>static\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    95
        | SuperM                   \<comment>\<open>super\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    96
        | IntVir                   \<comment>\<open>interface or virtual\<close>
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    97
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    98
record  sig =              \<comment>\<open>signature of a method, cf. 8.4.2\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
    99
          name ::"mname"   \<comment>\<open>acutally belongs to Decl.thy\<close>
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   100
          parTs::"ty list"        
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   101
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   102
translations
35431
8758fe1fc9f8 cleanup type translations;
wenzelm
parents: 35416
diff changeset
   103
  (type) "sig" <= (type) "\<lparr>name::mname,parTs::ty list\<rparr>"
8758fe1fc9f8 cleanup type translations;
wenzelm
parents: 35416
diff changeset
   104
  (type) "sig" <= (type) "\<lparr>name::mname,parTs::ty list,\<dots>::'a\<rparr>"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   105
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   106
\<comment>\<open>function codes for unary operations\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   107
datatype unop =  UPlus    \<comment> \<open>{\tt +} unary plus\<close> 
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   108
               | UMinus   \<comment> \<open>{\tt -} unary minus\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   109
               | UBitNot  \<comment> \<open>{\tt ~} bitwise NOT\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   110
               | UNot     \<comment> \<open>{\tt !} logical complement\<close>
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   111
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   112
\<comment>\<open>function codes for binary operations\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   113
datatype binop = Mul     \<comment> \<open>{\tt * }   multiplication\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   114
               | Div     \<comment> \<open>{\tt /}   division\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   115
               | Mod     \<comment> \<open>{\tt \%}   remainder\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   116
               | Plus    \<comment> \<open>{\tt +}   addition\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   117
               | Minus   \<comment> \<open>{\tt -}   subtraction\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   118
               | LShift  \<comment> \<open>{\tt <<}  left shift\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   119
               | RShift  \<comment> \<open>{\tt >>}  signed right shift\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   120
               | RShiftU \<comment> \<open>{\tt >>>} unsigned right shift\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   121
               | Less    \<comment> \<open>{\tt <}   less than\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   122
               | Le      \<comment> \<open>{\tt <=}  less than or equal\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   123
               | Greater \<comment> \<open>{\tt >}   greater than\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   124
               | Ge      \<comment> \<open>{\tt >=}  greater than or equal\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   125
               | Eq      \<comment> \<open>{\tt ==}  equal\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   126
               | Neq     \<comment> \<open>{\tt !=}  not equal\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   127
               | BitAnd  \<comment> \<open>{\tt \&}   bitwise AND\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   128
               | And     \<comment> \<open>{\tt \&}   boolean AND\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   129
               | BitXor  \<comment> \<open>{\texttt \^}   bitwise Xor\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   130
               | Xor     \<comment> \<open>{\texttt \^}   boolean Xor\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   131
               | BitOr   \<comment> \<open>{\tt |}   bitwise Or\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   132
               | Or      \<comment> \<open>{\tt |}   boolean Or\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   133
               | CondAnd \<comment> \<open>{\tt \&\&}  conditional And\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   134
               | CondOr  \<comment> \<open>{\tt ||}  conditional Or\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   135
text\<open>The boolean operators {\tt \&} and {\tt |} strictly evaluate both
13384
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13358
diff changeset
   136
of their arguments. The conditional operators {\tt \&\&} and {\tt ||} only 
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13358
diff changeset
   137
evaluate the second argument if the value of the whole expression isn't 
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13358
diff changeset
   138
allready determined by the first argument.
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13358
diff changeset
   139
e.g.: {\tt false \&\& e} e is not evaluated;  
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13358
diff changeset
   140
      {\tt true || e} e is not evaluated; 
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   141
\<close>
13358
c837ba4cfb62 fix latex output
schirmer
parents: 13345
diff changeset
   142
58310
91ea607a34d8 updated news
blanchet
parents: 58251
diff changeset
   143
datatype var
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   144
        = LVar lname \<comment>\<open>local variable (incl. parameters)\<close>
13384
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13358
diff changeset
   145
        | FVar qtname qtname bool expr vname ("{_,_,_}_.._"[10,10,10,85,99]90)
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   146
                     \<comment>\<open>class field\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   147
                     \<comment>\<open>@{term "{accC,statDeclC,stat}e..fn"}\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   148
                     \<comment>\<open>\<open>accC\<close>: accessing class (static class were\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   149
                     \<comment>\<open>the code is declared. Annotation only needed for\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   150
                     \<comment>\<open>evaluation to check accessibility)\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   151
                     \<comment>\<open>\<open>statDeclC\<close>: static declaration class of field\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   152
                     \<comment>\<open>\<open>stat\<close>: static or instance field?\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   153
                     \<comment>\<open>\<open>e\<close>: reference to object\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   154
                     \<comment>\<open>\<open>fn\<close>: field name\<close>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27226
diff changeset
   155
        | AVar expr expr ("_.[_]"[90,10   ]90)
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   156
                     \<comment>\<open>array component\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   157
                     \<comment>\<open>@{term "e1.[e2]"}: e1 array reference; e2 index\<close>
13384
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13358
diff changeset
   158
        | InsInitV stmt var 
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   159
                     \<comment>\<open>insertion of initialization before evaluation\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   160
                     \<comment>\<open>of var (technical term for smallstep semantics.)\<close>
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   161
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   162
and expr
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   163
        = NewC qtname         \<comment>\<open>class instance creation\<close>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27226
diff changeset
   164
        | NewA ty expr ("New _[_]"[99,10   ]85) 
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   165
                              \<comment>\<open>array creation\<close> 
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   166
        | Cast ty expr        \<comment>\<open>type cast\<close>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27226
diff changeset
   167
        | Inst expr ref_ty ("_ InstOf _"[85,99] 85)   
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   168
                              \<comment>\<open>instanceof\<close>     
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   169
        | Lit  val              \<comment>\<open>literal value, references not allowed\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   170
        | UnOp unop expr        \<comment>\<open>unary operation\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   171
        | BinOp binop expr expr \<comment>\<open>binary operation\<close>
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   172
        
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   173
        | Super               \<comment>\<open>special Super keyword\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   174
        | Acc  var            \<comment>\<open>variable access\<close>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27226
diff changeset
   175
        | Ass  var expr       ("_:=_"   [90,85   ]85)
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   176
                              \<comment>\<open>variable assign\<close> 
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   177
        | Cond expr expr expr ("_ ? _ : _" [85,85,80]80) \<comment>\<open>conditional\<close>  
13384
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13358
diff changeset
   178
        | Call qtname ref_ty inv_mode expr mname "(ty list)" "(expr list)"  
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13358
diff changeset
   179
            ("{_,_,_}_\<cdot>_'( {_}_')"[10,10,10,85,99,10,10]85) 
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   180
                    \<comment>\<open>method call\<close> 
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   181
                    \<comment>\<open>@{term "{accC,statT,mode}e\<cdot>mn({pTs}args)"} "\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   182
                    \<comment>\<open>\<open>accC\<close>: accessing class (static class were\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   183
                    \<comment>\<open>the call code is declared. Annotation only needed for\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   184
                    \<comment>\<open>evaluation to check accessibility)\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   185
                    \<comment>\<open>\<open>statT\<close>: static declaration class/interface of\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   186
                    \<comment>\<open>method\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   187
                    \<comment>\<open>\<open>mode\<close>: invocation mode\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   188
                    \<comment>\<open>\<open>e\<close>: reference to object\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   189
                    \<comment>\<open>\<open>mn\<close>: field name\<close>   
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   190
                    \<comment>\<open>\<open>pTs\<close>: types of parameters\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   191
                    \<comment>\<open>\<open>args\<close>: the actual parameters/arguments\<close> 
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   192
        | Methd qtname sig    \<comment>\<open>(folded) method (see below)\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   193
        | Body qtname stmt    \<comment>\<open>(unfolded) method body\<close>
13384
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13358
diff changeset
   194
        | InsInitE stmt expr  
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   195
                 \<comment>\<open>insertion of initialization before\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   196
                 \<comment>\<open>evaluation of expr (technical term for smallstep sem.)\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   197
        | Callee locals expr  \<comment>\<open>save callers locals in callee-Frame\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   198
                              \<comment>\<open>(technical term for smallstep semantics)\<close>
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   199
and  stmt
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   200
        = Skip                  \<comment>\<open>empty      statement\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   201
        | Expr  expr            \<comment>\<open>expression statement\<close>
13384
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13358
diff changeset
   202
        | Lab   jump stmt       ("_\<bullet> _" [      99,66]66)
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   203
                                \<comment>\<open>labeled statement; handles break\<close>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27226
diff changeset
   204
        | Comp  stmt stmt       ("_;; _"                  [      66,65]65)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27226
diff changeset
   205
        | If'   expr stmt stmt  ("If'(_') _ Else _"       [   80,79,79]70)
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27226
diff changeset
   206
        | Loop  label expr stmt ("_\<bullet> While'(_') _"        [   99,80,79]70)
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   207
        | Jmp jump              \<comment>\<open>break, continue, return\<close>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27226
diff changeset
   208
        | Throw expr
13384
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13358
diff changeset
   209
        | TryC  stmt qtname vname stmt ("Try _ Catch'(_ _') _"  [79,99,80,79]70)
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   210
             \<comment>\<open>@{term "Try c1 Catch(C vn) c2"}\<close> 
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   211
             \<comment>\<open>\<open>c1\<close>: block were exception may be thrown\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   212
             \<comment>\<open>\<open>C\<close>:  execption class to catch\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   213
             \<comment>\<open>\<open>vn\<close>: local name for exception used in \<open>c2\<close>\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   214
             \<comment>\<open>\<open>c2\<close>: block to execute when exception is cateched\<close>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 27226
diff changeset
   215
        | Fin  stmt  stmt        ("_ Finally _"               [      79,79]70)
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   216
        | FinA abopt stmt       \<comment>\<open>Save abruption of first statement\<close> 
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   217
                                \<comment>\<open>technical term  for smallstep sem.)\<close>
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   218
        | Init  qtname          \<comment>\<open>class initialization\<close>
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   219
58251
b13e5c3497f5 ported Bali to new datatypes
blanchet
parents: 58249
diff changeset
   220
datatype_compat var expr stmt
b13e5c3497f5 ported Bali to new datatypes
blanchet
parents: 58249
diff changeset
   221
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   222
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   223
text \<open>
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   224
The expressions Methd and Body are artificial program constructs, in the
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   225
sense that they are not used to define a concrete Bali program. In the 
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   226
operational semantic's they are "generated on the fly" 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   227
to decompose the task to define the behaviour of the Call expression. 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   228
They are crucial for the axiomatic semantics to give a syntactic hook to insert 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   229
some assertions (cf. AxSem.thy, Eval.thy). 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   230
The Init statement (to initialize a class on its first use) is 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   231
inserted in various places by the semantics. 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   232
Callee, InsInitV, InsInitE,FinA are only needed as intermediate steps in the
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   233
smallstep (transition) semantics (cf. Trans.thy). Callee is used to save the 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   234
local variables of the caller for method return. So ve avoid modelling a 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   235
frame stack.
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   236
The InsInitV/E terms are only used by the smallstep semantics to model the
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   237
intermediate steps of class-initialisation.
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   238
\<close>
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   239
 
41778
5f79a9e42507 modernized specifications;
wenzelm
parents: 37956
diff changeset
   240
type_synonym "term" = "(expr+stmt,var,expr list) sum3"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   241
translations
35431
8758fe1fc9f8 cleanup type translations;
wenzelm
parents: 35416
diff changeset
   242
  (type) "sig"   <= (type) "mname \<times> ty list"
8758fe1fc9f8 cleanup type translations;
wenzelm
parents: 35416
diff changeset
   243
  (type) "term"  <= (type) "(expr+stmt,var,expr list) sum3"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   244
35067
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   245
abbreviation this :: expr
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   246
  where "this == Acc (LVar This)"
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   247
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   248
abbreviation LAcc :: "vname \<Rightarrow> expr" ("!!")
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   249
  where "!!v == Acc (LVar (EName (VNam v)))"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   250
35067
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   251
abbreviation
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   252
  LAss :: "vname \<Rightarrow> expr \<Rightarrow>stmt" ("_:==_" [90,85] 85)
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   253
  where "v:==e == Expr (Ass (LVar (EName (VNam  v))) e)"
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   254
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   255
abbreviation
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   256
  Return :: "expr \<Rightarrow> stmt"
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   257
  where "Return e == Expr (Ass (LVar (EName Res)) e);; Jmp Ret" \<comment>\<open>\tt Res := e;; Jmp Ret\<close>
35067
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   258
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   259
abbreviation
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   260
  StatRef :: "ref_ty \<Rightarrow> expr"
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   261
  where "StatRef rt == Cast (RefT rt) (Lit Null)"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   262
  
37956
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   263
definition
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   264
  is_stmt :: "term \<Rightarrow> bool"
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   265
  where "is_stmt t = (\<exists>c. t=In1r c)"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   266
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   267
ML \<open>ML_Thms.bind_thms ("is_stmt_rews", sum3_instantiate @{context} @{thm is_stmt_def})\<close>
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   268
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   269
declare is_stmt_rews [simp]
13345
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   270
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   271
text \<open>
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: 13384
diff changeset
   272
  Here is some syntactic stuff to handle the injections of statements,
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: 13384
diff changeset
   273
  expressions, variables and expression lists into general terms.
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   274
\<close>
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: 13384
diff changeset
   275
35067
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   276
abbreviation (input)
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   277
  expr_inj_term :: "expr \<Rightarrow> term" ("\<langle>_\<rangle>\<^sub>e" 1000)
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   278
  where "\<langle>e\<rangle>\<^sub>e == In1l e"
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   279
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   280
abbreviation (input)
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   281
  stmt_inj_term :: "stmt \<Rightarrow> term" ("\<langle>_\<rangle>\<^sub>s" 1000)
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   282
  where "\<langle>c\<rangle>\<^sub>s == In1r c"
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: 13384
diff changeset
   283
35067
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   284
abbreviation (input)
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   285
  var_inj_term :: "var \<Rightarrow> term"  ("\<langle>_\<rangle>\<^sub>v" 1000)
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   286
  where "\<langle>v\<rangle>\<^sub>v == In2 v"
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   287
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   288
abbreviation (input)
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   289
  lst_inj_term :: "expr list \<Rightarrow> term" ("\<langle>_\<rangle>\<^sub>l" 1000)
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   290
  where "\<langle>es\<rangle>\<^sub>l == In3 es"
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: 13384
diff changeset
   291
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   292
text \<open>It seems to be more elegant to have an overloaded injection like the
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: 13384
diff changeset
   293
following.
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   294
\<close>
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: 13384
diff changeset
   295
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   296
class inj_term =
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   297
  fixes inj_term:: "'a \<Rightarrow> term" ("\<langle>_\<rangle>" 1000)
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: 13384
diff changeset
   298
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   299
text \<open>How this overloaded injections work can be seen in the theory 
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   300
\<open>DefiniteAssignment\<close>. Other big inductive relations on
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   301
terms defined in theories \<open>WellType\<close>, \<open>Eval\<close>, \<open>Evaln\<close> and
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   302
\<open>AxSem\<close> don't follow this convention right now, but introduce subtle 
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: 13384
diff changeset
   303
syntactic sugar in the relations themselves to make a distinction 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: 13384
diff changeset
   304
expressions, statements and so on. So unfortunately you will encounter a 
35067
af4c18c30593 modernized syntax translations, using mostly abbreviation/notation;
wenzelm
parents: 32960
diff changeset
   305
mixture of dealing with these injections. The abbreviations above are used
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: 13384
diff changeset
   306
as bridge between the different conventions.  
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   307
\<close>
13345
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   308
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   309
instantiation stmt :: inj_term
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   310
begin
13345
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   311
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   312
definition
37956
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   313
  stmt_inj_term_def: "\<langle>c::stmt\<rangle> = In1r c"
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   314
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   315
instance ..
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   316
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   317
end
13345
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   318
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   319
lemma stmt_inj_term_simp: "\<langle>c::stmt\<rangle> = In1r c"
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   320
by (simp add: stmt_inj_term_def)
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   321
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: 13384
diff changeset
   322
lemma  stmt_inj_term [iff]: "\<langle>x::stmt\<rangle> = \<langle>y\<rangle> \<equiv> x = y"
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: 13384
diff changeset
   323
  by (simp add: stmt_inj_term_simp)
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: 13384
diff changeset
   324
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   325
instantiation expr :: inj_term
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   326
begin
13345
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   327
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   328
definition
37956
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   329
  expr_inj_term_def: "\<langle>e::expr\<rangle> = In1l e"
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   330
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   331
instance ..
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   332
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   333
end
13345
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   334
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   335
lemma expr_inj_term_simp: "\<langle>e::expr\<rangle> = In1l e"
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   336
by (simp add: expr_inj_term_def)
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   337
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: 13384
diff changeset
   338
lemma expr_inj_term [iff]: "\<langle>x::expr\<rangle> = \<langle>y\<rangle> \<equiv> x = y"
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: 13384
diff changeset
   339
  by (simp add: expr_inj_term_simp)
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: 13384
diff changeset
   340
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   341
instantiation var :: inj_term
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   342
begin
13345
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   343
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   344
definition
37956
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   345
  var_inj_term_def: "\<langle>v::var\<rangle> = In2 v"
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   346
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   347
instance ..
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   348
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   349
end
13345
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   350
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   351
lemma var_inj_term_simp: "\<langle>v::var\<rangle> = In2 v"
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   352
by (simp add: var_inj_term_def)
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   353
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: 13384
diff changeset
   354
lemma var_inj_term [iff]: "\<langle>x::var\<rangle> = \<langle>y\<rangle> \<equiv> x = y"
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: 13384
diff changeset
   355
  by (simp add: var_inj_term_simp)
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: 13384
diff changeset
   356
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   357
class expr_of =
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   358
  fixes expr_of :: "'a \<Rightarrow> expr"
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   359
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   360
instantiation expr :: expr_of
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   361
begin
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   362
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   363
definition
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   364
  "expr_of = (\<lambda>(e::expr). e)"
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   365
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   366
instance ..
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   367
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   368
end 
13345
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   369
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   370
instantiation list :: (expr_of) inj_term
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   371
begin
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   372
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   373
definition
37956
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   374
  "\<langle>es::'a list\<rangle> = In3 (map expr_of es)"
35315
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   375
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   376
instance ..
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   377
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   378
end
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   379
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   380
lemma expr_list_inj_term_def:
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   381
  "\<langle>es::expr list\<rangle> \<equiv> In3 es"
fbdc860d87a3 dropped axclass
haftmann
parents: 35069
diff changeset
   382
  by (simp add: inj_term_list_def expr_of_expr_def)
13345
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   383
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   384
lemma expr_list_inj_term_simp: "\<langle>es::expr list\<rangle> = In3 es"
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   385
by (simp add: expr_list_inj_term_def)
bd611943cbc2 Fixed markup error in comment.
schirmer
parents: 13337
diff changeset
   386
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: 13384
diff changeset
   387
lemma expr_list_inj_term [iff]: "\<langle>x::expr list\<rangle> = \<langle>y\<rangle> \<equiv> x = y"
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: 13384
diff changeset
   388
  by (simp add: expr_list_inj_term_simp)
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: 13384
diff changeset
   389
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: 13384
diff changeset
   390
lemmas inj_term_simps = stmt_inj_term_simp expr_inj_term_simp var_inj_term_simp
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: 13384
diff changeset
   391
                        expr_list_inj_term_simp
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: 13384
diff changeset
   392
lemmas inj_term_sym_simps = stmt_inj_term_simp [THEN sym] 
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: 13384
diff changeset
   393
                            expr_inj_term_simp [THEN sym]
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: 13384
diff changeset
   394
                            var_inj_term_simp [THEN sym]
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: 13384
diff changeset
   395
                            expr_list_inj_term_simp [THEN sym]
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: 13384
diff changeset
   396
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: 13384
diff changeset
   397
lemma stmt_expr_inj_term [iff]: "\<langle>t::stmt\<rangle> \<noteq> \<langle>w::expr\<rangle>"
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: 13384
diff changeset
   398
  by (simp add: inj_term_simps)
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: 13384
diff changeset
   399
lemma expr_stmt_inj_term [iff]: "\<langle>t::expr\<rangle> \<noteq> \<langle>w::stmt\<rangle>"
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: 13384
diff changeset
   400
  by (simp add: inj_term_simps)
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: 13384
diff changeset
   401
lemma stmt_var_inj_term [iff]: "\<langle>t::stmt\<rangle> \<noteq> \<langle>w::var\<rangle>"
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: 13384
diff changeset
   402
  by (simp add: inj_term_simps)
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: 13384
diff changeset
   403
lemma var_stmt_inj_term [iff]: "\<langle>t::var\<rangle> \<noteq> \<langle>w::stmt\<rangle>"
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: 13384
diff changeset
   404
  by (simp add: inj_term_simps)
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: 13384
diff changeset
   405
lemma stmt_elist_inj_term [iff]: "\<langle>t::stmt\<rangle> \<noteq> \<langle>w::expr list\<rangle>"
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: 13384
diff changeset
   406
  by (simp add: inj_term_simps)
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: 13384
diff changeset
   407
lemma elist_stmt_inj_term [iff]: "\<langle>t::expr list\<rangle> \<noteq> \<langle>w::stmt\<rangle>"
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: 13384
diff changeset
   408
  by (simp add: inj_term_simps)
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: 13384
diff changeset
   409
lemma expr_var_inj_term [iff]: "\<langle>t::expr\<rangle> \<noteq> \<langle>w::var\<rangle>"
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: 13384
diff changeset
   410
  by (simp add: inj_term_simps)
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: 13384
diff changeset
   411
lemma var_expr_inj_term [iff]: "\<langle>t::var\<rangle> \<noteq> \<langle>w::expr\<rangle>"
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: 13384
diff changeset
   412
  by (simp add: inj_term_simps)
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: 13384
diff changeset
   413
lemma expr_elist_inj_term [iff]: "\<langle>t::expr\<rangle> \<noteq> \<langle>w::expr list\<rangle>"
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: 13384
diff changeset
   414
  by (simp add: inj_term_simps)
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: 13384
diff changeset
   415
lemma elist_expr_inj_term [iff]: "\<langle>t::expr list\<rangle> \<noteq> \<langle>w::expr\<rangle>"
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: 13384
diff changeset
   416
  by (simp add: inj_term_simps)
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: 13384
diff changeset
   417
lemma var_elist_inj_term [iff]: "\<langle>t::var\<rangle> \<noteq> \<langle>w::expr list\<rangle>"
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: 13384
diff changeset
   418
  by (simp add: inj_term_simps)
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: 13384
diff changeset
   419
lemma elist_var_inj_term [iff]: "\<langle>t::expr list\<rangle> \<noteq> \<langle>w::var\<rangle>"
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: 13384
diff changeset
   420
  by (simp add: inj_term_simps)
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: 13384
diff changeset
   421
13690
ac335b2f4a39 Inserted some extra paragraphs in large proofs to make tex run...
schirmer
parents: 13688
diff changeset
   422
lemma term_cases: "
ac335b2f4a39 Inserted some extra paragraphs in large proofs to make tex run...
schirmer
parents: 13688
diff changeset
   423
  \<lbrakk>\<And> v. P \<langle>v\<rangle>\<^sub>v; \<And> e. P \<langle>e\<rangle>\<^sub>e;\<And> c. P \<langle>c\<rangle>\<^sub>s;\<And> l. P \<langle>l\<rangle>\<^sub>l\<rbrakk>
ac335b2f4a39 Inserted some extra paragraphs in large proofs to make tex run...
schirmer
parents: 13688
diff changeset
   424
  \<Longrightarrow> P t"
ac335b2f4a39 Inserted some extra paragraphs in large proofs to make tex run...
schirmer
parents: 13688
diff changeset
   425
  apply (cases t)
58251
b13e5c3497f5 ported Bali to new datatypes
blanchet
parents: 58249
diff changeset
   426
  apply (rename_tac a, case_tac a)
13690
ac335b2f4a39 Inserted some extra paragraphs in large proofs to make tex run...
schirmer
parents: 13688
diff changeset
   427
  apply auto
ac335b2f4a39 Inserted some extra paragraphs in large proofs to make tex run...
schirmer
parents: 13688
diff changeset
   428
  done
ac335b2f4a39 Inserted some extra paragraphs in large proofs to make tex run...
schirmer
parents: 13688
diff changeset
   429
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   430
subsubsection \<open>Evaluation of unary operations\<close>
37956
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   431
primrec eval_unop :: "unop \<Rightarrow> val \<Rightarrow> val"
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   432
where
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   433
  "eval_unop UPlus v = Intg (the_Intg v)"
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   434
| "eval_unop UMinus v = Intg (- (the_Intg v))"
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   435
| "eval_unop UBitNot v = Intg 42"                \<comment> "FIXME: Not yet implemented"
37956
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   436
| "eval_unop UNot v = Bool (\<not> the_Bool 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: 13384
diff changeset
   437
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   438
subsubsection \<open>Evaluation of binary operations\<close>
37956
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   439
primrec eval_binop :: "binop \<Rightarrow> val \<Rightarrow> val \<Rightarrow> val"
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   440
where
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   441
  "eval_binop Mul     v1 v2 = Intg ((the_Intg v1) * (the_Intg v2))" 
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   442
| "eval_binop Div     v1 v2 = Intg ((the_Intg v1) div (the_Intg v2))"
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   443
| "eval_binop Mod     v1 v2 = Intg ((the_Intg v1) mod (the_Intg v2))"
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   444
| "eval_binop Plus    v1 v2 = Intg ((the_Intg v1) + (the_Intg v2))"
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   445
| "eval_binop Minus   v1 v2 = Intg ((the_Intg v1) - (the_Intg v2))"
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: 13384
diff changeset
   446
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   447
\<comment> "Be aware of the explicit coercion of the shift distance to nat"
37956
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   448
| "eval_binop LShift  v1 v2 = Intg ((the_Intg v1) *   (2^(nat (the_Intg v2))))"
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   449
| "eval_binop RShift  v1 v2 = Intg ((the_Intg v1) div (2^(nat (the_Intg v2))))"
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   450
| "eval_binop RShiftU v1 v2 = Intg 42" \<comment>"FIXME: Not yet implemented"
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: 13384
diff changeset
   451
37956
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   452
| "eval_binop Less    v1 v2 = Bool ((the_Intg v1) < (the_Intg v2))" 
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   453
| "eval_binop Le      v1 v2 = Bool ((the_Intg v1) \<le> (the_Intg v2))"
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   454
| "eval_binop Greater v1 v2 = Bool ((the_Intg v2) < (the_Intg v1))"
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   455
| "eval_binop Ge      v1 v2 = Bool ((the_Intg v2) \<le> (the_Intg v1))"
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: 13384
diff changeset
   456
37956
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   457
| "eval_binop Eq      v1 v2 = Bool (v1=v2)"
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   458
| "eval_binop Neq     v1 v2 = Bool (v1\<noteq>v2)"
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   459
| "eval_binop BitAnd  v1 v2 = Intg 42" \<comment> "FIXME: Not yet implemented"
37956
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   460
| "eval_binop And     v1 v2 = Bool ((the_Bool v1) \<and> (the_Bool v2))"
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   461
| "eval_binop BitXor  v1 v2 = Intg 42" \<comment> "FIXME: Not yet implemented"
37956
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   462
| "eval_binop Xor     v1 v2 = Bool ((the_Bool v1) \<noteq> (the_Bool v2))"
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   463
| "eval_binop BitOr   v1 v2 = Intg 42" \<comment> "FIXME: Not yet implemented"
37956
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   464
| "eval_binop Or      v1 v2 = Bool ((the_Bool v1) \<or> (the_Bool v2))"
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   465
| "eval_binop CondAnd v1 v2 = Bool ((the_Bool v1) \<and> (the_Bool v2))"
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   466
| "eval_binop CondOr  v1 v2 = Bool ((the_Bool v1) \<or> (the_Bool v2))"
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: 13384
diff changeset
   467
37956
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   468
definition
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   469
  need_second_arg :: "binop \<Rightarrow> val \<Rightarrow> bool" where
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   470
  "need_second_arg binop v1 = (\<not> ((binop=CondAnd \<and>  \<not> the_Bool v1) \<or>
ee939247b2fb modernized/unified some specifications;
wenzelm
parents: 35431
diff changeset
   471
                                 (binop=CondOr  \<and> the_Bool v1)))"
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   472
text \<open>@{term CondAnd} and @{term CondOr} only evalulate the second argument
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 58887
diff changeset
   473
 if the value isn't already determined by the first argument\<close>
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: 13384
diff changeset
   474
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: 13384
diff changeset
   475
lemma need_second_arg_CondAnd [simp]: "need_second_arg CondAnd (Bool b) = b" 
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: 13384
diff changeset
   476
by (simp add: need_second_arg_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: 13384
diff changeset
   477
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: 13384
diff changeset
   478
lemma need_second_arg_CondOr [simp]: "need_second_arg CondOr (Bool b) = (\<not> b)" 
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: 13384
diff changeset
   479
by (simp add: need_second_arg_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: 13384
diff changeset
   480
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: 13384
diff changeset
   481
lemma need_second_arg_strict[simp]: 
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: 13384
diff changeset
   482
 "\<lbrakk>binop\<noteq>CondAnd; binop\<noteq>CondOr\<rbrakk> \<Longrightarrow> need_second_arg binop b"
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: 13384
diff changeset
   483
by (cases binop) 
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: 13384
diff changeset
   484
   (simp_all add: need_second_arg_def)
62390
842917225d56 more canonical names
nipkow
parents: 62042
diff changeset
   485
end