src/HOL/MicroJava/DFA/LBVSpec.thy
author wenzelm
Sat, 18 Aug 2018 13:52:12 +0200
changeset 68761 8bb51b3de39f
parent 68386 98cf1c823c48
child 80914 d97fdabd9e2b
permissions -rw-r--r--
trim nodes_status: avoid potential memory leak;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42150
b0c0638c4aad tuned headers;
wenzelm
parents: 35416
diff changeset
     1
(*  Title:      HOL/MicroJava/DFA/LBVSpec.thy
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     2
    Author:     Gerwin Klein
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     3
    Copyright   1999 Technische Universitaet Muenchen
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     4
*)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     5
61361
8b5f00202e1a isabelle update_cartouches;
wenzelm
parents: 58886
diff changeset
     6
section \<open>The Lightweight Bytecode Verifier\<close>
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     7
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     8
theory LBVSpec
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
     9
imports SemilatAlg Opt
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    10
begin
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    11
42463
f270e3e18be5 modernized specifications;
wenzelm
parents: 42150
diff changeset
    12
type_synonym 's certificate = "'s list"   
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    13
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    14
primrec merge :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> nat \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's \<Rightarrow> 's" where
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    15
  "merge cert f r T pc []     x = x"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    16
| "merge cert f r T pc (s#ss) x = merge cert f r T pc ss (let (pc',s') = s in 
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    17
                                  if pc'=pc+1 then s' +_f x
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    18
                                  else if s' <=_r (cert!pc') then x
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    19
                                  else T)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    20
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    21
definition wtl_inst :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow>
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    22
             's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's" where
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    23
"wtl_inst cert f r T step pc s \<equiv> merge cert f r T pc (step pc s) (cert!(pc+1))"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    24
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    25
definition wtl_cert :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow>
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    26
             's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's" where
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    27
"wtl_cert cert f r T B step pc s \<equiv>
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    28
  if cert!pc = B then 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    29
    wtl_inst cert f r T step pc s
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    30
  else
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    31
    if s <=_r (cert!pc) then wtl_inst cert f r T step pc (cert!pc) else T"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    32
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    33
primrec wtl_inst_list :: "'a list \<Rightarrow> 's certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow>
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    34
                  's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's" where
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    35
  "wtl_inst_list []     cert f r T B step pc s = s"
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    36
| "wtl_inst_list (i#is) cert f r T B step pc s = 
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    37
    (let s' = wtl_cert cert f r T B step pc s in
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    38
      if s' = T \<or> s = T then T else wtl_inst_list is cert f r T B step (pc+1) s')"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    39
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    40
definition cert_ok :: "'s certificate \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow> 's set \<Rightarrow> bool" where
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    41
  "cert_ok cert n T B A \<equiv> (\<forall>i < n. cert!i \<in> A \<and> cert!i \<noteq> T) \<and> (cert!n = B)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    42
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33954
diff changeset
    43
