src/HOL/MicroJava/BV/JVM.thy
author berghofe
Wed, 07 Feb 2007 17:44:07 +0100
changeset 22271 51a80e238b29
parent 16417 9bc16273c2d4
child 26450 158b924b5153
permissions -rw-r--r--
Adapted to new inductive definition package.
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
12911
704713ca07ea new document
kleing
parents: 12516
diff changeset
     7
header {* \isaheader{Kildall for the JVM}\label{sec:JVM} *}
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
     8
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14045
diff changeset
     9
theory JVM imports Kildall Typing_Framework_JVM begin
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
    10
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    11
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    12
constdefs
13006
51c5f3f11d16 symbolized
kleing
parents: 12951
diff changeset
    13
  kiljvm :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> exception_table \<Rightarrow> 
51c5f3f11d16 symbolized
kleing
parents: 12951
diff changeset
    14
             instr list \<Rightarrow> state list \<Rightarrow> state list"
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    15
  "kiljvm G maxs maxr rT et bs ==
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    16
  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
    17
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    18
  wt_kil :: "jvm_prog \<Rightarrow> cname \<Rightarrow> ty list \<Rightarrow> ty \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    19
             exception_table \<Rightarrow> instr list \<Rightarrow> bool"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    20
  "wt_kil G C pTs rT mxs mxl et ins ==
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
    21
   check_bounded ins et \<and> 0 < size ins \<and> 
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    22
   (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
    23
        start  = OK first#(replicate (size ins - 1) (OK None));
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    24
        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
    25
    in \<forall>n < size ins. result!n \<noteq> Err)"
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    26
13006
51c5f3f11d16 symbolized
kleing
parents: 12951
diff changeset
    27
  wt_jvm_prog_kildall :: "jvm_prog \<Rightarrow> bool"
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    28
  "wt_jvm_prog_kildall G ==
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    29
  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
    30
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    31
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    32
theorem is_bcv_kiljvm:
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
    33
  "\<lbrakk> wf_prog wf_mb G; bounded (exec G maxs rT et bs) (size bs) \<rbrakk> \<Longrightarrow>
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    34
      is_bcv (JVMType.le G maxs maxr) Err (exec G maxs rT et bs)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    35
             (size bs) (states G maxs maxr) (kiljvm G maxs maxr rT et bs)"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    36
  apply (unfold kiljvm_def sl_triple_conv)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    37
  apply (rule is_bcv_kildall)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    38
       apply (simp (no_asm) add: sl_triple_conv [symmetric]) 
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13224
diff changeset
    39
       apply (force intro!: semilat_JVM_slI dest: wf_acyclic 
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13224
diff changeset
    40
	 simp add: symmetric sl_triple_conv)
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    41
      apply (simp (no_asm) add: JVM_le_unfold)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    42
      apply (blast intro!: order_widen wf_converse_subcls1_impl_acc_subtype
22271
51a80e238b29 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    43
                   dest: wf_subcls1 wfP_acyclicP wf_prog_ws_prog)
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    44
     apply (simp add: JVM_le_unfold)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    45
    apply (erule exec_pres_type)
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
    46
   apply assumption
14045
a34d89ce6097 Introduced distinction wf_prog vs. ws_prog
streckem
parents: 13224
diff changeset
    47
  apply (drule wf_prog_ws_prog, erule exec_mono, assumption)  
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    48
  done
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    49
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    50
lemma subset_replicate: "set (replicate n x) \<subseteq> {x}"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    51
  by (induct n) auto
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    52
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    53
lemma in_set_replicate:
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    54
  "x \<in> set (replicate n y) \<Longrightarrow> x = y"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    55
proof -
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    56
  assume "x \<in> set (replicate n y)"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    57
  also have "set (replicate n y) \<subseteq> {y}" by (rule subset_replicate)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    58
  finally have "x \<in> {y}" .
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    59
  thus ?thesis by simp
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    60
qed
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    61
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    62
theorem wt_kil_correct:
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    63
  assumes wf:  "wf_prog wf_mb G"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    64
  assumes C:   "is_class G C"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    65
  assumes pTs: "set pTs \<subseteq> types G"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    66
  
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    67
  assumes wtk: "wt_kil G C pTs rT maxs mxl et bs"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    68
  
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    69
  shows "\<exists>phi. wt_method G C pTs rT maxs mxl bs et phi"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    70
