src/HOL/MicroJava/BV/JVM.thy
author kleing
Sun, 16 Dec 2001 00:17:44 +0100
changeset 12516 d09d0f160888
parent 12230 b06cc3834ee5
child 12911 704713ca07ea
permissions -rw-r--r--
exceptions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
     1
(*  Title:      HOL/MicroJava/BV/JVM.thy
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
     2
    ID:         $Id$
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
     3
    Author:     Tobias Nipkow, Gerwin Klein
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
     4
    Copyright   2000 TUM
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
     5
*)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
     6
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
     7
header "Kildall for the JVM"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
     8
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
     9
theory JVM = Kildall_Lift + JVMType + Opt + Product + Typing_Framework_err +
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    10
             EffectMono + BVSpec:
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    11
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    12
constdefs
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    13
  exec :: "jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> exception_table \<Rightarrow> instr list \<Rightarrow> state step_type"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    14
  "exec G maxs rT et bs == 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    15
  err_step (\<lambda>pc. app (bs!pc) G maxs rT pc et) (\<lambda>pc. eff (bs!pc) G pc et)"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    16
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    17
  kiljvm :: "jvm_prog => nat => nat => ty => exception_table => 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    18
             instr list => state list => state list"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    19
  "kiljvm G maxs maxr rT et bs ==
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    20
  kildall (JVMType.le G maxs maxr) (JVMType.sup G maxs maxr) (exec G maxs rT et bs)"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    21
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    22
  wt_kil :: "jvm_prog \<Rightarrow> cname \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    23
             exception_table \<Rightarrow> instr list \<Rightarrow> bool"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    24
  "wt_kil G C pTs rT mxs mxl et ins ==
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    25
   bounded (exec G mxs rT et ins) (size ins) \<and> 0 < size ins \<and> 
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    26
   (let first  = Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err));
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11299
diff changeset
    27
        start  = OK first#(replicate (size ins - 1) (OK None));
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    28
        result = kiljvm G mxs (1+size pTs+mxl) rT et ins start
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    29
    in \<forall>n < size ins. result!n \<noteq> Err)"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    30
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    31
  wt_jvm_prog_kildall :: "jvm_prog => bool"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    32
  "wt_jvm_prog_kildall G ==
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    33
  wf_prog (\<lambda>G C (sig,rT,(maxs,maxl,b,et)). wt_kil G C (snd sig) rT maxs maxl et b) G"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    34
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    35
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    36
lemma special_ex_swap_lemma [iff]: 
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    37
  "(? X. (? n. X = A n & P n) & Q X) = (? n. Q(A n) & P n)"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    38
  by blast
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    39
10630
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    40
lemmas [iff del] = not_None_eq
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    41
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    42
lemma non_empty_succs: "succs i pc \<noteq> []"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    43
  by (cases i) auto
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    44
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    45
lemma non_empty:
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    46
  "non_empty (\<lambda>pc. eff (bs!pc) G pc et)" 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    47
  by (simp add: non_empty_def eff_def non_empty_succs)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    48
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    49
lemma listn_Cons_Suc [elim!]:
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    50
  "l#xs \<in> list n A \<Longrightarrow> (\<And>n'. n = Suc n' \<Longrightarrow> l \<in> A \<Longrightarrow> xs \<in> list n' A \<Longrightarrow> P) \<Longrightarrow> P"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    51
  by (cases n) auto
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    52
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    53
lemma listn_appendE [elim!]:
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    54
  "a@b \<in> list n A \<Longrightarrow> (\<And>n1 n2. n=n1+n2 \<Longrightarrow> a \<in> list n1 A \<Longrightarrow> b \<in> list n2 A \<Longrightarrow> P) \<Longrightarrow> P" 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    55
