src/HOL/MicroJava/BV/LBVSpec.thy
author kleing
Wed, 07 Jun 2000 17:14:19 +0200
changeset 9054 0e48e7d4d4f9
parent 9012 d1bd2144ab5d
child 9183 cd4494dc789a
permissions -rw-r--r--
minor tuning for pdf documents
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
     1
(*  Title:      HOL/MicroJava/BV/BVLightSpec.thy
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 {* Specification of the LBV *}
0e48e7d4d4f9 minor tuning for pdf documents
kleing
parents: 9012
diff changeset
     8
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
     9
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    10
theory LBVSpec = BVSpec:;
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    11
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    12
types
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    13
  certificate       = "state_type option list" 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    14
  class_certificate = "sig \\<Rightarrow> certificate"
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    15
  prog_certificate  = "cname \\<Rightarrow> class_certificate";
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    16
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    17
consts
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    18
 wtl_LS :: "[load_and_store,state_type,state_type,p_count,p_count] \\<Rightarrow> bool";
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    19
primrec
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    20
"wtl_LS (Load idx) s s' max_pc pc =
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    21
 (let (ST,LT) = s
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    22
  in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    23
  pc+1 < max_pc \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    24
  idx < length LT \\<and> 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    25
  (\\<exists>ts. (LT ! idx) = Some ts \\<and>  
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    26
   (ts # ST , LT) = s'))"
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    27
  
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    28
"wtl_LS (Store idx) s s' max_pc pc =
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    29
 (let (ST,LT) = s
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    30
  in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    31
  pc+1 < max_pc \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    32
  idx < length LT \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    33
  (\\<exists>ts ST'. ST = ts # ST' \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    34
   (ST' , LT[idx:=Some ts]) = s'))"
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    35
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    36
"wtl_LS (Bipush i) s s' max_pc pc =
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    37
	(let (ST,LT) = s
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    38
	 in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    39
	 pc+1 < max_pc \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    40
	 ((PrimT Integer) # ST , LT) = s')"
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    41
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    42
"wtl_LS (Aconst_null) s s' max_pc pc =
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    43
	(let (ST,LT) = s
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    44
	 in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    45
	 pc+1 < max_pc \\<and>
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    46
	 (NT # ST , LT) = s')";
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    47
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    48
consts
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    49
 wtl_MO	:: "[manipulate_object,jvm_prog,state_type,state_type,p_count,p_count] \\<Rightarrow> bool"; 
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    50
primrec
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    51
"wtl_MO (Getfield F C) G s s' max_pc pc =
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    52
	(let (ST,LT) = s
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    53
	 in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    54
	 pc+1 < max_pc \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    55
	 is_class G C \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    56
	 (\\<exists>T oT ST'. field (G,C) F = Some(C,T) \\<and>
8390
e5b618f6824e tuned for completeness of LBV
kleing
parents: 8245
diff changeset
    57
                          ST = oT # ST'  \\<and> 
e5b618f6824e tuned for completeness of LBV
kleing
parents: 8245
diff changeset
    58
		                  G \\<turnstile> oT \\<preceq> (Class C) \\<and>
e5b618f6824e tuned for completeness of LBV
kleing
parents: 8245
diff changeset
    59
		          (T # ST' , LT) = s'))"
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    60
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    61
"wtl_MO (Putfield F C) G s s' max_pc pc =
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    62
	(let (ST,LT) = s
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    63
	 in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    64
	 pc+1 < max_pc \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    65
	 is_class G C \\<and> 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    66
	 (\\<exists>T vT oT ST'.
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    67
             field (G,C) F = Some(C,T) \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    68
             ST = vT # oT # ST' \\<and> 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    69
             G \\<turnstile> oT \\<preceq> (Class C) \\<and>
8390
e5b618f6824e tuned for completeness of LBV
kleing
parents: 8245
diff changeset
    70
             G \\<turnstile> vT \\<preceq> T  \\<and>
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    71
             (ST' , LT) = s'))";
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    72
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    73
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    74
consts 
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    75
 wtl_CO	:: "[create_object,jvm_prog,state_type,state_type,p_count,p_count] \\<Rightarrow> bool"; 
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    76
primrec
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    77
"wtl_CO (New C) G s s' max_pc pc =
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    78
	(let (ST,LT) = s
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    79
	 in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    80
	 pc+1 < max_pc \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    81
	 is_class G C \\<and>
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    82
	 ((Class C) # ST , LT) = s')";
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    83
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    84
consts
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    85
 wtl_CH	:: "[check_object,jvm_prog,state_type,state_type,p_count,p_count] \\<Rightarrow> bool"; 
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    86
primrec
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    87
"wtl_CH (Checkcast C) G s s' max_pc pc = 
9054
0e48e7d4d4f9 minor tuning for pdf documents
kleing
parents: 9012
diff changeset
    88
	(let (ST,LT) = s 
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    89
	 in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    90
	 pc+1 < max_pc \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    91
	 is_class G C \\<and> 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    92
	 (\\<exists>rt ST'. ST = RefT rt # ST' \\<and>
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    93
                   (Class C # ST' , LT) = s'))";
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    94
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    95
consts 
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
    96
 wtl_OS	:: "[op_stack,state_type,state_type,p_count,p_count] \\<Rightarrow> bool"; 
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    97
primrec
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    98
"wtl_OS Pop s s' max_pc pc =
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    99
	(let (ST,LT) = s
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   100
	 in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   101
	 \\<exists>ts ST'. pc+1 < max_pc \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   102
		ST = ts # ST' \\<and> 	 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   103
		(ST' , LT) = s')"
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   104
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   105
"wtl_OS Dup s s' max_pc pc =
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   106
	(let (ST,LT) = s
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   107
	 in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   108
	 pc+1 < max_pc \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   109
	 (\\<exists>ts ST'. ST = ts # ST' \\<and> 	 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   110
		   (ts # ts # ST' , LT) = s'))"
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   111
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   112
"wtl_OS Dup_x1 s s' max_pc pc =
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   113
	(let (ST,LT) = s
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   114
	 in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   115
	 pc+1 < max_pc \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   116
	 (\\<exists>ts1 ts2 ST'. ST = ts1 # ts2 # ST' \\<and> 	 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   117
		        (ts1 # ts2 # ts1 # ST' , LT) = s'))"
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   118
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   119
"wtl_OS Dup_x2 s s' max_pc pc =
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   120
	(let (ST,LT) = s
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   121
	 in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   122
	 pc+1 < max_pc \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   123
	 (\\<exists>ts1 ts2 ts3 ST'. ST = ts1 # ts2 # ts3 # ST' \\<and> 	 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   124
		            (ts1 # ts2 # ts3 # ts1 # ST' , LT) = s'))"
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   125
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   126
"wtl_OS Swap s s' max_pc pc =
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   127
	(let (ST,LT) = s
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   128
	 in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   129
	 pc+1 < max_pc \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   130
	 (\\<exists>ts ts' ST'. ST = ts' # ts # ST' \\<and> 	 
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   131
		       (ts # ts' # ST' , LT) = s'))";
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   132
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   133
consts 
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   134
 wtl_BR	:: "[branch,jvm_prog,state_type,state_type,certificate,p_count,p_count] \\<Rightarrow> bool"; 
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   135
primrec
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   136
"wtl_BR (Ifcmpeq branch) G s s' cert max_pc pc =
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   137
	(let (ST,LT) = s
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   138
	 in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   139
	 pc+1 < max_pc \\<and> (nat(int pc+branch)) < max_pc \\<and> 
8390
e5b618f6824e tuned for completeness of LBV
kleing
parents: 8245
diff changeset
   140
	 (\\<exists>ts ts' ST'. ST = ts # ts' # ST' \\<and>
e5b618f6824e tuned for completeness of LBV
kleing
parents: 8245
diff changeset
   141
          ((\\<exists>p. ts = PrimT p \\<and> ts' = PrimT p) \\<or>
e5b618f6824e tuned for completeness of LBV
kleing
parents: 8245
diff changeset
   142
           (\\<exists>r r'. ts = RefT r \\<and> ts' = RefT r')) \\<and>
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   143
		       ((ST' , LT) = s') \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   144
            cert ! (nat(int pc+branch)) \\<noteq> None \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   145
		       G \\<turnstile> (ST' , LT) <=s the (cert ! (nat(int pc+branch)))))"
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   146
  
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   147
"wtl_BR (Goto branch) G s s' cert max_pc pc =
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   148
	((let (ST,LT) = s
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   149
	 in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   150
	 (nat(int pc+branch)) < max_pc \\<and> cert ! (nat(int pc+branch)) \\<noteq> None \\<and>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   151
	 G \\<turnstile> (ST , LT) <=s the (cert ! (nat(int pc+branch)))) \\<and>
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   152
   (cert ! (pc+1) = Some s'))";
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   153
  
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   154
consts
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   155
 wtl_MI :: "[meth_inv,jvm_prog,state_type,state_type,certificate,p_count,p_count] \\<Rightarrow> bool"; 
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   156
primrec
8390
e5b618f6824e tuned for completeness of LBV
kleing
parents: 8245
diff changeset
   157
"wtl_MI (Invoke mn fpTs) G s s' cert max_pc pc =
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   158
	(let (ST,LT) = s
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   159
	 in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   160
         pc+1 < max_pc \\<and>
8390
e5b618f6824e tuned for completeness of LBV
kleing
parents: 8245
diff changeset
   161
         (\\<exists>apTs X ST'. ST = (rev apTs) @ (X # ST') \\<and>
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   162
         length apTs = length fpTs \\<and>
8390
e5b618f6824e tuned for completeness of LBV
kleing
parents: 8245
diff changeset
   163
         (\\<exists>s''. cert ! (pc+1) = Some s'' \\<and> 
e5b618f6824e tuned for completeness of LBV
kleing
parents: 8245
diff changeset
   164
         ((s'' = s' \\<and> X = NT) \\<or>
e5b618f6824e tuned for completeness of LBV
kleing
parents: 8245
diff changeset
   165
          ((G \\<turnstile> s' <=s s'') \\<and> (\\<exists>C. X = Class C \\<and>  
e5b618f6824e tuned for completeness of LBV
kleing
parents: 8245
diff changeset
   166
          (\\<forall>(aT,fT)\\<in>set(zip apTs fpTs). G \\<turnstile> aT \\<preceq> fT) \\<and> 
e5b618f6824e tuned for completeness of LBV
kleing
parents: 8245
diff changeset
   167
          (\\<exists>D rT b. method (G,C) (mn,fpTs) = Some(D,rT,b) \\<and>
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   168
          (rT # ST' , LT) = s')))))))";
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   169
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   170
consts
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   171
 wtl_MR	:: "[meth_ret,jvm_prog,ty,state_type,state_type,certificate,p_count,p_count] \\<Rightarrow> bool";
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   172
primrec
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   173
  "wtl_MR Return G rT s s' cert max_pc pc = 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   174
	((let (ST,LT) = s
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   175
	 in
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   176
	 (\\<exists>T ST'. ST = T # ST' \\<and> G \\<turnstile> T \\<preceq> rT)) \\<and>
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   177
   (cert ! (pc+1) = Some s'))";
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   178
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   179
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   180
consts 
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   181
 wtl_inst :: "[instr,jvm_prog,ty,state_type,state_type,certificate,p_count,p_count] \\<Rightarrow> bool";
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   182
 primrec
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   183
 "wtl_inst (LS ins) G rT s s' cert max_pc pc = wtl_LS ins s s' max_pc pc"
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   184
 "wtl_inst (CO ins) G rT s s' cert max_pc pc = wtl_CO ins G s s' max_pc pc"
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   185
 "wtl_inst (MO ins) G rT s s' cert max_pc pc = wtl_MO ins G s s' max_pc pc"
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   186
 "wtl_inst (CH ins) G rT s s' cert max_pc pc = wtl_CH ins G s s' max_pc pc"
8390
e5b618f6824e tuned for completeness of LBV
kleing
parents: 8245
diff changeset
   187
 "wtl_inst (MI ins) G rT s s' cert max_pc pc = wtl_MI ins G s s' cert max_pc pc"
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   188
 "wtl_inst (MR ins) G rT s s' cert max_pc pc = wtl_MR ins G rT s s' cert max_pc pc"
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   189
 "wtl_inst (OS ins) G rT s s' cert max_pc pc = wtl_OS ins s s' max_pc pc"
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   190
 "wtl_inst (BR ins) G rT s s' cert max_pc pc = wtl_BR ins G s s' cert max_pc pc";
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   191
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   192
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   193
constdefs
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   194
wtl_inst_option :: "[instr,jvm_prog,ty,state_type,state_type,certificate,p_count,p_count] \\<Rightarrow> bool"
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   195
"wtl_inst_option i G rT s0 s1 cert max_pc pc \\<equiv>
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   196
     (case cert!pc of 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   197
          None     \\<Rightarrow> wtl_inst i G rT s0 s1 cert max_pc pc
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   198
        | Some s0' \\<Rightarrow> (G \\<turnstile> s0 <=s s0') \\<and>
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   199
                      wtl_inst i G rT s0' s1 cert max_pc pc)";
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   200
  
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   201
consts
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   202
 wtl_inst_list :: "[instr list,jvm_prog,ty,state_type,state_type,certificate,p_count,p_count] \\<Rightarrow> bool";
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   203
primrec
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   204
  "wtl_inst_list [] G rT s0 s2 cert max_pc pc = (s0 = s2)"
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   205
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   206
  "wtl_inst_list (instr#is) G rT s0 s2 cert max_pc pc =
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   207
     (\\<exists>s1. wtl_inst_option instr G rT s0 s1 cert max_pc pc \\<and> 
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   208
           wtl_inst_list is G rT s1 s2 cert max_pc (pc+1))";
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   209
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   210
constdefs
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   211
 wtl_method :: "[jvm_prog,cname,ty list,ty,nat,instr list,certificate] \\<Rightarrow> bool" 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   212
 "wtl_method G C pTs rT mxl ins cert \\<equiv> 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   213
	let max_pc = length ins 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   214
        in 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   215
	0 < max_pc \\<and>  
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   216
        (\\<exists>s2. wtl_inst_list ins G rT 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   217
                            ([],(Some(Class C))#((map Some pTs))@(replicate mxl None))
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   218
                            s2 cert max_pc 0)"
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   219
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   220
 wtl_jvm_prog :: "[jvm_prog,prog_certificate] \\<Rightarrow> bool"
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   221
 "wtl_jvm_prog G cert \\<equiv> 
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   222
    wf_prog (\\<lambda>G C (sig,rT,maxl,b). 
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   223
               wtl_method G C (snd sig) rT maxl b (cert C sig)) G"; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   224
9054
0e48e7d4d4f9 minor tuning for pdf documents
kleing
parents: 9012
diff changeset
   225
text {* \medskip *}
0e48e7d4d4f9 minor tuning for pdf documents
kleing
parents: 9012
diff changeset
   226
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   227
lemma rev_eq: "\\<lbrakk>length a = n; length x = n; rev a @ b # c = rev x @ y # z\\<rbrakk> \\<Longrightarrow> a = x \\<and> b = y \\<and> c = z";
9054
0e48e7d4d4f9 minor tuning for pdf documents
kleing
parents: 9012
diff changeset
   228
by auto;
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   229
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   230
lemma wtl_inst_unique: 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   231
"wtl_inst i G rT s0 s1 cert max_pc pc \\<longrightarrow>
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   232
 wtl_inst i G rT s0 s1' cert max_pc pc \\<longrightarrow> s1 = s1'" (is "?P i");
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   233
proof (induct i);
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   234
case LS;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   235
 show "?P (LS load_and_store)"; by (induct load_and_store) auto;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   236
case CO;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   237
 show "?P (CO create_object)"; by (induct create_object) auto;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   238
case MO;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   239
 show "?P (MO manipulate_object)"; by (induct manipulate_object) auto;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   240
case CH;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   241
 show "?P (CH check_object)"; by (induct check_object) auto;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   242
case MI;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   243
 show "?P (MI meth_inv)"; proof (induct meth_inv);
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   244
 case Invoke;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   245
  have "\\<exists>x y. s0 = (x,y)"; by (simp);
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   246
  thus "wtl_inst (MI (Invoke mname list)) G rT s0 s1 cert max_pc pc \\<longrightarrow>
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   247
        wtl_inst (MI (Invoke mname list)) G rT s0 s1' cert max_pc pc \\<longrightarrow>
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   248
        s1 = s1'";
9054
0e48e7d4d4f9 minor tuning for pdf documents
kleing
parents: 9012
diff changeset
   249
  proof elim;
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   250
    apply_end(clarsimp_tac, drule rev_eq, assumption+);
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   251
  qed auto;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   252
 qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   253
case MR;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   254
 show "?P (MR meth_ret)"; by (induct meth_ret) auto;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   255
case OS; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   256
 show "?P (OS op_stack)"; by (induct op_stack) auto;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   257
case BR;  
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   258
 show "?P (BR branch)"; by (induct branch) auto;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   259
qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   260
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   261
lemma wtl_inst_option_unique:
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   262
"\\<lbrakk>wtl_inst_option i G rT s0 s1 cert max_pc pc; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   263
  wtl_inst_option i G rT s0 s1' cert max_pc pc\\<rbrakk> \\<Longrightarrow> s1 = s1'";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   264
by (cases "cert!pc") (auto simp add: wtl_inst_unique wtl_inst_option_def);
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   265
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   266
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   267
lemma wtl_inst_list_unique: 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   268
"\\<forall> s0 pc. wtl_inst_list is G rT s0 s1 cert max_pc pc \\<longrightarrow> 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   269
 wtl_inst_list is G rT s0 s1' cert max_pc pc \\<longrightarrow> s1=s1'" (is "?P is");
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   270
proof (induct "?P" "is"); 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   271
  case Nil;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   272
  show "?P []"; by simp;
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   273
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   274
  case Cons;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   275
  show "?P (a # list)"; 
9054
0e48e7d4d4f9 minor tuning for pdf documents
kleing
parents: 9012
diff changeset
   276
  proof intro;
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   277
    fix s0; fix pc; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   278
    let "?o s0 s1" = "wtl_inst_option a G rT s0 s1 cert max_pc pc";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   279
    let "?l l s1 s2 pc" = "wtl_inst_list l G rT s1 s2 cert max_pc pc"; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   280
    assume a: "?l (a#list) s0 s1 pc";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   281
    assume b: "?l (a#list) s0 s1' pc";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   282
    with a;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   283
    show "s1 = s1'";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   284
      obtain s s' where   "?o s0 s" "?o s0 s'"
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   285
                  and l:  "?l list s s1 (Suc pc)"
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   286
                  and l': "?l list s' s1' (Suc pc)"; by auto;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   287
      have "s=s'"; by(rule wtl_inst_option_unique);
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   288
      with l l' Cons;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   289
      show ?thesis; by blast;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   290
    qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   291
  qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   292
qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   293
        
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   294
lemma wtl_partial:
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   295
"\\<forall> pc' pc s. 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   296
   wtl_inst_list is G rT s s' cert mpc pc \\<longrightarrow> \
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   297
   pc' < length is \\<longrightarrow> \
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   298
   (\\<exists> a b s1. a @ b = is \\<and> length a = pc' \\<and> \
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   299
              wtl_inst_list a G rT s  s1 cert mpc pc \\<and> \
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   300
              wtl_inst_list b G rT s1 s' cert mpc (pc+length a))" (is "?P is");
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   301
proof (induct "?P" "is");
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   302
  case Nil;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   303
  show "?P []"; by auto;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   304
  
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   305
  case Cons;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   306
  show "?P (a#list)";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   307
  proof (intro allI impI);
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   308
    fix pc' pc s;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   309
    assume length: "pc' < length (a # list)";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   310
    assume wtl:    "wtl_inst_list (a # list) G rT s s' cert mpc pc";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   311
    show "\\<exists> a' b s1. 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   312
            a' @ b = a#list \\<and> length a' = pc' \\<and> \
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   313
            wtl_inst_list a' G rT s  s1 cert mpc pc \\<and> \
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   314
            wtl_inst_list b G rT s1 s' cert mpc (pc+length a')"
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   315
        (is "\\<exists> a b s1. ?E a b s1");
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   316
    proof (cases "pc'");
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   317
      case 0;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   318
      with wtl;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   319
      have "?E [] (a#list) s"; by simp;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   320
      thus ?thesis; by blast;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   321
    next;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   322
      case Suc;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   323
      with wtl;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   324
      show ?thesis; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   325
        obtain s0 where wtlSuc: "wtl_inst_list list G rT s0 s' cert mpc (Suc pc)"
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   326
                  and   wtlOpt: "wtl_inst_option a G rT s s0 cert mpc pc"; by auto;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   327
        from Cons;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   328
        show ?thesis; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   329
          obtain a' b s1' 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   330
            where "a' @ b = list" "length a' = nat"
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   331
            and w:"wtl_inst_list a' G rT s0 s1' cert mpc (Suc pc)"
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   332
            and   "wtl_inst_list b G rT s1' s' cert mpc (Suc pc + length a')"; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   333
          proof (elim allE impE);
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   334
            from length Suc; show "nat < length list"; by simp; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   335
            from wtlSuc;     show "wtl_inst_list list G rT s0 s' cert mpc (Suc pc)"; .; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   336
          qed (elim exE conjE, auto);
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   337
          with Suc wtlOpt;          
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   338
          have "?E (a#a') b s1'"; by (auto simp del: split_paired_Ex);   
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   339
          thus ?thesis; by blast;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   340
        qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   341
      qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   342
    qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   343
  qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   344
qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   345
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   346
lemma "wtl_append1":
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   347
"\\<lbrakk>wtl_inst_list x G rT s0 s1 cert (length (x@y)) 0; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   348
  wtl_inst_list y G rT s1 s2 cert (length (x@y)) (length x)\\<rbrakk> \\<Longrightarrow>
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   349
  wtl_inst_list (x@y) G rT s0 s2 cert (length (x@y)) 0";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   350
proof -;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   351
  assume w:
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   352
    "wtl_inst_list x G rT s0 s1 cert (length (x@y)) 0"
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   353
    "wtl_inst_list y G rT s1 s2 cert (length (x@y)) (length x)";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   354
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   355
have
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   356
"\\<forall> pc s0. 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   357
 wtl_inst_list x G rT s0 s1 cert (pc+length (x@y)) pc \\<longrightarrow> 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   358
 wtl_inst_list y G rT s1 s2 cert (pc+length (x@y)) (pc+length x) \\<longrightarrow>
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   359
 wtl_inst_list (x@y) G rT s0 s2 cert (pc+length (x@y)) pc" (is "?P x");
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   360
proof (induct "?P" "x");
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   361
  case Nil;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   362
  show "?P []"; by simp;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   363
next;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   364
  case Cons;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   365
  show "?P (a#list)"; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   366
  proof intro;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   367
    fix pc s0;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   368
    assume y: 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   369
      "wtl_inst_list y G rT s1 s2 cert (pc + length ((a # list) @ y)) (pc + length (a # list))";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   370
    assume al: 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   371
      "wtl_inst_list (a # list) G rT s0 s1 cert (pc + length ((a # list) @ y)) pc";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   372
    thus "wtl_inst_list ((a # list) @ y) G rT s0 s2 cert (pc + length ((a # list) @ y)) pc";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   373
      obtain s' where 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   374
       a: "wtl_inst_option a G rT s0 s' cert (Suc pc + length (list@y)) pc" and
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   375
       l: "wtl_inst_list list G rT s' s1 cert (Suc pc + length (list@y)) (Suc pc)"; by auto;      
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   376
      with y Cons;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   377
      have "wtl_inst_list (list @ y) G rT s' s2 cert (Suc pc + length (list @ y)) (Suc pc)";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   378
        by elim (assumption, simp+);
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   379
      with a;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   380
      show ?thesis; by (auto simp del: split_paired_Ex);
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   381
    qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   382
  qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   383
qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   384
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   385
  with w;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   386
  show ?thesis; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   387
  proof elim;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   388
    from w; show "wtl_inst_list x G rT s0 s1 cert (0+length (x @ y)) 0"; by simp;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   389
  qed simp+;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   390
qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   391
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   392
lemma wtl_cons_appendl:
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   393
"\\<lbrakk>wtl_inst_list a G rT s0 s1 cert (length (a@i#b)) 0; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   394
  wtl_inst_option i G rT s1 s2 cert (length (a@i#b)) (length a); 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   395
  wtl_inst_list b G rT s2 s3 cert (length (a@i#b)) (Suc (length a))\\<rbrakk> \\<Longrightarrow> 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   396
  wtl_inst_list (a@i#b) G rT s0 s3 cert (length (a@i#b)) 0";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   397
proof -;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   398
  assume a: "wtl_inst_list a G rT s0 s1 cert (length (a@i#b)) 0";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   399
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   400
  assume "wtl_inst_option i G rT s1 s2 cert (length (a@i#b)) (length a)"
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   401
         "wtl_inst_list b G rT s2 s3 cert (length (a@i#b)) (Suc (length a))";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   402
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   403
  hence "wtl_inst_list (i#b) G rT s1 s3 cert (length (a@i#b)) (length a)";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   404
    by (auto simp del: split_paired_Ex);
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   405
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   406
  with a;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   407
  show ?thesis; by (rule wtl_append1);
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   408
qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   409
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   410
lemma "wtl_append":
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   411
"\\<lbrakk>wtl_inst_list a G rT s0 s1 cert (length (a@i#b)) 0; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   412
  wtl_inst_option i G rT s1 s2 cert (length (a@i#b)) (length a); 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   413
  wtl_inst_list b G rT s2 s3 cert (length (a@i#b)) (Suc (length a))\\<rbrakk> \\<Longrightarrow> 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   414
  wtl_inst_list (a@[i]) G rT s0 s2 cert (length (a@i#b)) 0";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   415
proof -;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   416
  assume a: "wtl_inst_list a G rT s0 s1 cert (length (a@i#b)) 0";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   417
  assume i: "wtl_inst_option i G rT s1 s2 cert (length (a@i#b)) (length a)";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   418
  assume b: "wtl_inst_list b G rT s2 s3 cert (length (a@i#b)) (Suc (length a))";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   419
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   420
  have "\\<forall> s0 pc. wtl_inst_list a G rT s0 s1 cert (pc+length (a@i#b)) pc \\<longrightarrow> 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   421
        wtl_inst_option i G rT s1 s2 cert (pc+length (a@i#b)) (pc + length a) \\<longrightarrow> 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   422
        wtl_inst_list b G rT s2 s3 cert (pc+length (a@i#b)) (Suc pc + length a) \\<longrightarrow> 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   423
          wtl_inst_list (a@[i]) G rT s0 s2 cert (pc+length (a@i#b)) pc" (is "?P a");
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   424
  proof (induct "?P" "a");
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   425
    case Nil;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   426
    show "?P []"; by (simp del: split_paired_Ex);
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   427
    case Cons;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   428
    show "?P (a#list)" (is "\\<forall>s0 pc. ?x s0 pc \\<longrightarrow> ?y s0 pc \\<longrightarrow> ?z s0 pc \\<longrightarrow> ?p s0 pc"); 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   429
    proof intro;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   430
      fix s0 pc;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   431
      assume y: "?y s0 pc";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   432
      assume z: "?z s0 pc";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   433
      assume "?x s0 pc";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   434
      thus "?p s0 pc";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   435
        obtain s0' where opt: "wtl_inst_option a G rT s0 s0' cert (pc + length ((a # list) @ i # b)) pc"
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   436
                    and list: "wtl_inst_list list G rT s0' s1 cert (Suc pc + length (list @ i # b)) (Suc pc)";
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   437
          by (auto simp del: split_paired_Ex);
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   438
        with y z Cons;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   439
        have "wtl_inst_list (list @ [i]) G rT s0' s2 cert (Suc pc + length (list @ i # b)) (Suc pc)"; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   440
        proof elim; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   441
          from list; show "wtl_inst_list list G rT s0' s1 cert (Suc pc + length (list @ i # b)) (Suc pc)"; .;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   442
        qed auto;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   443
        with opt;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   444
        show ?thesis; by (auto simp del: split_paired_Ex);
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   445
      qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   446
    qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   447
  qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   448
  with a i b;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   449
  show ?thesis; 
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   450
  proof elim;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   451
    from a; show "wtl_inst_list a G rT s0 s1 cert (0+length (a@i#b)) 0"; by simp;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   452
  qed auto;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   453
qed;
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   454
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   455
end;