proof -
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    71
  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
    72
                #(replicate (size bs - 1) (OK None))"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    73
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    74
  from wtk obtain maxr r where    
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
    75
    bounded: "check_bounded bs et" and
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    76
    result:  "r = kiljvm G maxs maxr rT et bs ?start" and
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    77
    success: "\<forall>n < size bs. r!n \<noteq> Err" and
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    78
    instrs:  "0 < size bs" and
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    79
    maxr:    "maxr = Suc (length pTs + mxl)" 
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
    80
    by (unfold wt_kil_def) simp
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    81
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
    82
  from bounded have "bounded (exec G maxs rT et bs) (size bs)"
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
    83
    by (unfold exec_def) (intro bounded_lift check_bounded_is_bounded)
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
    84
  with wf have bcv:
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    85
    "is_bcv (JVMType.le G maxs maxr) Err (exec G maxs rT et bs) 
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
    86
    (size bs) (states G maxs maxr) (kiljvm G maxs maxr rT et bs)"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    87
    by (rule is_bcv_kiljvm)
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
    88
    
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    89
  from C pTs instrs maxr
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    90
  have "?start \<in> list (length bs) (states G maxs maxr)"
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    91
    apply (unfold JVM_states_unfold)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    92
    apply (rule listI)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    93
    apply (auto intro: list_appendI dest!: in_set_replicate)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    94
    apply force
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    95
    done    
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
    96
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
    97
  with bcv success result have 
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
    98
    "\<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
    99
         ?start <=[JVMType.le G maxs maxr] ts \<and>
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   100
         wt_step (JVMType.le G maxs maxr) Err (exec G maxs rT et bs) ts"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   101
    by (unfold is_bcv_def) auto
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   102
  then obtain phi' where
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   103
    phi': "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
   104
    s: "?start <=[JVMType.le G maxs maxr] phi'" and
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   105
    w: "wt_step (JVMType.le G maxs maxr) Err (exec G maxs rT et bs) phi'"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   106
    by blast
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   107
  hence wt_err_step:
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   108
    "wt_err_step (sup_state_opt G) (exec G maxs rT et bs) phi'"
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   109
    by (simp add: wt_err_step_def exec_def JVM_le_Err_conv)
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   110
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   111
  from s have le: "JVMType.le G maxs maxr (?start ! 0) (phi'!0)"
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   112
    by (drule_tac p=0 in le_listD) (simp add: lesub_def)+
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   113
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   114
  from phi' have l: "size phi' = size bs" by simp  
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   115
  with instrs w have "phi' ! 0 \<noteq> Err" by (unfold wt_step_def) simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   116
  with instrs l have phi0: "OK (map ok_val phi' ! 0) = phi' ! 0"
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   117
    by (clarsimp simp add: not_Err_eq)  
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   118
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   119
  from phi' have "check_types G maxs maxr phi'" by(simp add: check_types_def)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   120
  also from w have "phi' = map OK (map ok_val phi')" 
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   121
    apply (clarsimp simp add: wt_step_def not_Err_eq) 
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   122
    apply (rule map_id [symmetric])
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   123
    apply (erule allE, erule impE, assumption)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   124
    apply clarsimp
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   125
    done    
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   126
  finally 
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   127
  have check_types:
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   128
    "check_types G maxs maxr (map OK (map ok_val phi'))" .
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   129
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   130
  from l bounded 
13067
kleing
parents: 13066
diff changeset
   131
  have "bounded (\<lambda>pc. eff (bs!pc) G pc et) (length phi')"
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   132
    by (simp add: exec_def check_bounded_is_bounded)  
13067
kleing
parents: 13066
diff changeset
   133
  hence bounded': "bounded (exec G maxs rT et bs) (length bs)"
