src/HOL/MicroJava/BV/LBVCorrect.thy
author kleing
Sun, 16 Dec 2001 00:17:44 +0100
changeset 12516 d09d0f160888
parent 12389 23bd419144eb
child 12911 704713ca07ea
permissions -rw-r--r--
exceptions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
     1
(*  Title:      HOL/MicroJava/BV/BVLCorrect.thy
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
     2
    ID:         $Id$
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
     3
    Author:     Gerwin Klein
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
     4
    Copyright   1999 Technische Universitaet Muenchen
9054
0e48e7d4d4f9 minor tuning for pdf documents
kleing
parents: 9012
diff changeset
     5
*)
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
     6
9054
0e48e7d4d4f9 minor tuning for pdf documents
kleing
parents: 9012
diff changeset
     7
header {* Correctness of the LBV *}
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
     8
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
     9
(* This theory is currently broken. The port to exceptions
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    10
  didn't make it into the Isabelle 2001 release. It is included for 
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    11
  documentation only. See Isabelle 99-2 for a working copy of this
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    12
  theory. *)
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    13
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    14
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
    15
theory LBVCorrect = BVSpec + LBVSpec:
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    16
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    17
lemmas [simp del] = split_paired_Ex split_paired_All
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    18
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    19
constdefs
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    20
fits :: "[method_type,instr list,jvm_prog,ty,state_type option,nat,nat,exception_table,certificate] => bool"
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    21
"fits phi is G rT s0 maxs maxr et cert == 
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9970
diff changeset
    22
  (\<forall>pc s1. pc < length is -->
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    23
    (wtl_inst_list (take pc is) G rT cert maxs maxr (length is) et 0 s0 = OK s1 -->
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9970
diff changeset
    24
    (case cert!pc of None   => phi!pc = s1
7164dc0d24d8 unsymbolized
kleing
parents: 9970
diff changeset
    25
                   | Some t => phi!pc = Some t)))"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    26
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    27
constdefs
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    28
make_phi :: "[instr list,jvm_prog,ty,state_type option,nat,nat,exception_table,certificate] => method_type"
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    29
"make_phi is G rT s0 maxs maxr et cert == 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    30
   map (\<lambda>pc. case cert!pc of 
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    31
               None   => ok_val (wtl_inst_list (take pc is) G rT cert maxs maxr (length is) et 0 s0) 
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9970
diff changeset
    32
             | Some t => Some t) [0..length is(]"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    33
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    34
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    35
lemma fitsD_None:
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    36
  "[|fits phi is G rT s0 mxs mxr et cert; pc < length is;
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    37
    wtl_inst_list (take pc is) G rT cert mxs mxr (length is) et 0 s0 = OK s1; 
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9970
diff changeset
    38
    cert ! pc = None|] ==> phi!pc = s1"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    39
  by (auto simp add: fits_def)
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    40
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    41
lemma fitsD_Some:
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    42
  "[|fits phi is G rT s0 mxs mxr et cert; pc < length is;
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    43
    wtl_inst_list (take pc is) G rT cert mxs mxr (length is) et 0 s0 = OK s1; 
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9970
diff changeset
    44
    cert ! pc = Some t|] ==> phi!pc = Some t"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    45
  by (auto simp add: fits_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    46
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    47
lemma make_phi_Some:
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    48
  "[| pc < length is; cert!pc = Some t |] ==> 
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    49
  make_phi is G rT s0 mxs mxr et cert ! pc = Some t"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    50
  by (simp add: make_phi_def)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    51
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    52
lemma make_phi_None:
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    53
  "[| pc < length is; cert!pc = None |] ==> 
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    54
  make_phi is G rT s0 mxs mxr et cert ! pc = 
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    55
  ok_val (wtl_inst_list (take pc is) G rT cert mxs mxr (length is) et 0 s0)"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    56
  by (simp add: make_phi_def)
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    57
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    58
lemma exists_phi:
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    59
  "\<exists>phi. fits phi is G rT s0 mxs mxr et cert"  
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    60
proof - 
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    61
  have "fits (make_phi is G rT s0 mxs mxr et cert) is G rT s0 mxs mxr et cert"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    62
    by (auto simp add: fits_def make_phi_Some make_phi_None 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    63
             split: option.splits) 
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    64
  thus ?thesis by fast
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    65
qed
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    66
  
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    67
lemma fits_lemma1:
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    68
  "[| wtl_inst_list is G rT cert mxs mxr (length is) et 0 s = OK s'; fits phi is G rT s mxs mxr et cert |]
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9970
diff changeset
    69
  ==> \<forall>pc t. pc < length is --> cert!pc = Some t --> phi!pc = Some t"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    70
proof (intro strip)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    71
  fix pc t 
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    72
  assume "wtl_inst_list is G rT cert mxs mxr (length is) et 0 s = OK s'"
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    73
  then obtain s'' where
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    74
    "wtl_inst_list (take pc is) G rT cert mxs mxr (length is) et 0 s = OK s''"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    75
    by (blast dest: wtl_take)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    76
  moreover
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    77
  assume "fits phi is G rT s mxs mxr et cert" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    78
         "pc < length is" 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    79
         "cert ! pc = Some t"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    80
  ultimately
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    81
  show "phi!pc = Some t" by (auto dest: fitsD_Some)
9580
kleing
parents: 9549
diff changeset
    82
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    83
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    84
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    85
lemma wtl_suc_pc:
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    86
 "[| wtl_inst_list is G rT cert mxs mxr (length is) et 0 s \<noteq> Err;
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    87
     wtl_inst_list (take pc is) G rT cert mxs mxr (length is) et 0 s = OK s';
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    88
     wtl_cert (is!pc) G rT s' cert mxs mxr (length is) et pc = OK s'';
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    89
     fits phi is G rT s mxs mxr et cert; Suc pc < length is |] ==>
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    90
  G \<turnstile> s'' <=' phi ! Suc pc"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    91