proof -
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    56
  have "\<And>n. a@b \<in> list n A \<Longrightarrow> \<exists>n1 n2. n=n1+n2 \<and> a \<in> list n1 A \<and> b \<in> list n2 A"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    57
    (is "\<And>n. ?list a n \<Longrightarrow> \<exists>n1 n2. ?P a n n1 n2")
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    58
  proof (induct a)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    59
    fix n assume "?list [] n"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    60
    hence "?P [] n 0 n" by simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    61
    thus "\<exists>n1 n2. ?P [] n n1 n2" by fast
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    62
  next
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    63
    fix n l ls
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    64
    assume "?list (l#ls) n"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    65
    then obtain n' where n: "n = Suc n'" "l \<in> A" and "ls@b \<in> list n' A" by fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    66
    assume "\<And>n. ls @ b \<in> list n A \<Longrightarrow> \<exists>n1 n2. n = n1 + n2 \<and> ls \<in> list n1 A \<and> b \<in> list n2 A"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    67
    hence "\<exists>n1 n2. n' = n1 + n2 \<and> ls \<in> list n1 A \<and> b \<in> list n2 A" .
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    68
    then obtain n1 n2 where "n' = n1 + n2" "ls \<in> list n1 A" "b \<in> list n2 A" by fast
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    69
    with n have "?P (l#ls) n (n1+1) n2" by simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    70
    thus "\<exists>n1 n2. ?P (l#ls) n n1 n2" by fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    71
  qed
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    72
  moreover
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    73
  assume "a@b \<in> list n A" "\<And>n1 n2. n=n1+n2 \<Longrightarrow> a \<in> list n1 A \<Longrightarrow> b \<in> list n2 A \<Longrightarrow> P"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    74
  ultimately
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    75
  show ?thesis by blast
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    76
qed
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    77
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    78
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    79
theorem exec_pres_type:
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    80
  "wf_prog wf_mb S ==> 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    81
  pres_type (exec S maxs rT et bs) (size bs) (states S maxs maxr)"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    82
  apply (unfold exec_def JVM_states_unfold)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    83
  apply (rule pres_type_lift)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    84
  apply clarify
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    85
  apply (case_tac s)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    86
   apply simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    87
   apply (drule effNone)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    88
   apply simp  
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    89
  apply (simp add: eff_def xcpt_eff_def norm_eff_def)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    90
  apply (case_tac "bs!p")
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    91
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    92
  apply (clarsimp simp add: not_Err_eq)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    93
  apply (drule listE_nth_in, assumption)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    94
  apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    95
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    96
  apply (fastsimp simp add: not_None_eq)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    97
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    98
  apply (fastsimp simp add: not_None_eq typeof_empty_is_type)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    99
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   100
  apply clarsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   101
  apply (erule disjE)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   102
   apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   103
  apply clarsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   104
  apply (rule_tac x="1" in exI)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   105
  apply fastsimp
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   106
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   107
  apply clarsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   108
  apply (erule disjE)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   109
   apply (fastsimp dest: field_fields fields_is_type)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   110
  apply simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   111
  apply (rule_tac x=1 in exI)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   112
  apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   113
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   114
  apply clarsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   115
  apply (erule disjE)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   116
   apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   117
  apply simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   118
  apply (rule_tac x=1 in exI)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   119
  apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   120
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   121
  apply clarsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   122
  apply (erule disjE)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   123
   apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   124
  apply clarsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   125
  apply (rule_tac x=1 in exI)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   126
  apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   127
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   128
  defer 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   129
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   130
  apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   131
  apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   132
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   133
  apply clarsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   134
  apply (rule_tac x="n'+2" in exI)  
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   135
  apply simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   136
  apply (drule listE_length)+
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   137
  apply fastsimp
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   138
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   139
  apply clarsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   140
  apply (rule_tac x="Suc (Suc (Suc (length ST)))" in exI)  
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   141
  apply simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   142
  apply (drule listE_length)+
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   143
  apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   144
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   145
  apply clarsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   146
  apply (rule_tac x="Suc (Suc (Suc (Suc (length ST))))" in exI)  
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   147
  apply simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   148
  apply (drule listE_length)+
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   149
  apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   150
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   151
  apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   152
  apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   153
  apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   154
  apply fastsimp
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   155
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   156
  apply clarsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   157
  apply (erule disjE)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   158
   apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   159
  apply clarsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   160
  apply (rule_tac x=1 in exI)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   161
  apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   162
  
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   163
  apply (erule disjE)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   164
   apply (clarsimp simp add: Un_subset_iff)  
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   165
   apply (drule method_wf_mdecl, assumption+)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   166
   apply (clarsimp simp add: wf_mdecl_def wf_mhead_def)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   167
   apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   168
  apply clarsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   169
  apply (rule_tac x=1 in exI)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   170
  apply fastsimp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   171
  done
