src/HOL/MicroJava/BV/LBVSpec.thy
author kleing
Wed, 20 Mar 2002 13:21:07 +0100
changeset 13062 4b1edf2f6bd2
parent 13006 51c5f3f11d16
child 13071 f538a1dba7ee
permissions -rw-r--r--
small refactoring for lbv with semilattices
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
     1
(*  Title:      HOL/MicroJava/BV/LBVSpec.thy
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
     2
    ID:         $Id$
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
     3
    Author:     Gerwin Klein
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
     4
    Copyright   1999 Technische Universitaet Muenchen
9054
0e48e7d4d4f9 minor tuning for pdf documents
kleing
parents: 9012
diff changeset
     5
*)
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
     6
12911
704713ca07ea new document
kleing
parents: 12516
diff changeset
     7
header {* \isaheader{The Lightweight Bytecode Verifier} *}
9054
0e48e7d4d4f9 minor tuning for pdf documents
kleing
parents: 9012
diff changeset
     8
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
     9
theory LBVSpec = SemilatAlg + Opt:
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    10
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    11
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    12
syntax
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    13
  "@lesuberropt" :: "'a option err \<Rightarrow> 'a ord \<Rightarrow> 'a option err \<Rightarrow> 'a option err ord"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    14
  ("_ \<le>|_ _" [50,50,51] 50) 
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    15
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    16
  "@superropt" :: "'a option err \<Rightarrow> 'a ebinop \<Rightarrow> 'a option err \<Rightarrow> 'a option err binop"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    17
  ("_ +|_ _" [50,50,51] 50)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    18
  
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    19
  "@supsuperropt" :: "'a option err list \<Rightarrow> 'a ebinop \<Rightarrow> 'a option err \<Rightarrow> 'a option err binop"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    20
  ("_ ++|_ _" [50,50,51] 50)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    21
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    22
translations
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    23
  "a \<le>|r b" == "a <=_(Err.le (Opt.le r)) b"  
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    24
  "a +|f b" == "a +_(Err.lift2 (Opt.sup f)) b"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    25
  "a ++|f b" == "a ++_(Err.lift2 (Opt.sup f)) b"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    26
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    27
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    28
types
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    29
  's certificate = "'s option list"   
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    30
  's steptype    = "'s option err step_type"
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    31
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    32
consts
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    33
merge :: "'s certificate \<Rightarrow> 's ebinop \<Rightarrow> 's ord \<Rightarrow> nat \<Rightarrow> 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    34
          (nat \<times> 's option err) list \<Rightarrow> 's option err \<Rightarrow> 's option err"
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
    35
primrec
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    36
"merge cert f r pc []     x = x"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    37
"merge cert f r pc (s#ss) x = (let (pc',s') = s in
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    38
                               merge cert f r pc ss (if pc'=pc+1 then (s' +|f x)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    39
                                                     else if s' \<le>|r (OK (cert!pc')) then x
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    40
                                                     else Err))"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    41
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    42
constdefs
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    43
wtl_inst :: "'s certificate \<Rightarrow> 's ebinop \<Rightarrow> 's ord \<Rightarrow> 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    44
             's steptype \<Rightarrow> nat \<Rightarrow> 's option \<Rightarrow> 's option err"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    45
"wtl_inst cert f r step pc s \<equiv> merge cert f r pc (step pc (OK s)) (OK (cert!(pc+1)))"
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    46
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    47
wtl_cert :: "'s certificate \<Rightarrow> 's ebinop \<Rightarrow> 's ord \<Rightarrow> 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    48
             's steptype \<Rightarrow> nat \<Rightarrow> 's option \<Rightarrow> 's option err"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    49
