src/HOL/Bali/WellType.thy
author wenzelm
Tue, 10 Jul 2007 23:29:43 +0200
changeset 23719 ccd9cb15c062
parent 21765 89275a3ed7be
child 23747 b07cff284683
permissions -rw-r--r--
more markup for inner and outer syntax; added enclose;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12857
a4386cc9b1c3 tuned header;
wenzelm
parents: 12854
diff changeset
     1
(*  Title:      HOL/Bali/WellType.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
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     4
*)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     5
header {* Well-typedness of Java programs *}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     6
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14981
diff changeset
     7
theory WellType imports DeclConcepts begin
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     8
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
     9
text {*
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    10
improvements over Java Specification 1.0:
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    11
\begin{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    12
\item methods of Object can be called upon references of interface or array type
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    13
\end{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    14
simplifications:
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    15
\begin{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    16
\item the type rules include all static checks on statements and expressions, 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    17
      e.g. definedness of names (of parameters, locals, fields, methods)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    18
\end{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    19
design issues:
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    20
\begin{itemize}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    21
\item unified type judgment for statements, variables, expressions, 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    22
      expression lists
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    23
\item statements are typed like expressions with dummy type Void
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    24
\item the typing rules take an extra argument that is capable of determining 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    25
  the dynamic type of objects. Therefore, they can be used for both 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    26
  checking static types and determining runtime types in transition semantics.
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
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    30
types	lenv
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: 13524
diff changeset
    31
	= "(lname, ty) table"  --{* local variables, including This and Result*}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    32
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    33
record env = 
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: 13524
diff changeset
    34
         prg:: "prog"    --{* program *}
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
    35
         cls:: "qtname"  --{* current package and class name *}
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
    36
         lcl:: "lenv"    --{* local environment *}     
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    37
  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    38
translations
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    39
  "lenv" <= (type) "(lname, ty) table"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    40
  "lenv" <= (type) "lname \<Rightarrow> ty option"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    41
  "env" <= (type) "\<lparr>prg::prog,cls::qtname,lcl::lenv\<rparr>"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    42
  "env" <= (type) "\<lparr>prg::prog,cls::qtname,lcl::lenv,\<dots>::'a\<rparr>"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    43
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    44
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    45
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    46
syntax
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: 13524
diff changeset
    47
  pkg :: "env \<Rightarrow> pname" --{* select the current package from an environment *}
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
    48
translations 
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    49
  "pkg e" == "pid (cls e)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    50
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    51
section "Static overloading: maximally specific methods "
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    52
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    53
types
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    54
  emhead = "ref_ty \<times> mhead"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    55
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: 13524
diff changeset
    56
--{* Some mnemotic selectors for emhead *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    57
constdefs 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    58
  "declrefT" :: "emhead \<Rightarrow> ref_ty"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    59
  "declrefT \<equiv> fst"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    60
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    61
  "mhd"     :: "emhead \<Rightarrow> mhead"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    62
  "mhd \<equiv> snd"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    63
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    64
lemma declrefT_simp[simp]:"declrefT (r,m) = r"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    65
by (simp add: declrefT_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    66
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    67
lemma mhd_simp[simp]:"mhd (r,m) = m"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    68
by (simp add: mhd_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    69
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    70
lemma static_mhd_simp[simp]: "static (mhd m) = is_static m"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    71
by (cases m) (simp add: member_is_static_simp mhd_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    72
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    73
lemma mhd_resTy_simp [simp]: "resTy (mhd m) = resTy m"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    74
by (cases m) simp
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    75
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    76
lemma mhd_is_static_simp [simp]: "is_static (mhd m) = is_static m"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    77
by (cases m) simp
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    78
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    79
lemma mhd_accmodi_simp [simp]: "accmodi (mhd m) = accmodi m"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    80
by (cases m) simp
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    81
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    82
consts
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    83
  cmheads        :: "prog \<Rightarrow> qtname \<Rightarrow> qtname \<Rightarrow> sig \<Rightarrow> emhead set"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    84
  Objectmheads   :: "prog \<Rightarrow> qtname           \<Rightarrow> sig \<Rightarrow> emhead set"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    85
  accObjectmheads:: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> emhead set"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    86
  mheads         :: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> emhead set"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    87
defs
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    88
 cmheads_def:
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    89
"cmheads G S C 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    90
  \<equiv> \<lambda>sig. (\<lambda>(Cls,mthd). (ClassT Cls,(mhead mthd))) ` o2s (accmethd G S C sig)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    91
  Objectmheads_def:
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    92
"Objectmheads G S  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    93
  \<equiv> \<lambda>sig. (\<lambda>(Cls,mthd). (ClassT Cls,(mhead mthd))) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    94
    ` o2s (filter_tab (\<lambda>sig m. accmodi m \<noteq> Private) (accmethd G S Object) sig)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    95
  accObjectmheads_def:
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    96
"accObjectmheads G S T
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    97
   \<equiv> if G\<turnstile>RefT T accessible_in (pid S)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    98
        then Objectmheads G S
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
    99
        else \<lambda>sig. {}"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   100
primrec
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   101
"mheads G S  NullT     = (\<lambda>sig. {})"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   102
"mheads G S (IfaceT I) = (\<lambda>sig. (\<lambda>(I,h).(IfaceT I,h)) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   103
                         ` accimethds G (pid S) I sig \<union> 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   104
                           accObjectmheads G S (IfaceT I) sig)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   105
"mheads G S (ClassT C) = cmheads G S C"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   106
"mheads G S (ArrayT T) = accObjectmheads G S (ArrayT T)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   107
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   108
constdefs
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: 13524
diff changeset
   109
  --{* applicable methods, cf. 15.11.2.1 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   110
  appl_methds   :: "prog \<Rightarrow> qtname \<Rightarrow>  ref_ty \<Rightarrow> sig \<Rightarrow> (emhead \<times> ty list)\<spacespace> set"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   111
 "appl_methds G S rt \<equiv> \<lambda> sig. 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   112
      {(mh,pTs') |mh pTs'. mh \<in> mheads G S rt \<lparr>name=name sig,parTs=pTs'\<rparr> \<and> 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   113
                           G\<turnstile>(parTs sig)[\<preceq>]pTs'}"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   114
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: 13524
diff changeset
   115
  --{* more specific methods, cf. 15.11.2.2 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   116
  more_spec     :: "prog \<Rightarrow> emhead \<times> ty list \<Rightarrow> emhead \<times> ty list \<Rightarrow> bool"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   117
 "more_spec G \<equiv> \<lambda>(mh,pTs). \<lambda>(mh',pTs'). G\<turnstile>pTs[\<preceq>]pTs'"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   118
(*more_spec G \<equiv>\<lambda>((d,h),pTs). \<lambda>((d',h'),pTs'). G\<turnstile>RefT d\<preceq>RefT d'\<and>G\<turnstile>pTs[\<preceq>]pTs'*)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   119
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: 13524
diff changeset
   120
  --{* maximally specific methods, cf. 15.11.2.2 *}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   121
   max_spec     :: "prog \<Rightarrow> qtname \<Rightarrow> ref_ty \<Rightarrow> sig \<Rightarrow> (emhead \<times> ty list)\<spacespace> set"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   122
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   123
 "max_spec G S rt sig \<equiv>{m. m \<in>appl_methds G S rt sig \<and>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   124
		      (\<forall>m'\<in>appl_methds G S rt sig. more_spec G m' m \<longrightarrow> m'=m)}"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   125
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   126
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   127
lemma max_spec2appl_meths: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   128
  "x \<in> max_spec G S T sig \<Longrightarrow> x \<in> appl_methds G S T sig" 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   129
by (auto simp: max_spec_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   130
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   131
lemma appl_methsD: "(mh,pTs')\<in>appl_methds G S T \<lparr>name=mn,parTs=pTs\<rparr> \<Longrightarrow>  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   132
   mh \<in> mheads G S T \<lparr>name=mn,parTs=pTs'\<rparr> \<and> G\<turnstile>pTs[\<preceq>]pTs'"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   133
by (auto simp: appl_methds_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   134
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   135
lemma max_spec2mheads: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   136
"max_spec G S rt \<lparr>name=mn,parTs=pTs\<rparr> = insert (mh, pTs') A 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   137
 \<Longrightarrow> mh \<in> mheads G S rt \<lparr>name=mn,parTs=pTs'\<rparr> \<and> G\<turnstile>pTs[\<preceq>]pTs'"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   138
apply (auto dest: equalityD2 subsetD max_spec2appl_meths appl_methsD)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   139
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   140
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   141
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   142
constdefs
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   143
  empty_dt :: "dyn_ty"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   144
 "empty_dt \<equiv> \<lambda>a. None"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   145
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   146
  invmode :: "('a::type)member_scheme \<Rightarrow> expr \<Rightarrow> inv_mode"
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   147
"invmode m e \<equiv> if is_static m 
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   148
                  then Static 
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   149
                  else if e=Super then SuperM else IntVir"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   150
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   151
lemma invmode_nonstatic [simp]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   152
  "invmode \<lparr>access=a,static=False,\<dots>=x\<rparr> (Acc (LVar e)) = IntVir"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   153
apply (unfold invmode_def)
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   154
apply (simp (no_asm) add: member_is_static_simp)
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   155
done
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   156
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   157
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   158
lemma invmode_Static_eq [simp]: "(invmode m e = Static) = is_static m"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   159
apply (unfold invmode_def)
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   160
apply (simp (no_asm))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   161
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   162
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   163
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   164
lemma invmode_IntVir_eq: "(invmode m e = IntVir) = (\<not>(is_static m) \<and> e\<noteq>Super)"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   165
apply (unfold invmode_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   166
apply (simp (no_asm))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   167
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   168
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   169
lemma Null_staticD: 
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   170
  "a'=Null \<longrightarrow> (is_static m) \<Longrightarrow> invmode m e = IntVir \<longrightarrow> a' \<noteq> Null"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   171
apply (clarsimp simp add: invmode_IntVir_eq)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   172
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   173
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: 13524
diff changeset
   174
section "Typing for unary operations"
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   175
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   176
consts unop_type :: "unop \<Rightarrow> prim_ty"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   177
primrec 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   178
"unop_type UPlus   = Integer"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   179
"unop_type UMinus  = Integer"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   180
"unop_type UBitNot = Integer"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   181
"unop_type UNot    = Boolean"    
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   182
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   183
consts wt_unop :: "unop \<Rightarrow> ty \<Rightarrow> bool"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   184
primrec
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   185
"wt_unop UPlus   t = (t = PrimT Integer)"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   186
"wt_unop UMinus  t = (t = PrimT Integer)"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   187
"wt_unop UBitNot t = (t = PrimT Integer)"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   188
"wt_unop UNot    t = (t = PrimT Boolean)"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   189
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: 13524
diff changeset
   190
section "Typing for binary operations"
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   191
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   192
consts binop_type :: "binop \<Rightarrow> prim_ty"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   193
primrec
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   194
"binop_type Mul      = Integer"     
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   195
"binop_type Div      = Integer"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   196
"binop_type Mod      = Integer"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   197
"binop_type Plus     = Integer"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   198
"binop_type Minus    = Integer"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   199
"binop_type LShift   = Integer"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   200
"binop_type RShift   = Integer"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   201
"binop_type RShiftU  = Integer"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   202
"binop_type Less     = Boolean"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   203
"binop_type Le       = Boolean"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   204
"binop_type Greater  = Boolean"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   205
"binop_type Ge       = Boolean"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   206
"binop_type Eq       = Boolean"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   207
"binop_type Neq      = Boolean"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   208
"binop_type BitAnd   = Integer"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   209
"binop_type And      = Boolean"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   210
"binop_type BitXor   = Integer"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   211
"binop_type Xor      = Boolean"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   212
"binop_type BitOr    = Integer"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   213
"binop_type Or       = Boolean"
13384
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   214
"binop_type CondAnd  = Boolean"
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   215
"binop_type CondOr   = Boolean"
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   216
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   217
consts wt_binop :: "prog \<Rightarrow> binop \<Rightarrow> ty \<Rightarrow> ty \<Rightarrow> bool"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   218
primrec
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   219
"wt_binop G Mul     t1 t2 = ((t1 = PrimT Integer) \<and> (t2 = PrimT Integer))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   220
"wt_binop G Div     t1 t2 = ((t1 = PrimT Integer) \<and> (t2 = PrimT Integer))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   221
"wt_binop G Mod     t1 t2 = ((t1 = PrimT Integer) \<and> (t2 = PrimT Integer))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   222
"wt_binop G Plus    t1 t2 = ((t1 = PrimT Integer) \<and> (t2 = PrimT Integer))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   223
"wt_binop G Minus   t1 t2 = ((t1 = PrimT Integer) \<and> (t2 = PrimT Integer))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   224
"wt_binop G LShift  t1 t2 = ((t1 = PrimT Integer) \<and> (t2 = PrimT Integer))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   225
"wt_binop G RShift  t1 t2 = ((t1 = PrimT Integer) \<and> (t2 = PrimT Integer))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   226
"wt_binop G RShiftU t1 t2 = ((t1 = PrimT Integer) \<and> (t2 = PrimT Integer))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   227
"wt_binop G Less    t1 t2 = ((t1 = PrimT Integer) \<and> (t2 = PrimT Integer))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   228
"wt_binop G Le      t1 t2 = ((t1 = PrimT Integer) \<and> (t2 = PrimT Integer))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   229
"wt_binop G Greater t1 t2 = ((t1 = PrimT Integer) \<and> (t2 = PrimT Integer))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   230
"wt_binop G Ge      t1 t2 = ((t1 = PrimT Integer) \<and> (t2 = PrimT Integer))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   231
"wt_binop G Eq      t1 t2 = (G\<turnstile>t1\<preceq>t2 \<or> G\<turnstile>t2\<preceq>t1)" 
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   232
"wt_binop G Neq     t1 t2 = (G\<turnstile>t1\<preceq>t2 \<or> G\<turnstile>t2\<preceq>t1)"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   233
"wt_binop G BitAnd  t1 t2 = ((t1 = PrimT Integer) \<and> (t2 = PrimT Integer))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   234
"wt_binop G And     t1 t2 = ((t1 = PrimT Boolean) \<and> (t2 = PrimT Boolean))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   235
"wt_binop G BitXor  t1 t2 = ((t1 = PrimT Integer) \<and> (t2 = PrimT Integer))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   236
"wt_binop G Xor     t1 t2 = ((t1 = PrimT Boolean) \<and> (t2 = PrimT Boolean))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   237
"wt_binop G BitOr   t1 t2 = ((t1 = PrimT Integer) \<and> (t2 = PrimT Integer))"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   238
"wt_binop G Or      t1 t2 = ((t1 = PrimT Boolean) \<and> (t2 = PrimT Boolean))"
13384
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   239
"wt_binop G CondAnd t1 t2 = ((t1 = PrimT Boolean) \<and> (t2 = PrimT Boolean))"
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   240
"wt_binop G CondOr  t1 t2 = ((t1 = PrimT Boolean) \<and> (t2 = PrimT Boolean))"
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   241
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: 13524
diff changeset
   242
section "Typing for terms"
13384
a34e38154413 Added conditional and (&&) and or (||).
schirmer
parents: 13337
diff changeset
   243
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   244
types tys  =        "ty + ty list"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   245
translations
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   246
  "tys"   <= (type) "ty + ty list"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   247
21765
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   248
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   249
inductive2 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   250
  wt :: "env \<Rightarrow> dyn_ty \<Rightarrow> [term,tys] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>_"  [51,51,51,51] 50)
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   251
  and wt_stmt :: "env \<Rightarrow> dyn_ty \<Rightarrow>  stmt       \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>\<surd>"  [51,51,51   ] 50)
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   252
  and ty_expr :: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>-_" [51,51,51,51] 50)
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   253
  and ty_var :: "env \<Rightarrow> dyn_ty \<Rightarrow> [var  ,ty ] \<Rightarrow> bool" ("_,_\<Turnstile>_\<Colon>=_" [51,51,51,51] 50)
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   254
  and ty_exprs :: "env \<Rightarrow> dyn_ty \<Rightarrow> [expr list, ty   list] \<Rightarrow> bool"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   255
    ("_,_\<Turnstile>_\<Colon>\<doteq>_" [51,51,51,51] 50)
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   256
where
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   257
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   258
  "E,dt\<Turnstile>s\<Colon>\<surd> \<equiv> E,dt\<Turnstile>In1r s\<Colon>Inl (PrimT Void)"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   259
| "E,dt\<Turnstile>e\<Colon>-T \<equiv> E,dt\<Turnstile>In1l e\<Colon>Inl T"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   260
| "E,dt\<Turnstile>e\<Colon>=T \<equiv> E,dt\<Turnstile>In2  e\<Colon>Inl T"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   261
| "E,dt\<Turnstile>e\<Colon>\<doteq>T \<equiv> E,dt\<Turnstile>In3  e\<Colon>Inr T"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   262
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   263
--{* well-typed statements *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   264
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   265
| Skip:					"E,dt\<Turnstile>Skip\<Colon>\<surd>"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   266
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   267
| Expr:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   268
					 E,dt\<Turnstile>Expr e\<Colon>\<surd>"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   269
  --{* cf. 14.6 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   270
| Lab:  "E,dt\<Turnstile>c\<Colon>\<surd> \<Longrightarrow>                   
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   271
                                         E,dt\<Turnstile>l\<bullet> c\<Colon>\<surd>" 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   272
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   273
| Comp:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   274
	  E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   275
					 E,dt\<Turnstile>c1;; c2\<Colon>\<surd>"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   276
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   277
  --{* cf. 14.8 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   278
| If:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-PrimT Boolean;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   279
	  E,dt\<Turnstile>c1\<Colon>\<surd>;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   280
	  E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   281
					 E,dt\<Turnstile>If(e) c1 Else c2\<Colon>\<surd>"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   282
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   283
  --{* cf. 14.10 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   284
| Loop:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-PrimT Boolean;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   285
	  E,dt\<Turnstile>c\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   286
					 E,dt\<Turnstile>l\<bullet> While(e) c\<Colon>\<surd>"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   287
  --{* cf. 14.13, 14.15, 14.16 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   288
| Jmp:                                   "E,dt\<Turnstile>Jmp jump\<Colon>\<surd>"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   289
21765
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   290
  --{* cf. 14.16 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   291
| Throw: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class tn;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   292
	  prg E\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   293
					 E,dt\<Turnstile>Throw e\<Colon>\<surd>"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   294
  --{* cf. 14.18 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   295
| Try:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; prg E\<turnstile>tn\<preceq>\<^sub>C SXcpt Throwable;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   296
	  lcl E (VName vn)=None; E \<lparr>lcl := lcl E(VName vn\<mapsto>Class tn)\<rparr>,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   297
          \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   298
					 E,dt\<Turnstile>Try c1 Catch(tn vn) c2\<Colon>\<surd>"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   299
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   300
  --{* cf. 14.18 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   301
| Fin:	"\<lbrakk>E,dt\<Turnstile>c1\<Colon>\<surd>; E,dt\<Turnstile>c2\<Colon>\<surd>\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   302
					 E,dt\<Turnstile>c1 Finally c2\<Colon>\<surd>"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   303
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   304
| Init:	"\<lbrakk>is_class (prg E) C\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   305
					 E,dt\<Turnstile>Init C\<Colon>\<surd>"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   306
  --{* @{term Init} is created on the fly during evaluation (see Eval.thy). 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   307
     The class isn't necessarily accessible from the points @{term Init} 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   308
     is called. Therefor we only demand @{term is_class} and not 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   309
     @{term is_acc_class} here. 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   310
   *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   311
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   312
--{* well-typed expressions *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   313
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   314
  --{* cf. 15.8 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   315
| NewC:	"\<lbrakk>is_acc_class (prg E) (pkg E) C\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   316
					 E,dt\<Turnstile>NewC C\<Colon>-Class C"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   317
  --{* cf. 15.9 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   318
| NewA:	"\<lbrakk>is_acc_type (prg E) (pkg E) T;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   319
	  E,dt\<Turnstile>i\<Colon>-PrimT Integer\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   320
					 E,dt\<Turnstile>New T[i]\<Colon>-T.[]"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   321
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   322
  --{* cf. 15.15 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   323
| Cast:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T; is_acc_type (prg E) (pkg E) T';
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   324
	  prg E\<turnstile>T\<preceq>? T'\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   325
					 E,dt\<Turnstile>Cast T' e\<Colon>-T'"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   326
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   327
  --{* cf. 15.19.2 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   328
| Inst:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT T; is_acc_type (prg E) (pkg E) (RefT T');
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   329
	  prg E\<turnstile>RefT T\<preceq>? RefT T'\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   330
					 E,dt\<Turnstile>e InstOf T'\<Colon>-PrimT Boolean"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   331
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   332
  --{* cf. 15.7.1 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   333
| Lit:	"\<lbrakk>typeof dt x = Some T\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   334
					 E,dt\<Turnstile>Lit x\<Colon>-T"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   335
21765
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   336
| UnOp: "\<lbrakk>E,dt\<Turnstile>e\<Colon>-Te; wt_unop unop Te; T=PrimT (unop_type unop)\<rbrakk> 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   337
          \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   338
	  E,dt\<Turnstile>UnOp unop e\<Colon>-T"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   339
  
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   340
| BinOp: "\<lbrakk>E,dt\<Turnstile>e1\<Colon>-T1; E,dt\<Turnstile>e2\<Colon>-T2; wt_binop (prg E) binop T1 T2; 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   341
           T=PrimT (binop_type binop)\<rbrakk> 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   342
           \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   343
	   E,dt\<Turnstile>BinOp binop e1 e2\<Colon>-T"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   344
  
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   345
  --{* cf. 15.10.2, 15.11.1 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   346
| Super: "\<lbrakk>lcl E This = Some (Class C); C \<noteq> Object;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   347
	  class (prg E) C = Some c\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   348
					 E,dt\<Turnstile>Super\<Colon>-Class (super c)"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   349
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   350
  --{* cf. 15.13.1, 15.10.1, 15.12 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   351
| Acc:	"\<lbrakk>E,dt\<Turnstile>va\<Colon>=T\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   352
					 E,dt\<Turnstile>Acc va\<Colon>-T"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   353
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   354
  --{* cf. 15.25, 15.25.1 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   355
| Ass:	"\<lbrakk>E,dt\<Turnstile>va\<Colon>=T; va \<noteq> LVar This;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   356
	  E,dt\<Turnstile>v \<Colon>-T';
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   357
	  prg E\<turnstile>T'\<preceq>T\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   358
					 E,dt\<Turnstile>va:=v\<Colon>-T'"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   359
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   360
  --{* cf. 15.24 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   361
| Cond:	"\<lbrakk>E,dt\<Turnstile>e0\<Colon>-PrimT Boolean;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   362
	  E,dt\<Turnstile>e1\<Colon>-T1; E,dt\<Turnstile>e2\<Colon>-T2;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   363
	  prg E\<turnstile>T1\<preceq>T2 \<and> T = T2  \<or>  prg E\<turnstile>T2\<preceq>T1 \<and> T = T1\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   364
					 E,dt\<Turnstile>e0 ? e1 : e2\<Colon>-T"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   365
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   366
  --{* cf. 15.11.1, 15.11.2, 15.11.3 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   367
| Call:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT statT;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   368
	  E,dt\<Turnstile>ps\<Colon>\<doteq>pTs;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   369
	  max_spec (prg E) (cls E) statT \<lparr>name=mn,parTs=pTs\<rparr> 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   370
            = {((statDeclT,m),pTs')}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   371
         \<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   372
		   E,dt\<Turnstile>{cls E,statT,invmode m e}e\<cdot>mn({pTs'}ps)\<Colon>-(resTy m)"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   373
21765
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   374
| Methd: "\<lbrakk>is_class (prg E) C;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   375
	  methd (prg E) C sig = Some m;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   376
	  E,dt\<Turnstile>Body (declclass m) (stmt (mbody (mthd m)))\<Colon>-T\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   377
					 E,dt\<Turnstile>Methd C sig\<Colon>-T"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   378
 --{* The class @{term C} is the dynamic class of the method call 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   379
    (cf. Eval.thy). 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   380
    It hasn't got to be directly accessible from the current package 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   381
    @{term "(pkg E)"}. 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   382
    Only the static class must be accessible (enshured indirectly by 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   383
    @{term Call}). 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   384
    Note that l is just a dummy value. It is only used in the smallstep 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   385
    semantics. To proof typesafety directly for the smallstep semantics 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   386
    we would have to assume conformance of l here!
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   387
  *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   388
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   389
| Body:	"\<lbrakk>is_class (prg E) D;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   390
	  E,dt\<Turnstile>blk\<Colon>\<surd>;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   391
	  (lcl E) Result = Some T;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   392
          is_type (prg E) T\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   393
   					 E,dt\<Turnstile>Body D blk\<Colon>-T"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   394
--{* The class @{term D} implementing the method must not directly be 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   395
     accessible  from the current package @{term "(pkg E)"}, but can also 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   396
     be indirectly accessible due to inheritance (enshured in @{term Call})
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   397
    The result type hasn't got to be accessible in Java! (If it is not 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   398
    accessible you can only assign it to Object).
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   399
    For dummy value l see rule @{term Methd}. 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   400
   *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   401
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   402
--{* well-typed variables *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   403
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   404
  --{* cf. 15.13.1 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   405
| LVar:	"\<lbrakk>lcl E vn = Some T; is_acc_type (prg E) (pkg E) T\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   406
					 E,dt\<Turnstile>LVar vn\<Colon>=T"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   407
  --{* cf. 15.10.1 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   408
| FVar:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class C; 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   409
	  accfield (prg E) (cls E) C fn = Some (statDeclC,f)\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   410
			 E,dt\<Turnstile>{cls E,statDeclC,is_static f}e..fn\<Colon>=(type f)"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   411
  --{* cf. 15.12 *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   412
| AVar:	"\<lbrakk>E,dt\<Turnstile>e\<Colon>-T.[]; 
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   413
	  E,dt\<Turnstile>i\<Colon>-PrimT Integer\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   414
					 E,dt\<Turnstile>e.[i]\<Colon>=T"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   415
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   416
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   417
--{* well-typed expression lists *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   418
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   419
  --{* cf. 15.11.??? *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   420
| Nil:					"E,dt\<Turnstile>[]\<Colon>\<doteq>[]"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   421
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   422
  --{* cf. 15.11.??? *}
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   423
| Cons:	"\<lbrakk>E,dt\<Turnstile>e \<Colon>-T;
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   424
	  E,dt\<Turnstile>es\<Colon>\<doteq>Ts\<rbrakk> \<Longrightarrow>
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   425
					 E,dt\<Turnstile>e#es\<Colon>\<doteq>T#Ts"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   426
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   427
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   428
syntax (* for purely static typing *)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   429
  wt_      :: "env \<Rightarrow> [term,tys] \<Rightarrow> bool" ("_|-_::_" [51,51,51] 50)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   430
  wt_stmt_ :: "env \<Rightarrow>  stmt       \<Rightarrow> bool" ("_|-_:<>" [51,51   ] 50)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   431
  ty_expr_ :: "env \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_|-_:-_" [51,51,51] 50)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   432
  ty_var_  :: "env \<Rightarrow> [var  ,ty ] \<Rightarrow> bool" ("_|-_:=_" [51,51,51] 50)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   433
  ty_exprs_:: "env \<Rightarrow> [expr list,
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   434
		     \<spacespace> ty   list] \<Rightarrow> bool" ("_|-_:#_" [51,51,51] 50)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   435
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   436
syntax (xsymbols)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   437
  wt_      :: "env \<Rightarrow> [term,tys] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>_"  [51,51,51] 50)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   438
  wt_stmt_ ::  "env \<Rightarrow>  stmt       \<Rightarrow> bool" ("_\<turnstile>_\<Colon>\<surd>"  [51,51   ] 50)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   439
  ty_expr_ :: "env \<Rightarrow> [expr ,ty ] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>-_" [51,51,51] 50)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   440
  ty_var_  :: "env \<Rightarrow> [var  ,ty ] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>=_" [51,51,51] 50)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   441
  ty_exprs_ :: "env \<Rightarrow> [expr list,
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   442
		    \<spacespace>  ty   list] \<Rightarrow> bool" ("_\<turnstile>_\<Colon>\<doteq>_" [51,51,51] 50)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   443
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   444
translations
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   445
	"E\<turnstile>t\<Colon> T" == "E,empty_dt\<Turnstile>t\<Colon> T"
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: 13524
diff changeset
   446
        "E\<turnstile>s\<Colon>\<surd>"  == "E\<turnstile>In1r s\<Colon>Inl (PrimT Void)"
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   447
	"E\<turnstile>e\<Colon>-T" == "E\<turnstile>In1l e\<Colon>Inl T"
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   448
	"E\<turnstile>e\<Colon>=T" == "E\<turnstile>In2  e\<Colon>Inl T"
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   449
	"E\<turnstile>e\<Colon>\<doteq>T" == "E\<turnstile>In3  e\<Colon>Inr T"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   450
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   451
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   452
declare not_None_eq [simp del] 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   453
declare split_if [split del] split_if_asm [split del]
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   454
declare split_paired_All [simp del] split_paired_Ex [simp del]
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   455
ML_setup {*
17876
b9c92f384109 change_claset/simpset;
wenzelm
parents: 16417
diff changeset
   456
change_simpset (fn ss => ss delloop "split_all_tac")
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   457
*}
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: 13524
diff changeset
   458
21765
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   459
inductive_cases2 wt_elim_cases [cases set]:
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   460
	"E,dt\<Turnstile>In2  (LVar vn)               \<Colon>T"
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   461
	"E,dt\<Turnstile>In2  ({accC,statDeclC,s}e..fn)\<Colon>T"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   462
	"E,dt\<Turnstile>In2  (e.[i])                 \<Colon>T"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   463
	"E,dt\<Turnstile>In1l (NewC C)                \<Colon>T"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   464
	"E,dt\<Turnstile>In1l (New T'[i])             \<Colon>T"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   465
	"E,dt\<Turnstile>In1l (Cast T' e)             \<Colon>T"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   466
	"E,dt\<Turnstile>In1l (e InstOf T')           \<Colon>T"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   467
	"E,dt\<Turnstile>In1l (Lit x)                 \<Colon>T"
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   468
        "E,dt\<Turnstile>In1l (UnOp unop e)           \<Colon>T"
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   469
        "E,dt\<Turnstile>In1l (BinOp binop e1 e2)     \<Colon>T"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   470
	"E,dt\<Turnstile>In1l (Super)                 \<Colon>T"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   471
	"E,dt\<Turnstile>In1l (Acc va)                \<Colon>T"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   472
	"E,dt\<Turnstile>In1l (Ass va v)              \<Colon>T"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   473
	"E,dt\<Turnstile>In1l (e0 ? e1 : e2)          \<Colon>T"
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   474
	"E,dt\<Turnstile>In1l ({accC,statT,mode}e\<cdot>mn({pT'}p))\<Colon>T"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   475
	"E,dt\<Turnstile>In1l (Methd C sig)           \<Colon>T"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   476
	"E,dt\<Turnstile>In1l (Body D blk)            \<Colon>T"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   477
	"E,dt\<Turnstile>In3  ([])                    \<Colon>Ts"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   478
	"E,dt\<Turnstile>In3  (e#es)                  \<Colon>Ts"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   479
	"E,dt\<Turnstile>In1r  Skip                   \<Colon>x"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   480
	"E,dt\<Turnstile>In1r (Expr e)                \<Colon>x"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   481
	"E,dt\<Turnstile>In1r (c1;; c2)               \<Colon>x"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   482
        "E,dt\<Turnstile>In1r (l\<bullet> c)                  \<Colon>x" 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   483
	"E,dt\<Turnstile>In1r (If(e) c1 Else c2)      \<Colon>x"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   484
	"E,dt\<Turnstile>In1r (l\<bullet> While(e) c)         \<Colon>x"
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: 13524
diff changeset
   485
        "E,dt\<Turnstile>In1r (Jmp jump)              \<Colon>x"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   486
	"E,dt\<Turnstile>In1r (Throw e)               \<Colon>x"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   487
	"E,dt\<Turnstile>In1r (Try c1 Catch(tn vn) c2)\<Colon>x"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   488
	"E,dt\<Turnstile>In1r (c1 Finally c2)         \<Colon>x"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   489
	"E,dt\<Turnstile>In1r (Init C)                \<Colon>x"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   490
declare not_None_eq [simp] 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   491
declare split_if [split] split_if_asm [split]
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   492
declare split_paired_All [simp] split_paired_Ex [simp]
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   493
ML_setup {*
17876
b9c92f384109 change_claset/simpset;
wenzelm
parents: 16417
diff changeset
   494
change_simpset (fn ss => ss addloop ("split_all_tac", split_all_tac));
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   495
*}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   496
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   497
lemma is_acc_class_is_accessible: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   498
  "is_acc_class G P C \<Longrightarrow> G\<turnstile>(Class C) accessible_in P"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   499
by (auto simp add: is_acc_class_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   500
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   501
lemma is_acc_iface_is_iface: "is_acc_iface G P I \<Longrightarrow> is_iface G I"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   502
by (auto simp add: is_acc_iface_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   503
13524
604d0f3622d6 *** empty log message ***
wenzelm
parents: 13462
diff changeset
   504
lemma is_acc_iface_Iface_is_accessible: 
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   505
  "is_acc_iface G P I \<Longrightarrow> G\<turnstile>(Iface I) accessible_in P"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   506
by (auto simp add: is_acc_iface_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   507
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   508
lemma is_acc_type_is_type: "is_acc_type G P T \<Longrightarrow> is_type G T"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   509
by (auto simp add: is_acc_type_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   510
13524
604d0f3622d6 *** empty log message ***
wenzelm
parents: 13462
diff changeset
   511
lemma is_acc_iface_is_accessible:
604d0f3622d6 *** empty log message ***
wenzelm
parents: 13462
diff changeset
   512
  "is_acc_type G P T \<Longrightarrow> G\<turnstile>T accessible_in P"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   513
by (auto simp add: is_acc_type_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   514
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   515
lemma wt_Methd_is_methd: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   516
  "E\<turnstile>In1l (Methd C sig)\<Colon>T \<Longrightarrow> is_methd (prg E) C sig"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   517
apply (erule_tac wt_elim_cases)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   518
apply clarsimp
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   519
apply (erule is_methdI, assumption)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   520
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   521
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: 13524
diff changeset
   522
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   523
text {* Special versions of some typing rules, better suited to pattern 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   524
        match the conclusion (no selectors in the conclusion)
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   525
*}
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   526
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   527
lemma wt_Call: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   528
"\<lbrakk>E,dt\<Turnstile>e\<Colon>-RefT statT; E,dt\<Turnstile>ps\<Colon>\<doteq>pTs;  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   529
  max_spec (prg E) (cls E) 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: 12858
diff changeset
   530
    = {((statDeclC,m),pTs')};rT=(resTy m);accC=cls E;
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   531
 mode = invmode m e\<rbrakk> \<Longrightarrow> E,dt\<Turnstile>{accC,statT,mode}e\<cdot>mn({pTs'}ps)\<Colon>-rT"
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   532
by (auto elim: wt.Call)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   533
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   534
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   535
lemma invocationTypeExpr_noClassD: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   536
"\<lbrakk> E\<turnstile>e\<Colon>-RefT statT\<rbrakk>
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   537
 \<Longrightarrow> (\<forall> statC. statT \<noteq> ClassT statC) \<longrightarrow> invmode m e \<noteq> SuperM"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   538
proof -
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   539
  assume wt: "E\<turnstile>e\<Colon>-RefT statT"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   540
  show ?thesis
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   541
  proof (cases "e=Super")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   542
    case True
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   543
    with wt obtain "C" where "statT = ClassT C" by (blast elim: wt_elim_cases)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   544
    then show ?thesis by blast
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   545
  next 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   546
    case False then show ?thesis 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   547
      by (auto simp add: invmode_def split: split_if_asm)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   548
  qed
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   549
qed
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   550
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   551
lemma wt_Super: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   552
"\<lbrakk>lcl E This = Some (Class C); C \<noteq> Object; class (prg E) C = Some c; D=super c\<rbrakk> 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   553
\<Longrightarrow> E,dt\<Turnstile>Super\<Colon>-Class D"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   554
by (auto elim: wt.Super)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   555
12925
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   556
lemma wt_FVar:	
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   557
"\<lbrakk>E,dt\<Turnstile>e\<Colon>-Class C; accfield (prg E) (cls E) C fn = Some (statDeclC,f);
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   558
  sf=is_static f; fT=(type f); accC=cls E\<rbrakk> 
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   559
\<Longrightarrow> E,dt\<Turnstile>{accC,statDeclC,sf}e..fn\<Colon>=fT"
99131847fb93 Added check for field/method access to operational semantics and proved the acesses valid.
schirmer
parents: 12858
diff changeset
   560
by (auto dest: wt.FVar)
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   561
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   562
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   563
lemma wt_init [iff]: "E,dt\<Turnstile>Init C\<Colon>\<surd> = is_class (prg E) C"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   564
by (auto elim: wt_elim_cases intro: "wt.Init")
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   565
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   566
declare wt.Skip [iff]
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   567
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   568
lemma wt_StatRef: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   569
  "is_acc_type (prg E) (pkg E) (RefT rt) \<Longrightarrow> E\<turnstile>StatRef rt\<Colon>-RefT rt"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   570
apply (rule wt.Cast)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   571
apply  (rule wt.Lit)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   572
apply   (simp (no_asm))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   573
apply  (simp (no_asm_simp))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   574
apply (rule cast.widen)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   575
apply (simp (no_asm))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   576
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   577
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   578
lemma wt_Inj_elim: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   579
  "\<And>E. E,dt\<Turnstile>t\<Colon>U \<Longrightarrow> case t of 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   580
                       In1 ec \<Rightarrow> (case ec of 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   581
                                    Inl e \<Rightarrow> \<exists>T. U=Inl T  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   582
                                  | Inr s \<Rightarrow> U=Inl (PrimT Void))  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   583
                     | In2 e \<Rightarrow> (\<exists>T. U=Inl T) 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   584
                     | In3 e \<Rightarrow> (\<exists>T. U=Inr T)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   585
apply (erule wt.induct)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   586
apply auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   587
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   588
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: 13524
diff changeset
   589
--{* In the special syntax to distinguish the typing judgements for expressions, 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   590
     statements, variables and expression lists the kind of term corresponds
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   591
     to the kind of type in the end e.g. An statement (injection @{term In3} 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   592
    into terms, always has type void (injection @{term Inl} into the generalised
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   593
    types. The following simplification procedures establish these kinds 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: 13524
diff changeset
   594
    correlation. 
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   595
 *}
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   596
21765
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   597
lemma wt_expr_eq: "E,dt\<Turnstile>In1l t\<Colon>U = (\<exists>T. U=Inl T \<and> E,dt\<Turnstile>t\<Colon>-T)"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   598
  by (auto, frule wt_Inj_elim, auto)
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   599
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   600
lemma wt_var_eq: "E,dt\<Turnstile>In2 t\<Colon>U = (\<exists>T. U=Inl T \<and> E,dt\<Turnstile>t\<Colon>=T)"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   601
  by (auto, frule wt_Inj_elim, auto)
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   602
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   603
lemma wt_exprs_eq: "E,dt\<Turnstile>In3 t\<Colon>U = (\<exists>Ts. U=Inr Ts \<and> E,dt\<Turnstile>t\<Colon>\<doteq>Ts)"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   604
  by (auto, frule wt_Inj_elim, auto)
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   605
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   606
lemma wt_stmt_eq: "E,dt\<Turnstile>In1r t\<Colon>U = (U=Inl(PrimT Void)\<and>E,dt\<Turnstile>t\<Colon>\<surd>)"
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   607
  by (auto, frule wt_Inj_elim, auto, frule wt_Inj_elim, auto)
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   608
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   609
ML {*
21765
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   610
fun wt_fun name lhs =
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   611
let
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   612
  fun is_Inj (Const (inj,_) $ _) = true
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   613
    | is_Inj _                   = false
21765
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   614
  fun pred (_ $ _ $ _ $ _ $ x) = is_Inj x
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   615
in
13462
56610e2ba220 sane interface for simprocs;
wenzelm
parents: 13384
diff changeset
   616
  cond_simproc name lhs pred (thm name)
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   617
end
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   618
21765
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   619
val wt_expr_proc  = wt_fun "wt_expr_eq"  "E,dt\<Turnstile>In1l t\<Colon>U";
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   620
val wt_var_proc   = wt_fun "wt_var_eq"   "E,dt\<Turnstile>In2 t\<Colon>U";
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   621
val wt_exprs_proc = wt_fun "wt_exprs_eq" "E,dt\<Turnstile>In3 t\<Colon>U";
89275a3ed7be Adapted to new inductive definition package.
berghofe
parents: 17876
diff changeset
   622
val wt_stmt_proc  = wt_fun "wt_stmt_eq"  "E,dt\<Turnstile>In1r t\<Colon>U";
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   623
*}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   624
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   625
ML {*
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   626
Addsimprocs [wt_expr_proc,wt_var_proc,wt_exprs_proc,wt_stmt_proc]
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   627
*}
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   628
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: 13524
diff changeset
   629
lemma wt_elim_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: 13524
diff changeset
   630
  "\<lbrakk>E,dt\<Turnstile>In1l (BinOp binop e1 e2)\<Colon>T;
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   631
    \<And>T1 T2 T3.
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   632
       \<lbrakk>E,dt\<Turnstile>e1\<Colon>-T1; E,dt\<Turnstile>e2\<Colon>-T2; wt_binop (prg E) binop T1 T2;
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   633
          E,dt\<Turnstile>(if b then In1l e2 else In1r Skip)\<Colon>T3;
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   634
          T = Inl (PrimT (binop_type binop))\<rbrakk>
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   635
       \<Longrightarrow> P\<rbrakk>
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   636
\<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: 13524
diff changeset
   637
apply (erule wt_elim_cases)
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   638
apply (cases 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: 13524
diff changeset
   639
apply auto
a0b16d42d489 "Definite Assignment Analysis" included, with proof of correctness. Large adjustments of type safety proof and soundness proof of the axiomatic semantics were necessary. Completeness proof of the loop rule of the axiomatic semantic was altered. So the additional polymorphic variants of some rules could be removed.
schirmer
parents: 13524
diff changeset
   640
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: 13524
diff changeset
   641
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   642
lemma Inj_eq_lemma [simp]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   643
  "(\<forall>T. (\<exists>T'. T = Inj T' \<and> P T') \<longrightarrow> Q T) = (\<forall>T'. P T' \<longrightarrow> Q (Inj T'))"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   644
by auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   645
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   646
(* unused *)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   647
lemma single_valued_tys_lemma [rule_format (no_asm)]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   648
  "\<forall>S T. G\<turnstile>S\<preceq>T \<longrightarrow> G\<turnstile>T\<preceq>S \<longrightarrow> S = T \<Longrightarrow> E,dt\<Turnstile>t\<Colon>T \<Longrightarrow>  
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   649
     G = prg E \<longrightarrow> (\<forall>T'. E,dt\<Turnstile>t\<Colon>T' \<longrightarrow> T  = T')"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   650
apply (cases "E", erule wt.induct)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   651
apply (safe del: disjE)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   652
apply (simp_all (no_asm_use) split del: split_if_asm)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   653
apply (safe del: disjE)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   654
(* 17 subgoals *)
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   655
apply (tactic {* ALLGOALS (fn i => if i = 11 then EVERY'[thin_tac "?E,dt\<Turnstile>e0\<Colon>-PrimT Boolean", thin_tac "?E,dt\<Turnstile>e1\<Colon>-?T1", thin_tac "?E,dt\<Turnstile>e2\<Colon>-?T2"] i else thin_tac "All ?P" i) *})
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   656
(*apply (safe del: disjE elim!: wt_elim_cases)*)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   657
apply (tactic {*ALLGOALS (eresolve_tac (thms "wt_elim_cases"))*})
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   658
apply (simp_all (no_asm_use) split del: split_if_asm)
13337
f75dfc606ac7 Added unary and binary operations like (+,-,<, ...); Added smallstep semantics (no proofs about it yet).
schirmer
parents: 12925
diff changeset
   659
apply (erule_tac [12] V = "All ?P" in thin_rl) (* Call *)
12854
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   660
apply ((blast del: equalityCE dest: sym [THEN trans])+)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   661
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   662
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   663
(* unused *)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   664
lemma single_valued_tys: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   665
"ws_prog (prg E) \<Longrightarrow> single_valued {(t,T). E,dt\<Turnstile>t\<Colon>T}"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   666
apply (unfold single_valued_def)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   667
apply clarsimp
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   668
apply (rule single_valued_tys_lemma)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   669
apply (auto intro!: widen_antisym)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   670
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   671
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   672
lemma typeof_empty_is_type [rule_format (no_asm)]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   673
  "typeof (\<lambda>a. None) v = Some T \<longrightarrow> is_type G T"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   674
apply (rule val.induct)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   675
apply    	auto
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   676
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   677
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   678
(* unused *)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   679
lemma typeof_is_type [rule_format (no_asm)]: 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   680
 "(\<forall>a. v \<noteq> Addr a) \<longrightarrow> (\<exists>T. typeof dt v = Some T \<and> is_type G T)"
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   681
apply (rule val.induct)
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   682
prefer 5 
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   683
apply     fast
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   684
apply  (simp_all (no_asm))
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   685
done
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   686
00d4a435777f Isabelle/Bali sources;
schirmer
parents:
diff changeset
   687
end