10630
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
   172
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
   173
lemmas [iff] = not_None_eq
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
   174
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   175
lemma map_fst_eq:
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   176
  "map fst (map (\<lambda>z. (f z, x z)) a) = map fst (map (\<lambda>z. (f z, y z)) a)"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   177
  by (induct a) auto
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   178
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   179
lemma succs_stable_eff:
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   180
  "succs_stable (sup_state_opt G) (\<lambda>pc. eff (bs!pc) G pc et)"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   181
  apply (unfold succs_stable_def eff_def xcpt_eff_def)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   182
  apply (simp add: map_fst_eq)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   183
  done
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   184
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   185
lemma sup_state_opt_unfold:
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   186
  "sup_state_opt G \<equiv> Opt.le (Product.le (Listn.le (subtype G)) (Listn.le (Err.le (subtype G))))"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   187
  by (simp add: sup_state_opt_def sup_state_def sup_loc_def sup_ty_opt_def)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   188
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   189
constdefs
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   190
  opt_states :: "'c prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> (ty list \<times> ty err list) option set"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   191
  "opt_states G maxs maxr \<equiv> opt (\<Union>{list n (types G) |n. n \<le> maxs} \<times> list maxr (err (types G)))"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   192
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   193
lemma app_mono:
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   194
  "app_mono (sup_state_opt G) (\<lambda>pc. app (bs!pc) G maxs rT pc et) (length bs) (opt_states G maxs maxr)"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   195
  by (unfold app_mono_def lesub_def) (blast intro: EffectMono.app_mono)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   196
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   197
lemma le_list_appendI:
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   198
  "\<And>b c d. a <=[r] b \<Longrightarrow> c <=[r] d \<Longrightarrow> a@c <=[r] b@d"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   199
apply (induct a)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   200
 apply simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   201