proof -
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    92
  
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    93
  assume all:  "wtl_inst_list is G rT cert mxs mxr (length is) et 0 s \<noteq> Err"
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    94
  assume fits: "fits phi is G rT s mxs mxr et cert"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    95
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    96
  assume wtl:  "wtl_inst_list (take pc is) G rT cert mxs mxr (length is) et 0 s = OK s'" and
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
    97
         wtc:  "wtl_cert (is!pc) G rT s' cert mxs mxr (length is) et pc = OK s''" and
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    98
         pc:   "Suc pc < length is"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    99
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   100
  hence wts: 
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   101
    "wtl_inst_list (take (Suc pc) is) G rT cert mxs mxr (length is) et 0 s = OK s''"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   102
    by (rule wtl_Suc)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   103
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   104
  from all
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   105
  have app: 
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   106
  "wtl_inst_list (take (Suc pc) is@drop (Suc pc) is) G rT cert mxs mxr (length is) et 0 s \<noteq> Err"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   107
    by simp
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   108
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   109
  from pc 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   110
  have "0 < length (drop (Suc pc) is)" 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   111
    by simp
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   112
  then obtain l ls where
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   113
    "drop (Suc pc) is = l#ls"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   114
    by (auto simp add: neq_Nil_conv simp del: length_drop)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   115
  with app wts pc
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   116
  obtain x where 
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   117
    "wtl_cert l G rT s'' cert mxs mxr (length is) et (Suc pc) = OK x"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   118
    by (auto simp add: wtl_append min_def simp del: append_take_drop_id)
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   119
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9970
diff changeset
   120
  hence c1: "!!t. cert!Suc pc = Some t ==> G \<turnstile> s'' <=' cert!Suc pc"
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   121
    by (simp add: wtl_cert_def split: split_if_asm)
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   122
  moreover
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   123
  from fits pc wts
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9970
diff changeset
   124
  have c2: "!!t. cert!Suc pc = Some t ==> phi!Suc pc = cert!Suc pc"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   125
    by - (drule fitsD_Some, auto)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   126
  moreover
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   127
  from fits pc wts
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   128
  have c3: "cert!Suc pc = None ==> phi!Suc pc = s''"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   129
    by (rule fitsD_None)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   130
  ultimately
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   131
  show ?thesis by (cases "cert!Suc pc", auto)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   132
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   133
9376
c32c5696ec2a flat instruction set
kleing
parents: 9260
diff changeset
   134
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   135
lemma wtl_fits_wt:
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   136
  "[| wtl_inst_list is G rT cert mxs mxr (length is) et 0 s \<noteq> Err; 
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   137
      fits phi is G rT s mxs mxr et cert; pc < length is |] ==>
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   138
   wt_instr (is!pc) G rT phi mxs (length is) et pc"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   139
