src/HOL/MicroJava/BV/JVM.thy
author oheimb
Tue, 12 Jun 2001 14:11:00 +0200
changeset 11372 648795477bb5
parent 11299 1d3d110c456e
child 11701 3d51fbf81c17
permissions -rw-r--r--
corrected xsymbol/HTML syntax
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
     1
(*  Title:      HOL/BCV/JVM.thy
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
     2
    ID:         $Id$
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
     3
    Author:     Tobias Nipkow
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
*)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
     7
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
     8
header "Kildall for the JVM"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
     9
11229
f417841385b7 Got rid of is_dfa
nipkow
parents: 11186
diff changeset
    10
theory JVM = Kildall + JVMType + Opt + Product + Typing_Framework_err +
f417841385b7 Got rid of is_dfa
nipkow
parents: 11186
diff changeset
    11
             StepMono + BVSpec:
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    12
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    13
constdefs
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    14
  exec :: "jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> instr list \<Rightarrow> nat \<Rightarrow> state \<Rightarrow> state"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    15
  "exec G maxs rT bs == err_step (\<lambda>pc. app (bs!pc) G maxs rT) (\<lambda>pc. step (bs!pc) G)"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    16
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    17
  kiljvm :: "jvm_prog => nat => nat => ty => instr list => state list => state list"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    18
  "kiljvm G maxs maxr rT bs ==
11298
acd83fa66e92 simplified defs and proofs a little
nipkow
parents: 11229
diff changeset
    19
  kildall (JVMType.le G maxs maxr) (JVMType.sup G maxs maxr) (exec G maxs rT bs) (\<lambda>pc. succs (bs!pc) pc)"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    20
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    21
  wt_kil :: "jvm_prog \<Rightarrow> cname \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> instr list \<Rightarrow> bool"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    22
  "wt_kil G C pTs rT mxs mxl ins ==
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    23
   bounded (\<lambda>n. succs (ins!n) n) (size ins) \<and> 0 < size ins \<and> 
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    24
   (let first  = Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err));
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    25
        start  = OK first#(replicate (size ins-1) (OK None));
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    26
        result = kiljvm G mxs (1+size pTs+mxl) rT ins start
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    27
    in \<forall>n < size ins. result!n \<noteq> Err)"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    28
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    29
  wt_jvm_prog_kildall :: "jvm_prog => bool"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    30
  "wt_jvm_prog_kildall G ==
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    31
  wf_prog (\<lambda>G C (sig,rT,(maxs,maxl,b)). wt_kil G C (snd sig) rT maxs maxl b) G"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    32
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    33
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    34
lemma special_ex_swap_lemma [iff]: 
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    35
  "(? X. (? n. X = A n & P n) & Q X) = (? n. Q(A n) & P n)"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    36
  by blast
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    37
10630
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    38
lemmas [iff del] = not_None_eq
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    39
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    40
theorem exec_pres_type:
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    41
  "[| wf_prog wf_mb S |] ==> 
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    42
      pres_type (exec S maxs rT bs) (size bs) (states S maxs maxr)"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    43
 apply (unfold pres_type_def list_def step_def JVM_states_unfold)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    44
 apply (clarify elim!: option_map_in_optionI lift_in_errI)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    45
 apply (simp add: exec_def err_step_def lift_def split: err.split)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    46
 apply (simp add: step_def option_map_def split: option.splits)  
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    47
 apply clarify
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    48
 apply (case_tac "bs!p")
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    49
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    50
 apply fastsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    51
 apply fastsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    52
 apply fastsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    53
 apply fastsimp
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    54
 apply clarsimp
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    55
 defer
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    56
 apply fastsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    57
 apply fastsimp
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    58
 apply clarsimp
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    59
 defer
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    60
 apply fastsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    61
 apply fastsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    62
 apply fastsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    63
 apply fastsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    64
 apply fastsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    65
 apply fastsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    66
 apply fastsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    67
 apply fastsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    68
 apply fastsimp
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    69
 defer
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    70
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    71
 (* Invoke *)
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    72
 apply (simp add: Un_subset_iff)
