src/HOL/MicroJava/BV/BVSpec.thy
author haftmann
Fri, 11 Jun 2010 17:14:01 +0200
changeset 37406 982f3e02f3c4
parent 33954 1bc3b688548c
child 58886 8a6cac7c7247
permissions -rw-r--r--
modernized specifications
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/BV/BVSpec.thy
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
     2
    Author:     Cornelia Pusch, Gerwin Klein
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     3
    Copyright   1999 Technische Universitaet Muenchen
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     4
*)
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
     5
12911
704713ca07ea new document
kleing
parents: 12516
diff changeset
     6
header {* \isaheader{The Bytecode Verifier}\label{sec:BVSpec} *}
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9559
diff changeset
     7
27680
5a557a5afc48 added explicit root theory; some tuning
haftmann
parents: 26453
diff changeset
     8
theory BVSpec
5a557a5afc48 added explicit root theory; some tuning
haftmann
parents: 26453
diff changeset
     9
imports Effect
5a557a5afc48 added explicit root theory; some tuning
haftmann
parents: 26453
diff changeset
    10
begin
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    11
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    12
text {*
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    13
  This theory contains a specification of the BV. The specification
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    14
  describes correct typings of method bodies; it corresponds 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    15
  to type \emph{checking}.
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    16
*}
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    17
37406
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    18
definition
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    19
  -- "The program counter will always be inside the method:"
37406
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    20
  check_bounded :: "instr list \<Rightarrow> exception_table \<Rightarrow> bool" where
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    21
  "check_bounded ins et \<longleftrightarrow>
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    22
  (\<forall>pc < length ins. \<forall>pc' \<in> set (succs (ins!pc) pc). pc' < length ins) \<and>
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    23
                     (\<forall>e \<in> set et. fst (snd (snd e)) < length ins)"
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    24
37406
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    25
definition
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    26
  -- "The method type only contains declared classes:"
37406
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    27
  check_types :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> JVMType.state list \<Rightarrow> bool" where
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    28
  "check_types G mxs mxr phi \<longleftrightarrow> set phi \<subseteq> states G mxs mxr"
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    29
37406
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    30
definition
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    31
  -- "An instruction is welltyped if it is applicable and its effect"
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    32
  -- "is compatible with the type at all successor instructions:"
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    33
  wt_instr :: "[instr,jvm_prog,ty,method_type,nat,p_count,
37406
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    34
                exception_table,p_count] \<Rightarrow> bool" where
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    35
  "wt_instr i G rT phi mxs max_pc et pc \<longleftrightarrow>
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    36
  app i G mxs rT pc et (phi!pc) \<and>
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    37
  (\<forall>(pc',s') \<in> set (eff i G pc et (phi!pc)). pc' < max_pc \<and> G \<turnstile> s' <=' phi!pc')"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
    38
37406
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    39
definition
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    40
  -- {* The type at @{text "pc=0"} conforms to the method calling convention: *}
37406
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    41
  wt_start :: "[jvm_prog,cname,ty list,nat,method_type] \<Rightarrow> bool" where
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    42
  "wt_start G C pTs mxl phi \<longleftrightarrow>
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    43
  G \<turnstile> Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err)) <=' phi!0"
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    44
37406
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    45
definition
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    46
  -- "A method is welltyped if the body is not empty, if execution does not"
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    47
  -- "leave the body, if the method type covers all instructions and mentions"
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    48
  -- "declared classes only, if the method calling convention is respected, and"
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    49
  -- "if all instructions are welltyped."
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    50
  wt_method :: "[jvm_prog,cname,ty list,ty,nat,nat,instr list,
37406
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    51
                 exception_table,method_type] \<Rightarrow> bool" where
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    52
  "wt_method G C pTs rT mxs mxl ins et phi \<longleftrightarrow>
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    53
  (let max_pc = length ins in
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    54
  0 < max_pc \<and> 
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    55
  length phi = length ins \<and>
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    56
  check_bounded ins et \<and> 
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    57
  check_types G mxs (1+length pTs+mxl) (map OK phi) \<and>
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    58
  wt_start G C pTs mxl phi \<and>
37406
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    59
  (\<forall>pc. pc<max_pc \<longrightarrow> wt_instr (ins!pc) G rT phi mxs max_pc et pc))"
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
    60
37406
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    61
definition
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    62
  -- "A program is welltyped if it is wellformed and all methods are welltyped"