proof -
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   140
  assume fits: "fits phi is G rT s mxs mxr et cert"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   141
  assume pc:  "pc < length is" and
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   142
         wtl: "wtl_inst_list is G rT cert mxs mxr (length is) et 0 s \<noteq> Err"        
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   143
  then obtain s' s'' where
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   144
    w: "wtl_inst_list (take pc is) G rT cert mxs mxr (length is) et 0 s = OK s'" and
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   145
    c: "wtl_cert (is!pc) G rT s' cert mxs mxr (length is) et pc = OK s''"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   146
    by - (drule wtl_all, auto)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   147
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   148
  from fits wtl pc have cert_Some: 
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9970
diff changeset
   149
    "!!t pc. [| pc < length is; cert!pc = Some t |] ==> phi!pc = Some t"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   150
    by (auto dest: fits_lemma1)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   151
  
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   152
  from fits wtl pc have cert_None: "cert!pc = None ==> phi!pc = s'"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   153
    by - (drule fitsD_None)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   154
  
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   155
  from pc c cert_None cert_Some
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   156
  have wti: "wtl_inst (is ! pc) G rT (phi!pc) cert mxs mxr (length is) et pc = OK s''"
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   157
    by (auto simp add: wtl_cert_def split: split_if_asm option.splits)
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   158
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   159
  -- "we now must show that @{text wt_instr} holds; the definition 
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   160
      of @{text wt_instr} is: @{thm [display] wt_instr_def[no_vars]}"
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   161
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   162
  { fix pc' r
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   163
    assume pc': "(pc',r) \<in> set (eff (is!pc) G pc et (phi!pc))"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   164
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   165
    with wti have less: "pc' < length is" by (simp add: wtl_inst_OK) blast
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   166
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   167
    have "G \<turnstile> r <=' phi ! pc'" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   168
    proof (cases "pc' = Suc pc")
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   169
      case False          
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   170
      with wti pc'
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   171
      have G: "G \<turnstile> r <=' cert ! pc'" by (simp add: wtl_inst_OK) blast
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   172
      hence "cert!pc' = None ==> r = None" by simp
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   173
      hence "cert!pc' = None ==> ?thesis" by simp
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   174
      moreover {
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   175
        fix t assume "cert!pc' = Some t"
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   176
        with less have "phi!pc' = cert!pc'" by (simp add: cert_Some)
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   177
        with G have ?thesis by simp
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   178
      }
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   179
      ultimately show ?thesis by blast      
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   180
    next
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   181
      case True
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   182
      with pc' wti
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   183
      have "G \<turnstile> r <=' s''"  sorry
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   184
      also
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   185
      from wtl w fits c pc 
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   186
      have "Suc pc < length is ==> G \<turnstile> s'' <=' phi ! Suc pc" 
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   187
        by - (rule wtl_suc_pc)
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   188
      with True less
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   189
      have "G \<turnstile> s'' <=' phi ! Suc pc" by blast
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   190
      finally show ?thesis by (simp add: True)
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   191
    qed
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   192
  }  
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   193
  with wti show ?thesis
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
   194
    by (auto simp add: wtl_inst_OK wt_instr_def)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   195
qed
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   196
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   197
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   198
    
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   199
lemma fits_first:
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   200
  "[| 0 < length is; wtl_inst_list is G rT cert mxs mxr (length is) et 0 s \<noteq> Err; 
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   201
      fits phi is G rT s mxs mxr et cert |] ==> 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   202
  G \<turnstile> s <=' phi ! 0"
9580
kleing
parents: 9549
diff changeset
   203
proof -
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   204
  assume wtl:  "wtl_inst_list is G rT cert mxs mxr (length is) et 0 s \<noteq> Err"
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   205
  assume fits: "fits phi is G rT s mxs mxr et cert"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   206
  assume pc:   "0 < length is"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   207
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   208
  from wtl
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   209
  have wt0: "wtl_inst_list (take 0 is) G rT cert mxs mxr (length is) et 0 s = OK s"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   210
    by simp
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   211
  
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   212
  with fits pc
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   213
  have "cert!0 = None ==> phi!0 = s"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   214
    by (rule fitsD_None)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   215
  moreover    
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   216
  from fits pc wt0
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9970
diff changeset
   217
  have "!!t. cert!0 = Some t ==> phi!0 = cert!0"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   218
    by - (drule fitsD_Some, auto)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   219
  moreover
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   220
  from pc
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   221
  obtain x xs where "is = x#xs" 