10630
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    73
 apply (drule method_wf_mdecl, assumption+)
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    74
 apply (simp add: wf_mdecl_def wf_mhead_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    75
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    76
 (* Getfield *)
10630
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    77
 apply (rule_tac fn = "(vname,cname)" in fields_is_type)
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    78
 defer
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    79
 apply assumption+
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    80
 apply (simp add: field_def)
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    81
 apply (drule map_of_SomeD)
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    82
 apply (rule map_of_SomeI)
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    83
 apply (auto simp add: unique_fields)
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    84
 done
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    85
10630
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    86
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    87
lemmas [iff] = not_None_eq
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    88
f89c3fc4fde1 updated for is_class changes
kleing
parents: 10612
diff changeset
    89
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    90
theorem exec_mono:
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    91
  "wf_prog wf_mb G ==>
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10657
diff changeset
    92
  mono (JVMType.le G maxs maxr) (exec G maxs rT bs) (size bs) (states G maxs maxr)"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    93
  apply (unfold mono_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    94
  apply clarify
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    95
  apply (unfold lesub_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    96
  apply (case_tac t)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    97
   apply (simp add: JVM_le_unfold Err.le_def exec_def err_step_def lift_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    98
  apply (case_tac s)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    99
   apply (simp add: JVM_le_unfold Err.le_def exec_def err_step_def lift_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   100
  apply (simp add: JVM_le_convert exec_def err_step_def lift_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   101
  apply (simp add: JVM_le_unfold Err.le_def exec_def err_step_def lift_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   102
  apply (rule conjI)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   103
   apply clarify
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   104
   apply (rule step_mono, assumption+)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   105
  apply (rule impI)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   106
  apply (erule contrapos_nn)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   107
  apply (rule app_mono, assumption+)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   108
  done
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   109
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   110
theorem semilat_JVM_slI:
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10657
diff changeset
   111
  "[| wf_prog wf_mb G |] ==> semilat (JVMType.sl G maxs maxr)"
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10657
diff changeset
   112
  apply (unfold JVMType.sl_def stk_esl_def reg_sl_def)
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   113
  apply (rule semilat_opt)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   114
  apply (rule err_semilat_Product_esl)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   115
  apply (rule err_semilat_upto_esl)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   116
  apply (rule err_semilat_JType_esl, assumption+)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   117
  apply (rule err_semilat_eslI)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   118
  apply (rule semilat_Listn_sl)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   119
  apply (rule err_semilat_JType_esl, assumption+)  
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   120
  done
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   121
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   122
lemma sl_triple_conv:
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10657
diff changeset
   123
  "JVMType.sl G maxs maxr == 
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10657
diff changeset
   124
  (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
   125
  by (simp (no_asm) add: states_def JVMType.le_def JVMType.sup_def)
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   126
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   127
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   128
theorem is_bcv_kiljvm:
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   129
  "[| wf_prog wf_mb G; bounded (\<lambda>pc. succs (bs!pc) pc) (size bs) |] ==> 
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10657
diff changeset
   130
      is_bcv (JVMType.le G maxs maxr) Err (exec G maxs rT bs) (\<lambda>pc. succs (bs!pc) pc)
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   131
             (size bs) (states G maxs maxr) (kiljvm G maxs maxr rT bs)"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   132
  apply (unfold kiljvm_def sl_triple_conv)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   133
  apply (rule is_bcv_kildall)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   134
       apply (simp (no_asm) add: sl_triple_conv [symmetric]) 
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   135
       apply (force intro!: semilat_JVM_slI dest: wf_acyclic simp add: symmetric sl_triple_conv)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   136
      apply (simp (no_asm) add: JVM_le_unfold)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   137
      apply (blast intro!: order_widen wf_converse_subcls1_impl_acc_subtype
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   138
                   dest: wf_subcls1 wf_acyclic) 
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   139
     apply (simp add: JVM_le_unfold)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   140
    apply (erule exec_pres_type)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   141
    apply assumption
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   142
  apply (erule exec_mono)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   143
  done
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   144
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   145
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   146
theorem wt_kil_correct:
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   147
  "[| wt_kil G C pTs rT maxs mxl bs; wf_prog wf_mb G; 
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   148
      is_class G C; \<forall>x \<in> set pTs. is_type G x |]
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   149
  ==> \<exists>phi. wt_method G C pTs rT maxs mxl bs phi"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   150
proof -
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   151
  let ?start = "OK (Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err)))
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   152
                #(replicate (size bs-1) (OK None))"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   153
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   154
  assume wf:      "wf_prog wf_mb G"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   155
  assume isclass: "is_class G C"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   156
  assume istype:  "\<forall>x \<in> set pTs. is_type G x"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   157
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   158
  assume "wt_kil G C pTs rT maxs mxl bs"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   159
  then obtain maxr r where    
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   160
    bounded: "bounded (\<lambda>pc. succs (bs!pc) pc) (size bs)" and
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   161
    result:  "r = kiljvm G maxs maxr rT bs ?start" and
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   162
    success: "\<forall>n < size bs. r!n \<noteq> Err" and
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   163
    instrs:  "0 < size bs" and
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   164
    maxr:    "maxr = Suc (length pTs + mxl)" 
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   165
    by (unfold wt_kil_def) simp
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   166
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   167
  { fix pc
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   168
    have "succs (bs!pc) pc \<noteq> []"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   169
      by (cases "bs!pc") auto
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   170
  }
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   171
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   172
  hence non_empty: "non_empty (\<lambda>pc. succs (bs!pc) pc)"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   173
    by (unfold non_empty_def) blast
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   174
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   175
  from wf bounded
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   176
  have bcv:
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10657
diff changeset
   177
    "is_bcv (JVMType.le G maxs maxr) Err (exec G maxs rT bs) (\<lambda>pc. succs (bs!pc) pc)
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   178
            (size bs) (states G maxs maxr) (kiljvm G maxs maxr rT bs)"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   179
    by (rule is_bcv_kiljvm)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   180
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   181
  { fix l x
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   182
    have "set (replicate l x) \<subseteq> {x}"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   183
      by (cases "0 < l") simp+
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   184
  } note subset_replicate = this
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   185
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   186
  from istype
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   187
  have "set pTs \<subseteq> types G"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   188
    by auto
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   189
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10812
diff changeset
   190
  hence "OK ` set pTs \<subseteq> err (types G)"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   191
    by auto
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   192
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   193
  with instrs maxr isclass 
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   194
  have "?start \<in> list (length bs) (states G maxs maxr)"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   195
    apply (unfold list_def JVM_states_unfold)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   196
    apply simp
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   197
    apply (rule conjI)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   198
     apply (simp add: Un_subset_iff)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   199
     apply (rule_tac B = "{Err}" in subset_trans)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   200
      apply (simp add: subset_replicate)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   201
     apply simp
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   202
    apply (rule_tac B = "{OK None}" in subset_trans)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   203
     apply (simp add: subset_replicate)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   204
    apply simp
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   205
    done
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   206
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   207
  with bcv success result 
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   208
  have 
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   209
    "\<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
   210
         ?start <=[JVMType.le G maxs maxr] ts \<and>
11299
1d3d110c456e welltyping -> wt_step
nipkow
parents: 11298
diff changeset
   211
         wt_step (JVMType.le G maxs maxr) Err (exec G maxs rT bs) (\<lambda>pc. succs (bs ! pc) pc) ts"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   212
    by (unfold is_bcv_def) auto
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   213
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   214
  then obtain phi' where
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   215
    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
   216
    s: "?start <=[JVMType.le G maxs maxr] phi'" and
11299
1d3d110c456e welltyping -> wt_step
nipkow
parents: 11298
diff changeset
   217
    w: "wt_step (JVMType.le G maxs maxr) Err (exec G maxs rT bs) (\<lambda>pc. succs (bs ! pc) pc) phi'"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   218
    by blast
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   219
   
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   220
  hence dynamic:
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   221
    "dynamic_wt (sup_state_opt G) (exec G maxs rT bs) (\<lambda>pc. succs (bs ! pc) pc) phi'"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   222
    by (simp add: dynamic_wt_def exec_def JVM_le_Err_conv)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   223
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   224
  from s
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10657
diff changeset
   225
  have le: "JVMType.le G maxs maxr (?start ! 0) (phi'!0)"    
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   226
    by (drule_tac p=0 in le_listD) (simp add: lesub_def)+
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   227
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   228
  from l
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   229
  have l: "size phi' = size bs"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   230
    by simp
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   231
  
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   232
  with instrs w
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   233
  have "phi' ! 0 \<noteq> Err"
11299
1d3d110c456e welltyping -> wt_step
nipkow
parents: 11298
diff changeset
   234
    by (unfold wt_step_def) simp
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   235
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   236
  with instrs l
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   237
  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
   238
    by (clarsimp simp add: not_Err_eq)
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   239
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   240
  from l bounded
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   241
  have "bounded (\<lambda>pc. succs (bs ! pc) pc) (length phi')"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   242
    by simp
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   243
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   244
  with dynamic non_empty
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   245
  have "static_wt (sup_state_opt G) (\<lambda>pc. app (bs!pc) G maxs rT) (\<lambda>pc. step (bs!pc) G) 
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   246
                                    (\<lambda>pc. succs (bs!pc) pc) (map ok_val phi')"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   247
    by (auto intro: dynamic_imp_static simp add: exec_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   248
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   249
  with instrs l le bounded
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   250
  have "wt_method G C pTs rT maxs mxl bs (map ok_val phi')"
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   251
    apply (unfold wt_method_def static_wt_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   252
    apply simp
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   253
    apply (rule conjI)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   254
     apply (unfold wt_start_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   255
     apply (rule JVM_le_convert [THEN iffD1])
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   256
     apply (simp (no_asm) add: phi0)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   257
    apply clarify
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   258
    apply (erule allE, erule impE, assumption)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   259
    apply (elim conjE)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   260
    apply (clarsimp simp add: lesub_def wt_instr_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   261
    apply (unfold bounded_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   262
    apply blast    
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   263
    done
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   264
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   265
  thus ?thesis by blast
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   266
qed
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   267
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   268
10657
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   269
(* there's still one easy, and one nontrivial (but provable) sorry in here  *)
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   270
(*
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   271
theorem wt_kil_complete:
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   272
  "[| wt_method G C pTs rT maxs mxl bs phi; wf_prog wf_mb G; 
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   273
      length phi = length bs; is_class G C; \<forall>x \<in> set pTs. is_type G x |]
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   274
  ==> wt_kil G C pTs rT maxs mxl bs"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   275
proof -
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   276
  assume wf:      "wf_prog wf_mb G"  
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   277
  assume isclass: "is_class G C"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   278
  assume istype:  "\<forall>x \<in> set pTs. is_type G x"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   279
  assume length:  "length phi = length bs"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   280
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   281
  assume "wt_method G C pTs rT maxs mxl bs phi"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   282
  then obtain
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   283
    instrs:   "0 < length bs" and
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   284
    wt_start: "wt_start G C pTs mxl phi" and
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   285
    wt_ins:   "\<forall>pc. pc < length bs \<longrightarrow> 
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   286
                    wt_instr (bs ! pc) G rT phi maxs (length bs) pc"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   287
    by (unfold wt_method_def) simp
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   288
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   289
  let ?succs = "\<lambda>pc. succs (bs!pc) pc"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   290
  let ?step  = "\<lambda>pc. step (bs!pc) G"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   291
  let ?app   = "\<lambda>pc. app (bs!pc) G maxs rT"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   292
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   293
  from wt_ins
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   294
  have bounded: "bounded ?succs (size bs)"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   295
    by (unfold wt_instr_def bounded_def) blast
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   296
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   297
  from wt_ins
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   298
  have "static_wt (sup_state_opt G) ?app ?step ?succs phi"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   299
    apply (unfold static_wt_def wt_instr_def lesub_def)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   300
    apply (simp (no_asm) only: length)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   301
    apply blast
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   302
    done
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   303
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   304
  with bounded
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   305
  have "dynamic_wt (sup_state_opt G) (err_step ?app ?step) ?succs (map OK phi)"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   306
    by - (erule static_imp_dynamic, simp add: length)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   307
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   308
  hence dynamic:
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   309
    "dynamic_wt (sup_state_opt G) (exec G maxs rT bs) ?succs (map OK phi)"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   310
    by (unfold exec_def)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   311
 
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   312
  let ?maxr = "1+size pTs+mxl"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   313
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   314
  from wf bounded
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   315
  have is_bcv: 
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10657
diff changeset
   316
    "is_bcv (JVMType.le G maxs ?maxr) Err (exec G maxs rT bs) ?succs 
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   317
            (size bs) (states G maxs ?maxr) (kiljvm G maxs ?maxr rT bs)"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   318
    by (rule is_bcv_kiljvm)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   319
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   320
  let ?start = "OK (Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err)))
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   321
                #(replicate (size bs-1) (OK None))"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   322
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   323
  { fix l x
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   324
    have "set (replicate l x) \<subseteq> {x}"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   325
      by (cases "0 < l") simp+
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   326
  } note subset_replicate = this
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   327
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   328
  from istype
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   329
  have "set pTs \<subseteq> types G"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   330
    by auto
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   331
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10812
diff changeset
   332
  hence "OK ` set pTs \<subseteq> err (types G)"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   333
    by auto
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   334
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   335
  with instrs isclass 
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   336
  have start:
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   337
    "?start \<in> list (length bs) (states G maxs ?maxr)"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   338
    apply (unfold list_def JVM_states_unfold)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   339
    apply simp
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   340
    apply (rule conjI)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   341
     apply (simp add: Un_subset_iff)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   342
     apply (rule_tac B = "{Err}" in subset_trans)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   343
      apply (simp add: subset_replicate)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   344
     apply simp
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   345
    apply (rule_tac B = "{OK None}" in subset_trans)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   346
     apply (simp add: subset_replicate)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   347
    apply simp
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   348
    done
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   349
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   350
  let ?phi = "map OK phi"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   351
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   352
  have 1: "?phi \<in> list (length bs) (states G maxs ?maxr)" sorry
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   353
10657
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   354
  have 2: "?start <=[le G maxs ?maxr] ?phi"
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   355
  proof -
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   356
    { fix n
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   357
      from wt_start
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   358
      have "G \<turnstile> ok_val (?start!0) <=' phi!0"
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   359
        by (simp add: wt_start_def)
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   360
      moreover
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   361
      from instrs length
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   362
      have "0 < length phi" by simp
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   363
      ultimately
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   364
      have "le G maxs ?maxr (?start!0) (?phi!0)"
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   365
        by (simp add: JVM_le_Err_conv Err.le_def lesub_def)
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   366
      moreover
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   367
      { fix n'
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   368
        have "le G maxs ?maxr (OK None) (?phi!n)"
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   369
          by (auto simp add: JVM_le_Err_conv Err.le_def lesub_def 
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   370
            split: err.splits)        
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   371
        hence "[| n = Suc n'; n < length ?start |] 
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   372
          ==> le G maxs ?maxr (?start!n) (?phi!n)"
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   373
          by simp
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   374
      }
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   375
      ultimately
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   376
      have "n < length ?start ==> le G maxs ?maxr (?start!n) (?phi!n)"
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   377
        by - (cases n, blast+)
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   378
    }
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   379
    thus ?thesis sorry
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   380
  qed         
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   381
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   382
  from dynamic
11299
1d3d110c456e welltyping -> wt_step
nipkow
parents: 11298
diff changeset
   383
  have "wt_step (le G maxs ?maxr) Err (exec G maxs rT bs) ?succs ?phi"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   384
    by (simp add: dynamic_wt_def JVM_le_Err_conv)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   385
  
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   386
  with start 1 2 is_bcv
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   387
  have "\<forall>p. p < length bs \<longrightarrow> kiljvm G maxs ?maxr rT bs ?start ! p \<noteq> Err"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   388
    by (unfold is_bcv_def) blast
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   389
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   390
  with bounded instrs
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   391
  show "wt_kil G C pTs rT maxs mxl bs"
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   392
    by (unfold wt_kil_def) simp
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   393
qed
10657
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   394
*)
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   395
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   396
lemma is_type_pTs:
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   397
  "[| 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
   398
      t \<in> set (snd sig) |]
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   399
  ==> is_type G t"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   400
proof -
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   401
  assume "wf_prog wf_mb G" 
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   402
         "(C,S,fs,mdecls) \<in> set G"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   403
         "(sig,rT,code) \<in> set mdecls"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   404
  hence "wf_mdecl wf_mb G C (sig,rT,code)"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   405
    by (unfold wf_prog_def wf_cdecl_def) auto
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   406
  hence "\<forall>t \<in> set (snd sig). is_type G t" 
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   407
    by (unfold wf_mdecl_def wf_mhead_def) auto
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   408
  moreover
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   409
  assume "t \<in> set (snd sig)"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   410
  ultimately
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   411
  show ?thesis by blast
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   412
qed
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   413
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   414
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   415
theorem jvm_kildall_correct:
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   416
  "wt_jvm_prog_kildall G ==> \<exists>Phi. wt_jvm_prog G Phi"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   417
proof -  
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   418
  assume wtk: "wt_jvm_prog_kildall G"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   419
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   420
  then obtain wf_mb where
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   421
    wf: "wf_prog wf_mb G"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   422
    by (auto simp add: wt_jvm_prog_kildall_def)
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   423
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   424
  let ?Phi = "\<lambda>C sig. let (C,rT,(maxs,maxl,ins)) = the (method (G,C) sig) in 
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   425
              SOME phi. wt_method G C (snd sig) rT maxs maxl ins phi"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   426
   
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   427
  { fix C S fs mdecls sig rT code
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   428
    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
   429
    with wf
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   430
    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
   431
      by (simp add: methd is_type_pTs)
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   432
  } note this [simp]
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   433
 
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   434
  from wtk
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   435
  have "wt_jvm_prog G ?Phi"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   436
    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
   437
    apply clarsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   438
    apply (drule bspec, assumption)
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   439
    apply (unfold wf_mdecl_def)
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   440
    apply clarsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   441
    apply (drule bspec, assumption)
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   442
    apply clarsimp
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   443
    apply (drule wt_kil_correct [OF _ wf])
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   444
    apply (auto intro: someI)
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   445
    done
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   446
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   447
  thus ?thesis by blast
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   448
qed
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   449
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   450
end