37406
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    63
  wt_jvm_prog :: "[jvm_prog,prog_type] \<Rightarrow> bool" where
982f3e02f3c4 modernized specifications
haftmann
parents: 33954
diff changeset
    64
  "wt_jvm_prog G phi \<longleftrightarrow>
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    65
  wf_prog (\<lambda>G C (sig,rT,(maxs,maxl,b,et)).
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    66
           wt_method G C (snd sig) rT maxs maxl b et (phi C sig)) G"
9559
kleing
parents: 9549
diff changeset
    67
kleing
parents: 9549
diff changeset
    68
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    69
lemma check_boundedD:
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    70
  "\<lbrakk> check_bounded ins et; pc < length ins; 
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    71
    (pc',s') \<in> set (eff (ins!pc) G pc et s)  \<rbrakk> \<Longrightarrow> 
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    72
  pc' < length ins"
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    73
  apply (unfold eff_def)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    74
  apply simp
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    75
  apply (unfold check_bounded_def)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    76
  apply clarify
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    77
  apply (erule disjE)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    78
   apply blast
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    79
  apply (erule allE, erule impE, assumption)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    80
  apply (unfold xcpt_eff_def)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    81
  apply clarsimp    
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    82
  apply (drule xcpt_names_in_et)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    83
  apply clarify
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    84
  apply (drule bspec, assumption)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    85
  apply simp
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    86
  done
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    87
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
    88
lemma wt_jvm_progD:
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
    89
  "wt_jvm_prog G phi \<Longrightarrow> (\<exists>wt. wf_prog wt G)"
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    90
  by (unfold wt_jvm_prog_def, blast)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
    91
10629
d790faef9c07 removed two intermediate comments
oheimb
parents: 10612
diff changeset
    92
lemma wt_jvm_prog_impl_wt_instr:
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
    93
  "\<lbrakk> wt_jvm_prog G phi; is_class G C;
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
    94
      method (G,C) sig = Some (C,rT,maxs,maxl,ins,et); pc < length ins \<rbrakk> 
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    95
  \<Longrightarrow> wt_instr (ins!pc) G rT (phi C sig) maxs (length ins) et pc"
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    96
  by (unfold wt_jvm_prog_def, drule method_wf_mdecl, 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    97
      simp, simp, simp add: wf_mdecl_def wt_method_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
    98
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
    99
text {*
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   100
  We could leave out the check @{term "pc' < max_pc"} in the 
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   101
  definition of @{term wt_instr} in the context of @{term wt_method}.
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   102
*}
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   103
lemma wt_instr_def2:
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   104
  "\<lbrakk> wt_jvm_prog G Phi; is_class G C;
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   105
      method (G,C) sig = Some (C,rT,maxs,maxl,ins,et); pc < length ins; 
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   106
      i = ins!pc; phi = Phi C sig; max_pc = length ins \<rbrakk> 
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   107
  \<Longrightarrow> wt_instr i G rT phi maxs max_pc et pc =
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   108
     (app i G maxs rT pc et (phi!pc) \<and>
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   109
     (\<forall>(pc',s') \<in> set (eff i G pc et (phi!pc)). G \<turnstile> s' <=' phi!pc'))"
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   110
apply (simp add: wt_instr_def)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   111
apply (unfold wt_jvm_prog_def)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   112
apply (drule method_wf_mdecl)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   113
apply (simp, simp, simp add: wf_mdecl_def wt_method_def)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   114
apply (auto dest: check_boundedD)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   115
done
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13006
diff changeset
   116
10629
d790faef9c07 removed two intermediate comments
oheimb
parents: 10612
diff changeset
   117
lemma wt_jvm_prog_impl_wt_start:
13006
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   118
  "\<lbrakk> wt_jvm_prog G phi; is_class G C;
51c5f3f11d16 symbolized
kleing
parents: 12911
diff changeset
   119
      method (G,C) sig = Some (C,rT,maxs,maxl,ins,et) \<rbrakk> \<Longrightarrow> 
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   120
  0 < (length ins) \<and> wt_start G C (snd sig) maxl (phi C sig)"
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   121
  by (unfold wt_jvm_prog_def, drule method_wf_mdecl, 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   122
      simp, simp, simp add: wf_mdecl_def wt_method_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   123
8011
d14c4e9e9c8e *** empty log message ***
nipkow
parents:
diff changeset
   124
end