12389
23bd419144eb eliminated old use of intro/elim method;
wenzelm
parents: 11085
diff changeset
   222
    by (auto simp add: neq_Nil_conv)
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   223
  with wtl
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   224
  obtain s' where
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   225
    "wtl_cert x G rT s cert mxs mxr (length is) et 0 = OK s'"
12389
23bd419144eb eliminated old use of intro/elim method;
wenzelm
parents: 11085
diff changeset
   226
    by auto
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   227
  hence 
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9970
diff changeset
   228
    "!!t. cert!0 = Some t ==> G \<turnstile> s <=' cert!0"
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   229
    by (simp add: wtl_cert_def split: split_if_asm)
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   230
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   231
  ultimately
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   232
  show ?thesis
12389
23bd419144eb eliminated old use of intro/elim method;
wenzelm
parents: 11085
diff changeset
   233
    by (cases "cert!0") auto
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   234
qed
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   235
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   236
  
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   237
lemma wtl_method_correct:
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   238
"wtl_method G C pTs rT mxs mxl et ins cert ==> \<exists> phi. wt_method G C pTs rT mxs mxl ins et phi"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   239
proof (unfold wtl_method_def, simp only: Let_def, elim conjE)
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
   240
  let "?s0" = "Some ([], OK (Class C) # map OK pTs @ replicate mxl Err)"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   241
  assume pc:  "0 < length ins"
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   242
  assume wtl: "wtl_inst_list ins G rT cert mxs mxl (length ins) et 0 ?s0 \<noteq> Err"
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   243
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   244
  obtain phi where fits: "fits phi ins G rT ?s0 mxs mxl et cert"    
9941
fe05af7ec816 renamed atts: rulify to rule_format, elimify to elim_format;
wenzelm
parents: 9906
diff changeset
   245
    by (rule exists_phi [elim_format]) blast
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   246
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   247
  with wtl
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   248
  have allpc:
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   249
    "\<forall>pc. pc < length ins --> wt_instr (ins ! pc) G rT phi mxs (length ins) et pc"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   250
    by (blast intro: wtl_fits_wt)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   251
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   252
  from pc wtl fits
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   253
  have "wt_start G C pTs mxl phi"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   254
    by (unfold wt_start_def) (rule fits_first)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   255
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   256
  with pc allpc 
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   257
  show ?thesis by (auto simp add: wt_method_def)
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   258
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   259
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   260
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   261
theorem wtl_correct:
10628
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   262
  "wtl_jvm_prog G cert ==> \<exists> Phi. wt_jvm_prog G Phi"
11085
b830bf10bf71 tuned for 99-2 release
kleing
parents: 10812
diff changeset
   263
proof -  
10628
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   264
  assume wtl: "wtl_jvm_prog G cert"
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   265
12516
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   266
  let ?Phi = "\<lambda>C sig. let (C,rT,(maxs,maxl,ins,et)) = the (method (G,C) sig) in 
d09d0f160888 exceptions
kleing
parents: 12389
diff changeset
   267
              SOME phi. wt_method G C (snd sig) rT maxs maxl ins et phi"
10628
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   268
   
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   269
  { fix C S fs mdecls sig rT code
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   270
    assume "(C,S,fs,mdecls) \<in> set G" "(sig,rT,code) \<in> set mdecls"
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   271
    moreover
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   272
    from wtl obtain wf_mb where "wf_prog wf_mb G" 
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   273
      by (auto simp add: wtl_jvm_prog_def)
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   274
    ultimately
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   275
    have "method (G,C) sig = Some (C,rT,code)"
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   276
      by (simp add: methd)
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   277
  } note this [simp]
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   278
 
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   279
  from wtl
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   280
  have "wt_jvm_prog G ?Phi"
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   281
    apply (clarsimp simp add: wt_jvm_prog_def wtl_jvm_prog_def wf_prog_def wf_cdecl_def)
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   282
    apply (drule bspec, assumption)
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   283
    apply (clarsimp simp add: wf_mdecl_def)
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   284
    apply (drule bspec, assumption)
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   285
    apply (clarsimp dest!: wtl_method_correct)
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   286
    apply (rule someI, assumption)
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   287
    done
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   288
10628
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   289
  thus ?thesis
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   290
    by blast
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   291
qed   
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10612
diff changeset
   292
      
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10628
diff changeset
   293
end