definition bottom :: "'a ord \<Rightarrow> 'a \<Rightarrow> bool" where
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    44
  "bottom r B \<equiv> \<forall>x. B <=_r x"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    45
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    46
locale lbv = Semilat +
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    47
  fixes T :: "'a" ("\<top>") 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    48
  fixes B :: "'a" ("\<bottom>") 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    49
  fixes step :: "'a step_type" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    50
  assumes top: "top r \<top>"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    51
  assumes T_A: "\<top> \<in> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    52
  assumes bot: "bottom r \<bottom>" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    53
  assumes B_A: "\<bottom> \<in> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    54
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    55
  fixes merge :: "'a certificate \<Rightarrow> nat \<Rightarrow> (nat \<times> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    56
  defines mrg_def: "merge cert \<equiv> LBVSpec.merge cert f r \<top>"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    57
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    58
  fixes wti :: "'a certificate \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    59
  defines wti_def: "wti cert \<equiv> wtl_inst cert f r \<top> step"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    60
 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    61
  fixes wtc :: "'a certificate \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    62
  defines wtc_def: "wtc cert \<equiv> wtl_cert cert f r \<top> \<bottom> step"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    63
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    64
  fixes wtl :: "'b list \<Rightarrow> 'a certificate \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    65
  defines wtl_def: "wtl ins cert \<equiv> wtl_inst_list ins cert f r \<top> \<bottom> step"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    66
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    67
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    68
lemma (in lbv) wti:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    69
  "wti c pc s \<equiv> merge c pc (step pc s) (c!(pc+1))"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    70
  by (simp add: wti_def mrg_def wtl_inst_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    71
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    72
lemma (in lbv) wtc: 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    73
  "wtc c pc s \<equiv> if c!pc = \<bottom> then wti c pc s else if s <=_r c!pc then wti c pc (c!pc) else \<top>"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    74
  by (unfold wtc_def wti_def wtl_cert_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    75
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    76
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    77
lemma cert_okD1 [intro?]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    78
  "cert_ok c n T B A \<Longrightarrow> pc < n \<Longrightarrow> c!pc \<in> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    79
  by (unfold cert_ok_def) fast
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    80
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    81
lemma cert_okD2 [intro?]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    82
  "cert_ok c n T B A \<Longrightarrow> c!n = B"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    83
  by (simp add: cert_ok_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    84
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    85
lemma cert_okD3 [intro?]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    86
  "cert_ok c n T B A \<Longrightarrow> B \<in> A \<Longrightarrow> pc < n \<Longrightarrow> c!Suc pc \<in> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    87
  by (drule Suc_leI) (auto simp add: le_eq_less_or_eq dest: cert_okD1 cert_okD2)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    88
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    89
lemma cert_okD4 [intro?]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    90
  "cert_ok c n T B A \<Longrightarrow> pc < n \<Longrightarrow> c!pc \<noteq> T"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    91
  by (simp add: cert_ok_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    92
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    93
declare Let_def [simp]
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    94
58886
8a6cac7c7247 modernized header;
wenzelm
parents: 54863
diff changeset
    95
subsection "more semilattice lemmas"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    96
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    97
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    98
lemma (in lbv) sup_top [simp, elim]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
    99
  assumes x: "x \<in> A" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   100
  shows "x +_f \<top> = \<top>"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   101
proof -
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   102
  from top have "x +_f \<top> <=_r \<top>" ..
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   103
  moreover from x T_A have "\<top> <=_r x +_f \<top>" ..
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   104
  ultimately show ?thesis ..
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   105
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   106
  
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   107
lemma (in lbv) plusplussup_top [simp, elim]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   108
  "set xs \<subseteq> A \<Longrightarrow> xs ++_f \<top> = \<top>"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   109
  by (induct xs) auto
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   110
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   111
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   112
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   113
lemma (in Semilat) pp_ub1':
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   114
  assumes S: "snd`set S \<subseteq> A" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   115
  assumes y: "y \<in> A" and ab: "(a, b) \<in> set S" 
68386
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
   116
  shows "b <=_r map snd [(p', t') \<leftarrow> S . p' = a] ++_f y"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   117
proof -
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   118
  from S have "\<forall>(x,y) \<in> set S. y \<in> A" by auto
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   119
  with semilat y ab show ?thesis by - (rule ub1')
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   120
qed 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   121
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   122
lemma (in lbv) bottom_le [simp, intro]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   123
  "\<bottom> <=_r x"
63258
576fb8068ba6 tuned proofs;
wenzelm
parents: 62390
diff changeset
   124
  using bot by (simp add: bottom_def)
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   125
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   126
lemma (in lbv) le_bottom [simp]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   127
  "x <=_r \<bottom> = (x = \<bottom>)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   128
  by (blast intro: antisym_r)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   129
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   130
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   131
58886
8a6cac7c7247 modernized header;
wenzelm
parents: 54863
diff changeset
   132
subsection "merge"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   133
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   134
lemma (in lbv) merge_Nil [simp]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   135
  "merge c pc [] x = x" by (simp add: mrg_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   136
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   137
lemma (in lbv) merge_Cons [simp]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   138
  "merge c pc (l#ls) x = merge c pc ls (if fst l=pc+1 then snd l +_f x
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   139
                                        else if snd l <=_r (c!fst l) then x
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   140
                                        else \<top>)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   141
  by (simp add: mrg_def split_beta)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   142
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   143
lemma (in lbv) merge_Err [simp]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   144
  "snd`set ss \<subseteq> A \<Longrightarrow> merge c pc ss \<top> = \<top>"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   145
  by (induct ss) auto
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   146
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   147
lemma (in lbv) merge_not_top:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   148
  "\<And>x. snd`set ss \<subseteq> A \<Longrightarrow> merge c pc ss x \<noteq> \<top> \<Longrightarrow> 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   149
  \<forall>(pc',s') \<in> set ss. (pc' \<noteq> pc+1 \<longrightarrow> s' <=_r (c!pc'))"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   150
  (is "\<And>x. ?set ss \<Longrightarrow> ?merge ss x \<Longrightarrow> ?P ss")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   151
proof (induct ss)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   152
  show "?P []" by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   153
next
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   154
  fix x ls l
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   155
  assume "?set (l#ls)" then obtain set: "snd`set ls \<subseteq> A" by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   156
  assume merge: "?merge (l#ls) x" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   157
  moreover
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   158
  obtain pc' s' where l: "l = (pc',s')" by (cases l)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   159
  ultimately
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   160
  obtain x' where merge': "?merge ls x'" by simp 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   161
  assume "\<And>x. ?set ls \<Longrightarrow> ?merge ls x \<Longrightarrow> ?P ls" hence "?P ls" using set merge' .
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   162
  moreover
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   163
  from merge set
62390
842917225d56 more canonical names
nipkow
parents: 61361
diff changeset
   164
  have "pc' \<noteq> pc+1 \<longrightarrow> s' <=_r (c!pc')" by (simp add: l split: if_split_asm)
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   165
  ultimately
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   166
  show "?P (l#ls)" by (simp add: l)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   167
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   168
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   169
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   170
lemma (in lbv) merge_def:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   171
  shows 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   172
  "\<And>x. x \<in> A \<Longrightarrow> snd`set ss \<subseteq> A \<Longrightarrow>
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   173
  merge c pc ss x = 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   174
  (if \<forall>(pc',s') \<in> set ss. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc' then
68386
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
   175
    map snd [(p',t') \<leftarrow> ss. p'=pc+1] ++_f x
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   176
  else \<top>)" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   177
  (is "\<And>x. _ \<Longrightarrow> _ \<Longrightarrow> ?merge ss x = ?if ss x" is "\<And>x. _ \<Longrightarrow> _ \<Longrightarrow> ?P ss x")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   178
proof (induct ss)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   179
  fix x show "?P [] x" by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   180
next 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   181
  fix x assume x: "x \<in> A" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   182
  fix l::"nat \<times> 'a" and ls  
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   183
  assume "snd`set (l#ls) \<subseteq> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   184
  then obtain l: "snd l \<in> A" and ls: "snd`set ls \<subseteq> A" by auto
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   185
  assume "\<And>x. x \<in> A \<Longrightarrow> snd`set ls \<subseteq> A \<Longrightarrow> ?P ls x" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   186
  hence IH: "\<And>x. x \<in> A \<Longrightarrow> ?P ls x" using ls by iprover
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   187
  obtain pc' s' where [simp]: "l = (pc',s')" by (cases l)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   188
  hence "?merge (l#ls) x = ?merge ls 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   189
    (if pc'=pc+1 then s' +_f x else if s' <=_r c!pc' then x else \<top>)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   190
    (is "?merge (l#ls) x = ?merge ls ?if'")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   191
    by simp 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   192
  also have "\<dots> = ?if ls ?if'" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   193
  proof -
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   194
    from l have "s' \<in> A" by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   195
    with x have "s' +_f x \<in> A" by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   196
    with x T_A have "?if' \<in> A" by auto
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   197
    hence "?P ls ?if'" by (rule IH) thus ?thesis by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   198
  qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   199
  also have "\<dots> = ?if (l#ls) x"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   200
    proof (cases "\<forall>(pc', s')\<in>set (l#ls). pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   201
      case True
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   202
      hence "\<forall>(pc', s')\<in>set ls. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'" by auto
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   203
      moreover
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   204
      from True have 
68386
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
   205
        "map snd [(p',t')\<leftarrow>ls . p'=pc+1] ++_f ?if' = 
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
   206
        (map snd [(p',t')\<leftarrow>l#ls . p'=pc+1] ++_f x)"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   207
        by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   208
      ultimately
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   209
      show ?thesis using True by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   210
    next
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   211
      case False 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   212
      moreover
68386
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
   213
      from ls have "set (map snd [(p', t')\<leftarrow>ls . p' = Suc pc]) \<subseteq> A" by auto
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   214
      ultimately show ?thesis by auto
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   215
    qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   216
  finally show "?P (l#ls) x" .
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   217
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   218
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   219
lemma (in lbv) merge_not_top_s:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   220
  assumes x:  "x \<in> A" and ss: "snd`set ss \<subseteq> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   221
  assumes m:  "merge c pc ss x \<noteq> \<top>"
68386
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
   222
  shows "merge c pc ss x = (map snd [(p',t') \<leftarrow> ss. p'=pc+1] ++_f x)"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   223
proof -
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   224
  from ss m have "\<forall>(pc',s') \<in> set ss. (pc' \<noteq> pc+1 \<longrightarrow> s' <=_r c!pc')" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   225
    by (rule merge_not_top)
62390
842917225d56 more canonical names
nipkow
parents: 61361
diff changeset
   226
  with x ss m show ?thesis by - (drule merge_def, auto split: if_split_asm)
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   227
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   228
58886
8a6cac7c7247 modernized header;
wenzelm
parents: 54863
diff changeset
   229
subsection "wtl-inst-list"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   230
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   231
lemmas [iff] = not_Err_eq
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   232
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   233
lemma (in lbv) wtl_Nil [simp]: "wtl [] c pc s = s" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   234
  by (simp add: wtl_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   235
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   236
lemma (in lbv) wtl_Cons [simp]: 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   237
  "wtl (i#is) c pc s = 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   238
  (let s' = wtc c pc s in if s' = \<top> \<or> s = \<top> then \<top> else wtl is c (pc+1) s')"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   239
  by (simp add: wtl_def wtc_def)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   240
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   241
lemma (in lbv) wtl_Cons_not_top:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   242
  "wtl (i#is) c pc s \<noteq> \<top> = 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   243
  (wtc c pc s \<noteq> \<top> \<and> s \<noteq> T \<and> wtl is c (pc+1) (wtc c pc s) \<noteq> \<top>)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   244
  by (auto simp del: split_paired_Ex)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   245
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   246
lemma (in lbv) wtl_top [simp]:  "wtl ls c pc \<top> = \<top>"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   247
  by (cases ls) auto
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   248
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   249
lemma (in lbv) wtl_not_top:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   250
  "wtl ls c pc s \<noteq> \<top> \<Longrightarrow> s \<noteq> \<top>"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   251
  by (cases "s=\<top>") auto
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   252
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   253
lemma (in lbv) wtl_append [simp]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   254
  "\<And>pc s. wtl (a@b) c pc s = wtl b c (pc+length a) (wtl a c pc s)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   255
  by (induct a) auto
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   256
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   257
lemma (in lbv) wtl_take:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   258
  "wtl is c pc s \<noteq> \<top> \<Longrightarrow> wtl (take pc' is) c pc s \<noteq> \<top>"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   259
  (is "?wtl is \<noteq> _ \<Longrightarrow> _")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   260
proof -
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   261
  assume "?wtl is \<noteq> \<top>"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   262
  hence "?wtl (take pc' is @ drop pc' is) \<noteq> \<top>" by simp  
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   263
  thus ?thesis by (auto dest!: wtl_not_top simp del: append_take_drop_id)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   264
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   265
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   266
lemma take_Suc:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   267
  "\<forall>n. n < length l \<longrightarrow> take (Suc n) l = (take n l)@[l!n]" (is "?P l")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   268
proof (induct l)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   269
  show "?P []" by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   270
next
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   271
  fix x xs assume IH: "?P xs"  
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   272
  show "?P (x#xs)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   273
  proof (intro strip)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   274
    fix n assume "n < length (x#xs)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   275
    with IH show "take (Suc n) (x # xs) = take n (x # xs) @ [(x # xs) ! n]" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   276
      by (cases n, auto)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   277
  qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   278
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   279
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   280
lemma (in lbv) wtl_Suc:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   281
  assumes suc: "pc+1 < length is"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   282
  assumes wtl: "wtl (take pc is) c 0 s \<noteq> \<top>"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   283
  shows "wtl (take (pc+1) is) c 0 s = wtc c pc (wtl (take pc is) c 0 s)"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   284
proof -
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   285
  from suc have "take (pc+1) is=(take pc is)@[is!pc]" by (simp add: take_Suc)
54863
82acc20ded73 prefer more canonical names for lemmas on min/max
haftmann
parents: 42463
diff changeset
   286
  with suc wtl show ?thesis by (simp add: min.absorb2)
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   287
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   288
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   289
lemma (in lbv) wtl_all:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   290
  assumes all: "wtl is c 0 s \<noteq> \<top>" (is "?wtl is \<noteq> _") 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   291
  assumes pc:  "pc < length is"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   292
  shows  "wtc c pc (wtl (take pc is) c 0 s) \<noteq> \<top>"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   293
proof -
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   294
  from pc have "0 < length (drop pc is)" by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   295
  then  obtain i r where Cons: "drop pc is = i#r" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   296
    by (auto simp add: neq_Nil_conv simp del: length_drop drop_eq_Nil)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   297
  hence "i#r = drop pc is" ..
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   298
  with all have take: "?wtl (take pc is@i#r) \<noteq> \<top>" by simp 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   299
  from pc have "is!pc = drop pc is ! 0" by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   300
  with Cons have "is!pc = i" by simp
54863
82acc20ded73 prefer more canonical names for lemmas on min/max
haftmann
parents: 42463
diff changeset
   301
  with take pc show ?thesis by (auto simp add: min.absorb2)
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   302
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   303
58886
8a6cac7c7247 modernized header;
wenzelm
parents: 54863
diff changeset
   304
subsection "preserves-type"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   305
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   306
lemma (in lbv) merge_pres:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   307
  assumes s0: "snd`set ss \<subseteq> A" and x: "x \<in> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   308
  shows "merge c pc ss x \<in> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   309
proof -
68386
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
   310
  from s0 have "set (map snd [(p', t')\<leftarrow>ss . p'=pc+1]) \<subseteq> A" by auto
98cf1c823c48 Keep filter input syntax
nipkow
parents: 68249
diff changeset
   311
  with x  have "(map snd [(p', t')\<leftarrow>ss . p'=pc+1] ++_f x) \<in> A"
33954
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   312
    by (auto intro!: plusplus_closed semilat)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   313
  with s0 x show ?thesis by (simp add: merge_def T_A)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   314
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   315
  
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   316
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   317
lemma pres_typeD2:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   318
  "pres_type step n A \<Longrightarrow> s \<in> A \<Longrightarrow> p < n \<Longrightarrow> snd`set (step p s) \<subseteq> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   319
  by auto (drule pres_typeD)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   320
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   321
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   322
lemma (in lbv) wti_pres [intro?]:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   323
  assumes pres: "pres_type step n A" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   324
  assumes cert: "c!(pc+1) \<in> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   325
  assumes s_pc: "s \<in> A" "pc < n"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   326
  shows "wti c pc s \<in> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   327
proof -
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   328
  from pres s_pc have "snd`set (step pc s) \<subseteq> A" by (rule pres_typeD2)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   329
  with cert show ?thesis by (simp add: wti merge_pres)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   330
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   331
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   332
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   333
lemma (in lbv) wtc_pres:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   334
  assumes pres: "pres_type step n A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   335
  assumes cert: "c!pc \<in> A" and cert': "c!(pc+1) \<in> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   336
  assumes s: "s \<in> A" and pc: "pc < n"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   337
  shows "wtc c pc s \<in> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   338
proof -
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   339
  have "wti c pc s \<in> A" using pres cert' s pc ..
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   340
  moreover have "wti c pc (c!pc) \<in> A" using pres cert' cert pc ..
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   341
  ultimately show ?thesis using T_A by (simp add: wtc) 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   342
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   343
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   344
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   345
lemma (in lbv) wtl_pres:
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   346
  assumes pres: "pres_type step (length is) A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   347
  assumes cert: "cert_ok c (length is) \<top> \<bottom> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   348
  assumes s:    "s \<in> A" 
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   349
  assumes all:  "wtl is c 0 s \<noteq> \<top>"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   350
  shows "pc < length is \<Longrightarrow> wtl (take pc is) c 0 s \<in> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   351
  (is "?len pc \<Longrightarrow> ?wtl pc \<in> A")
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   352
proof (induct pc)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   353
  from s show "?wtl 0 \<in> A" by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   354
next
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   355
  fix n assume IH: "Suc n < length is"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   356
  then have n: "n < length is" by simp  
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   357
  from IH have n1: "n+1 < length is" by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   358
  assume prem: "n < length is \<Longrightarrow> ?wtl n \<in> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   359
  have "wtc c n (?wtl n) \<in> A"
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   360
  using pres _ _ _ n
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   361
  proof (rule wtc_pres)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   362
    from prem n show "?wtl n \<in> A" .
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   363
    from cert n show "c!n \<in> A" by (rule cert_okD1)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   364
    from cert n1 show "c!(n+1) \<in> A" by (rule cert_okD1)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   365
  qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   366
  also
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   367
  from all n have "?wtl n \<noteq> \<top>" by - (rule wtl_take)
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   368
  with n1 have "wtc c n (?wtl n) = ?wtl (n+1)" by (rule wtl_Suc [symmetric])
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   369
  finally  show "?wtl (Suc n) \<in> A" by simp
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   370
qed
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   371
1bc3b688548c backported parts of abstract byte code verifier from AFP/Jinja
haftmann
parents:
diff changeset
   372
end