"wtl_cert cert f r step pc s \<equiv>
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    50
  case cert!pc of 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    51
    None    \<Rightarrow> wtl_inst cert f r step pc s
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    52
  | Some s' \<Rightarrow> if OK s \<le>|r (OK (Some s')) then wtl_inst cert f r step pc (Some s') else Err"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    53
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
    54
consts 
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    55
wtl_inst_list :: "'a list \<Rightarrow> 's certificate \<Rightarrow> 's ebinop \<Rightarrow> 's ord \<Rightarrow> 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    56
                  's steptype \<Rightarrow> nat \<Rightarrow> 's option \<Rightarrow> 's option err"
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    57
primrec
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    58
"wtl_inst_list []      cert f r step pc s = OK s"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    59
"wtl_inst_list (i#ins) cert f r step pc s = 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    60
  (let s' = wtl_cert cert f r step pc s in
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    61
   strict (wtl_inst_list ins cert f r step (pc+1)) s')"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    62
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    63
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    64
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
    65
constdefs
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    66
  cert_ok :: "'s certificate \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> bool"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    67
  "cert_ok cert n A \<equiv> \<forall>i < n. cert!i \<in> opt A"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    68
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    69
lemma cert_okD:
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    70
  "cert_ok cert n A \<Longrightarrow> pc < n \<Longrightarrow> cert!pc \<in> opt A"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    71
  by (unfold cert_ok_def) fast
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    72
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    73
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    74
declare Let_def [simp]
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    75
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    76
section "more semilattice lemmas"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    77
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    78
(*
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    79
lemma sup_top [simp]:
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    80
  assumes sl:  "semilat (A,r,f)" 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    81
  assumes set: "x \<in> A"  "t \<in> A"  
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    82
  assumes top: "top r t" 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    83
  shows "x +_f t = t"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    84
proof -
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    85
  from sl have "order r" ..
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    86
  moreover from top have "x +_f t <=_r t" ..
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    87
  moreover from sl set have "t <=_r x +_f t" by simp 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    88
  ultimately show ?thesis by (rule order_antisym)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    89
qed
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    90
  
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    91
lemma plusplussup_top:
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    92
  "semilat (A,r,f) \<Longrightarrow> top r Err \<Longrightarrow> set xs \<subseteq> A \<Longrightarrow> Err \<in> A \<Longrightarrow> xs ++_f Err = Err"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    93
  by (induct xs) auto
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    94
*)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    95
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    96
lemma err_semilatDorderI [simp, intro]:
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    97
  "err_semilat (A,r,f) \<Longrightarrow> order r"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    98
  apply (simp add: Err.sl_def)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
    99
  apply (rule order_le_err [THEN iffD1])
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   100
  apply (rule semilatDorderI)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   101
  apply assumption
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   102
  done
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   103
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   104
lemma err_opt_semilat [simp,elim]:
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   105
  "err_semilat (A,r,f) \<Longrightarrow> semilat (err (opt A), Err.le (Opt.le r), lift2 (Opt.sup f))"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   106
proof -
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   107
  assume "err_semilat (A,r,f)"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   108
  hence "err_semilat (Opt.esl (A,r,f))" ..
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   109
  thus ?thesis by (simp add: Err.sl_def Opt.esl_def)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   110
qed
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   111
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   112
lemma plusplus_erropt_Err [simp]:
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   113
  "xs ++|f Err = Err"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   114
  by (induct xs) auto
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   115
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   116
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   117
section "merge"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   118
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   119
lemma merge_Err [simp]:
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   120
  "merge cert f r pc ss Err = Err"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   121
  by (induct ss) auto
8245
6acc80f7f36f lightweight bytecode verifier with correctness proof
kleing
parents:
diff changeset
   122
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   123
lemma merge_ok:
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   124
  "\<And>x. merge cert f r pc ss x = OK s \<Longrightarrow> 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   125
  \<forall>(pc',s') \<in> set ss. (pc' \<noteq> pc+1 \<longrightarrow> s' \<le>|r (OK (cert!pc')))"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   126
  (is "\<And>x. ?merge ss x \<Longrightarrow> ?P ss")
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   127
proof (induct ss)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   128
  show "?P []" by simp
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   129
next
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   130
  fix x l ls assume merge: "?merge (l#ls) x"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   131
  moreover
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   132
  obtain pc' s' where [simp]: "l = (pc',s')" by (cases l)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   133
  ultimately
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   134
  obtain x' where "?merge ls x'" by simp  
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   135
  assume "\<And>x. ?merge ls x \<Longrightarrow> ?P ls" hence "?P ls" .
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   136
  moreover
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   137
  from merge
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   138
  have "pc' \<noteq> pc+1 \<longrightarrow> s' \<le>|r (OK (cert!pc'))" by (simp split: split_if_asm)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   139
  ultimately
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   140
  show "?P (l#ls)" by simp
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   141
qed
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   142
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   143
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   144
lemma merge_def:
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   145
  assumes semi: "err_semilat (A,r,f)" 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   146
  shows 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   147
  "\<And>x. x \<in> err (opt A) \<Longrightarrow> \<forall>(pc', s') \<in> set ss. s' \<in> err (opt A) \<Longrightarrow>
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   148
  merge cert f r pc ss x = 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   149
  (if \<forall>(pc',s') \<in> set ss. (pc' \<noteq> pc+1 \<longrightarrow> s' \<le>|r (OK (cert!pc'))) then
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   150
    map snd [(p',t') \<in> ss. p'=pc+1] ++|f x
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   151
  else Err)" 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   152
  (is "\<And>x. _ \<Longrightarrow> _ \<Longrightarrow> ?merge ss x = ?if ss x" is "\<And>x. _ \<Longrightarrow> _ \<Longrightarrow> ?P ss x")
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   153
proof (induct ss)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   154
  fix x show "?P [] x" by simp
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   155
next 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   156
  fix x assume x: "x \<in> err (opt A)" 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   157
  fix l::"nat \<times> 'a option err" and ls  
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   158
  assume "\<forall>(pc', s') \<in> set (l#ls). s' \<in> err (opt A)"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   159
  then obtain l: "snd l \<in> err (opt A)" and ls: "\<forall>(pc', s') \<in> set ls. s' \<in> err (opt A)" by auto
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   160
  assume "\<And>x. x \<in> err (opt A) \<Longrightarrow> \<forall>(pc',s') \<in> set ls. s' \<in> err (opt A) \<Longrightarrow> ?P ls x" 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   161
  hence IH: "\<And>x. x \<in> err (opt A) \<Longrightarrow> ?P ls x" .
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   162
  obtain pc' s' where [simp]: "l = (pc',s')" by (cases l)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   163
  hence "?merge (l#ls) x = ?merge ls 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   164
    (if pc'=pc+1 then s' +|f x else if s' \<le>|r (OK (cert!pc')) then x else Err)"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   165
    (is "?merge (l#ls) x = ?merge ls ?if'")
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   166
    by simp 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   167
  also have "\<dots> = ?if ls ?if'" 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   168
  proof -
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   169
    from l have "s' \<in> err (opt A)" by simp
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   170
    with x semi have "(s' +|f x) \<in> err (opt A)" by (fast intro: closedD)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   171
    with x have "?if' \<in> err (opt A)" by auto
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   172
    hence "?P ls ?if'" by (rule IH) thus ?thesis by simp
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   173
  qed
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   174
  also have "\<dots> = ?if (l#ls) x"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   175
    proof (cases "\<forall>(pc', s')\<in>set (l#ls). pc' \<noteq> pc + 1 \<longrightarrow> s' \<le>|r OK (cert ! pc')")
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   176
      case True
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   177
      hence "\<forall>(pc', s')\<in>set ls. pc' \<noteq> pc + 1 \<longrightarrow> s' \<le>|r (OK (cert ! pc'))" by auto
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   178
      moreover
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   179
      from True have 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   180
        "map snd [(p',t')\<in>ls . p'=pc+1] ++|f ?if' = 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   181
        (map snd [(p',t')\<in>l#ls . p'=pc+1] ++|f x)"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   182
        by simp
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   183
      ultimately
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   184
      show ?thesis using True by simp
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   185
    next
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   186
      case False thus ?thesis by auto
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   187
    qed
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   188
  finally show "?P (l#ls) x" .
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   189
qed
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   190
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   191
lemma merge_ok_s:
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   192
  assumes sl: "err_semilat (A,r,f)" 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   193
  assumes x:  "x \<in> err (opt A)"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   194
  assumes ss: "\<forall>(pc', s') \<in> set ss. s' \<in> err (opt A)"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   195
  assumes m:  "merge cert f r pc ss x = OK s"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   196
  shows "merge cert f r pc ss x = (map snd [(p',t') \<in> ss. p'=pc+1] ++|f x)"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   197
proof -
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   198
  from m have "\<forall>(pc',s') \<in> set ss. (pc' \<noteq> pc+1 \<longrightarrow> s' \<le>|r (OK (cert!pc')))" 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   199
    by (rule merge_ok)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   200
  with sl x ss m show ?thesis by - (drule merge_def, auto split: split_if_asm)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   201
qed
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   202
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   203
section "wtl-inst-list"
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9757
diff changeset
   204
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10638
diff changeset
   205
lemmas [iff] = not_Err_eq
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10638
diff changeset
   206
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   207
lemma wtl_Cons:
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   208
  "wtl_inst_list (i#is) cert f r step pc s \<noteq> Err = 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   209
  (\<exists>s'. wtl_cert cert f r step pc s = OK s' \<and> 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   210
        wtl_inst_list is cert f r step (pc+1) s' \<noteq> Err)"
10812
ead84e90bfeb merged semilattice orders with <=' from Convert.thy (now defined in JVMType.thy)
kleing
parents: 10638
diff changeset
   211
  by (auto simp del: split_paired_Ex)
9549
40d64cb4f4e6 BV and LBV specified in terms of app and step functions
kleing
parents: 9376
diff changeset
   212
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   213
lemma wtl_append:
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   214
"\<forall>s pc. (wtl_inst_list (a@b) cert f r step pc s = OK s') =
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   215
   (\<exists>s''. wtl_inst_list a cert f r step pc s = OK s'' \<and> 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   216
          wtl_inst_list b cert f r step (pc+length a) s'' = OK s')" 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   217
  (is "\<forall>s pc. ?C s pc a" is "?P a")
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   218
proof (induct ?P a)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   219
  show "?P []" by simp
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   220
next
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   221
  fix x xs  assume IH: "?P xs"
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   222
  show "?P (x#xs)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   223
  proof (intro allI)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   224
    fix s pc 
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   225
    show "?C s pc (x#xs)"
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   226
    proof (cases "wtl_cert cert f r step pc s")
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   227
      case Err thus ?thesis by simp
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   228
    next
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   229
      case (OK s0)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   230
      with IH have "?C s0 (pc+1) xs" by blast
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   231
      thus ?thesis by (simp!)
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   232
    qed
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   233
  qed
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   234
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   235
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   236
lemma wtl_take:
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   237
  "wtl_inst_list is cert f r step pc s = OK s'' \<Longrightarrow>
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   238
   \<exists>s'. wtl_inst_list (take pc' is) cert f r step pc s = OK s'"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   239
  (is "?wtl is = _ \<Longrightarrow> _")
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   240
proof -
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   241
  assume "?wtl is = OK s''"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   242
  hence "?wtl (take pc' is @ drop pc' is) = OK s''" by simp
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   243
  thus ?thesis by (auto simp add: wtl_append simp del: append_take_drop_id)
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   244
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   245
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   246
lemma take_Suc:
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   247
  "\<forall>n. n < length l --> take (Suc n) l = (take n l)@[l!n]" (is "?P l")
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   248
proof (induct l)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   249
  show "?P []" by simp
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   250
next
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   251
  fix x xs assume IH: "?P xs"  
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   252
  show "?P (x#xs)"
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   253
  proof (intro strip)
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   254
    fix n assume "n < length (x#xs)"
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   255
    with IH show "take (Suc n) (x # xs) = take n (x # xs) @ [(x # xs) ! n]" 
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   256
      by (cases n, auto)
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   257
  qed
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   258
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   259
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   260
lemma wtl_Suc:
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   261
  assumes "wtl_inst_list (take pc is) cert f r step 0 s = OK s'"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   262
           "wtl_cert cert f r step pc s' = OK s''" and
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   263
      suc: "pc+1 < length is"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   264
  shows "wtl_inst_list (take (pc+1) is) cert f r step 0 s = OK s''"    
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   265
proof -
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   266
  from suc have "take (pc+1) is=(take pc is)@[is!pc]" by (simp add: take_Suc)
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   267
  thus ?thesis by (simp! add: wtl_append min_def)
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   268
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   269
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   270
lemma wtl_all:
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   271
  assumes
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   272
  all: "wtl_inst_list is cert f r step 0 s \<noteq> Err" (is "?wtl is \<noteq> _") and
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   273
  pc:  "pc < length is"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   274
  shows
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   275
  "\<exists>s' s''. wtl_inst_list (take pc is) cert f r step 0 s = OK s' \<and> 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   276
            wtl_cert cert f r step pc s' = OK s''"
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   277
proof -
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   278
  from pc have "0 < length (drop pc is)" by simp
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   279
  then  obtain i r where Cons: "drop pc is = i#r" 
9757
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   280
    by (auto simp add: neq_Nil_conv simp del: length_drop)
1024a2d80ac0 functional LBV style, dead code, type safety -> Isar
kleing
parents: 9664
diff changeset
   281
  hence "i#r = drop pc is" ..
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   282
  with all have take: "?wtl (take pc is@i#r) \<noteq> Err" by simp 
12516
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   283
  from pc have "is!pc = drop pc is ! 0" by simp
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   284
  with Cons have "is!pc = i" by simp
d09d0f160888 exceptions
kleing
parents: 11085
diff changeset
   285
  with take pc show ?thesis by (auto simp add: wtl_append min_def)
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   286
qed
9012
d1bd2144ab5d switched to Isar proofs
kleing
parents: 8390
diff changeset
   287
13062
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   288
section "preserves-type"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   289
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   290
lemma merge_pres:
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   291
  assumes semi: "err_semilat (A,r,f)"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   292
  assumes s0_A: "\<forall>(pc', s')\<in>set ss. s' \<in> err (opt A)"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   293
  assumes x_A:  "x \<in> err (opt A)"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   294
  assumes merge:"merge cert f r pc ss x = OK s'"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   295
  shows "s' \<in> opt A"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   296
proof -
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   297
  from s0_A
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   298
  have "set (map snd [(p', t')\<in>ss . p'=pc+1]) \<subseteq> err (opt A)" by auto
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   299
  with semi x_A
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   300
  have "(map snd [(p', t')\<in>ss . p'=pc+1] ++|f x) \<in> err (opt A)"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   301
    by (auto intro!: plusplus_closed)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   302
  also {
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   303
    note merge
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   304
    also from semi x_A s0_A
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   305
    have "merge cert f r pc ss x = 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   306
      (if \<forall>(pc', s')\<in>set ss. pc' \<noteq> pc + 1 \<longrightarrow> s' \<le>|r (OK (cert!pc'))
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   307
      then map snd [(p', t')\<in>ss . p'=pc+1] ++|f x else Err)"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   308
      by (rule merge_def)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   309
    finally have "(map snd [(p', t')\<in>ss . p'=pc+1] ++|f x) = OK s'" 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   310
      by (simp split: split_if_asm)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   311
  } 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   312
  finally show ?thesis by simp
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   313
qed
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   314
  
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   315
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   316
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   317
lemma wtl_inst_pres [intro?]:
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   318
  assumes semi: "err_semilat (A,r,f)"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   319
  assumes pres: "pres_type step n (err (opt A))" 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   320
  assumes cert: "cert!(pc+1) \<in> opt A"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   321
  assumes s_A:  "s \<in> opt A"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   322
  assumes pc:   "pc < n"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   323
  assumes wtl:  "wtl_inst cert f r step pc s = OK s'"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   324
  shows "s' \<in> opt A"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   325
proof -
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   326
  from pres pc s_A
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   327
  have "\<forall>(q, s')\<in>set (step pc (OK s)). s' \<in> err (opt A)"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   328
    by (unfold pres_type_def) blast
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   329
  moreover
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   330
  from cert have "OK (cert!(pc+1)) \<in> err (opt A)" by simp
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   331
  moreover
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   332
  from wtl
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   333
  have "merge cert f r pc (step pc (OK s)) (OK (cert!(pc+1))) = OK s'"  
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   334
    by (unfold wtl_inst_def) simp
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   335
  ultimately
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   336
  show "s' \<in> opt A" using semi by - (rule merge_pres)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   337
qed
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   338
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   339
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   340
lemma wtl_cert_pres:
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   341
  assumes "err_semilat (A,r,f)"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   342
  assumes "pres_type step n (err (opt A))"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   343
  assumes "cert!pc \<in> opt A" and "cert!(pc+1) \<in> opt A"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   344
  assumes "s \<in> opt A" and "pc < n"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   345
  assumes wtc: "wtl_cert cert f r step pc s = OK s'"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   346
  shows "s' \<in> opt A"            
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   347
proof -
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   348
  have "wtl_inst cert f r step pc s = OK s' \<Longrightarrow> s' \<in> opt A" ..
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   349
  moreover
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   350
  have "wtl_inst cert f r step pc (cert!pc) = OK s' \<Longrightarrow> s' \<in> opt A" ..
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   351
  ultimately
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   352
  show ?thesis using wtc
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   353
    by (unfold wtl_cert_def) (simp split: option.splits split_if_asm)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   354
qed
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   355
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   356
lemma wtl_inst_list_pres:
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   357
  assumes semi: "err_semilat (A,r,f)"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   358
  assumes pres: "pres_type step (length is) (err (opt A))"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   359
  assumes cert: "cert_ok cert (length is) A"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   360
  assumes s:    "s \<in> opt A" 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   361
  assumes all:  "wtl_inst_list is cert f r step 0 s \<noteq> Err"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   362
  shows "\<And>s'. pc < length is \<Longrightarrow> wtl_inst_list (take pc is) cert f r step 0 s = OK s'
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   363
         \<Longrightarrow> s' \<in> opt A" (is "\<And>s'. ?len pc \<Longrightarrow> ?wtl pc s' \<Longrightarrow> ?A s'")
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   364
proof (induct pc)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   365
  from s show "\<And>s'. ?wtl 0 s' \<Longrightarrow> ?A s'" by simp
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   366
next
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   367
  fix s' n
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   368
  assume "Suc n < length is" 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   369
  then obtain "n < length is" by simp  
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   370
  with all obtain s1 s2 where
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   371
    "wtl_inst_list (take n is) cert f r step 0 s = OK s1" 
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   372
    "wtl_cert cert f r step n s1 = OK s2"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   373
    by - (drule wtl_all, auto)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   374
    
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   375
  assume "?wtl (Suc n) s'"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   376
  moreover
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   377
  have n1: "n+1 < length is" by simp
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   378
  hence "?wtl (n+1) s2" by - (rule wtl_Suc)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   379
  ultimately
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   380
  have [simp]: "s2 = s'" by simp
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   381
  
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   382
  assume IH: "\<And>s'. n < length is \<Longrightarrow> ?wtl n s' \<Longrightarrow> s' \<in> opt A"
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   383
  hence "s1 \<in> opt A" .
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   384
  moreover
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   385
  from cert have "cert!n \<in> opt A" by (rule cert_okD)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   386
  moreover
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   387
  from cert n1 have "cert!(n+1) \<in> opt A" by (rule cert_okD)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   388
  ultimately
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   389
  have "s2 \<in> opt A" using semi pres by - (rule wtl_cert_pres)
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   390
  thus "s' \<in> opt A" by simp
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   391
qed
4b1edf2f6bd2 small refactoring for lbv with semilattices
kleing
parents: 13006
diff changeset
   392
10042
7164dc0d24d8 unsymbolized
kleing
parents: 9757
diff changeset
   393
9183
cd4494dc789a tuned for ProofGeneral 3.2
kleing
parents: 9054
diff changeset
   394
end