kleing
parents: 13066
diff changeset
   134
    by (auto intro: bounded_lift simp add: exec_def l)
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   135
  with wt_err_step
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   136
  have "wt_app_eff (sup_state_opt G) (\<lambda>pc. app (bs!pc) G maxs rT pc et) 
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   137
                   (\<lambda>pc. eff (bs!pc) G pc et) (map ok_val phi')"
13067
kleing
parents: 13066
diff changeset
   138
    by (auto intro: wt_err_imp_wt_app_eff simp add: l exec_def)
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   139
  with instrs l le bounded bounded' check_types maxr
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   140
  have "wt_method G C pTs rT maxs mxl bs et (map ok_val phi')"
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   141
    apply (unfold wt_method_def wt_app_eff_def)
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   142
    apply simp
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   143
    apply (rule conjI)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   144
     apply (unfold wt_start_def)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   145
     apply (rule JVM_le_convert [THEN iffD1])
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   146
     apply (simp (no_asm) add: phi0)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   147
    apply clarify
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   148
    apply (erule allE, erule impE, assumption)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   149
    apply (elim conjE)
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   150
    apply (clarsimp simp add: lesub_def wt_instr_def)
13067
kleing
parents: 13066
diff changeset
   151
    apply (simp add: exec_def)
kleing
parents: 13066
diff changeset
   152
    apply (drule bounded_err_stepD, assumption+)
kleing
parents: 13066
diff changeset
   153
    apply blast
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   154
    done
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   155
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   156
  thus ?thesis by blast
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   157
qed
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   158
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   159
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   160
theorem wt_kil_complete:
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   161
  assumes wf:  "wf_prog wf_mb G"  
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   162
  assumes C:   "is_class G C"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   163
  assumes pTs: "set pTs \<subseteq> types G"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   164
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   165
  assumes wtm: "wt_method G C pTs rT maxs mxl bs et phi"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   166
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   167
  shows "wt_kil G C pTs rT maxs mxl et bs"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   168
proof -
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   169
  let ?mxr = "1+size pTs+mxl"
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   170
  
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   171
  from wtm obtain
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   172
    instrs:   "0 < length bs" and
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   173
    len:      "length phi = length bs" and
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   174
    bounded:  "check_bounded bs et" and
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   175
    ck_types: "check_types G maxs ?mxr (map OK phi)" and
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   176
    wt_start: "wt_start G C pTs mxl phi" and
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   177
    wt_ins:   "\<forall>pc. pc < length bs \<longrightarrow> 
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   178
                    wt_instr (bs ! pc) G rT phi maxs (length bs) et pc"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   179
    by (unfold wt_method_def) simp
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   180
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   181
  from ck_types len
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   182
  have istype_phi: 
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   183
    "map OK phi \<in> list (length bs) (states G maxs (1+size pTs+mxl))"
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   184
    by (auto simp add: check_types_def intro!: listI)
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   185
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   186
  let ?eff  = "\<lambda>pc. eff (bs!pc) G pc et"
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   187
  let ?app   = "\<lambda>pc. app (bs!pc) G maxs rT pc et"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   188
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   189
  from bounded
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   190
  have bounded_exec: "bounded (exec G maxs rT et bs) (size bs)"
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   191
    by (unfold exec_def) (intro bounded_lift check_bounded_is_bounded)
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   192
 
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   193
  from wt_ins
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   194
  have "wt_app_eff (sup_state_opt G) ?app ?eff phi"
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   195
    apply (unfold wt_app_eff_def wt_instr_def lesub_def)
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   196
    apply (simp (no_asm) only: len)
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   197
    apply blast
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   198
    done
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   199
  with bounded_exec
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   200
  have "wt_err_step (sup_state_opt G) (err_step (size phi) ?app ?eff) (map OK phi)"
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   201
    by - (erule wt_app_eff_imp_wt_err,simp add: exec_def len)
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   202
  hence wt_err:
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   203
    "wt_err_step (sup_state_opt G) (exec G maxs rT et bs) (map OK phi)"
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   204
    by (unfold exec_def) (simp add: len)
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   205
 
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   206
  from wf bounded_exec
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   207
  have is_bcv: 
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   208
    "is_bcv (JVMType.le G maxs ?mxr) Err (exec G maxs rT et bs) 
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   209
            (size bs) (states G maxs ?mxr) (kiljvm G maxs ?mxr rT et bs)"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   210
    by (rule is_bcv_kiljvm)
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   211
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   212
  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
   213
                #(replicate (size bs - 1) (OK None))"
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   214
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   215
  from C pTs instrs
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   216
  have start: "?start \<in> list (length bs) (states G maxs ?mxr)"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   217
    apply (unfold JVM_states_unfold)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   218
    apply (rule listI)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   219
    apply (auto intro!: list_appendI dest!: in_set_replicate)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   220
    apply force
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   221
    done    
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   222
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   223
  let ?phi = "map OK phi"  
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   224
  have less_phi: "?start <=[JVMType.le G maxs ?mxr] ?phi"
10657
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   225
  proof -
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   226
    from len instrs
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   227
    have "length ?start = length (map OK phi)" by simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   228
    moreover
10657
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   229
    { fix n
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   230
      from wt_start
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   231
      have "G \<turnstile> ok_val (?start!0) <=' phi!0"
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   232
        by (simp add: wt_start_def)
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   233
      moreover
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   234
      from instrs len
10657
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   235
      have "0 < length phi" by simp
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   236
      ultimately
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   237
      have "JVMType.le G maxs ?mxr (?start!0) (?phi!0)"
10657
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   238
        by (simp add: JVM_le_Err_conv Err.le_def lesub_def)
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   239
      moreover
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   240
      { fix n'
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   241
        have "JVMType.le G maxs ?mxr (OK None) (?phi!n)"
10657
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   242
          by (auto simp add: JVM_le_Err_conv Err.le_def lesub_def 
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   243
            split: err.splits)        
13006
51c5f3f11d16 symbolized
kleing
parents: 12951
diff changeset
   244
        hence "\<lbrakk> n = Suc n'; n < length ?start \<rbrakk> 
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   245
          \<Longrightarrow> JVMType.le G maxs ?mxr (?start!n) (?phi!n)"
10657
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   246
          by simp
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   247
      }
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   248
      ultimately
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   249
      have "n < length ?start \<Longrightarrow> (?start!n) <=_(JVMType.le G maxs ?mxr) (?phi!n)"
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   250
        by (unfold lesub_def) (cases n, blast+)
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   251
    } 
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   252
    ultimately show ?thesis by (rule le_listI)
10657
7e5d659899bf removed sorry proof
kleing
parents: 10651
diff changeset
   253
  qed         
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   254
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   255
  from wt_err
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   256
  have "wt_step (JVMType.le G maxs ?mxr) Err (exec G maxs rT et bs) ?phi"
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   257
    by (simp add: wt_err_step_def JVM_le_Err_conv)  
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   258
  with start istype_phi less_phi is_bcv
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   259
  have "\<forall>p. p < length bs \<longrightarrow> kiljvm G maxs ?mxr rT et bs ?start ! p \<noteq> Err"
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   260
    by (unfold is_bcv_def) auto
13066
b57d926d1de2 cleanup + simpler monotonicity
kleing
parents: 13006
diff changeset
   261
  with bounded instrs
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   262
  show "wt_kil G C pTs rT maxs mxl et bs" by (unfold wt_kil_def) simp
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   263
qed
10651
bb3a81a005f7 completeness (unfinished)
kleing
parents: 10637
diff changeset
   264
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   265
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   266
theorem jvm_kildall_sound_complete:
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   267
  "wt_jvm_prog_kildall G = (\<exists>Phi. wt_jvm_prog G Phi)"
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   268
proof 
12516
d09d0f160888 exceptions
kleing
parents: 12230
diff changeset
   269
  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
   270
              SOME phi. wt_method G C (snd sig) rT maxs maxl ins et phi"
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   271
  
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   272
  assume "wt_jvm_prog_kildall G"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   273
  hence "wt_jvm_prog G ?Phi"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   274
    apply (unfold wt_jvm_prog_def wt_jvm_prog_kildall_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   275
    apply (erule jvm_prog_lift)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   276
    apply (auto dest!: wt_kil_correct intro: someI)
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   277
    done
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   278
  thus "\<exists>Phi. wt_jvm_prog G Phi" by fast
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   279
next
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   280
  assume "\<exists>Phi. wt_jvm_prog G Phi"
13224
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   281
  thus "wt_jvm_prog_kildall G"
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   282
    apply (clarify)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   283
    apply (unfold wt_jvm_prog_def wt_jvm_prog_kildall_def)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   284
    apply (erule jvm_prog_lift)
6f0928a942d1 LBV instantiantion refactored, streamlined
kleing
parents: 13214
diff changeset
   285
    apply (auto intro: wt_kil_complete)
13214
2aa33ed5f526 wt_method now checks bounded+types ==> wt_kildall <=> wt_method
kleing
parents: 13074
diff changeset
   286
    done
10637
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   287
qed
41d309b48afe kildall ==> wt_method for whole program
kleing
parents: 10630
diff changeset
   288
10592
fc0b575a2cf7 BCV Integration
kleing
parents:
diff changeset
   289
end