apply (case_tac b)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   202
apply auto
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   203
done
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   204
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   205
lemma le_listI:
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   206
  "length a = length b \<Longrightarrow> (\<And>n. n < length a \<Longrightarrow> a!n <=_r b!n) \<Longrightarrow> a <=[r] b"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   207
  apply (unfold lesub_def Listn.le_def)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   208
  apply (simp add: list_all2_conv_all_nth)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   209
  done
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   210
  
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   211
lemma eff_mono:
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   212
  "\<lbrakk>p < length bs; s <=_(sup_state_opt G) t; app (bs!p) G maxs rT pc et t\<rbrakk>
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   213
  \<Longrightarrow> eff (bs!p) G p et s <=|sup_state_opt G| eff (bs!p) G p et t"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   214
  apply (unfold eff_def)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   215
  apply (rule le_list_appendI)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   216
   apply (simp add: norm_eff_def)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   217
   apply (rule le_listI)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   218
    apply simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   219
   apply simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   220
   apply (simp add: lesub_def)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   221
   apply (case_tac s)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   222
    apply simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   223
   apply (simp del: split_paired_All split_paired_Ex)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   224
   apply (elim exE conjE)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   225
   apply simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   226
   apply (drule eff'_mono, assumption)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   227
   apply assumption
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   228
  apply (simp add: xcpt_eff_def)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   229
  apply (rule le_listI)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   230
    apply simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   231
  apply simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   232
  apply (simp add: lesub_def)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   233
  apply (case_tac s)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   234
   apply simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   235
  apply simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   236
  apply (case_tac t)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   237
   apply simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   238
  apply (clarsimp simp add: sup_state_conv)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   239
  done
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   240
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   241
lemma order_sup_state_opt:
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   242
  "wf_prog wf_mb G \<Longrightarrow> order (sup_state_opt G)"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   243
  by (unfold sup_state_opt_unfold) (blast dest: acyclic_subcls1 order_widen)
10630
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
   244
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   245
theorem exec_mono:
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   246
  "wf_prog wf_mb G ==>
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   247
  mono (JVMType.le G maxs maxr) (exec G maxs rT et bs) (size bs) (states G maxs maxr)"  
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   248
  apply (unfold exec_def JVM_le_unfold JVM_states_unfold)  
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   249
  apply (rule mono_lift)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   250
     apply (fold sup_state_opt_unfold opt_states_def)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   251
     apply (erule order_sup_state_opt)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   252
    apply (rule succs_stable_eff)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   253
   apply (rule app_mono)
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   254
  apply clarify
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   255
  apply (rule eff_mono)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   256
  apply assumption+
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   257
  done
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   258
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   259
theorem semilat_JVM_slI:
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   260
  "wf_prog wf_mb G ==> semilat (JVMType.sl G maxs maxr)"
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10657
diff changeset
   261
  apply (unfold JVMType.sl_def stk_esl_def reg_sl_def)
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   262
  apply (rule semilat_opt)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   263
  apply (rule err_semilat_Product_esl)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   264
  apply (rule err_semilat_upto_esl)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   265
  apply (rule err_semilat_JType_esl, assumption+)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   266
  apply (rule err_semilat_eslI)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   267
  apply (rule semilat_Listn_sl)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   268
  apply (rule err_semilat_JType_esl, assumption+)  
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   269
  done
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   270
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   271
lemma sl_triple_conv:
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10657
diff changeset
   272
  "JVMType.sl G maxs maxr == 
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10657
diff changeset
   273
  (states G maxs maxr, JVMType.le G maxs maxr, JVMType.sup G maxs maxr)"
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10657
diff changeset
   274
  by (simp (no_asm) add: states_def JVMType.le_def JVMType.sup_def)
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   275
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   276
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   277
theorem is_bcv_kiljvm:
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   278
  "[| wf_prog wf_mb G; bounded (exec G maxs rT et bs) (size bs) |] ==> 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   279
      is_bcv (JVMType.le G maxs maxr) Err (exec G maxs rT et bs)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   280
             (size bs) (states G maxs maxr) (kiljvm G maxs maxr rT et bs)"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   281
  apply (unfold kiljvm_def sl_triple_conv)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   282
  apply (rule is_bcv_kildall)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   283
       apply (simp (no_asm) add: sl_triple_conv [symmetric]) 
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   284
       apply (force intro!: semilat_JVM_slI dest: wf_acyclic simp add: symmetric sl_triple_conv)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   285
      apply (simp (no_asm) add: JVM_le_unfold)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   286
      apply (blast intro!: order_widen wf_converse_subcls1_impl_acc_subtype
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   287
                   dest: wf_subcls1 wf_acyclic) 
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   288
     apply (simp add: JVM_le_unfold)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   289
    apply (erule exec_pres_type)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   290
    apply assumption
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   291
  apply (erule exec_mono)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   292
  done
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   293
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   294
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   295
theorem wt_kil_correct:
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   296
  "[| wt_kil G C pTs rT maxs mxl et bs; wf_prog wf_mb G; 
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   297
      is_class G C; \<forall>x \<in> set pTs. is_type G x |]
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   298
  ==> \<exists>phi. wt_method G C pTs rT maxs mxl bs et phi"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   299
