src/HOL/MicroJava/BV/LBVSpec.thy
author kleing
Sun, 03 Mar 2002 16:59:08 +0100
changeset 13006 51c5f3f11d16
parent 12978 16cc829b9c65
child 13062 4b1edf2f6bd2
permissions -rw-r--r--
symbolized
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/LBVSpec.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
12911
704713ca07ea new document
kleing
parents: 12516
diff changeset
     7
header {* \isaheader{The Lightweight Bytecode Verifier} *}
9054
0e48e7d4d4f9 minor tuning for pdf documents
kleing
parents: 9012
diff changeset
     8
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
     9
theory LBVSpec = Effect + Kildall:
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    10
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    11
text {* 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    12
  The Lightweight Bytecode Verifier with exceptions has not 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    13
  made it completely into the Isabelle 2001 release. Currently 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    14
  there is only the specification itself available. The proofs of
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    15
  soundness and completeness are broken (they still need to be
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    16
  ported to the exception version). Both theories are included
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    17
  for documentation (but they don't work for this specification), 
12978
16cc829b9c65 included LBVSpec again
kleing
parents: 12911
diff changeset
    18
  please see the Isabelle 99-2 release for a working copy, or
16cc829b9c65 included LBVSpec again
kleing
parents: 12911
diff changeset
    19
  \url{http://isabelle.in.tum.de/verificard} for the most recent
16cc829b9c65 included LBVSpec again
kleing
parents: 12911
diff changeset
    20
  development of \mJava.
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    21
*}
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    22
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    23
types
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    24
  certificate       = "state_type option list" 
13006
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
    25
  class_certificate = "sig \<Rightarrow> certificate"
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
    26
  prog_certificate  = "cname \<Rightarrow> class_certificate"
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    27
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    28
consts
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    29
  merge :: "[jvm_prog, p_count, nat, nat, p_count, certificate, (nat \<times> (state_type option)) list, state] \<Rightarrow> state"
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    30
primrec
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    31
  "merge G pc mxs mxr max_pc cert []     x = x"
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    32
  "merge G pc mxs mxr max_pc cert (s#ss) x = (let (pc',s') = s in 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    33
                                       if pc' < max_pc \<and> pc' = pc+1 then 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    34
                                         merge G pc mxs mxr max_pc cert ss (OK s' +_(sup G mxs mxr) x) 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    35
                                       else if pc' < max_pc \<and> G \<turnstile> s' <=' cert!pc' then 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    36
                                         merge G pc mxs mxr max_pc cert ss x
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    37
                                       else Err)"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    38
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    39
constdefs 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    40
  wtl_inst :: "[instr,jvm_prog,ty,state_type option,certificate,nat,nat,p_count,exception_table,p_count] 
13006
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
    41
               \<Rightarrow> state_type option err" 
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    42
  "wtl_inst i G rT s cert maxs maxr max_pc et pc == 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    43
     if app i G maxs rT pc et s then 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    44
       merge G pc maxs maxr max_pc cert (eff i G pc et s) (OK (cert!(pc+1)))
10628
4ea7f3e8e471 fixed and cleaned up wt[l]_jvm_prog proofs
kleing
parents: 10592
diff changeset
    45
     else Err"
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    46
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    47
  wtl_cert :: "[instr,jvm_prog,ty,state_type option,certificate,nat,nat,p_count,exception_table,p_count] 
13006
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
    48
               \<Rightarrow> state_type option err"  
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    49
  "wtl_cert i G rT s cert maxs maxr max_pc et pc ==
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    50
     case cert!pc of
13006
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
    51
        None    \<Rightarrow> wtl_inst i G rT s cert maxs maxr max_pc et pc
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
    52
      | Some s' \<Rightarrow> if G \<turnstile> s <=' (Some s') then 
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    53
                    wtl_inst i G rT (Some s') cert maxs maxr max_pc et pc 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    54
                  else Err"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    55
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    56
consts 
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    57
  wtl_inst_list :: "[instr list,jvm_prog,ty,certificate,nat,nat,p_count,exception_table,p_count,
13006
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
    58
                     state_type option] \<Rightarrow> state_type option err"
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    59
primrec
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    60
  "wtl_inst_list []     G rT cert maxs maxr max_pc et pc s = OK s"
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    61
  "wtl_inst_list (i#is) G rT cert maxs maxr max_pc et pc s = 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    62
    (let s' = wtl_cert i G rT s cert maxs maxr max_pc et pc in
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    63
     strict (wtl_inst_list is G rT cert maxs maxr max_pc et (pc+1)) s')"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    64
              
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    65
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    66
constdefs
13006
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
    67
 wtl_method :: "[jvm_prog,cname,ty list,ty,nat,nat,exception_table,instr list,certificate] \<Rightarrow> bool"  
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    68
 "wtl_method G C pTs rT mxs mxl et ins cert ==  
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    69
	let max_pc = length ins  
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    70
  in 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    71
  0 < max_pc \<and>   
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    72
  wtl_inst_list ins G rT cert mxs mxl max_pc et 0 
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
    73
                (Some ([],(OK (Class C))#((map OK pTs))@(replicate mxl Err))) \<noteq> Err"
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    74
13006
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
    75
 wtl_jvm_prog :: "[jvm_prog,prog_certificate] \<Rightarrow> bool" 
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9757
diff changeset
    76
 "wtl_jvm_prog G cert ==  
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    77
  wf_prog (\<lambda>G C (sig,rT,maxs,maxl,b,et). wtl_method G C (snd sig) rT maxs maxl et b (cert C sig)) G"
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    78
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9757
diff changeset
    79
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10638
diff changeset
    80
lemmas [iff] = not_Err_eq
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10638
diff changeset
    81
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    82
lemma if_eq_cases:
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    83
  "(P \<Longrightarrow> x = z) \<Longrightarrow> (\<not>P \<Longrightarrow> y = z) \<Longrightarrow> (if P then x else y) = z"
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    84
  by simp
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    85
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    86
lemma merge_def:
13006
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
    87
  "\<And>x. merge G pc mxs mxr max_pc cert ss x = 
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    88
  (if \<forall>(pc',s') \<in> set ss. pc' < max_pc \<and> (pc' \<noteq> pc+1 \<longrightarrow> G \<turnstile> s' <=' cert!pc') then
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    89
    map (OK \<circ> snd) [(p',t') \<in> ss. p'=pc+1] ++_(sup G mxs mxr) x
13006
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
    90
  else Err)" (is "\<And>x. ?merge ss x = ?if ss x" is "\<And>x. ?P ss x")
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    91
proof (induct ss)
13006
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
    92
  show "\<And>x. ?P [] x" by simp
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    93
next
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    94
  have OK_snd: "(\<lambda>u. OK (snd u)) = OK \<circ> snd" by (rule ext) auto
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    95
  fix x ss and s::"nat \<times> (state_type option)"
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    96
  assume IH: "\<And>x. ?P ss x"
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    97
  obtain pc' s' where s: "s = (pc',s')" by (cases s)  
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    98
  hence "?merge (s#ss) x = ?merge ((pc',s')#ss) x" by hypsubst (rule refl)
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    99
  also
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   100
  have "\<dots> = (if pc' < max_pc \<and> pc' = pc+1 then 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   101
               ?merge ss (OK s' +_(sup G mxs mxr) x)
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   102
             else if pc' < max_pc \<and> G \<turnstile> s' <=' cert!pc' then 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   103
               ?merge ss x
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   104
             else Err)" 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   105
    (is "_ = (if ?pc' then ?merge ss (_ +_?f _) else if ?G then _ else _)")
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   106
    by simp 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   107
  also 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   108
  let "if ?all ss then _ else _" = "?if ss x"    
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   109
  have "\<dots> = ?if ((pc',s')#ss) x"
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   110
  proof (cases "?pc'")    
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   111
    case True
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   112
    hence "?all (((pc',s')#ss)) = (pc+1 < max_pc \<and> ?all ss)" by simp
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   113
    with True
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   114
    have "?if ss (OK s' +_?f x) = ?if ((pc',s')#ss) x" by (auto simp add: OK_snd)
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   115
    hence "?merge ss (OK s' +_?f x) = ?if ((pc',s')#ss) x" by (simp only: IH)
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   116
    with True show ?thesis by (fast intro: if_eq_cases)
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   117
  next
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   118
    case False
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   119
    have "(if ?G then ?merge ss x else Err) = ?if ((pc',s')#ss) x" 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   120
    proof (cases ?G)
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   121
      assume G: ?G with False
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   122
      have "?if ss x = ?if ((pc',s')#ss) x" by (auto simp add: OK_snd)
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   123
      hence "?merge ss x = ?if ((pc',s')#ss) x" by (simp only: IH)
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   124
      with G show ?thesis by (fast intro: if_eq_cases)
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   125
    next
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   126
      assume G: "\<not>?G"
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   127
      with False have "Err = ?if ((pc',s')#ss) x" by auto
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   128
      with G show ?thesis by (fast intro: if_eq_cases)
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   129
    qed
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   130
    with False show ?thesis by (fast intro: if_eq_cases)
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   131
  qed
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   132
  also from s have "\<dots> = ?if (s#ss) x" by hypsubst (rule refl)
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   133
  finally show "?P (s#ss) x" .
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   134
qed
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   135
  
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9757
diff changeset
   136
10496
f2d304bdf3cc BCV integration (first step)
kleing
parents: 10042
diff changeset
   137
lemma wtl_inst_OK:
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   138
"(wtl_inst i G rT s cert mxs mxr max_pc et pc = OK s') = (
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   139
 app i G mxs rT pc et s \<and> 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   140
  (\<forall>(pc',r) \<in> set (eff i G pc et s). 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   141
  pc' < max_pc \<and> (pc' \<noteq> pc+1 \<longrightarrow> G \<turnstile> r <=' cert!pc')) \<and> 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   142
  map (OK \<circ> snd) [(p',t') \<in> (eff i G pc et s). p'=pc+1] 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   143
  ++_(sup G mxs mxr) (OK (cert!(pc+1))) = OK s')"
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   144
  by (auto simp add: wtl_inst_def merge_def split: split_if_asm)
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   145
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   146
lemma wtl_Cons:
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   147
  "wtl_inst_list (i#is) G rT cert maxs maxr max_pc et pc s \<noteq> Err = 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   148
  (\<exists>s'. wtl_cert i G rT s cert maxs maxr max_pc et pc = OK s' \<and> 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   149
        wtl_inst_list is G rT cert maxs maxr max_pc et (pc+1) s' \<noteq> Err)"
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10638
diff changeset
   150
  by (auto simp del: split_paired_Ex)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   151
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   152
lemma wtl_append:
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   153
"\<forall> s pc. (wtl_inst_list (a@b) G rT cert mxs mxr mpc et pc s = OK s') =
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   154
   (\<exists>s''. wtl_inst_list a G rT cert mxs mxr mpc et pc s = OK s'' \<and> 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   155
          wtl_inst_list b G rT cert mxs mxr mpc et (pc+length a) s'' = OK s')" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   156
  (is "\<forall> s pc. ?C s pc a" is "?P a")
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   157
proof (induct ?P a)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   158
  show "?P []" by simp
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   159
next
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   160
  fix x xs  assume IH: "?P xs"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   161
  show "?P (x#xs)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   162
  proof (intro allI)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   163
    fix s pc 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   164
    show "?C s pc (x#xs)"
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   165
    proof (cases "wtl_cert x G rT s cert mxs mxr mpc et pc")
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   166
      case Err thus ?thesis by simp
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   167
    next
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   168
      fix s0 assume "wtl_cert x G rT s cert mxs mxr mpc et pc = OK s0"      
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   169
      with IH have "?C s0 (Suc pc) xs" by blast
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   170
      thus ?thesis by (simp!)
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   171
    qed
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   172
  qed
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   173
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   174
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   175
lemma wtl_take:
13006
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
   176
  "wtl_inst_list is G rT cert mxs mxr mpc et pc s = OK s'' \<Longrightarrow>
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   177
   \<exists>s'. wtl_inst_list (take pc' is) G rT cert mxs mxr mpc et pc s = OK s'"
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   178
proof -
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   179
  assume "wtl_inst_list is G rT cert mxs mxr mpc et pc s = OK s''"
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   180
  hence "wtl_inst_list (take pc' is @ drop pc' is) G rT cert mxs mxr mpc et pc s = OK s''"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   181
    by simp
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   182
  thus ?thesis by (auto simp add: wtl_append simp del: append_take_drop_id)
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   183
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   184
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   185
lemma take_Suc:
13006
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
   186
  "\<forall>n. n < length l \<longrightarrow> take (Suc n) l = (take n l)@[l!n]" (is "?P l")
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   187
proof (induct l)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   188
  show "?P []" by simp
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   189
next
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   190
  fix x xs assume IH: "?P xs"  
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   191
  show "?P (x#xs)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   192
  proof (intro strip)
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   193
    fix n assume "n < length (x#xs)"
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   194
    with IH show "take (Suc n) (x # xs) = take n (x # xs) @ [(x # xs) ! n]" 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   195
      by (cases n, auto)
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   196
  qed
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   197
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   198
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   199
lemma wtl_Suc:
13006
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
   200
 "\<lbrakk> wtl_inst_list (take pc is) G rT cert maxs maxr (length is) et 0 s = OK s'; 
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   201
     wtl_cert (is!pc) G rT s' cert maxs maxr (length is) et pc = OK s'';
13006
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
   202
     Suc pc < length is \<rbrakk> \<Longrightarrow>
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   203
  wtl_inst_list (take (Suc pc) is) G rT cert maxs maxr (length is) et 0 s = OK s''" 
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   204
proof -
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   205
  assume "wtl_inst_list (take pc is) G rT cert maxs maxr (length is) et 0 s = OK s'"
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   206
         "wtl_cert (is!pc) G rT s' cert maxs maxr (length is) et pc = OK s''"
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   207
         "Suc pc < length is"
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   208
  hence "take (Suc pc) is = (take pc is)@[is!pc]" by (simp add: take_Suc)
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   209
  thus ?thesis by (simp! add: wtl_append min_def)
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   210
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   211
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   212
lemma wtl_all:
13006
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
   213
  "\<lbrakk> wtl_inst_list is G rT cert maxs maxr (length is) et 0 s \<noteq> Err;
51c5f3f11d16 symbolized
kleing
parents: 12978
diff changeset
   214
      pc < length is \<rbrakk> \<Longrightarrow> 
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   215
   \<exists>s' s''. wtl_inst_list (take pc is) G rT cert maxs maxr (length is) et 0 s = OK s' \<and> 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   216
            wtl_cert (is!pc) G rT s' cert maxs maxr (length is) et pc = OK s''"
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   217
proof -
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   218
  assume all: "wtl_inst_list is G rT cert maxs maxr (length is) et 0 s \<noteq> Err"
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   219
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   220
  assume pc: "pc < length is"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   221
  hence  "0 < length (drop pc is)" by simp
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   222
  then  obtain i r where Cons: "drop pc is = i#r" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   223
    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
   224
  hence "i#r = drop pc is" ..
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   225
  with all have take: 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   226
    "wtl_inst_list (take pc is@i#r) G rT cert maxs maxr (length is) et 0 s \<noteq> Err"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   227
    by simp
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   228
 
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   229
  from pc have "is!pc = drop pc is ! 0" by simp
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   230
  with Cons have "is!pc = i" by simp
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   231
  with take pc show ?thesis by (auto simp add: wtl_append min_def)
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   232
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   233
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9757
diff changeset
   234
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   235
end