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