proof -
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   300
  let ?start = "OK (Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err)))
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11299
diff changeset
   301
                #(replicate (size bs - 1) (OK None))"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   302
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   303
  assume wf:      "wf_prog wf_mb G"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   304
  assume isclass: "is_class G C"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   305
  assume istype:  "\<forall>x \<in> set pTs. is_type G x"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   306
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   307
  assume "wt_kil G C pTs rT maxs mxl et bs"
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   308
  then obtain maxr r where    
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   309
    bounded: "bounded (exec G maxs rT et bs) (size bs)" and
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   310
    result:  "r = kiljvm G maxs maxr rT et bs ?start" and
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   311
    success: "\<forall>n < size bs. r!n \<noteq> Err" and
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   312
    instrs:  "0 < size bs" and
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   313
    maxr:    "maxr = Suc (length pTs + mxl)" 
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   314
    by (unfold wt_kil_def) simp
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   315
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   316
  from wf bounded
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   317
  have bcv:
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   318
    "is_bcv (JVMType.le G maxs maxr) Err (exec G maxs rT et bs) 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   319
            (size bs) (states G maxs maxr) (kiljvm G maxs maxr rT et bs)"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   320
    by (rule is_bcv_kiljvm)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   321
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   322
  { fix l x have "set (replicate l x) \<subseteq> {x}" by (cases "0 < l") simp+
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   323
  } note subset_replicate = this
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   324
  from istype have "set pTs \<subseteq> types G" by auto
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   325
  hence "OK ` set pTs \<subseteq> err (types G)" by auto
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   326
  with instrs maxr isclass 
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   327
  have "?start \<in> list (length bs) (states G maxs maxr)"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   328
    apply (unfold list_def JVM_states_unfold)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   329
    apply simp
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   330
    apply (rule conjI)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   331
     apply (simp add: Un_subset_iff)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   332
     apply (rule_tac B = "{Err}" in subset_trans)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   333
      apply (simp add: subset_replicate)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   334
     apply simp
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   335
    apply (rule_tac B = "{OK None}" in subset_trans)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   336
     apply (simp add: subset_replicate)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   337
    apply simp
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   338
    done
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   339
  with bcv success result have 
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   340
    "\<exists>ts\<in>list (length bs) (states G maxs maxr).
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10657
diff changeset
   341
         ?start <=[JVMType.le G maxs maxr] ts \<and>
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   342
         wt_step (JVMType.le G maxs maxr) Err (exec G maxs rT et bs) ts"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   343
    by (unfold is_bcv_def) auto
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   344
  then obtain phi' where
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   345
    l: "phi' \<in> list (length bs) (states G maxs maxr)" and
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10657
diff changeset
   346
    s: "?start <=[JVMType.le G maxs maxr] phi'" and
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   347
    w: "wt_step (JVMType.le G maxs maxr) Err (exec G maxs rT et bs) phi'"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   348
    by blast
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   349
  hence dynamic:
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   350
    "dynamic_wt (sup_state_opt G) (exec G maxs rT et bs) phi'"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   351
    by (simp add: dynamic_wt_def exec_def JVM_le_Err_conv)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   352
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   353
  from s have le: "JVMType.le G maxs maxr (?start ! 0) (phi'!0)"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   354
    by (drule_tac p=0 in le_listD) (simp add: lesub_def)+
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   355
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   356
  from l have l: "size phi' = size bs" by simp  
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   357
  with instrs w have "phi' ! 0 \<noteq> Err" by (unfold wt_step_def) simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   358
  with instrs l have phi0: "OK (map ok_val phi' ! 0) = phi' ! 0"
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10657
diff changeset
   359
    by (clarsimp simp add: not_Err_eq)
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   360
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   361
  from l bounded 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   362
  have bounded': "bounded (\<lambda>pc. eff (bs!pc) G pc et) (length phi')"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   363
    by (simp add: exec_def bounded_lift)  
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   364
  with dynamic
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   365
  have "static_wt (sup_state_opt G) (\<lambda>pc. app (bs!pc) G maxs rT pc et) 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   366
                  (\<lambda>pc. eff (bs!pc) G pc et) (map ok_val phi')"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   367
    by (auto intro: dynamic_imp_static simp add: exec_def non_empty)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   368
  with instrs l le bounded'
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   369
  have "wt_method G C pTs rT maxs mxl bs et (map ok_val phi')"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   370
    apply (unfold wt_method_def static_wt_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   371
    apply simp
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   372
    apply (rule conjI)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   373
     apply (unfold wt_start_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   374
     apply (rule JVM_le_convert [THEN iffD1])
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   375
     apply (simp (no_asm) add: phi0)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   376
    apply clarify
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   377
    apply (erule allE, erule impE, assumption)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   378
    apply (elim conjE)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   379
    apply (clarsimp simp add: lesub_def wt_instr_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   380
    apply (unfold bounded_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   381
    apply blast    
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   382
    done
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   383
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   384
  thus ?thesis by blast
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   385
qed
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   386
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   387
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   388
theorem wt_kil_complete:
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   389
  "[| wt_method G C pTs rT maxs mxl bs et phi; wf_prog wf_mb G; 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   390
      length phi = length bs; is_class G C; \<forall>x \<in> set pTs. is_type G x;
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   391
      map OK phi \<in> list (length bs) (states G maxs (1+size pTs+mxl)) |]
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   392
  ==> wt_kil G C pTs rT maxs mxl et bs"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   393
proof -
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   394
  assume wf: "wf_prog wf_mb G"  
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   395
  assume isclass: "is_class G C"
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   396
  assume istype: "\<forall>x \<in> set pTs. is_type G x"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   397
  assume length: "length phi = length bs"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   398
  assume istype_phi: "map OK phi \<in> list (length bs) (states G maxs (1+size pTs+mxl))"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   399
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   400
  assume "wt_method G C pTs rT maxs mxl bs et phi"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   401
  then obtain
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   402
    instrs:   "0 < length bs" and
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   403
    wt_start: "wt_start G C pTs mxl phi" and
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   404
    wt_ins:   "\<forall>pc. pc < length bs \<longrightarrow> 
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   405
                    wt_instr (bs ! pc) G rT phi maxs (length bs) et pc"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   406
    by (unfold wt_method_def) simp
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   407
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   408
  let ?eff  = "\<lambda>pc. eff (bs!pc) G pc et"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   409
  let ?app   = "\<lambda>pc. app (bs!pc) G maxs rT pc et"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   410
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   411
  have bounded_eff: "bounded ?eff (size bs)"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   412
  proof (unfold bounded_def, clarify)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   413
    fix pc pc' s s' assume "pc < length bs"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   414
    with wt_ins have "wt_instr (bs!pc) G rT phi maxs (length bs) et pc" by fast
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   415
    then obtain "\<forall>(pc',s') \<in> set (?eff pc (phi!pc)). pc' < length bs"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   416
      by (unfold wt_instr_def) fast
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   417
    hence "\<forall>pc' \<in> set (map fst (?eff pc (phi!pc))). pc' < length bs" by auto
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   418
    also 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   419
    note succs_stable_eff 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   420
    hence "map fst (?eff pc (phi!pc)) = map fst (?eff pc s)" 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   421
      by (rule succs_stableD)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   422
    finally have "\<forall>(pc',s') \<in> set (?eff pc s). pc' < length bs" by auto
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   423
    moreover assume "(pc',s') \<in> set (?eff pc s)"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   424
    ultimately show "pc' < length bs" by blast
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   425
  qed
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   426
  hence bounded_exec: "bounded (exec G maxs rT et bs) (size bs)" 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   427
    by (simp add: exec_def bounded_lift)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   428
 
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   429
  from wt_ins
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   430
  have "static_wt (sup_state_opt G) ?app ?eff phi"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   431
    apply (unfold static_wt_def wt_instr_def lesub_def)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   432
    apply (simp (no_asm) only: length)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   433
    apply blast
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   434
    done
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   435
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   436
  with bounded_eff
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   437
  have "dynamic_wt (sup_state_opt G) (err_step ?app ?eff) (map OK phi)"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   438
    by - (erule static_imp_dynamic, simp add: length)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   439
  hence dynamic:
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   440
    "dynamic_wt (sup_state_opt G) (exec G maxs rT et bs) (map OK phi)"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   441
    by (unfold exec_def)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   442
 
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   443
  let ?maxr = "1+size pTs+mxl"
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   444
  from wf bounded_exec
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   445
  have is_bcv: 
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   446
    "is_bcv (JVMType.le G maxs ?maxr) Err (exec G maxs rT et bs) 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   447
            (size bs) (states G maxs ?maxr) (kiljvm G maxs ?maxr rT et bs)"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   448
    by (rule is_bcv_kiljvm)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   449
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   450
  let ?start = "OK (Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err)))
11701
3d51fbf81c17 sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents: 11299
diff changeset
   451
                #(replicate (size bs - 1) (OK None))"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   452
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   453
  { fix l x have "set (replicate l x) \<subseteq> {x}" by (cases "0 < l") simp+
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   454
  } note subset_replicate = this
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   455
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   456
  from istype have "set pTs \<subseteq> types G" by auto
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   457
  hence "OK ` set pTs \<subseteq> err (types G)" by auto
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   458
  with instrs isclass have start:
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   459
    "?start \<in> list (length bs) (states G maxs ?maxr)"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   460
    apply (unfold list_def JVM_states_unfold)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   461
    apply simp
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   462
    apply (rule conjI)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   463
     apply (simp add: Un_subset_iff)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   464
     apply (rule_tac B = "{Err}" in subset_trans)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   465
      apply (simp add: subset_replicate)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   466
     apply simp
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   467
    apply (rule_tac B = "{OK None}" in subset_trans)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   468
     apply (simp add: subset_replicate)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   469
    apply simp
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   470
    done
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   471
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   472
  let ?phi = "map OK phi"  
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   473
  have less_phi: "?start <=[JVMType.le G maxs ?maxr] ?phi"
10657
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   474
  proof -
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   475
    from length instrs
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   476
    have "length ?start = length (map OK phi)" by simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   477
    moreover
10657
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   478
    { fix n
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   479
      from wt_start
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   480
      have "G \<turnstile> ok_val (?start!0) <=' phi!0"
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   481
        by (simp add: wt_start_def)
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   482
      moreover
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   483
      from instrs length
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   484
      have "0 < length phi" by simp
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   485
      ultimately
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   486
      have "JVMType.le G maxs ?maxr (?start!0) (?phi!0)"
10657
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   487
        by (simp add: JVM_le_Err_conv Err.le_def lesub_def)
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   488
      moreover
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   489
      { fix n'
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   490
        have "JVMType.le G maxs ?maxr (OK None) (?phi!n)"
10657
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   491
          by (auto simp add: JVM_le_Err_conv Err.le_def lesub_def 
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   492
            split: err.splits)        
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   493
        hence "[| n = Suc n'; n < length ?start |] 
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   494
          ==> JVMType.le G maxs ?maxr (?start!n) (?phi!n)"
10657
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   495
          by simp
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   496
      }
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   497
      ultimately
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   498
      have "n < length ?start ==> (?start!n) <=_(JVMType.le G maxs ?maxr) (?phi!n)"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   499
        by (unfold lesub_def) (cases n, blast+)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   500
    } 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   501
    ultimately show ?thesis by (rule le_listI)
10657
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   502
  qed         
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   503
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   504
  from dynamic
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   505
  have "wt_step (JVMType.le G maxs ?maxr) Err (exec G maxs rT et bs) ?phi"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   506
    by (simp add: dynamic_wt_def JVM_le_Err_conv)  
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   507
  with start istype_phi less_phi is_bcv
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   508
  have "\<forall>p. p < length bs \<longrightarrow> kiljvm G maxs ?maxr rT et bs ?start ! p \<noteq> Err"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   509
    by (unfold is_bcv_def) auto
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   510
  with bounded_exec instrs
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   511
  show "wt_kil G C pTs rT maxs mxl et bs" by (unfold wt_kil_def) simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   512
qed
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   513
text {*
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   514
  The above theorem @{text wt_kil_complete} is all nice'n shiny except
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   515
  for one assumption: @{term "map OK phi \<in> list (length bs) (states G maxs (1+size pTs+mxl))"}  
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   516
  It does not hold for all @{text phi} that satisfy @{text wt_method}.
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   517
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   518
  The assumption states mainly that all entries in @{text phi} are legal
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   519
  types in the program context, that the stack size is bounded by @{text maxs},
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   520
  and that the register sizes are exactly @{term "1+size pTs+mxl"}. 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   521
  The BV specification, i.e.~@{text wt_method}, only gives us this
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   522
  property for \emph{reachable} code. For unreachable code, 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   523
  e.g.~unused registers may contain arbitrary garbage. Even the stack
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   524
  and register sizes can be different from the rest of the program (as 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   525
  long as they are consistent inside each chunk of unreachable code).
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   526
  
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   527
  All is not lost, though: for each @{text phi} that satisfies 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   528
  @{text wt_method} there is a @{text phi'} that also satisfies 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   529
  @{text wt_method} and that additionally satisfies our assumption.
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   530
  The construction is quite easy: the entries for reachable code
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   531
  are the same in @{text phi} and @{text phi'}, the entries for
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   532
  unreachable code are all @{text None} in @{text phi'} (as it would
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   533
  be produced by Kildall's algorithm). 
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   534
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   535
  Although this is proved easily by comment, it requires some more
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   536
  overhead (i.e.~talking about reachable instructions) if you try
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   537
  it the hard way. Thus it is missing here for the time being.
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   538
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   539
  The other direction (@{text wt_kil_correct}) can be lifted to
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   540
  programs without problems:
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   541
*}
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   542
lemma is_type_pTs:
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   543
  "[| wf_prog wf_mb G; (C,S,fs,mdecls) \<in> set G; (sig,rT,code) \<in> set mdecls; 
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   544
      t \<in> set (snd sig) |]
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   545
  ==> is_type G t"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   546
proof -
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   547
  assume "wf_prog wf_mb G" 
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   548
         "(C,S,fs,mdecls) \<in> set G"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   549
         "(sig,rT,code) \<in> set mdecls"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   550
  hence "wf_mdecl wf_mb G C (sig,rT,code)"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   551
    by (unfold wf_prog_def wf_cdecl_def) auto
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   552
  hence "\<forall>t \<in> set (snd sig). is_type G t" 
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   553
    by (unfold wf_mdecl_def wf_mhead_def) auto
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   554
  moreover
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   555
  assume "t \<in> set (snd sig)"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   556
  ultimately
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   557
  show ?thesis by blast
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   558
qed
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   559
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   560
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   561
theorem jvm_kildall_correct:
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   562
  "wt_jvm_prog_kildall G ==> \<exists>Phi. wt_jvm_prog G Phi"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   563
proof -  
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   564
  assume wtk: "wt_jvm_prog_kildall G"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   565
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   566
  then obtain wf_mb where
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   567
    wf: "wf_prog wf_mb G"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   568
    by (auto simp add: wt_jvm_prog_kildall_def)
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   569
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   570
  let ?Phi = "\<lambda>C sig. let (C,rT,(maxs,maxl,ins,et)) = the (method (G,C) sig) in 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   571
              SOME phi. wt_method G C (snd sig) rT maxs maxl ins et phi"
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   572
   
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   573
  { fix C S fs mdecls sig rT code
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   574
    assume "(C,S,fs,mdecls) \<in> set G" "(sig,rT,code) \<in> set mdecls"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   575
    with wf
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   576
    have "method (G,C) sig = Some (C,rT,code) \<and> is_class G C \<and> (\<forall>t \<in> set (snd sig). is_type G t)"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   577
      by (simp add: methd is_type_pTs)
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   578
  } note this [simp]
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   579
 
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   580
  from wtk
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   581
  have "wt_jvm_prog G ?Phi"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   582
    apply (unfold wt_jvm_prog_def wt_jvm_prog_kildall_def wf_prog_def wf_cdecl_def)
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   583
    apply clarsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   584
    apply (drule bspec, assumption)
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   585
    apply (unfold wf_mdecl_def)
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   586
    apply clarsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   587
    apply (drule bspec, assumption)
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   588
    apply clarsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   589
    apply (drule wt_kil_correct [OF _ wf])
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   590
    apply (auto intro: someI)
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   591
    done
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   592
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   593
  thus ?thesis by blast
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   594
qed
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   595
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   596
end