backported parts of abstract byte code verifier from AFP/Jinja
authorhaftmann
Tue, 24 Nov 2009 14:37:23 +0100
changeset 33954 1bc3b688548c
parent 33930 6a973bd43949
child 33955 fff6f11b1f09
backported parts of abstract byte code verifier from AFP/Jinja
src/HOL/IsaMakefile
src/HOL/MicroJava/BV/Altern.thy
src/HOL/MicroJava/BV/BVExample.thy
src/HOL/MicroJava/BV/BVNoTypeError.thy
src/HOL/MicroJava/BV/BVSpec.thy
src/HOL/MicroJava/BV/BVSpecTypeSafe.thy
src/HOL/MicroJava/BV/Correct.thy
src/HOL/MicroJava/BV/Effect.thy
src/HOL/MicroJava/BV/EffectMono.thy
src/HOL/MicroJava/BV/Err.thy
src/HOL/MicroJava/BV/JType.thy
src/HOL/MicroJava/BV/JVM.thy
src/HOL/MicroJava/BV/JVMType.thy
src/HOL/MicroJava/BV/Kildall.thy
src/HOL/MicroJava/BV/LBVComplete.thy
src/HOL/MicroJava/BV/LBVCorrect.thy
src/HOL/MicroJava/BV/LBVJVM.thy
src/HOL/MicroJava/BV/LBVSpec.thy
src/HOL/MicroJava/BV/Listn.thy
src/HOL/MicroJava/BV/Opt.thy
src/HOL/MicroJava/BV/Product.thy
src/HOL/MicroJava/BV/Semilat.thy
src/HOL/MicroJava/BV/SemilatAlg.thy
src/HOL/MicroJava/BV/Typing_Framework.thy
src/HOL/MicroJava/BV/Typing_Framework_JVM.thy
src/HOL/MicroJava/BV/Typing_Framework_err.thy
src/HOL/MicroJava/Comp/AuxLemmas.thy
src/HOL/MicroJava/Comp/CorrComp.thy
src/HOL/MicroJava/Comp/CorrCompTp.thy
src/HOL/MicroJava/Comp/DefsComp.thy
src/HOL/MicroJava/Comp/Index.thy
src/HOL/MicroJava/Comp/LemmasComp.thy
src/HOL/MicroJava/Comp/NatCanonify.thy
src/HOL/MicroJava/Comp/TranslComp.thy
src/HOL/MicroJava/Comp/TranslCompTp.thy
src/HOL/MicroJava/Comp/TypeInf.thy
src/HOL/MicroJava/DFA/Abstract_BV.thy
src/HOL/MicroJava/DFA/Err.thy
src/HOL/MicroJava/DFA/Kildall.thy
src/HOL/MicroJava/DFA/LBVComplete.thy
src/HOL/MicroJava/DFA/LBVCorrect.thy
src/HOL/MicroJava/DFA/LBVSpec.thy
src/HOL/MicroJava/DFA/Listn.thy
src/HOL/MicroJava/DFA/Opt.thy
src/HOL/MicroJava/DFA/Product.thy
src/HOL/MicroJava/DFA/Semilat.thy
src/HOL/MicroJava/DFA/SemilatAlg.thy
src/HOL/MicroJava/DFA/Semilattices.thy
src/HOL/MicroJava/DFA/Typing_Framework.thy
src/HOL/MicroJava/DFA/Typing_Framework_err.thy
src/HOL/MicroJava/J/Conform.thy
src/HOL/MicroJava/J/Decl.thy
src/HOL/MicroJava/J/Eval.thy
src/HOL/MicroJava/J/Example.thy
src/HOL/MicroJava/J/Exceptions.thy
src/HOL/MicroJava/J/JBasis.thy
src/HOL/MicroJava/J/JListExample.thy
src/HOL/MicroJava/J/JTypeSafe.thy
src/HOL/MicroJava/J/State.thy
src/HOL/MicroJava/J/SystemClasses.thy
src/HOL/MicroJava/J/Term.thy
src/HOL/MicroJava/J/Type.thy
src/HOL/MicroJava/J/TypeRel.thy
src/HOL/MicroJava/J/Value.thy
src/HOL/MicroJava/J/WellForm.thy
src/HOL/MicroJava/J/WellType.thy
src/HOL/MicroJava/JVM/JVMDefensive.thy
src/HOL/MicroJava/JVM/JVMExceptions.thy
src/HOL/MicroJava/JVM/JVMExec.thy
src/HOL/MicroJava/JVM/JVMExecInstr.thy
src/HOL/MicroJava/JVM/JVMInstructions.thy
src/HOL/MicroJava/JVM/JVMListExample.thy
src/HOL/MicroJava/JVM/JVMState.thy
src/HOL/MicroJava/MicroJava.thy
src/HOL/MicroJava/ROOT.ML
src/HOL/MicroJava/document/introduction.tex
src/HOL/MicroJava/document/root.bib
src/HOL/MicroJava/document/root.tex
--- a/src/HOL/IsaMakefile	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/IsaMakefile	Tue Nov 24 14:37:23 2009 +0100
@@ -848,20 +848,20 @@
   MicroJava/J/JListExample.thy MicroJava/JVM/JVMExec.thy		\
   MicroJava/JVM/JVMInstructions.thy MicroJava/JVM/JVMState.thy		\
   MicroJava/JVM/JVMExecInstr.thy MicroJava/JVM/JVMListExample.thy	\
-  MicroJava/JVM/JVMExceptions.thy MicroJava/BV/BVSpec.thy		\
+  MicroJava/JVM/JVMExceptions.thy MicroJava/DFA/Abstract_BV.thy		\
+  MicroJava/DFA/Err.thy MicroJava/DFA/Kildall.thy			\
+  MicroJava/DFA/LBVComplete.thy MicroJava/DFA/LBVCorrect.thy		\
+  MicroJava/DFA/LBVSpec.thy MicroJava/DFA/Listn.thy			\
+  MicroJava/DFA/Opt.thy MicroJava/DFA/Product.thy			\
+  MicroJava/DFA/SemilatAlg.thy MicroJava/DFA/Semilat.thy		\
+  MicroJava/DFA/Semilattices.thy MicroJava/DFA/Typing_Framework_err.thy	\
+  MicroJava/DFA/Typing_Framework.thy MicroJava/BV/BVSpec.thy		\
   MicroJava/BV/BVSpecTypeSafe.thy MicroJava/BV/Correct.thy		\
-  MicroJava/BV/Err.thy MicroJava/BV/JType.thy MicroJava/BV/JVM.thy	\
-  MicroJava/BV/JVMType.thy MicroJava/BV/Kildall.thy			\
-  MicroJava/BV/LBVSpec.thy MicroJava/BV/Listn.thy MicroJava/BV/Opt.thy	\
-  MicroJava/BV/Product.thy MicroJava/BV/Semilat.thy			\
+  MicroJava/BV/JType.thy MicroJava/BV/JVM.thy MicroJava/BV/JVMType.thy	\
   MicroJava/BV/Effect.thy MicroJava/BV/EffectMono.thy			\
-  MicroJava/BV/Typing_Framework.thy					\
-  MicroJava/BV/Typing_Framework_err.thy					\
   MicroJava/BV/Typing_Framework_JVM.thy MicroJava/BV/BVExample.thy	\
-  MicroJava/BV/LBVSpec.thy MicroJava/BV/LBVCorrect.thy			\
-  MicroJava/BV/LBVComplete.thy MicroJava/BV/LBVJVM.thy			\
-  MicroJava/document/root.bib MicroJava/document/root.tex		\
-  MicroJava/document/introduction.tex
+  MicroJava/BV/LBVJVM.thy MicroJava/document/root.bib			\
+  MicroJava/document/root.tex MicroJava/document/introduction.tex
 	@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL MicroJava
 
 
--- a/src/HOL/MicroJava/BV/Altern.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/BV/Altern.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -1,15 +1,12 @@
 (*  Title:      HOL/MicroJava/BV/Altern.thy
-    ID:         $Id$
     Author:     Martin Strecker
 *)
 
-
-(* Alternative definition of well-typing of bytecode, 
-   used in compiler type correctness proof *)
+header {* Alternative definition of well-typing of bytecode,  used in compiler type correctness proof *}
 
-
-theory Altern imports BVSpec begin
-
+theory Altern
+imports BVSpec
+begin
 
 constdefs
   check_type :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> JVMType.state \<Rightarrow> bool"
--- a/src/HOL/MicroJava/BV/BVExample.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/BV/BVExample.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -66,27 +66,28 @@
 
 text {* The subclass releation spelled out: *}
 lemma subcls1:
-  "subcls1 E = (\<lambda>C D. (C, D) \<in> {(list_name,Object), (test_name,Object), (Xcpt NullPointer, Object),
-                (Xcpt ClassCast, Object), (Xcpt OutOfMemory, Object)})"
+  "subcls1 E = {(list_name,Object), (test_name,Object), (Xcpt NullPointer, Object),
+                (Xcpt ClassCast, Object), (Xcpt OutOfMemory, Object)}"
 apply (simp add: subcls1_def2)
 apply (simp add: name_defs class_defs system_defs E_def class_def)
-apply (auto simp: expand_fun_eq)
+apply (simp add: Sigma_def)
+apply auto
 done
 
 text {* The subclass relation is acyclic; hence its converse is well founded: *}
 lemma notin_rtrancl:
-  "r\<^sup>*\<^sup>* a b \<Longrightarrow> a \<noteq> b \<Longrightarrow> (\<And>y. \<not> r a y) \<Longrightarrow> False"
-  by (auto elim: converse_rtranclpE)
+  "(a, b) \<in> r\<^sup>* \<Longrightarrow> a \<noteq> b \<Longrightarrow> (\<And>y. (a, y) \<notin> r) \<Longrightarrow> False"
+  by (auto elim: converse_rtranclE)
 
-lemma acyclic_subcls1_E: "acyclicP (subcls1 E)"
-  apply (rule acyclicI [to_pred])
+lemma acyclic_subcls1_E: "acyclic (subcls1 E)"
+  apply (rule acyclicI)
   apply (simp add: subcls1)
-  apply (auto dest!: tranclpD)
+  apply (auto dest!: tranclD)
   apply (auto elim!: notin_rtrancl simp add: name_defs distinct_classes)
   done
 
-lemma wf_subcls1_E: "wfP ((subcls1 E)\<inverse>\<inverse>)"
-  apply (rule finite_acyclic_wf_converse [to_pred])
+lemma wf_subcls1_E: "wf ((subcls1 E)\<inverse>)"
+  apply (rule finite_acyclic_wf_converse)
   apply (simp add: subcls1 del: insert_iff)
   apply (rule acyclic_subcls1_E)
   done  
--- a/src/HOL/MicroJava/BV/BVNoTypeError.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/BV/BVNoTypeError.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -4,7 +4,9 @@
 
 header {* \isaheader{Welltyped Programs produce no Type Errors} *}
 
-theory BVNoTypeError imports "../JVM/JVMDefensive" BVSpecTypeSafe begin
+theory BVNoTypeError
+imports "../JVM/JVMDefensive" BVSpecTypeSafe
+begin
 
 text {*
   Some simple lemmas about the type testing functions of the
--- a/src/HOL/MicroJava/BV/BVSpec.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/BV/BVSpec.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -1,8 +1,6 @@
 (*  Title:      HOL/MicroJava/BV/BVSpec.thy
-    ID:         $Id$
     Author:     Cornelia Pusch, Gerwin Klein
     Copyright   1999 Technische Universitaet Muenchen
-
 *)
 
 header {* \isaheader{The Bytecode Verifier}\label{sec:BVSpec} *}
--- a/src/HOL/MicroJava/BV/Correct.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/BV/Correct.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -1,15 +1,13 @@
-
 (*  Title:      HOL/MicroJava/BV/Correct.thy
-    ID:         $Id$
     Author:     Cornelia Pusch, Gerwin Klein
     Copyright   1999 Technische Universitaet Muenchen
-
-The invariant for the type safety proof.
 *)
 
 header {* \isaheader{BV Type Safety Invariant} *}
 
-theory Correct imports BVSpec "../JVM/JVMExec" begin
+theory Correct
+imports BVSpec "../JVM/JVMExec"
+begin
 
 constdefs
   approx_val :: "[jvm_prog,aheap,val,ty err] \<Rightarrow> bool"
--- a/src/HOL/MicroJava/BV/Effect.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/BV/Effect.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -9,7 +9,6 @@
 imports JVMType "../JVM/JVMExceptions"
 begin
 
-
 types
   succ_type = "(p_count \<times> state_type option) list"
 
--- a/src/HOL/MicroJava/BV/EffectMono.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/BV/EffectMono.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -1,13 +1,13 @@
 (*  Title:      HOL/MicroJava/BV/EffMono.thy
-    ID:         $Id$
     Author:     Gerwin Klein
     Copyright   2000 Technische Universitaet Muenchen
 *)
 
 header {* \isaheader{Monotonicity of eff and app} *}
 
-theory EffectMono imports Effect begin
-
+theory EffectMono
+imports Effect
+begin
 
 lemma PrimT_PrimT: "(G \<turnstile> xb \<preceq> PrimT p) = (xb = PrimT p)"
   by (auto elim: widen.cases)
--- a/src/HOL/MicroJava/BV/Err.thy	Wed Dec 02 12:04:07 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,353 +0,0 @@
-(*  Title:      HOL/MicroJava/BV/Err.thy
-    ID:         $Id$
-    Author:     Tobias Nipkow
-    Copyright   2000 TUM
-
-The error type
-*)
-
-header {* \isaheader{The Error Type} *}
-
-theory Err
-imports Semilat
-begin
-
-datatype 'a err = Err | OK 'a
-
-types 'a ebinop = "'a \<Rightarrow> 'a \<Rightarrow> 'a err"
-      'a esl =    "'a set * 'a ord * 'a ebinop"
-
-consts
-  ok_val :: "'a err \<Rightarrow> 'a"
-primrec
-  "ok_val (OK x) = x"
-
-constdefs
- lift :: "('a \<Rightarrow> 'b err) \<Rightarrow> ('a err \<Rightarrow> 'b err)"
-"lift f e == case e of Err \<Rightarrow> Err | OK x \<Rightarrow> f x"
-
- lift2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c err) \<Rightarrow> 'a err \<Rightarrow> 'b err \<Rightarrow> 'c err"
-"lift2 f e1 e2 ==
- case e1 of Err  \<Rightarrow> Err
-          | OK x \<Rightarrow> (case e2 of Err \<Rightarrow> Err | OK y \<Rightarrow> f x y)"
-
- le :: "'a ord \<Rightarrow> 'a err ord"
-"le r e1 e2 ==
-        case e2 of Err \<Rightarrow> True |
-                   OK y \<Rightarrow> (case e1 of Err \<Rightarrow> False | OK x \<Rightarrow> x <=_r y)"
-
- sup :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a err \<Rightarrow> 'b err \<Rightarrow> 'c err)"
-"sup f == lift2(%x y. OK(x +_f y))"
-
- err :: "'a set \<Rightarrow> 'a err set"
-"err A == insert Err {x . ? y:A. x = OK y}"
-
- esl :: "'a sl \<Rightarrow> 'a esl"
-"esl == %(A,r,f). (A,r, %x y. OK(f x y))"
-
- sl :: "'a esl \<Rightarrow> 'a err sl"
-"sl == %(A,r,f). (err A, le r, lift2 f)"
-
-syntax
- err_semilat :: "'a esl \<Rightarrow> bool"
-translations
-"err_semilat L" == "semilat(Err.sl L)"
-
-
-consts
-  strict  :: "('a \<Rightarrow> 'b err) \<Rightarrow> ('a err \<Rightarrow> 'b err)"
-primrec
-  "strict f Err    = Err"
-  "strict f (OK x) = f x"
-
-lemma strict_Some [simp]: 
-  "(strict f x = OK y) = (\<exists> z. x = OK z \<and> f z = OK y)"
-  by (cases x, auto)
-
-lemma not_Err_eq:
-  "(x \<noteq> Err) = (\<exists>a. x = OK a)" 
-  by (cases x) auto
-
-lemma not_OK_eq:
-  "(\<forall>y. x \<noteq> OK y) = (x = Err)"
-  by (cases x) auto  
-
-lemma unfold_lesub_err:
-  "e1 <=_(le r) e2 == le r e1 e2"
-  by (simp add: lesub_def)
-
-lemma le_err_refl:
-  "!x. x <=_r x \<Longrightarrow> e <=_(Err.le r) e"
-apply (unfold lesub_def Err.le_def)
-apply (simp split: err.split)
-done 
-
-lemma le_err_trans [rule_format]:
-  "order r \<Longrightarrow> e1 <=_(le r) e2 \<longrightarrow> e2 <=_(le r) e3 \<longrightarrow> e1 <=_(le r) e3"
-apply (unfold unfold_lesub_err le_def)
-apply (simp split: err.split)
-apply (blast intro: order_trans)
-done
-
-lemma le_err_antisym [rule_format]:
-  "order r \<Longrightarrow> e1 <=_(le r) e2 \<longrightarrow> e2 <=_(le r) e1 \<longrightarrow> e1=e2"
-apply (unfold unfold_lesub_err le_def)
-apply (simp split: err.split)
-apply (blast intro: order_antisym)
-done 
-
-lemma OK_le_err_OK:
-  "(OK x <=_(le r) OK y) = (x <=_r y)"
-  by (simp add: unfold_lesub_err le_def)
-
-lemma order_le_err [iff]:
-  "order(le r) = order r"
-apply (rule iffI)
- apply (subst Semilat.order_def)
- apply (blast dest: order_antisym OK_le_err_OK [THEN iffD2]
-              intro: order_trans OK_le_err_OK [THEN iffD1])
-apply (subst Semilat.order_def)
-apply (blast intro: le_err_refl le_err_trans le_err_antisym
-             dest: order_refl)
-done 
-
-lemma le_Err [iff]:  "e <=_(le r) Err"
-  by (simp add: unfold_lesub_err le_def)
-
-lemma Err_le_conv [iff]:
- "Err <=_(le r) e  = (e = Err)"
-  by (simp add: unfold_lesub_err le_def  split: err.split)
-
-lemma le_OK_conv [iff]:
-  "e <=_(le r) OK x  =  (? y. e = OK y & y <=_r x)"
-  by (simp add: unfold_lesub_err le_def split: err.split)
-
-lemma OK_le_conv:
- "OK x <=_(le r) e  =  (e = Err | (? y. e = OK y & x <=_r y))"
-  by (simp add: unfold_lesub_err le_def split: err.split)
-
-lemma top_Err [iff]: "top (le r) Err";
-  by (simp add: top_def)
-
-lemma OK_less_conv [rule_format, iff]:
-  "OK x <_(le r) e = (e=Err | (? y. e = OK y & x <_r y))"
-  by (simp add: lesssub_def lesub_def le_def split: err.split)
-
-lemma not_Err_less [rule_format, iff]:
-  "~(Err <_(le r) x)"
-  by (simp add: lesssub_def lesub_def le_def split: err.split)
-
-lemma semilat_errI [intro]:
-  assumes semilat: "semilat (A, r, f)"
-  shows "semilat(err A, Err.le r, lift2(%x y. OK(f x y)))"
-  apply(insert semilat)
-  apply (unfold semilat_Def closed_def plussub_def lesub_def 
-    lift2_def Err.le_def err_def)
-  apply (simp split: err.split)
-  done
-
-lemma err_semilat_eslI_aux:
-  assumes semilat: "semilat (A, r, f)"
-  shows "err_semilat(esl(A,r,f))"
-  apply (unfold sl_def esl_def)
-  apply (simp add: semilat_errI[OF semilat])
-  done
-
-lemma err_semilat_eslI [intro, simp]:
- "\<And>L. semilat L \<Longrightarrow> err_semilat(esl L)"
-by(simp add: err_semilat_eslI_aux split_tupled_all)
-
-lemma acc_err [simp, intro!]:  "acc r \<Longrightarrow> acc(le r)"
-apply (unfold acc_def lesub_def le_def lesssub_def)
-apply (simp add: wfP_eq_minimal split: err.split)
-apply clarify
-apply (case_tac "Err : Q")
- apply blast
-apply (erule_tac x = "{a . OK a : Q}" in allE)
-apply (case_tac "x")
- apply fast
-apply blast
-done 
-
-lemma Err_in_err [iff]: "Err : err A"
-  by (simp add: err_def)
-
-lemma Ok_in_err [iff]: "(OK x : err A) = (x:A)"
-  by (auto simp add: err_def)
-
-section {* lift *}
-
-lemma lift_in_errI:
-  "\<lbrakk> e : err S; !x:S. e = OK x \<longrightarrow> f x : err S \<rbrakk> \<Longrightarrow> lift f e : err S"
-apply (unfold lift_def)
-apply (simp split: err.split)
-apply blast
-done 
-
-lemma Err_lift2 [simp]: 
-  "Err +_(lift2 f) x = Err"
-  by (simp add: lift2_def plussub_def)
-
-lemma lift2_Err [simp]: 
-  "x +_(lift2 f) Err = Err"
-  by (simp add: lift2_def plussub_def split: err.split)
-
-lemma OK_lift2_OK [simp]:
-  "OK x +_(lift2 f) OK y = x +_f y"
-  by (simp add: lift2_def plussub_def split: err.split)
-
-
-section {* sup *}
-
-lemma Err_sup_Err [simp]:
-  "Err +_(Err.sup f) x = Err"
-  by (simp add: plussub_def Err.sup_def Err.lift2_def)
-
-lemma Err_sup_Err2 [simp]:
-  "x +_(Err.sup f) Err = Err"
-  by (simp add: plussub_def Err.sup_def Err.lift2_def split: err.split)
-
-lemma Err_sup_OK [simp]:
-  "OK x +_(Err.sup f) OK y = OK(x +_f y)"
-  by (simp add: plussub_def Err.sup_def Err.lift2_def)
-
-lemma Err_sup_eq_OK_conv [iff]:
-  "(Err.sup f ex ey = OK z) = (? x y. ex = OK x & ey = OK y & f x y = z)"
-apply (unfold Err.sup_def lift2_def plussub_def)
-apply (rule iffI)
- apply (simp split: err.split_asm)
-apply clarify
-apply simp
-done
-
-lemma Err_sup_eq_Err [iff]:
-  "(Err.sup f ex ey = Err) = (ex=Err | ey=Err)"
-apply (unfold Err.sup_def lift2_def plussub_def)
-apply (simp split: err.split)
-done 
-
-section {* semilat (err A) (le r) f *}
-
-lemma semilat_le_err_Err_plus [simp]:
-  "\<lbrakk> x: err A; semilat(err A, le r, f) \<rbrakk> \<Longrightarrow> Err +_f x = Err"
-  by (blast intro: Semilat.le_iff_plus_unchanged [OF Semilat.intro, THEN iffD1]
-                   Semilat.le_iff_plus_unchanged2 [OF Semilat.intro, THEN iffD1])
-
-lemma semilat_le_err_plus_Err [simp]:
-  "\<lbrakk> x: err A; semilat(err A, le r, f) \<rbrakk> \<Longrightarrow> x +_f Err = Err"
-  by (blast intro: Semilat.le_iff_plus_unchanged [OF Semilat.intro, THEN iffD1]
-                   Semilat.le_iff_plus_unchanged2 [OF Semilat.intro, THEN iffD1])
-
-lemma semilat_le_err_OK1:
-  "\<lbrakk> x:A; y:A; semilat(err A, le r, f); OK x +_f OK y = OK z \<rbrakk> 
-  \<Longrightarrow> x <=_r z";
-apply (rule OK_le_err_OK [THEN iffD1])
-apply (erule subst)
-apply (simp add: Semilat.ub1 [OF Semilat.intro])
-done
-
-lemma semilat_le_err_OK2:
-  "\<lbrakk> x:A; y:A; semilat(err A, le r, f); OK x +_f OK y = OK z \<rbrakk> 
-  \<Longrightarrow> y <=_r z"
-apply (rule OK_le_err_OK [THEN iffD1])
-apply (erule subst)
-apply (simp add: Semilat.ub2 [OF Semilat.intro])
-done
-
-lemma eq_order_le:
-  "\<lbrakk> x=y; order r \<rbrakk> \<Longrightarrow> x <=_r y"
-apply (unfold Semilat.order_def)
-apply blast
-done
-
-lemma OK_plus_OK_eq_Err_conv [simp]:
-  assumes "x:A" and "y:A" and "semilat(err A, le r, fe)"
-  shows "((OK x) +_fe (OK y) = Err) = (~(? z:A. x <=_r z & y <=_r z))"
-proof -
-  have plus_le_conv3: "\<And>A x y z f r. 
-    \<lbrakk> semilat (A,r,f); x +_f y <=_r z; x:A; y:A; z:A \<rbrakk> 
-    \<Longrightarrow> x <=_r z \<and> y <=_r z"
-    by (rule Semilat.plus_le_conv [OF Semilat.intro, THEN iffD1])
-  from prems show ?thesis
-  apply (rule_tac iffI)
-   apply clarify
-   apply (drule OK_le_err_OK [THEN iffD2])
-   apply (drule OK_le_err_OK [THEN iffD2])
-   apply (drule Semilat.lub [OF Semilat.intro, of _ _ _ "OK x" _ "OK y"])
-        apply assumption
-       apply assumption
-      apply simp
-     apply simp
-    apply simp
-   apply simp
-  apply (case_tac "(OK x) +_fe (OK y)")
-   apply assumption
-  apply (rename_tac z)
-  apply (subgoal_tac "OK z: err A")
-  apply (drule eq_order_le)
-    apply (erule Semilat.orderI [OF Semilat.intro])
-   apply (blast dest: plus_le_conv3) 
-  apply (erule subst)
-  apply (blast intro: Semilat.closedI [OF Semilat.intro] closedD)
-  done 
-qed
-
-section {* semilat (err(Union AS)) *}
-
-(* FIXME? *)
-lemma all_bex_swap_lemma [iff]:
-  "(!x. (? y:A. x = f y) \<longrightarrow> P x) = (!y:A. P(f y))"
-  by blast
-
-lemma closed_err_Union_lift2I: 
-  "\<lbrakk> !A:AS. closed (err A) (lift2 f); AS ~= {}; 
-      !A:AS.!B:AS. A~=B \<longrightarrow> (!a:A.!b:B. a +_f b = Err) \<rbrakk> 
-  \<Longrightarrow> closed (err(Union AS)) (lift2 f)"
-apply (unfold closed_def err_def)
-apply simp
-apply clarify
-apply simp
-apply fast
-done 
-
-text {* 
-  If @{term "AS = {}"} the thm collapses to
-  @{prop "order r & closed {Err} f & Err +_f Err = Err"}
-  which may not hold 
-*}
-lemma err_semilat_UnionI:
-  "\<lbrakk> !A:AS. err_semilat(A, r, f); AS ~= {}; 
-      !A:AS.!B:AS. A~=B \<longrightarrow> (!a:A.!b:B. ~ a <=_r b & a +_f b = Err) \<rbrakk> 
-  \<Longrightarrow> err_semilat(Union AS, r, f)"
-apply (unfold semilat_def sl_def)
-apply (simp add: closed_err_Union_lift2I)
-apply (rule conjI)
- apply blast
-apply (simp add: err_def)
-apply (rule conjI)
- apply clarify
- apply (rename_tac A a u B b)
- apply (case_tac "A = B")
-  apply simp
- apply simp
-apply (rule conjI)
- apply clarify
- apply (rename_tac A a u B b)
- apply (case_tac "A = B")
-  apply simp
- apply simp
-apply clarify
-apply (rename_tac A ya yb B yd z C c a b)
-apply (case_tac "A = B")
- apply (case_tac "A = C")
-  apply simp
- apply (rotate_tac -1)
- apply simp
-apply (rotate_tac -1)
-apply (case_tac "B = C")
- apply simp
-apply (rotate_tac -1)
-apply simp
-done 
-
-end
--- a/src/HOL/MicroJava/BV/JType.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/BV/JType.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -1,12 +1,13 @@
 (*  Title:      HOL/MicroJava/BV/JType.thy
-    ID:         $Id$
     Author:     Tobias Nipkow, Gerwin Klein
     Copyright   2000 TUM
 *)
 
 header {* \isaheader{The Java Type System as Semilattice} *}
 
-theory JType imports "../J/WellForm" Err begin
+theory JType
+imports "../DFA/Semilattices" "../J/WellForm"
+begin
 
 constdefs
   super :: "'a prog \<Rightarrow> cname \<Rightarrow> cname"
@@ -34,7 +35,7 @@
 
   is_ty :: "'c prog \<Rightarrow> ty \<Rightarrow> bool"
   "is_ty G T == case T of PrimT P \<Rightarrow> True | RefT R \<Rightarrow>
-               (case R of NullT \<Rightarrow> True | ClassT C \<Rightarrow> (subcls1 G)^** C Object)"
+               (case R of NullT \<Rightarrow> True | ClassT C \<Rightarrow> (C, Object) \<in> (subcls1 G)^*)"
 
 
 translations
@@ -78,7 +79,7 @@
     have "R \<noteq> ClassT Object \<Longrightarrow> ?thesis"
      by (auto simp add: is_ty_def is_class_def split_tupled_all
                elim!: subcls1.cases
-               elim: converse_rtranclpE
+               elim: converse_rtranclE
                split: ref_ty.splits)
     ultimately    
     show ?thesis by blast
@@ -86,7 +87,7 @@
 qed
 
 lemma order_widen:
-  "acyclicP (subcls1 G) \<Longrightarrow> order (subtype G)"
+  "acyclic (subcls1 G) \<Longrightarrow> order (subtype G)"
   apply (unfold Semilat.order_def lesub_def subtype_def)
   apply (auto intro: widen_trans)
   apply (case_tac x)
@@ -102,16 +103,16 @@
   apply (case_tac ref_tya)
    apply simp
   apply simp
-  apply (auto dest: acyclic_impl_antisym_rtrancl [to_pred] antisymD)
+  apply (auto dest: acyclic_impl_antisym_rtrancl antisymD)
   done
 
 lemma wf_converse_subcls1_impl_acc_subtype:
-  "wfP ((subcls1 G)^--1) \<Longrightarrow> acc (subtype G)"
+  "wf ((subcls1 G)^-1) \<Longrightarrow> acc (subtype G)"
 apply (unfold Semilat.acc_def lesssub_def)
-apply (drule_tac p = "inf ((subcls1 G)^--1) op \<noteq>" in wfP_subset)
+apply (drule_tac p = "((subcls1 G)^-1) - Id" in wf_subset)
  apply auto
-apply (drule wfP_trancl)
-apply (simp add: wfP_eq_minimal)
+apply (drule wf_trancl)
+apply (simp add: wf_eq_minimal)
 apply clarify
 apply (unfold lesub_def subtype_def)
 apply (rename_tac M T) 
@@ -146,20 +147,20 @@
 apply (case_tac t)
  apply simp
 apply simp
-apply (insert rtranclp_r_diff_Id [symmetric, standard, of "subcls1 G"])
+apply (insert rtrancl_r_diff_Id [symmetric, standard, of "subcls1 G"])
 apply simp
-apply (erule rtranclp.cases)
+apply (erule rtrancl.cases)
  apply blast
-apply (drule rtranclp_converseI)
-apply (subgoal_tac "(inf (subcls1 G) op \<noteq>)^--1 = (inf ((subcls1 G)^--1) op \<noteq>)")
+apply (drule rtrancl_converseI)
+apply (subgoal_tac "(subcls1 G - Id)^-1 = (subcls1 G)^-1 - Id")
  prefer 2
- apply (simp add: converse_meet)
+ apply (simp add: converse_Int) apply safe[1]
 apply simp
-apply (blast intro: rtranclp_into_tranclp2)
-done 
+apply (blast intro: rtrancl_into_trancl2)
+done
 
 lemma closed_err_types:
-  "\<lbrakk> ws_prog G; single_valuedP (subcls1 G); acyclicP (subcls1 G) \<rbrakk> 
+  "\<lbrakk> ws_prog G; single_valued (subcls1 G); acyclic (subcls1 G) \<rbrakk> 
   \<Longrightarrow> closed (err (types G)) (lift2 (sup G))"
   apply (unfold closed_def plussub_def lift2_def sup_def)
   apply (auto split: err.split)
@@ -171,13 +172,13 @@
 
 
 lemma sup_subtype_greater:
-  "\<lbrakk> ws_prog G; single_valuedP (subcls1 G); acyclicP (subcls1 G);
+  "\<lbrakk> ws_prog G; single_valued (subcls1 G); acyclic (subcls1 G);
       is_type G t1; is_type G t2; sup G t1 t2 = OK s \<rbrakk> 
   \<Longrightarrow> subtype G t1 s \<and> subtype G t2 s"
 proof -
   assume ws_prog:       "ws_prog G"
-  assume single_valued: "single_valuedP (subcls1 G)"
-  assume acyclic:       "acyclicP (subcls1 G)"
+  assume single_valued: "single_valued (subcls1 G)"
+  assume acyclic:       "acyclic (subcls1 G)"
  
   { fix c1 c2
     assume is_class: "is_class G c1" "is_class G c2"
@@ -188,7 +189,7 @@
       by (blast intro: subcls_C_Object)
     with ws_prog single_valued
     obtain u where
-      "is_lub ((subcls1 G)^** ) c1 c2 u"
+      "is_lub ((subcls1 G)^* ) c1 c2 u"
       by (blast dest: single_valued_has_lubs)
     moreover
     note acyclic
@@ -210,14 +211,14 @@
 qed
 
 lemma sup_subtype_smallest:
-  "\<lbrakk> ws_prog G; single_valuedP (subcls1 G); acyclicP (subcls1 G);
+  "\<lbrakk> ws_prog G; single_valued (subcls1 G); acyclic (subcls1 G);
       is_type G a; is_type G b; is_type G c; 
       subtype G a c; subtype G b c; sup G a b = OK d \<rbrakk>
   \<Longrightarrow> subtype G d c"
 proof -
   assume ws_prog:       "ws_prog G"
-  assume single_valued: "single_valuedP (subcls1 G)"
-  assume acyclic:       "acyclicP (subcls1 G)"
+  assume single_valued: "single_valued (subcls1 G)"
+  assume acyclic:       "acyclic (subcls1 G)"
 
   { fix c1 c2 D
     assume is_class: "is_class G c1" "is_class G c2"
@@ -229,7 +230,7 @@
       by (blast intro: subcls_C_Object)
     with ws_prog single_valued
     obtain u where
-      lub: "is_lub ((subcls1 G)^** ) c1 c2 u"
+      lub: "is_lub ((subcls1 G)^*) c1 c2 u"
       by (blast dest: single_valued_has_lubs)   
     with acyclic
     have "exec_lub (subcls1 G) (super G) c1 c2 = u"
@@ -260,12 +261,12 @@
            split: ty.splits ref_ty.splits)
 
 lemma err_semilat_JType_esl_lemma:
-  "\<lbrakk> ws_prog G; single_valuedP (subcls1 G); acyclicP (subcls1 G) \<rbrakk> 
+  "\<lbrakk> ws_prog G; single_valued (subcls1 G); acyclic (subcls1 G) \<rbrakk> 
   \<Longrightarrow> err_semilat (esl G)"
 proof -
   assume ws_prog:   "ws_prog G"
-  assume single_valued: "single_valuedP (subcls1 G)"
-  assume acyclic:   "acyclicP (subcls1 G)"
+  assume single_valued: "single_valued (subcls1 G)"
+  assume acyclic:   "acyclic (subcls1 G)"
   
   hence "order (subtype G)"
     by (rule order_widen)
@@ -275,10 +276,10 @@
     by (rule closed_err_types)
   moreover
 
-  from ws_prog single_valued acyclic 
+  from ws_prog single_valued acyclic
   have
-    "(\<forall>x\<in>err (types G). \<forall>y\<in>err (types G). x <=_(le (subtype G)) x +_(lift2 (sup G)) y) \<and> 
-     (\<forall>x\<in>err (types G). \<forall>y\<in>err (types G). y <=_(le (subtype G)) x +_(lift2 (sup G)) y)"
+    "(\<forall>x\<in>err (types G). \<forall>y\<in>err (types G). x <=_(Err.le (subtype G)) x +_(lift2 (sup G)) y) \<and> 
+     (\<forall>x\<in>err (types G). \<forall>y\<in>err (types G). y <=_(Err.le (subtype G)) x +_(lift2 (sup G)) y)"
     by (auto simp add: lesub_def plussub_def Err.le_def lift2_def sup_subtype_greater split: err.split)
 
   moreover
@@ -286,18 +287,18 @@
   from ws_prog single_valued acyclic 
   have
     "\<forall>x\<in>err (types G). \<forall>y\<in>err (types G). \<forall>z\<in>err (types G). 
-    x <=_(le (subtype G)) z \<and> y <=_(le (subtype G)) z \<longrightarrow> x +_(lift2 (sup G)) y <=_(le (subtype G)) z"
+    x <=_(Err.le (subtype G)) z \<and> y <=_(Err.le (subtype G)) z \<longrightarrow> x +_(lift2 (sup G)) y <=_(Err.le (subtype G)) z"
     by (unfold lift2_def plussub_def lesub_def Err.le_def)
        (auto intro: sup_subtype_smallest sup_exists split: err.split)
 
   ultimately
   
   show ?thesis
-    by (unfold esl_def semilat_def sl_def) auto
+    by (unfold esl_def semilat_def Err.sl_def) auto
 qed
 
 lemma single_valued_subcls1:
-  "ws_prog G \<Longrightarrow> single_valuedP (subcls1 G)"
+  "ws_prog G \<Longrightarrow> single_valued (subcls1 G)"
   by (auto simp add: ws_prog_def unique_def single_valued_def
     intro: subcls1I elim!: subcls1.cases)
 
--- a/src/HOL/MicroJava/BV/JVM.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/BV/JVM.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -5,8 +5,9 @@
 
 header {* \isaheader{Kildall for the JVM}\label{sec:JVM} *}
 
-theory JVM imports Kildall Typing_Framework_JVM begin
-
+theory JVM
+imports Typing_Framework_JVM
+begin
 
 constdefs
   kiljvm :: "jvm_prog \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> exception_table \<Rightarrow> 
@@ -39,7 +40,7 @@
          simp add: symmetric sl_triple_conv)
       apply (simp (no_asm) add: JVM_le_unfold)
       apply (blast intro!: order_widen wf_converse_subcls1_impl_acc_subtype
-                   dest: wf_subcls1 wfP_acyclicP wf_prog_ws_prog)
+                   dest: wf_subcls1 wf_acyclic wf_prog_ws_prog)
      apply (simp add: JVM_le_unfold)
     apply (erule exec_pres_type)
    apply assumption
--- a/src/HOL/MicroJava/BV/JVMType.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/BV/JVMType.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -1,13 +1,13 @@
 (*  Title:      HOL/MicroJava/BV/JVM.thy
-    ID:         $Id$
     Author:     Gerwin Klein
     Copyright   2000 TUM
-
 *)
 
 header {* \isaheader{The JVM Type System as Semilattice} *}
 
-theory JVMType imports Opt Product Listn JType begin
+theory JVMType
+imports JType
+begin
 
 types
   locvars_type = "ty err list"
--- a/src/HOL/MicroJava/BV/Kildall.thy	Wed Dec 02 12:04:07 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,498 +0,0 @@
-(*  Title:      HOL/MicroJava/BV/Kildall.thy
-    ID:         $Id$
-    Author:     Tobias Nipkow, Gerwin Klein
-    Copyright   2000 TUM
-
-Kildall's algorithm
-*)
-
-header {* \isaheader{Kildall's Algorithm}\label{sec:Kildall} *}
-
-theory Kildall
-imports SemilatAlg While_Combinator
-begin
-
-
-consts
- iter :: "'s binop \<Rightarrow> 's step_type \<Rightarrow>
-          's list \<Rightarrow> nat set \<Rightarrow> 's list \<times> nat set"
- propa :: "'s binop \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's list \<Rightarrow> nat set \<Rightarrow> 's list * nat set"
-
-primrec
-"propa f []      ss w = (ss,w)"
-"propa f (q'#qs) ss w = (let (q,t) = q';
-                             u = t +_f ss!q;
-                             w' = (if u = ss!q then w else insert q w)
-                         in propa f qs (ss[q := u]) w')"
-
-defs iter_def:
-"iter f step ss w ==
- while (%(ss,w). w \<noteq> {})
-       (%(ss,w). let p = SOME p. p \<in> w
-                 in propa f (step p (ss!p)) ss (w-{p}))
-       (ss,w)"
-
-constdefs
- unstables :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> nat set"
-"unstables r step ss == {p. p < size ss \<and> \<not>stable r step ss p}"
-
- kildall :: "'s ord \<Rightarrow> 's binop \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> 's list"
-"kildall r f step ss == fst(iter f step ss (unstables r step ss))"
-
-consts merges :: "'s binop \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's list \<Rightarrow> 's list"
-primrec
-"merges f []      ss = ss"
-"merges f (p'#ps) ss = (let (p,s) = p' in merges f ps (ss[p := s +_f ss!p]))"
-
-
-lemmas [simp] = Let_def Semilat.le_iff_plus_unchanged [OF Semilat.intro, symmetric]
-
-
-lemma (in Semilat) nth_merges:
- "\<And>ss. \<lbrakk>p < length ss; ss \<in> list n A; \<forall>(p,t)\<in>set ps. p<n \<and> t\<in>A \<rbrakk> \<Longrightarrow>
-  (merges f ps ss)!p = map snd [(p',t') \<leftarrow> ps. p'=p] ++_f ss!p"
-  (is "\<And>ss. \<lbrakk>_; _; ?steptype ps\<rbrakk> \<Longrightarrow> ?P ss ps")
-proof (induct ps)
-  show "\<And>ss. ?P ss []" by simp
-
-  fix ss p' ps'
-  assume ss: "ss \<in> list n A"
-  assume l:  "p < length ss"
-  assume "?steptype (p'#ps')"
-  then obtain a b where
-    p': "p'=(a,b)" and ab: "a<n" "b\<in>A" and ps': "?steptype ps'"
-    by (cases p') auto
-  assume "\<And>ss. p< length ss \<Longrightarrow> ss \<in> list n A \<Longrightarrow> ?steptype ps' \<Longrightarrow> ?P ss ps'"
-  from this [OF _ _ ps'] have IH: "\<And>ss. ss \<in> list n A \<Longrightarrow> p < length ss \<Longrightarrow> ?P ss ps'" .
-
-  from ss ab
-  have "ss[a := b +_f ss!a] \<in> list n A" by (simp add: closedD)
-  moreover
-  from calculation l
-  have "p < length (ss[a := b +_f ss!a])" by simp
-  ultimately
-  have "?P (ss[a := b +_f ss!a]) ps'" by (rule IH)
-  with p' l
-  show "?P ss (p'#ps')" by simp
-qed
-
-
-(** merges **)
-
-lemma length_merges [rule_format, simp]:
-  "\<forall>ss. size(merges f ps ss) = size ss"
-  by (induct_tac ps, auto)
-
-
-lemma (in Semilat) merges_preserves_type_lemma:
-shows "\<forall>xs. xs \<in> list n A \<longrightarrow> (\<forall>(p,x) \<in> set ps. p<n \<and> x\<in>A)
-          \<longrightarrow> merges f ps xs \<in> list n A"
-apply (insert closedI)
-apply (unfold closed_def)
-apply (induct_tac ps)
- apply simp
-apply clarsimp
-done
-
-lemma (in Semilat) merges_preserves_type [simp]:
- "\<lbrakk> xs \<in> list n A; \<forall>(p,x) \<in> set ps. p<n \<and> x\<in>A \<rbrakk>
-  \<Longrightarrow> merges f ps xs \<in> list n A"
-by (simp add: merges_preserves_type_lemma)
-
-lemma (in Semilat) merges_incr_lemma:
- "\<forall>xs. xs \<in> list n A \<longrightarrow> (\<forall>(p,x)\<in>set ps. p<size xs \<and> x \<in> A) \<longrightarrow> xs <=[r] merges f ps xs"
-apply (induct_tac ps)
- apply simp
-apply simp
-apply clarify
-apply (rule order_trans)
-  apply simp
- apply (erule list_update_incr)
-  apply simp
- apply simp
-apply (blast intro!: listE_set intro: closedD listE_length [THEN nth_in])
-done
-
-lemma (in Semilat) merges_incr:
- "\<lbrakk> xs \<in> list n A; \<forall>(p,x)\<in>set ps. p<size xs \<and> x \<in> A \<rbrakk> 
-  \<Longrightarrow> xs <=[r] merges f ps xs"
-  by (simp add: merges_incr_lemma)
-
-
-lemma (in Semilat) merges_same_conv [rule_format]:
- "(\<forall>xs. xs \<in> list n A \<longrightarrow> (\<forall>(p,x)\<in>set ps. p<size xs \<and> x\<in>A) \<longrightarrow> 
-     (merges f ps xs = xs) = (\<forall>(p,x)\<in>set ps. x <=_r xs!p))"
-  apply (induct_tac ps)
-   apply simp
-  apply clarsimp
-  apply (rename_tac p x ps xs)
-  apply (rule iffI)
-   apply (rule context_conjI)
-    apply (subgoal_tac "xs[p := x +_f xs!p] <=[r] xs")
-     apply (drule_tac p = p in le_listD)
-      apply simp
-     apply simp
-    apply (erule subst, rule merges_incr)
-       apply (blast intro!: listE_set intro: closedD listE_length [THEN nth_in])
-      apply clarify
-      apply (rule conjI)
-       apply simp
-       apply (blast dest: boundedD)
-      apply blast
-   apply clarify
-   apply (erule allE)
-   apply (erule impE)
-    apply assumption
-   apply (drule bspec)
-    apply assumption
-   apply (simp add: le_iff_plus_unchanged [THEN iffD1] list_update_same_conv [THEN iffD2])
-   apply blast
-  apply clarify 
-  apply (simp add: le_iff_plus_unchanged [THEN iffD1] list_update_same_conv [THEN iffD2])
-  done
-
-
-lemma (in Semilat) list_update_le_listI [rule_format]:
-  "set xs <= A \<longrightarrow> set ys <= A \<longrightarrow> xs <=[r] ys \<longrightarrow> p < size xs \<longrightarrow>  
-   x <=_r ys!p \<longrightarrow> x\<in>A \<longrightarrow> xs[p := x +_f xs!p] <=[r] ys"
-  apply(insert semilat)
-  apply (unfold Listn.le_def lesub_def semilat_def)
-  apply (simp add: list_all2_conv_all_nth nth_list_update)
-  done
-
-lemma (in Semilat) merges_pres_le_ub:
-  assumes "set ts <= A" and "set ss <= A"
-    and "\<forall>(p,t)\<in>set ps. t <=_r ts!p \<and> t \<in> A \<and> p < size ts" and "ss <=[r] ts"
-  shows "merges f ps ss <=[r] ts"
-proof -
-  { fix t ts ps
-    have
-    "\<And>qs. \<lbrakk>set ts <= A; \<forall>(p,t)\<in>set ps. t <=_r ts!p \<and> t \<in> A \<and> p< size ts \<rbrakk> \<Longrightarrow>
-    set qs <= set ps  \<longrightarrow> 
-    (\<forall>ss. set ss <= A \<longrightarrow> ss <=[r] ts \<longrightarrow> merges f qs ss <=[r] ts)"
-    apply (induct_tac qs)
-     apply simp
-    apply (simp (no_asm_simp))
-    apply clarify
-    apply (rotate_tac -2)
-    apply simp
-    apply (erule allE, erule impE, erule_tac [2] mp)
-     apply (drule bspec, assumption)
-     apply (simp add: closedD)
-    apply (drule bspec, assumption)
-    apply (simp add: list_update_le_listI)
-    done 
-  } note this [dest]
-  
-  from prems show ?thesis by blast
-qed
-
-
-(** propa **)
-
-
-lemma decomp_propa:
-  "\<And>ss w. (\<forall>(q,t)\<in>set qs. q < size ss) \<Longrightarrow> 
-   propa f qs ss w = 
-   (merges f qs ss, {q. \<exists>t. (q,t)\<in>set qs \<and> t +_f ss!q \<noteq> ss!q} Un w)"
-  apply (induct qs)
-   apply simp   
-  apply (simp (no_asm))
-  apply clarify  
-  apply simp
-  apply (rule conjI) 
-   apply blast
-  apply (simp add: nth_list_update)
-  apply blast
-  done 
-
-(** iter **)
-
-lemma (in Semilat) stable_pres_lemma:
-shows "\<lbrakk>pres_type step n A; bounded step n; 
-     ss \<in> list n A; p \<in> w; \<forall>q\<in>w. q < n; 
-     \<forall>q. q < n \<longrightarrow> q \<notin> w \<longrightarrow> stable r step ss q; q < n; 
-     \<forall>s'. (q,s') \<in> set (step p (ss ! p)) \<longrightarrow> s' +_f ss ! q = ss ! q; 
-     q \<notin> w \<or> q = p \<rbrakk> 
-  \<Longrightarrow> stable r step (merges f (step p (ss!p)) ss) q"
-  apply (unfold stable_def)
-  apply (subgoal_tac "\<forall>s'. (q,s') \<in> set (step p (ss!p)) \<longrightarrow> s' : A")
-   prefer 2
-   apply clarify
-   apply (erule pres_typeD)
-    prefer 3 apply assumption
-    apply (rule listE_nth_in)
-     apply assumption
-    apply simp
-   apply simp
-  apply simp
-  apply clarify
-  apply (subst nth_merges)
-       apply simp
-       apply (blast dest: boundedD)
-      apply assumption
-     apply clarify
-     apply (rule conjI)
-      apply (blast dest: boundedD)
-     apply (erule pres_typeD)
-       prefer 3 apply assumption
-      apply simp
-     apply simp
-apply(subgoal_tac "q < length ss")
-prefer 2 apply simp
-  apply (frule nth_merges [of q _ _ "step p (ss!p)"]) (* fixme: why does method subst not work?? *)
-apply assumption
-  apply clarify
-  apply (rule conjI)
-   apply (blast dest: boundedD)
-  apply (erule pres_typeD)
-     prefer 3 apply assumption
-    apply simp
-   apply simp
-  apply (drule_tac P = "\<lambda>x. (a, b) \<in> set (step q x)" in subst)
-   apply assumption
-
- apply (simp add: plusplus_empty)
- apply (cases "q \<in> w")
-  apply simp
-  apply (rule ub1')
-     apply (rule semilat)
-    apply clarify
-    apply (rule pres_typeD)
-       apply assumption
-      prefer 3 apply assumption
-     apply (blast intro: listE_nth_in dest: boundedD)
-    apply (blast intro: pres_typeD dest: boundedD)
-   apply (blast intro: listE_nth_in dest: boundedD)
-  apply assumption
-
- apply simp
- apply (erule allE, erule impE, assumption, erule impE, assumption)
- apply (rule order_trans)
-   apply simp
-  defer
- apply (rule pp_ub2)(*
-    apply assumption*)
-   apply simp
-   apply clarify
-   apply simp
-   apply (rule pres_typeD)
-      apply assumption
-     prefer 3 apply assumption
-    apply (blast intro: listE_nth_in dest: boundedD)
-   apply (blast intro: pres_typeD dest: boundedD)
-  apply (blast intro: listE_nth_in dest: boundedD)
- apply blast
- done
-
-
-lemma (in Semilat) merges_bounded_lemma:
- "\<lbrakk> mono r step n A; bounded step n; 
-    \<forall>(p',s') \<in> set (step p (ss!p)). s' \<in> A; ss \<in> list n A; ts \<in> list n A; p < n; 
-    ss <=[r] ts; \<forall>p. p < n \<longrightarrow> stable r step ts p \<rbrakk> 
-  \<Longrightarrow> merges f (step p (ss!p)) ss <=[r] ts" 
-  apply (unfold stable_def)
-  apply (rule merges_pres_le_ub)
-     apply simp
-    apply simp
-   prefer 2 apply assumption
-
-  apply clarsimp
-  apply (drule boundedD, assumption+)
-  apply (erule allE, erule impE, assumption)
-  apply (drule bspec, assumption)
-  apply simp
-
-  apply (drule monoD [of _ _ _ _ p "ss!p"  "ts!p"])
-     apply assumption
-    apply simp
-   apply (simp add: le_listD)
-  
-  apply (drule lesub_step_typeD, assumption) 
-  apply clarify
-  apply (drule bspec, assumption)
-  apply simp
-  apply (blast intro: order_trans)
-  done
-
-lemma termination_lemma:
-  assumes semilat: "semilat (A, r, f)"
-  shows "\<lbrakk> ss \<in> list n A; \<forall>(q,t)\<in>set qs. q<n \<and> t\<in>A; p\<in>w \<rbrakk> \<Longrightarrow> 
-  ss <[r] merges f qs ss \<or> 
-  merges f qs ss = ss \<and> {q. \<exists>t. (q,t)\<in>set qs \<and> t +_f ss!q \<noteq> ss!q} Un (w-{p}) < w" (is "PROP ?P")
-proof -
-  interpret Semilat A r f using assms by (rule Semilat.intro)
-  show "PROP ?P" apply(insert semilat)
-    apply (unfold lesssub_def)
-    apply (simp (no_asm_simp) add: merges_incr)
-    apply (rule impI)
-    apply (rule merges_same_conv [THEN iffD1, elim_format]) 
-    apply assumption+
-    defer
-    apply (rule sym, assumption)
-    defer apply simp
-    apply (subgoal_tac "\<forall>q t. \<not>((q, t) \<in> set qs \<and> t +_f ss ! q \<noteq> ss ! q)")
-    apply (blast intro!: psubsetI elim: equalityE)
-    apply clarsimp
-    apply (drule bspec, assumption) 
-    apply (drule bspec, assumption)
-    apply clarsimp
-    done
-qed
-
-lemma iter_properties[rule_format]:
-  assumes semilat: "semilat (A, r, f)"
-  shows "\<lbrakk> acc r ; pres_type step n A; mono r step n A;
-     bounded step n; \<forall>p\<in>w0. p < n; ss0 \<in> list n A;
-     \<forall>p<n. p \<notin> w0 \<longrightarrow> stable r step ss0 p \<rbrakk> \<Longrightarrow>
-   iter f step ss0 w0 = (ss',w')
-   \<longrightarrow>
-   ss' \<in> list n A \<and> stables r step ss' \<and> ss0 <=[r] ss' \<and>
-   (\<forall>ts\<in>list n A. ss0 <=[r] ts \<and> stables r step ts \<longrightarrow> ss' <=[r] ts)"
-  (is "PROP ?P")
-proof -
-  interpret Semilat A r f using assms by (rule Semilat.intro)
-  show "PROP ?P" apply(insert semilat)
-apply (unfold iter_def stables_def)
-apply (rule_tac P = "%(ss,w).
- ss \<in> list n A \<and> (\<forall>p<n. p \<notin> w \<longrightarrow> stable r step ss p) \<and> ss0 <=[r] ss \<and>
- (\<forall>ts\<in>list n A. ss0 <=[r] ts \<and> stables r step ts \<longrightarrow> ss <=[r] ts) \<and>
- (\<forall>p\<in>w. p < n)" and
- r = "{(ss',ss) . ss <[r] ss'} <*lex*> finite_psubset"
-       in while_rule)
-
--- "Invariant holds initially:"
-apply (simp add:stables_def)
-
--- "Invariant is preserved:"
-apply(simp add: stables_def split_paired_all)
-apply(rename_tac ss w)
-apply(subgoal_tac "(SOME p. p \<in> w) \<in> w")
- prefer 2; apply (fast intro: someI)
-apply(subgoal_tac "\<forall>(q,t) \<in> set (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))). q < length ss \<and> t \<in> A")
- prefer 2
- apply clarify
- apply (rule conjI)
-  apply(clarsimp, blast dest!: boundedD)
- apply (erule pres_typeD)
-  prefer 3
-  apply assumption
-  apply (erule listE_nth_in)
-  apply simp
- apply simp
-apply (subst decomp_propa)
- apply fast
-apply simp
-apply (rule conjI)
- apply (rule merges_preserves_type)
- apply blast
- apply clarify
- apply (rule conjI)
-  apply(clarsimp, fast dest!: boundedD)
- apply (erule pres_typeD)
-  prefer 3
-  apply assumption
-  apply (erule listE_nth_in)
-  apply blast
- apply blast
-apply (rule conjI)
- apply clarify
- apply (blast intro!: stable_pres_lemma)
-apply (rule conjI)
- apply (blast intro!: merges_incr intro: le_list_trans)
-apply (rule conjI)
- apply clarsimp
- apply (blast intro!: merges_bounded_lemma)
-apply (blast dest!: boundedD)
-
-
--- "Postcondition holds upon termination:"
-apply(clarsimp simp add: stables_def split_paired_all)
-
--- "Well-foundedness of the termination relation:"
-apply (rule wf_lex_prod)
- apply (insert orderI [THEN acc_le_listI])
- apply (simp add: acc_def lesssub_def wfP_wf_eq [symmetric])
-apply (rule wf_finite_psubset) 
-
--- "Loop decreases along termination relation:"
-apply(simp add: stables_def split_paired_all)
-apply(rename_tac ss w)
-apply(subgoal_tac "(SOME p. p \<in> w) \<in> w")
- prefer 2; apply (fast intro: someI)
-apply(subgoal_tac "\<forall>(q,t) \<in> set (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))). q < length ss \<and> t \<in> A")
- prefer 2
- apply clarify
- apply (rule conjI)
-  apply(clarsimp, blast dest!: boundedD)
- apply (erule pres_typeD)
-  prefer 3
-  apply assumption
-  apply (erule listE_nth_in)
-  apply blast
- apply blast
-apply (subst decomp_propa)
- apply blast
-apply clarify
-apply (simp del: listE_length
-    add: lex_prod_def finite_psubset_def 
-         bounded_nat_set_is_finite)
-apply (rule termination_lemma)
-apply assumption+
-defer
-apply assumption
-apply clarsimp
-done
-
-qed
-
-lemma kildall_properties:
-assumes semilat: "semilat (A, r, f)"
-shows "\<lbrakk> acc r; pres_type step n A; mono r step n A;
-     bounded step n; ss0 \<in> list n A \<rbrakk> \<Longrightarrow>
-  kildall r f step ss0 \<in> list n A \<and>
-  stables r step (kildall r f step ss0) \<and>
-  ss0 <=[r] kildall r f step ss0 \<and>
-  (\<forall>ts\<in>list n A. ss0 <=[r] ts \<and> stables r step ts \<longrightarrow>
-                 kildall r f step ss0 <=[r] ts)"
-  (is "PROP ?P")
-proof -
-  interpret Semilat A r f using assms by (rule Semilat.intro)
-  show "PROP ?P"
-apply (unfold kildall_def)
-apply(case_tac "iter f step ss0 (unstables r step ss0)")
-apply(simp)
-apply (rule iter_properties)
-apply (simp_all add: unstables_def stable_def)
-apply (rule semilat)
-done
-qed
-
-lemma is_bcv_kildall:
-assumes semilat: "semilat (A, r, f)"
-shows "\<lbrakk> acc r; top r T; pres_type step n A; bounded step n; mono r step n A \<rbrakk>
-  \<Longrightarrow> is_bcv r T step n A (kildall r f step)"
-  (is "PROP ?P")
-proof -
-  interpret Semilat A r f using assms by (rule Semilat.intro)
-  show "PROP ?P"
-apply(unfold is_bcv_def wt_step_def)
-apply(insert semilat kildall_properties[of A])
-apply(simp add:stables_def)
-apply clarify
-apply(subgoal_tac "kildall r f step ss \<in> list n A")
- prefer 2 apply (simp(no_asm_simp))
-apply (rule iffI)
- apply (rule_tac x = "kildall r f step ss" in bexI) 
-  apply (rule conjI)
-   apply (blast)
-  apply (simp  (no_asm_simp))
- apply(assumption)
-apply clarify
-apply(subgoal_tac "kildall r f step ss!p <=_r ts!p")
- apply simp
-apply (blast intro!: le_listD less_lengthI)
-done
-qed
-
-end
--- a/src/HOL/MicroJava/BV/LBVComplete.thy	Wed Dec 02 12:04:07 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,379 +0,0 @@
-(*  Title:      HOL/MicroJava/BV/LBVComplete.thy
-    ID:         $Id$
-    Author:     Gerwin Klein
-    Copyright   2000 Technische Universitaet Muenchen
-*)
-
-header {* \isaheader{Completeness of the LBV} *}
-
-theory LBVComplete
-imports LBVSpec Typing_Framework
-begin
-
-constdefs
-  is_target :: "['s step_type, 's list, nat] \<Rightarrow> bool" 
-  "is_target step phi pc' \<equiv>
-     \<exists>pc s'. pc' \<noteq> pc+1 \<and> pc < length phi \<and> (pc',s') \<in> set (step pc (phi!pc))"
-
-  make_cert :: "['s step_type, 's list, 's] \<Rightarrow> 's certificate"
-  "make_cert step phi B \<equiv> 
-     map (\<lambda>pc. if is_target step phi pc then phi!pc else B) [0..<length phi] @ [B]"
-
-lemma [code]:
-  "is_target step phi pc' =
-  list_ex (\<lambda>pc. pc' \<noteq> pc+1 \<and> pc' mem (map fst (step pc (phi!pc)))) [0..<length phi]"
-by (force simp: list_ex_iff is_target_def mem_iff)
-
-
-locale lbvc = lbv + 
-  fixes phi :: "'a list" ("\<phi>")
-  fixes c   :: "'a list" 
-  defines cert_def: "c \<equiv> make_cert step \<phi> \<bottom>"
-
-  assumes mono: "mono r step (length \<phi>) A"
-  assumes pres: "pres_type step (length \<phi>) A" 
-  assumes phi:  "\<forall>pc < length \<phi>. \<phi>!pc \<in> A \<and> \<phi>!pc \<noteq> \<top>"
-  assumes bounded: "bounded step (length \<phi>)"
-
-  assumes B_neq_T: "\<bottom> \<noteq> \<top>" 
-
-
-lemma (in lbvc) cert: "cert_ok c (length \<phi>) \<top> \<bottom> A"
-proof (unfold cert_ok_def, intro strip conjI)  
-  note [simp] = make_cert_def cert_def nth_append 
-
-  show "c!length \<phi> = \<bottom>" by simp
-
-  fix pc assume pc: "pc < length \<phi>" 
-  from pc phi B_A show "c!pc \<in> A" by simp
-  from pc phi B_neq_T show "c!pc \<noteq> \<top>" by simp
-qed
-
-lemmas [simp del] = split_paired_Ex
-
-
-lemma (in lbvc) cert_target [intro?]:
-  "\<lbrakk> (pc',s') \<in> set (step pc (\<phi>!pc));
-      pc' \<noteq> pc+1; pc < length \<phi>; pc' < length \<phi> \<rbrakk>
-  \<Longrightarrow> c!pc' = \<phi>!pc'"
-  by (auto simp add: cert_def make_cert_def nth_append is_target_def)
-
-
-lemma (in lbvc) cert_approx [intro?]:
-  "\<lbrakk> pc < length \<phi>; c!pc \<noteq> \<bottom> \<rbrakk>
-  \<Longrightarrow> c!pc = \<phi>!pc"
-  by (auto simp add: cert_def make_cert_def nth_append)
-
-
-lemma (in lbv) le_top [simp, intro]:
-  "x <=_r \<top>"
-  by (insert top) simp
-  
-
-lemma (in lbv) merge_mono:
-  assumes less:  "ss2 <=|r| ss1"
-  assumes x:     "x \<in> A"
-  assumes ss1:   "snd`set ss1 \<subseteq> A"
-  assumes ss2:   "snd`set ss2 \<subseteq> A"
-  shows "merge c pc ss2 x <=_r merge c pc ss1 x" (is "?s2 <=_r ?s1")
-proof-
-  have "?s1 = \<top> \<Longrightarrow> ?thesis" by simp
-  moreover {
-    assume merge: "?s1 \<noteq> T" 
-    from x ss1 have "?s1 =
-      (if \<forall>(pc', s')\<in>set ss1. pc' \<noteq> pc + 1 \<longrightarrow> s' <=_r c!pc'
-      then (map snd [(p', t') \<leftarrow> ss1 . p'=pc+1]) ++_f x
-      else \<top>)" 
-      by (rule merge_def)  
-    with merge obtain
-      app: "\<forall>(pc',s')\<in>set ss1. pc' \<noteq> pc+1 \<longrightarrow> s' <=_r c!pc'" 
-           (is "?app ss1") and
-      sum: "(map snd [(p',t') \<leftarrow> ss1 . p' = pc+1] ++_f x) = ?s1" 
-           (is "?map ss1 ++_f x = _" is "?sum ss1 = _")
-      by (simp split: split_if_asm)
-    from app less 
-    have "?app ss2" by (blast dest: trans_r lesub_step_typeD)
-    moreover {
-      from ss1 have map1: "set (?map ss1) \<subseteq> A" by auto
-      with x have "?sum ss1 \<in> A" by (auto intro!: plusplus_closed semilat)
-      with sum have "?s1 \<in> A" by simp
-      moreover    
-      have mapD: "\<And>x ss. x \<in> set (?map ss) \<Longrightarrow> \<exists>p. (p,x) \<in> set ss \<and> p=pc+1" by auto
-      from x map1 
-      have "\<forall>x \<in> set (?map ss1). x <=_r ?sum ss1"
-        by clarify (rule pp_ub1)
-      with sum have "\<forall>x \<in> set (?map ss1). x <=_r ?s1" by simp
-      with less have "\<forall>x \<in> set (?map ss2). x <=_r ?s1"
-        by (fastsimp dest!: mapD lesub_step_typeD intro: trans_r)
-      moreover 
-      from map1 x have "x <=_r (?sum ss1)" by (rule pp_ub2)
-      with sum have "x <=_r ?s1" by simp
-      moreover 
-      from ss2 have "set (?map ss2) \<subseteq> A" by auto
-      ultimately
-      have "?sum ss2 <=_r ?s1" using x by - (rule pp_lub)
-    }
-    moreover
-    from x ss2 have 
-      "?s2 =
-      (if \<forall>(pc', s')\<in>set ss2. pc' \<noteq> pc + 1 \<longrightarrow> s' <=_r c!pc'
-      then map snd [(p', t') \<leftarrow> ss2 . p' = pc + 1] ++_f x
-      else \<top>)" 
-      by (rule merge_def)
-    ultimately have ?thesis by simp
-  }
-  ultimately show ?thesis by (cases "?s1 = \<top>") auto
-qed
-
-
-lemma (in lbvc) wti_mono:
-  assumes less: "s2 <=_r s1"
-  assumes pc:   "pc < length \<phi>" 
-  assumes s1:   "s1 \<in> A"
-  assumes s2:   "s2 \<in> A"
-  shows "wti c pc s2 <=_r wti c pc s1" (is "?s2' <=_r ?s1'")
-proof -
-  from mono pc s2 less have "step pc s2 <=|r| step pc s1" by (rule monoD)
-  moreover
-  from cert B_A pc have "c!Suc pc \<in> A" by (rule cert_okD3)
-  moreover 
-  from pres s1 pc
-  have "snd`set (step pc s1) \<subseteq> A" by (rule pres_typeD2)
-  moreover
-  from pres s2 pc
-  have "snd`set (step pc s2) \<subseteq> A" by (rule pres_typeD2)
-  ultimately
-  show ?thesis by (simp add: wti merge_mono)
-qed 
-
-lemma (in lbvc) wtc_mono:
-  assumes less: "s2 <=_r s1"
-  assumes pc:   "pc < length \<phi>" 
-  assumes s1:   "s1 \<in> A"
-  assumes s2:   "s2 \<in> A"
-  shows "wtc c pc s2 <=_r wtc c pc s1" (is "?s2' <=_r ?s1'")
-proof (cases "c!pc = \<bottom>")
-  case True 
-  moreover from less pc s1 s2 have "wti c pc s2 <=_r wti c pc s1" by (rule wti_mono)
-  ultimately show ?thesis by (simp add: wtc)
-next
-  case False
-  have "?s1' = \<top> \<Longrightarrow> ?thesis" by simp
-  moreover {
-    assume "?s1' \<noteq> \<top>" 
-    with False have c: "s1 <=_r c!pc" by (simp add: wtc split: split_if_asm)
-    with less have "s2 <=_r c!pc" ..
-    with False c have ?thesis by (simp add: wtc)
-  }
-  ultimately show ?thesis by (cases "?s1' = \<top>") auto
-qed
-
-
-lemma (in lbv) top_le_conv [simp]:
-  "\<top> <=_r x = (x = \<top>)"
-  by (insert semilat) (simp add: top top_le_conv) 
-
-lemma (in lbv) neq_top [simp, elim]:
-  "\<lbrakk> x <=_r y; y \<noteq> \<top> \<rbrakk> \<Longrightarrow> x \<noteq> \<top>"
-  by (cases "x = T") auto
-
-
-lemma (in lbvc) stable_wti:
-  assumes stable:  "stable r step \<phi> pc"
-  assumes pc:      "pc < length \<phi>"
-  shows "wti c pc (\<phi>!pc) \<noteq> \<top>"
-proof -
-  let ?step = "step pc (\<phi>!pc)"
-  from stable 
-  have less: "\<forall>(q,s')\<in>set ?step. s' <=_r \<phi>!q" by (simp add: stable_def)
-  
-  from cert B_A pc 
-  have cert_suc: "c!Suc pc \<in> A" by (rule cert_okD3)
-  moreover  
-  from phi pc have "\<phi>!pc \<in> A" by simp
-  from pres this pc 
-  have stepA: "snd`set ?step \<subseteq> A" by (rule pres_typeD2) 
-  ultimately
-  have "merge c pc ?step (c!Suc pc) =
-    (if \<forall>(pc',s')\<in>set ?step. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'
-    then map snd [(p',t') \<leftarrow> ?step.p'=pc+1] ++_f c!Suc pc
-    else \<top>)" unfolding mrg_def by (rule lbv.merge_def [OF lbvc.axioms(1), OF lbvc_axioms])
-  moreover {
-    fix pc' s' assume s': "(pc', s') \<in> set ?step" and suc_pc: "pc' \<noteq> pc+1"
-    with less have "s' <=_r \<phi>!pc'" by auto
-    also 
-    from bounded pc s' have "pc' < length \<phi>" by (rule boundedD)
-    with s' suc_pc pc have "c!pc' = \<phi>!pc'" ..
-    hence "\<phi>!pc' = c!pc'" ..
-    finally have "s' <=_r c!pc'" .
-  } hence "\<forall>(pc',s')\<in>set ?step. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'" by auto
-  moreover
-  from pc have "Suc pc = length \<phi> \<or> Suc pc < length \<phi>" by auto
-  hence "map snd [(p',t') \<leftarrow> ?step.p'=pc+1] ++_f c!Suc pc \<noteq> \<top>" 
-         (is "?map ++_f _ \<noteq> _")
-  proof (rule disjE)
-    assume pc': "Suc pc = length \<phi>"
-    with cert have "c!Suc pc = \<bottom>" by (simp add: cert_okD2)
-    moreover 
-    from pc' bounded pc 
-    have "\<forall>(p',t')\<in>set ?step. p'\<noteq>pc+1" by clarify (drule boundedD, auto)
-    hence "[(p',t') \<leftarrow> ?step.p'=pc+1] = []" by (blast intro: filter_False) 
-    hence "?map = []" by simp
-    ultimately show ?thesis by (simp add: B_neq_T)  
-  next
-    assume pc': "Suc pc < length \<phi>"
-    from pc' phi have "\<phi>!Suc pc \<in> A" by simp
-    moreover note cert_suc
-    moreover from stepA 
-    have "set ?map \<subseteq> A" by auto
-    moreover
-    have "\<And>s. s \<in> set ?map \<Longrightarrow> \<exists>t. (Suc pc, t) \<in> set ?step" by auto
-    with less have "\<forall>s' \<in> set ?map. s' <=_r \<phi>!Suc pc" by auto
-    moreover
-    from pc' have "c!Suc pc <=_r \<phi>!Suc pc" 
-      by (cases "c!Suc pc = \<bottom>") (auto dest: cert_approx)
-    ultimately
-    have "?map ++_f c!Suc pc <=_r \<phi>!Suc pc" by (rule pp_lub)
-    moreover
-    from pc' phi have "\<phi>!Suc pc \<noteq> \<top>" by simp
-    ultimately
-    show ?thesis by auto
-  qed
-  ultimately
-  have "merge c pc ?step (c!Suc pc) \<noteq> \<top>" by simp
-  thus ?thesis by (simp add: wti)  
-qed
-
-lemma (in lbvc) wti_less:
-  assumes stable:  "stable r step \<phi> pc"
-  assumes suc_pc:   "Suc pc < length \<phi>"
-  shows "wti c pc (\<phi>!pc) <=_r \<phi>!Suc pc" (is "?wti <=_r _")
-proof -
-  let ?step = "step pc (\<phi>!pc)"
-
-  from stable 
-  have less: "\<forall>(q,s')\<in>set ?step. s' <=_r \<phi>!q" by (simp add: stable_def)
-   
-  from suc_pc have pc: "pc < length \<phi>" by simp
-  with cert B_A have cert_suc: "c!Suc pc \<in> A" by (rule cert_okD3)
-  moreover  
-  from phi pc have "\<phi>!pc \<in> A" by simp
-  with pres pc have stepA: "snd`set ?step \<subseteq> A" by - (rule pres_typeD2)
-  moreover
-  from stable pc have "?wti \<noteq> \<top>" by (rule stable_wti)
-  hence "merge c pc ?step (c!Suc pc) \<noteq> \<top>" by (simp add: wti)
-  ultimately
-  have "merge c pc ?step (c!Suc pc) =
-    map snd [(p',t')\<leftarrow> ?step.p'=pc+1] ++_f c!Suc pc" by (rule merge_not_top_s) 
-  hence "?wti = \<dots>" (is "_ = (?map ++_f _)" is "_ = ?sum") by (simp add: wti)
-  also {
-    from suc_pc phi have "\<phi>!Suc pc \<in> A" by simp
-    moreover note cert_suc
-    moreover from stepA have "set ?map \<subseteq> A" by auto
-    moreover
-    have "\<And>s. s \<in> set ?map \<Longrightarrow> \<exists>t. (Suc pc, t) \<in> set ?step" by auto
-    with less have "\<forall>s' \<in> set ?map. s' <=_r \<phi>!Suc pc" by auto
-    moreover
-    from suc_pc have "c!Suc pc <=_r \<phi>!Suc pc"
-      by (cases "c!Suc pc = \<bottom>") (auto dest: cert_approx)
-    ultimately
-    have "?sum <=_r \<phi>!Suc pc" by (rule pp_lub)
-  }
-  finally show ?thesis .
-qed
-
-lemma (in lbvc) stable_wtc:
-  assumes stable:  "stable r step phi pc"
-  assumes pc:      "pc < length \<phi>"
-  shows "wtc c pc (\<phi>!pc) \<noteq> \<top>"
-proof -
-  from stable pc have wti: "wti c pc (\<phi>!pc) \<noteq> \<top>" by (rule stable_wti)
-  show ?thesis
-  proof (cases "c!pc = \<bottom>")
-    case True with wti show ?thesis by (simp add: wtc)
-  next
-    case False
-    with pc have "c!pc = \<phi>!pc" ..    
-    with False wti show ?thesis by (simp add: wtc)
-  qed
-qed
-
-lemma (in lbvc) wtc_less:
-  assumes stable: "stable r step \<phi> pc"
-  assumes suc_pc: "Suc pc < length \<phi>"
-  shows "wtc c pc (\<phi>!pc) <=_r \<phi>!Suc pc" (is "?wtc <=_r _")
-proof (cases "c!pc = \<bottom>")
-  case True
-  moreover from stable suc_pc have "wti c pc (\<phi>!pc) <=_r \<phi>!Suc pc"
-    by (rule wti_less)
-  ultimately show ?thesis by (simp add: wtc)
-next
-  case False
-  from suc_pc have pc: "pc < length \<phi>" by simp
-  with stable have "?wtc \<noteq> \<top>" by (rule stable_wtc)
-  with False have "?wtc = wti c pc (c!pc)" 
-    by (unfold wtc) (simp split: split_if_asm)
-  also from pc False have "c!pc = \<phi>!pc" .. 
-  finally have "?wtc = wti c pc (\<phi>!pc)" .
-  also from stable suc_pc have "wti c pc (\<phi>!pc) <=_r \<phi>!Suc pc" by (rule wti_less)
-  finally show ?thesis .
-qed
-
-
-lemma (in lbvc) wt_step_wtl_lemma:
-  assumes wt_step: "wt_step r \<top> step \<phi>"
-  shows "\<And>pc s. pc+length ls = length \<phi> \<Longrightarrow> s <=_r \<phi>!pc \<Longrightarrow> s \<in> A \<Longrightarrow> s\<noteq>\<top> \<Longrightarrow>
-                wtl ls c pc s \<noteq> \<top>"
-  (is "\<And>pc s. _ \<Longrightarrow> _ \<Longrightarrow> _ \<Longrightarrow> _ \<Longrightarrow> ?wtl ls pc s \<noteq> _")
-proof (induct ls)
-  fix pc s assume "s\<noteq>\<top>" thus "?wtl [] pc s \<noteq> \<top>" by simp
-next
-  fix pc s i ls
-  assume "\<And>pc s. pc+length ls=length \<phi> \<Longrightarrow> s <=_r \<phi>!pc \<Longrightarrow> s \<in> A \<Longrightarrow> s\<noteq>\<top> \<Longrightarrow> 
-                  ?wtl ls pc s \<noteq> \<top>"
-  moreover
-  assume pc_l: "pc + length (i#ls) = length \<phi>"
-  hence suc_pc_l: "Suc pc + length ls = length \<phi>" by simp
-  ultimately
-  have IH: "\<And>s. s <=_r \<phi>!Suc pc \<Longrightarrow> s \<in> A \<Longrightarrow> s \<noteq> \<top> \<Longrightarrow> ?wtl ls (Suc pc) s \<noteq> \<top>" .
-
-  from pc_l obtain pc: "pc < length \<phi>" by simp
-  with wt_step have stable: "stable r step \<phi> pc" by (simp add: wt_step_def)
-  from this pc have wt_phi: "wtc c pc (\<phi>!pc) \<noteq> \<top>" by (rule stable_wtc)
-  assume s_phi: "s <=_r \<phi>!pc"
-  from phi pc have phi_pc: "\<phi>!pc \<in> A" by simp
-  assume s: "s \<in> A"
-  with s_phi pc phi_pc have wt_s_phi: "wtc c pc s <=_r wtc c pc (\<phi>!pc)" by (rule wtc_mono)
-  with wt_phi have wt_s: "wtc c pc s \<noteq> \<top>" by simp
-  moreover
-  assume s': "s \<noteq> \<top>" 
-  ultimately
-  have "ls = [] \<Longrightarrow> ?wtl (i#ls) pc s \<noteq> \<top>" by simp
-  moreover {
-    assume "ls \<noteq> []" 
-    with pc_l have suc_pc: "Suc pc < length \<phi>" by (auto simp add: neq_Nil_conv)
-    with stable have "wtc c pc (phi!pc) <=_r \<phi>!Suc pc" by (rule wtc_less)
-    with wt_s_phi have "wtc c pc s <=_r \<phi>!Suc pc" by (rule trans_r)      
-    moreover
-    from cert suc_pc have "c!pc \<in> A" "c!(pc+1) \<in> A" 
-      by (auto simp add: cert_ok_def)
-    from pres this s pc have "wtc c pc s \<in> A" by (rule wtc_pres)
-    ultimately
-    have "?wtl ls (Suc pc) (wtc c pc s) \<noteq> \<top>" using IH wt_s by blast
-    with s' wt_s have "?wtl (i#ls) pc s \<noteq> \<top>" by simp
-  }
-  ultimately show "?wtl (i#ls) pc s \<noteq> \<top>" by (cases ls) blast+
-qed
-
-  
-theorem (in lbvc) wtl_complete:
-  assumes wt: "wt_step r \<top> step \<phi>"
-    and s: "s <=_r \<phi>!0" "s \<in> A" "s \<noteq> \<top>"
-    and len: "length ins = length phi"
-  shows "wtl ins c 0 s \<noteq> \<top>"
-proof -
-  from len have "0+length ins = length phi" by simp
-  from wt this s show ?thesis by (rule wt_step_wtl_lemma)
-qed
-
-end
--- a/src/HOL/MicroJava/BV/LBVCorrect.thy	Wed Dec 02 12:04:07 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,223 +0,0 @@
-(*
-    ID:         $Id$
-    Author:     Gerwin Klein
-    Copyright   1999 Technische Universitaet Muenchen
-*)
-
-header {* \isaheader{Correctness of the LBV} *}
-
-theory LBVCorrect
-imports LBVSpec Typing_Framework
-begin
-
-locale lbvs = lbv +
-  fixes s0  :: 'a ("s\<^sub>0")
-  fixes c   :: "'a list"
-  fixes ins :: "'b list"
-  fixes phi :: "'a list" ("\<phi>")
-  defines phi_def:
-  "\<phi> \<equiv> map (\<lambda>pc. if c!pc = \<bottom> then wtl (take pc ins) c 0 s0 else c!pc) 
-       [0..<length ins]"
-
-  assumes bounded: "bounded step (length ins)"
-  assumes cert: "cert_ok c (length ins) \<top> \<bottom> A"
-  assumes pres: "pres_type step (length ins) A"
-
-
-lemma (in lbvs) phi_None [intro?]:
-  "\<lbrakk> pc < length ins; c!pc = \<bottom> \<rbrakk> \<Longrightarrow> \<phi> ! pc = wtl (take pc ins) c 0 s0"
-  by (simp add: phi_def)
-
-lemma (in lbvs) phi_Some [intro?]:
-  "\<lbrakk> pc < length ins; c!pc \<noteq> \<bottom> \<rbrakk> \<Longrightarrow> \<phi> ! pc = c ! pc"
-  by (simp add: phi_def)
-
-lemma (in lbvs) phi_len [simp]:
-  "length \<phi> = length ins"
-  by (simp add: phi_def)
-
-
-lemma (in lbvs) wtl_suc_pc:
-  assumes all: "wtl ins c 0 s\<^sub>0 \<noteq> \<top>" 
-  assumes pc:  "pc+1 < length ins"
-  shows "wtl (take (pc+1) ins) c 0 s0 \<le>\<^sub>r \<phi>!(pc+1)"
-proof -
-  from all pc
-  have "wtc c (pc+1) (wtl (take (pc+1) ins) c 0 s0) \<noteq> T" by (rule wtl_all)
-  with pc show ?thesis by (simp add: phi_def wtc split: split_if_asm)
-qed
-
-
-lemma (in lbvs) wtl_stable:
-  assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>" 
-  assumes s0:  "s0 \<in> A" 
-  assumes pc:  "pc < length ins" 
-  shows "stable r step \<phi> pc"
-proof (unfold stable_def, clarify)
-  fix pc' s' assume step: "(pc',s') \<in> set (step pc (\<phi> ! pc))" 
-                      (is "(pc',s') \<in> set (?step pc)")
-  
-  from bounded pc step have pc': "pc' < length ins" by (rule boundedD)
-
-  from wtl have tkpc: "wtl (take pc ins) c 0 s0 \<noteq> \<top>" (is "?s1 \<noteq> _") by (rule wtl_take)
-  from wtl have s2: "wtl (take (pc+1) ins) c 0 s0 \<noteq> \<top>" (is "?s2 \<noteq> _") by (rule wtl_take)
-  
-  from wtl pc have wt_s1: "wtc c pc ?s1 \<noteq> \<top>" by (rule wtl_all)
-
-  have c_Some: "\<forall>pc t. pc < length ins \<longrightarrow> c!pc \<noteq> \<bottom> \<longrightarrow> \<phi>!pc = c!pc" 
-    by (simp add: phi_def)
-  from pc have c_None: "c!pc = \<bottom> \<Longrightarrow> \<phi>!pc = ?s1" ..
-
-  from wt_s1 pc c_None c_Some
-  have inst: "wtc c pc ?s1  = wti c pc (\<phi>!pc)"
-    by (simp add: wtc split: split_if_asm)
-
-  from pres cert s0 wtl pc have "?s1 \<in> A" by (rule wtl_pres)
-  with pc c_Some cert c_None
-  have "\<phi>!pc \<in> A" by (cases "c!pc = \<bottom>") (auto dest: cert_okD1)
-  with pc pres
-  have step_in_A: "snd`set (?step pc) \<subseteq> A" by (auto dest: pres_typeD2)
-
-  show "s' <=_r \<phi>!pc'" 
-  proof (cases "pc' = pc+1")
-    case True
-    with pc' cert
-    have cert_in_A: "c!(pc+1) \<in> A" by (auto dest: cert_okD1)
-    from True pc' have pc1: "pc+1 < length ins" by simp
-    with tkpc have "?s2 = wtc c pc ?s1" by - (rule wtl_Suc)
-    with inst 
-    have merge: "?s2 = merge c pc (?step pc) (c!(pc+1))" by (simp add: wti)
-    also    
-    from s2 merge have "\<dots> \<noteq> \<top>" (is "?merge \<noteq> _") by simp
-    with cert_in_A step_in_A
-    have "?merge = (map snd [(p',t') \<leftarrow> ?step pc. p'=pc+1] ++_f (c!(pc+1)))"
-      by (rule merge_not_top_s) 
-    finally
-    have "s' <=_r ?s2" using step_in_A cert_in_A True step 
-      by (auto intro: pp_ub1')
-    also 
-    from wtl pc1 have "?s2 <=_r \<phi>!(pc+1)" by (rule wtl_suc_pc)
-    also note True [symmetric]
-    finally show ?thesis by simp    
-  next
-    case False
-    from wt_s1 inst
-    have "merge c pc (?step pc) (c!(pc+1)) \<noteq> \<top>" by (simp add: wti)
-    with step_in_A
-    have "\<forall>(pc', s')\<in>set (?step pc). pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'" 
-      by - (rule merge_not_top)
-    with step False 
-    have ok: "s' <=_r c!pc'" by blast
-    moreover
-    from ok
-    have "c!pc' = \<bottom> \<Longrightarrow> s' = \<bottom>" by simp
-    moreover
-    from c_Some pc'
-    have "c!pc' \<noteq> \<bottom> \<Longrightarrow> \<phi>!pc' = c!pc'" by auto
-    ultimately
-    show ?thesis by (cases "c!pc' = \<bottom>") auto 
-  qed
-qed
-
-  
-lemma (in lbvs) phi_not_top:
-  assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>"
-  assumes pc:  "pc < length ins"
-  shows "\<phi>!pc \<noteq> \<top>"
-proof (cases "c!pc = \<bottom>")
-  case False with pc
-  have "\<phi>!pc = c!pc" ..
-  also from cert pc have "\<dots> \<noteq> \<top>" by (rule cert_okD4)
-  finally show ?thesis .
-next
-  case True with pc
-  have "\<phi>!pc = wtl (take pc ins) c 0 s0" ..
-  also from wtl have "\<dots> \<noteq> \<top>" by (rule wtl_take)
-  finally show ?thesis .
-qed
-
-lemma (in lbvs) phi_in_A:
-  assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>"
-  assumes s0:  "s0 \<in> A"
-  shows "\<phi> \<in> list (length ins) A"
-proof -
-  { fix x assume "x \<in> set \<phi>"
-    then obtain xs ys where "\<phi> = xs @ x # ys" 
-      by (auto simp add: in_set_conv_decomp)
-    then obtain pc where pc: "pc < length \<phi>" and x: "\<phi>!pc = x"
-      by (simp add: that [of "length xs"] nth_append)
-    
-    from pres cert wtl s0 pc
-    have "wtl (take pc ins) c 0 s0 \<in> A" by (auto intro!: wtl_pres)
-    moreover
-    from pc have "pc < length ins" by simp
-    with cert have "c!pc \<in> A" ..
-    ultimately
-    have "\<phi>!pc \<in> A" using pc by (simp add: phi_def)
-    hence "x \<in> A" using x by simp
-  } 
-  hence "set \<phi> \<subseteq> A" ..
-  thus ?thesis by (unfold list_def) simp
-qed
-
-
-lemma (in lbvs) phi0:
-  assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>"
-  assumes 0:   "0 < length ins"
-  shows "s0 <=_r \<phi>!0"
-proof (cases "c!0 = \<bottom>")
-  case True
-  with 0 have "\<phi>!0 = wtl (take 0 ins) c 0 s0" ..
-  moreover have "wtl (take 0 ins) c 0 s0 = s0" by simp
-  ultimately have "\<phi>!0 = s0" by simp
-  thus ?thesis by simp
-next
-  case False
-  with 0 have "phi!0 = c!0" ..
-  moreover 
-  from wtl have "wtl (take 1 ins) c 0 s0 \<noteq> \<top>"  by (rule wtl_take)
-  with 0 False 
-  have "s0 <=_r c!0" by (auto simp add: neq_Nil_conv wtc split: split_if_asm)
-  ultimately
-  show ?thesis by simp
-qed
-
-
-theorem (in lbvs) wtl_sound:
-  assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>" 
-  assumes s0: "s0 \<in> A" 
-  shows "\<exists>ts. wt_step r \<top> step ts"
-proof -
-  have "wt_step r \<top> step \<phi>"
-  proof (unfold wt_step_def, intro strip conjI)
-    fix pc assume "pc < length \<phi>"
-    then have pc: "pc < length ins" by simp
-    with wtl show "\<phi>!pc \<noteq> \<top>" by (rule phi_not_top)
-    from wtl s0 pc show "stable r step \<phi> pc" by (rule wtl_stable)
-  qed
-  thus ?thesis ..
-qed
-
-
-theorem (in lbvs) wtl_sound_strong:
-  assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>" 
-  assumes s0: "s0 \<in> A" 
-  assumes nz: "0 < length ins"
-  shows "\<exists>ts \<in> list (length ins) A. wt_step r \<top> step ts \<and> s0 <=_r ts!0"
-proof -
-  from wtl s0 have "\<phi> \<in> list (length ins) A" by (rule phi_in_A)
-  moreover
-  have "wt_step r \<top> step \<phi>"
-  proof (unfold wt_step_def, intro strip conjI)
-    fix pc assume "pc < length \<phi>"
-    then have pc: "pc < length ins" by simp
-    with wtl show "\<phi>!pc \<noteq> \<top>" by (rule phi_not_top)
-    from wtl s0 pc show "stable r step \<phi> pc" by (rule wtl_stable)
-  qed
-  moreover
-  from wtl nz have "s0 <=_r \<phi>!0" by (rule phi0)
-  ultimately
-  show ?thesis by fast
-qed
-
-end
--- a/src/HOL/MicroJava/BV/LBVJVM.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/BV/LBVJVM.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -1,5 +1,4 @@
 (*  Title:      HOL/MicroJava/BV/JVM.thy
-    ID:         $Id$
     Author:     Tobias Nipkow, Gerwin Klein
     Copyright   2000 TUM
 *)
@@ -7,7 +6,7 @@
 header {* \isaheader{LBV for the JVM}\label{sec:JVM} *}
 
 theory LBVJVM
-imports LBVCorrect LBVComplete Typing_Framework_JVM
+imports Typing_Framework_JVM
 begin
 
 types prog_cert = "cname \<Rightarrow> sig \<Rightarrow> JVMType.state list"
--- a/src/HOL/MicroJava/BV/LBVSpec.thy	Wed Dec 02 12:04:07 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,381 +0,0 @@
-(*  Title:      HOL/MicroJava/BV/LBVSpec.thy
-    Author:     Gerwin Klein
-    Copyright   1999 Technische Universitaet Muenchen
-*)
-
-header {* \isaheader{The Lightweight Bytecode Verifier} *}
-
-theory LBVSpec
-imports SemilatAlg Opt
-begin
-
-types
-  's certificate = "'s list"   
-
-consts
-merge :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> nat \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's \<Rightarrow> 's"
-primrec
-"merge cert f r T pc []     x = x"
-"merge cert f r T pc (s#ss) x = merge cert f r T pc ss (let (pc',s') = s in 
-                                  if pc'=pc+1 then s' +_f x
-                                  else if s' <=_r (cert!pc') then x
-                                  else T)"
-
-constdefs
-wtl_inst :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow>
-             's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's"
-"wtl_inst cert f r T step pc s \<equiv> merge cert f r T pc (step pc s) (cert!(pc+1))"
-
-wtl_cert :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow>
-             's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's"
-"wtl_cert cert f r T B step pc s \<equiv>
-  if cert!pc = B then 
-    wtl_inst cert f r T step pc s
-  else
-    if s <=_r (cert!pc) then wtl_inst cert f r T step pc (cert!pc) else T"
-
-consts 
-wtl_inst_list :: "'a list \<Rightarrow> 's certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow>
-                  's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's"
-primrec
-"wtl_inst_list []     cert f r T B step pc s = s"
-"wtl_inst_list (i#is) cert f r T B step pc s = 
-    (let s' = wtl_cert cert f r T B step pc s in
-      if s' = T \<or> s = T then T else wtl_inst_list is cert f r T B step (pc+1) s')"
-
-constdefs
-  cert_ok :: "'s certificate \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow> 's set \<Rightarrow> bool"
-  "cert_ok cert n T B A \<equiv> (\<forall>i < n. cert!i \<in> A \<and> cert!i \<noteq> T) \<and> (cert!n = B)"
-
-constdefs
-  bottom :: "'a ord \<Rightarrow> 'a \<Rightarrow> bool"
-  "bottom r B \<equiv> \<forall>x. B <=_r x"
-
-
-locale lbv = Semilat +
-  fixes T :: "'a" ("\<top>") 
-  fixes B :: "'a" ("\<bottom>") 
-  fixes step :: "'a step_type" 
-  assumes top: "top r \<top>"
-  assumes T_A: "\<top> \<in> A"
-  assumes bot: "bottom r \<bottom>" 
-  assumes B_A: "\<bottom> \<in> A"
-
-  fixes merge :: "'a certificate \<Rightarrow> nat \<Rightarrow> (nat \<times> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a"
-  defines mrg_def: "merge cert \<equiv> LBVSpec.merge cert f r \<top>"
-
-  fixes wti :: "'a certificate \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a"
-  defines wti_def: "wti cert \<equiv> wtl_inst cert f r \<top> step"
- 
-  fixes wtc :: "'a certificate \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a"
-  defines wtc_def: "wtc cert \<equiv> wtl_cert cert f r \<top> \<bottom> step"
-
-  fixes wtl :: "'b list \<Rightarrow> 'a certificate \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a"
-  defines wtl_def: "wtl ins cert \<equiv> wtl_inst_list ins cert f r \<top> \<bottom> step"
-
-
-lemma (in lbv) wti:
-  "wti c pc s \<equiv> merge c pc (step pc s) (c!(pc+1))"
-  by (simp add: wti_def mrg_def wtl_inst_def)
-
-lemma (in lbv) wtc: 
-  "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>"
-  by (unfold wtc_def wti_def wtl_cert_def)
-
-
-lemma cert_okD1 [intro?]:
-  "cert_ok c n T B A \<Longrightarrow> pc < n \<Longrightarrow> c!pc \<in> A"
-  by (unfold cert_ok_def) fast
-
-lemma cert_okD2 [intro?]:
-  "cert_ok c n T B A \<Longrightarrow> c!n = B"
-  by (simp add: cert_ok_def)
-
-lemma cert_okD3 [intro?]:
-  "cert_ok c n T B A \<Longrightarrow> B \<in> A \<Longrightarrow> pc < n \<Longrightarrow> c!Suc pc \<in> A"
-  by (drule Suc_leI) (auto simp add: le_eq_less_or_eq dest: cert_okD1 cert_okD2)
-
-lemma cert_okD4 [intro?]:
-  "cert_ok c n T B A \<Longrightarrow> pc < n \<Longrightarrow> c!pc \<noteq> T"
-  by (simp add: cert_ok_def)
-
-declare Let_def [simp]
-
-section "more semilattice lemmas"
-
-
-lemma (in lbv) sup_top [simp, elim]:
-  assumes x: "x \<in> A" 
-  shows "x +_f \<top> = \<top>"
-proof -
-  from top have "x +_f \<top> <=_r \<top>" ..
-  moreover from x T_A have "\<top> <=_r x +_f \<top>" ..
-  ultimately show ?thesis ..
-qed
-  
-lemma (in lbv) plusplussup_top [simp, elim]:
-  "set xs \<subseteq> A \<Longrightarrow> xs ++_f \<top> = \<top>"
-  by (induct xs) auto
-
-
-
-lemma (in Semilat) pp_ub1':
-  assumes S: "snd`set S \<subseteq> A" 
-  assumes y: "y \<in> A" and ab: "(a, b) \<in> set S" 
-  shows "b <=_r map snd [(p', t') \<leftarrow> S . p' = a] ++_f y"
-proof -
-  from S have "\<forall>(x,y) \<in> set S. y \<in> A" by auto
-  with semilat y ab show ?thesis by - (rule ub1')
-qed 
-
-lemma (in lbv) bottom_le [simp, intro]:
-  "\<bottom> <=_r x"
-  by (insert bot) (simp add: bottom_def)
-
-lemma (in lbv) le_bottom [simp]:
-  "x <=_r \<bottom> = (x = \<bottom>)"
-  by (blast intro: antisym_r)
-
-
-
-section "merge"
-
-lemma (in lbv) merge_Nil [simp]:
-  "merge c pc [] x = x" by (simp add: mrg_def)
-
-lemma (in lbv) merge_Cons [simp]:
-  "merge c pc (l#ls) x = merge c pc ls (if fst l=pc+1 then snd l +_f x
-                                        else if snd l <=_r (c!fst l) then x
-                                        else \<top>)"
-  by (simp add: mrg_def split_beta)
-
-lemma (in lbv) merge_Err [simp]:
-  "snd`set ss \<subseteq> A \<Longrightarrow> merge c pc ss \<top> = \<top>"
-  by (induct ss) auto
-
-lemma (in lbv) merge_not_top:
-  "\<And>x. snd`set ss \<subseteq> A \<Longrightarrow> merge c pc ss x \<noteq> \<top> \<Longrightarrow> 
-  \<forall>(pc',s') \<in> set ss. (pc' \<noteq> pc+1 \<longrightarrow> s' <=_r (c!pc'))"
-  (is "\<And>x. ?set ss \<Longrightarrow> ?merge ss x \<Longrightarrow> ?P ss")
-proof (induct ss)
-  show "?P []" by simp
-next
-  fix x ls l
-  assume "?set (l#ls)" then obtain set: "snd`set ls \<subseteq> A" by simp
-  assume merge: "?merge (l#ls) x" 
-  moreover
-  obtain pc' s' where l: "l = (pc',s')" by (cases l)
-  ultimately
-  obtain x' where merge': "?merge ls x'" by simp 
-  assume "\<And>x. ?set ls \<Longrightarrow> ?merge ls x \<Longrightarrow> ?P ls" hence "?P ls" using set merge' .
-  moreover
-  from merge set
-  have "pc' \<noteq> pc+1 \<longrightarrow> s' <=_r (c!pc')" by (simp add: l split: split_if_asm)
-  ultimately
-  show "?P (l#ls)" by (simp add: l)
-qed
-
-
-lemma (in lbv) merge_def:
-  shows 
-  "\<And>x. x \<in> A \<Longrightarrow> snd`set ss \<subseteq> A \<Longrightarrow>
-  merge c pc ss x = 
-  (if \<forall>(pc',s') \<in> set ss. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc' then
-    map snd [(p',t') \<leftarrow> ss. p'=pc+1] ++_f x
-  else \<top>)" 
-  (is "\<And>x. _ \<Longrightarrow> _ \<Longrightarrow> ?merge ss x = ?if ss x" is "\<And>x. _ \<Longrightarrow> _ \<Longrightarrow> ?P ss x")
-proof (induct ss)
-  fix x show "?P [] x" by simp
-next 
-  fix x assume x: "x \<in> A" 
-  fix l::"nat \<times> 'a" and ls  
-  assume "snd`set (l#ls) \<subseteq> A"
-  then obtain l: "snd l \<in> A" and ls: "snd`set ls \<subseteq> A" by auto
-  assume "\<And>x. x \<in> A \<Longrightarrow> snd`set ls \<subseteq> A \<Longrightarrow> ?P ls x" 
-  hence IH: "\<And>x. x \<in> A \<Longrightarrow> ?P ls x" using ls by iprover
-  obtain pc' s' where [simp]: "l = (pc',s')" by (cases l)
-  hence "?merge (l#ls) x = ?merge ls 
-    (if pc'=pc+1 then s' +_f x else if s' <=_r c!pc' then x else \<top>)"
-    (is "?merge (l#ls) x = ?merge ls ?if'")
-    by simp 
-  also have "\<dots> = ?if ls ?if'" 
-  proof -
-    from l have "s' \<in> A" by simp
-    with x have "s' +_f x \<in> A" by simp
-    with x T_A have "?if' \<in> A" by auto
-    hence "?P ls ?if'" by (rule IH) thus ?thesis by simp
-  qed
-  also have "\<dots> = ?if (l#ls) x"
-    proof (cases "\<forall>(pc', s')\<in>set (l#ls). pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'")
-      case True
-      hence "\<forall>(pc', s')\<in>set ls. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'" by auto
-      moreover
-      from True have 
-        "map snd [(p',t')\<leftarrow>ls . p'=pc+1] ++_f ?if' = 
-        (map snd [(p',t')\<leftarrow>l#ls . p'=pc+1] ++_f x)"
-        by simp
-      ultimately
-      show ?thesis using True by simp
-    next
-      case False 
-      moreover
-      from ls have "set (map snd [(p', t')\<leftarrow>ls . p' = Suc pc]) \<subseteq> A" by auto
-      ultimately show ?thesis by auto
-    qed
-  finally show "?P (l#ls) x" .
-qed
-
-lemma (in lbv) merge_not_top_s:
-  assumes x:  "x \<in> A" and ss: "snd`set ss \<subseteq> A"
-  assumes m:  "merge c pc ss x \<noteq> \<top>"
-  shows "merge c pc ss x = (map snd [(p',t') \<leftarrow> ss. p'=pc+1] ++_f x)"
-proof -
-  from ss m have "\<forall>(pc',s') \<in> set ss. (pc' \<noteq> pc+1 \<longrightarrow> s' <=_r c!pc')" 
-    by (rule merge_not_top)
-  with x ss m show ?thesis by - (drule merge_def, auto split: split_if_asm)
-qed
-
-section "wtl-inst-list"
-
-lemmas [iff] = not_Err_eq
-
-lemma (in lbv) wtl_Nil [simp]: "wtl [] c pc s = s" 
-  by (simp add: wtl_def)
-
-lemma (in lbv) wtl_Cons [simp]: 
-  "wtl (i#is) c pc s = 
-  (let s' = wtc c pc s in if s' = \<top> \<or> s = \<top> then \<top> else wtl is c (pc+1) s')"
-  by (simp add: wtl_def wtc_def)
-
-lemma (in lbv) wtl_Cons_not_top:
-  "wtl (i#is) c pc s \<noteq> \<top> = 
-  (wtc c pc s \<noteq> \<top> \<and> s \<noteq> T \<and> wtl is c (pc+1) (wtc c pc s) \<noteq> \<top>)"
-  by (auto simp del: split_paired_Ex)
-
-lemma (in lbv) wtl_top [simp]:  "wtl ls c pc \<top> = \<top>"
-  by (cases ls) auto
-
-lemma (in lbv) wtl_not_top:
-  "wtl ls c pc s \<noteq> \<top> \<Longrightarrow> s \<noteq> \<top>"
-  by (cases "s=\<top>") auto
-
-lemma (in lbv) wtl_append [simp]:
-  "\<And>pc s. wtl (a@b) c pc s = wtl b c (pc+length a) (wtl a c pc s)"
-  by (induct a) auto
-
-lemma (in lbv) wtl_take:
-  "wtl is c pc s \<noteq> \<top> \<Longrightarrow> wtl (take pc' is) c pc s \<noteq> \<top>"
-  (is "?wtl is \<noteq> _ \<Longrightarrow> _")
-proof -
-  assume "?wtl is \<noteq> \<top>"
-  hence "?wtl (take pc' is @ drop pc' is) \<noteq> \<top>" by simp  
-  thus ?thesis by (auto dest!: wtl_not_top simp del: append_take_drop_id)
-qed
-
-lemma take_Suc:
-  "\<forall>n. n < length l \<longrightarrow> take (Suc n) l = (take n l)@[l!n]" (is "?P l")
-proof (induct l)
-  show "?P []" by simp
-next
-  fix x xs assume IH: "?P xs"  
-  show "?P (x#xs)"
-  proof (intro strip)
-    fix n assume "n < length (x#xs)"
-    with IH show "take (Suc n) (x # xs) = take n (x # xs) @ [(x # xs) ! n]" 
-      by (cases n, auto)
-  qed
-qed
-
-lemma (in lbv) wtl_Suc:
-  assumes suc: "pc+1 < length is"
-  assumes wtl: "wtl (take pc is) c 0 s \<noteq> \<top>"
-  shows "wtl (take (pc+1) is) c 0 s = wtc c pc (wtl (take pc is) c 0 s)"
-proof -
-  from suc have "take (pc+1) is=(take pc is)@[is!pc]" by (simp add: take_Suc)
-  with suc wtl show ?thesis by (simp add: min_max.inf_absorb2)
-qed
-
-lemma (in lbv) wtl_all:
-  assumes all: "wtl is c 0 s \<noteq> \<top>" (is "?wtl is \<noteq> _") 
-  assumes pc:  "pc < length is"
-  shows  "wtc c pc (wtl (take pc is) c 0 s) \<noteq> \<top>"
-proof -
-  from pc have "0 < length (drop pc is)" by simp
-  then  obtain i r where Cons: "drop pc is = i#r" 
-    by (auto simp add: neq_Nil_conv simp del: length_drop drop_eq_Nil)
-  hence "i#r = drop pc is" ..
-  with all have take: "?wtl (take pc is@i#r) \<noteq> \<top>" by simp 
-  from pc have "is!pc = drop pc is ! 0" by simp
-  with Cons have "is!pc = i" by simp
-  with take pc show ?thesis by (auto simp add: min_max.inf_absorb2)
-qed
-
-section "preserves-type"
-
-lemma (in lbv) merge_pres:
-  assumes s0: "snd`set ss \<subseteq> A" and x: "x \<in> A"
-  shows "merge c pc ss x \<in> A"
-proof -
-  from s0 have "set (map snd [(p', t')\<leftarrow>ss . p'=pc+1]) \<subseteq> A" by auto
-  with x  have "(map snd [(p', t')\<leftarrow>ss . p'=pc+1] ++_f x) \<in> A"
-    by (auto intro!: plusplus_closed semilat)
-  with s0 x show ?thesis by (simp add: merge_def T_A)
-qed
-  
-
-lemma pres_typeD2:
-  "pres_type step n A \<Longrightarrow> s \<in> A \<Longrightarrow> p < n \<Longrightarrow> snd`set (step p s) \<subseteq> A"
-  by auto (drule pres_typeD)
-
-
-lemma (in lbv) wti_pres [intro?]:
-  assumes pres: "pres_type step n A" 
-  assumes cert: "c!(pc+1) \<in> A"
-  assumes s_pc: "s \<in> A" "pc < n"
-  shows "wti c pc s \<in> A"
-proof -
-  from pres s_pc have "snd`set (step pc s) \<subseteq> A" by (rule pres_typeD2)
-  with cert show ?thesis by (simp add: wti merge_pres)
-qed
-
-
-lemma (in lbv) wtc_pres:
-  assumes pres: "pres_type step n A"
-  assumes cert: "c!pc \<in> A" and cert': "c!(pc+1) \<in> A"
-  assumes s: "s \<in> A" and pc: "pc < n"
-  shows "wtc c pc s \<in> A"
-proof -
-  have "wti c pc s \<in> A" using pres cert' s pc ..
-  moreover have "wti c pc (c!pc) \<in> A" using pres cert' cert pc ..
-  ultimately show ?thesis using T_A by (simp add: wtc) 
-qed
-
-
-lemma (in lbv) wtl_pres:
-  assumes pres: "pres_type step (length is) A"
-  assumes cert: "cert_ok c (length is) \<top> \<bottom> A"
-  assumes s:    "s \<in> A" 
-  assumes all:  "wtl is c 0 s \<noteq> \<top>"
-  shows "pc < length is \<Longrightarrow> wtl (take pc is) c 0 s \<in> A"
-  (is "?len pc \<Longrightarrow> ?wtl pc \<in> A")
-proof (induct pc)
-  from s show "?wtl 0 \<in> A" by simp
-next
-  fix n assume IH: "Suc n < length is"
-  then have n: "n < length is" by simp  
-  from IH have n1: "n+1 < length is" by simp
-  assume prem: "n < length is \<Longrightarrow> ?wtl n \<in> A"
-  have "wtc c n (?wtl n) \<in> A"
-  using pres _ _ _ n
-  proof (rule wtc_pres)
-    from prem n show "?wtl n \<in> A" .
-    from cert n show "c!n \<in> A" by (rule cert_okD1)
-    from cert n1 show "c!(n+1) \<in> A" by (rule cert_okD1)
-  qed
-  also
-  from all n have "?wtl n \<noteq> \<top>" by - (rule wtl_take)
-  with n1 have "wtc c n (?wtl n) = ?wtl (n+1)" by (rule wtl_Suc [symmetric])
-  finally  show "?wtl (Suc n) \<in> A" by simp
-qed
-
-end
--- a/src/HOL/MicroJava/BV/Listn.thy	Wed Dec 02 12:04:07 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,544 +0,0 @@
-(*  Title:      HOL/MicroJava/BV/Listn.thy
-    Author:     Tobias Nipkow
-    Copyright   2000 TUM
-
-Lists of a fixed length
-*)
-
-header {* \isaheader{Fixed Length Lists} *}
-
-theory Listn
-imports Err
-begin
-
-constdefs
-
- list :: "nat \<Rightarrow> 'a set \<Rightarrow> 'a list set"
-"list n A == {xs. length xs = n & set xs <= A}"
-
- le :: "'a ord \<Rightarrow> ('a list)ord"
-"le r == list_all2 (%x y. x <=_r y)"
-
-syntax "@lesublist" :: "'a list \<Rightarrow> 'a ord \<Rightarrow> 'a list \<Rightarrow> bool"
-       ("(_ /<=[_] _)" [50, 0, 51] 50)
-syntax "@lesssublist" :: "'a list \<Rightarrow> 'a ord \<Rightarrow> 'a list \<Rightarrow> bool"
-       ("(_ /<[_] _)" [50, 0, 51] 50)
-translations
- "x <=[r] y" == "x <=_(Listn.le r) y"
- "x <[r] y"  == "x <_(Listn.le r) y"
-
-constdefs
- map2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list"
-"map2 f == (%xs ys. map (split f) (zip xs ys))"
-
-syntax "@plussublist" :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b list \<Rightarrow> 'c list"
-       ("(_ /+[_] _)" [65, 0, 66] 65)
-translations  "x +[f] y" == "x +_(map2 f) y"
-
-consts coalesce :: "'a err list \<Rightarrow> 'a list err"
-primrec
-"coalesce [] = OK[]"
-"coalesce (ex#exs) = Err.sup (op #) ex (coalesce exs)"
-
-constdefs
- sl :: "nat \<Rightarrow> 'a sl \<Rightarrow> 'a list sl"
-"sl n == %(A,r,f). (list n A, le r, map2 f)"
-
- sup :: "('a \<Rightarrow> 'b \<Rightarrow> 'c err) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list err"
-"sup f == %xs ys. if size xs = size ys then coalesce(xs +[f] ys) else Err"
-
- upto_esl :: "nat \<Rightarrow> 'a esl \<Rightarrow> 'a list esl"
-"upto_esl m == %(A,r,f). (Union{list n A |n. n <= m}, le r, sup f)"
-
-lemmas [simp] = set_update_subsetI
-
-lemma unfold_lesub_list:
-  "xs <=[r] ys == Listn.le r xs ys"
-  by (simp add: lesub_def)
-
-lemma Nil_le_conv [iff]:
-  "([] <=[r] ys) = (ys = [])"
-apply (unfold lesub_def Listn.le_def)
-apply simp
-done
-
-lemma Cons_notle_Nil [iff]: 
-  "~ x#xs <=[r] []"
-apply (unfold lesub_def Listn.le_def)
-apply simp
-done
-
-
-lemma Cons_le_Cons [iff]:
-  "x#xs <=[r] y#ys = (x <=_r y & xs <=[r] ys)"
-apply (unfold lesub_def Listn.le_def)
-apply simp
-done
-
-lemma Cons_less_Conss [simp]:
-  "order r \<Longrightarrow> 
-  x#xs <_(Listn.le r) y#ys = 
-  (x <_r y & xs <=[r] ys  |  x = y & xs <_(Listn.le r) ys)"
-apply (unfold lesssub_def)
-apply blast
-done  
-
-lemma list_update_le_cong:
-  "\<lbrakk> i<size xs; xs <=[r] ys; x <=_r y \<rbrakk> \<Longrightarrow> xs[i:=x] <=[r] ys[i:=y]";
-apply (unfold unfold_lesub_list)
-apply (unfold Listn.le_def)
-apply (simp add: list_all2_conv_all_nth nth_list_update)
-done
-
-
-lemma le_listD:
-  "\<lbrakk> xs <=[r] ys; p < size xs \<rbrakk> \<Longrightarrow> xs!p <=_r ys!p"
-apply (unfold Listn.le_def lesub_def)
-apply (simp add: list_all2_conv_all_nth)
-done
-
-lemma le_list_refl:
-  "!x. x <=_r x \<Longrightarrow> xs <=[r] xs"
-apply (unfold unfold_lesub_list)
-apply (simp add: Listn.le_def list_all2_conv_all_nth)
-done
-
-lemma le_list_trans:
-  "\<lbrakk> order r; xs <=[r] ys; ys <=[r] zs \<rbrakk> \<Longrightarrow> xs <=[r] zs"
-apply (unfold unfold_lesub_list)
-apply (simp add: Listn.le_def list_all2_conv_all_nth)
-apply clarify
-apply simp
-apply (blast intro: order_trans)
-done
-
-lemma le_list_antisym:
-  "\<lbrakk> order r; xs <=[r] ys; ys <=[r] xs \<rbrakk> \<Longrightarrow> xs = ys"
-apply (unfold unfold_lesub_list)
-apply (simp add: Listn.le_def list_all2_conv_all_nth)
-apply (rule nth_equalityI)
- apply blast
-apply clarify
-apply simp
-apply (blast intro: order_antisym)
-done
-
-lemma order_listI [simp, intro!]:
-  "order r \<Longrightarrow> order(Listn.le r)"
-apply (subst Semilat.order_def)
-apply (blast intro: le_list_refl le_list_trans le_list_antisym
-             dest: order_refl)
-done
-
-
-lemma lesub_list_impl_same_size [simp]:
-  "xs <=[r] ys \<Longrightarrow> size ys = size xs"  
-apply (unfold Listn.le_def lesub_def)
-apply (simp add: list_all2_conv_all_nth)
-done 
-
-lemma lesssub_list_impl_same_size:
-  "xs <_(Listn.le r) ys \<Longrightarrow> size ys = size xs"
-apply (unfold lesssub_def)
-apply auto
-done  
-
-lemma le_list_appendI:
-  "\<And>b c d. a <=[r] b \<Longrightarrow> c <=[r] d \<Longrightarrow> a@c <=[r] b@d"
-apply (induct a)
- apply simp
-apply (case_tac b)
-apply auto
-done
-
-lemma le_listI:
-  "length a = length b \<Longrightarrow> (\<And>n. n < length a \<Longrightarrow> a!n <=_r b!n) \<Longrightarrow> a <=[r] b"
-  apply (unfold lesub_def Listn.le_def)
-  apply (simp add: list_all2_conv_all_nth)
-  done
-
-lemma listI:
-  "\<lbrakk> length xs = n; set xs <= A \<rbrakk> \<Longrightarrow> xs : list n A"
-apply (unfold list_def)
-apply blast
-done
-
-lemma listE_length [simp]:
-   "xs : list n A \<Longrightarrow> length xs = n"
-apply (unfold list_def)
-apply blast
-done 
-
-lemma less_lengthI:
-  "\<lbrakk> xs : list n A; p < n \<rbrakk> \<Longrightarrow> p < length xs"
-  by simp
-
-lemma listE_set [simp]:
-  "xs : list n A \<Longrightarrow> set xs <= A"
-apply (unfold list_def)
-apply blast
-done 
-
-lemma list_0 [simp]:
-  "list 0 A = {[]}"
-apply (unfold list_def)
-apply auto
-done 
-
-lemma in_list_Suc_iff: 
-  "(xs : list (Suc n) A) = (\<exists>y\<in> A. \<exists>ys\<in> list n A. xs = y#ys)"
-apply (unfold list_def)
-apply (case_tac "xs")
-apply auto
-done 
-
-lemma Cons_in_list_Suc [iff]:
-  "(x#xs : list (Suc n) A) = (x\<in> A & xs : list n A)";
-apply (simp add: in_list_Suc_iff)
-done 
-
-lemma list_not_empty:
-  "\<exists>a. a\<in> A \<Longrightarrow> \<exists>xs. xs : list n A";
-apply (induct "n")
- apply simp
-apply (simp add: in_list_Suc_iff)
-apply blast
-done
-
-
-lemma nth_in [rule_format, simp]:
-  "!i n. length xs = n \<longrightarrow> set xs <= A \<longrightarrow> i < n \<longrightarrow> (xs!i) : A"
-apply (induct "xs")
- apply simp
-apply (simp add: nth_Cons split: nat.split)
-done
-
-lemma listE_nth_in:
-  "\<lbrakk> xs : list n A; i < n \<rbrakk> \<Longrightarrow> (xs!i) : A"
-  by auto
-
-
-lemma listn_Cons_Suc [elim!]:
-  "l#xs \<in> list n A \<Longrightarrow> (\<And>n'. n = Suc n' \<Longrightarrow> l \<in> A \<Longrightarrow> xs \<in> list n' A \<Longrightarrow> P) \<Longrightarrow> P"
-  by (cases n) auto
-
-lemma listn_appendE [elim!]:
-  "a@b \<in> list n A \<Longrightarrow> (\<And>n1 n2. n=n1+n2 \<Longrightarrow> a \<in> list n1 A \<Longrightarrow> b \<in> list n2 A \<Longrightarrow> P) \<Longrightarrow> P" 
-proof -
-  have "\<And>n. a@b \<in> list n A \<Longrightarrow> \<exists>n1 n2. n=n1+n2 \<and> a \<in> list n1 A \<and> b \<in> list n2 A"
-    (is "\<And>n. ?list a n \<Longrightarrow> \<exists>n1 n2. ?P a n n1 n2")
-  proof (induct a)
-    fix n assume "?list [] n"
-    hence "?P [] n 0 n" by simp
-    thus "\<exists>n1 n2. ?P [] n n1 n2" by fast
-  next
-    fix n l ls
-    assume "?list (l#ls) n"
-    then obtain n' where n: "n = Suc n'" "l \<in> A" and list_n': "ls@b \<in> list n' A" by fastsimp
-    assume "\<And>n. ls @ b \<in> list n A \<Longrightarrow> \<exists>n1 n2. n = n1 + n2 \<and> ls \<in> list n1 A \<and> b \<in> list n2 A"
-    hence "\<exists>n1 n2. n' = n1 + n2 \<and> ls \<in> list n1 A \<and> b \<in> list n2 A" by this (rule list_n')
-    then obtain n1 n2 where "n' = n1 + n2" "ls \<in> list n1 A" "b \<in> list n2 A" by fast
-    with n have "?P (l#ls) n (n1+1) n2" by simp
-    thus "\<exists>n1 n2. ?P (l#ls) n n1 n2" by fastsimp
-  qed
-  moreover
-  assume "a@b \<in> list n A" "\<And>n1 n2. n=n1+n2 \<Longrightarrow> a \<in> list n1 A \<Longrightarrow> b \<in> list n2 A \<Longrightarrow> P"
-  ultimately
-  show ?thesis by blast
-qed
-
-
-lemma listt_update_in_list [simp, intro!]:
-  "\<lbrakk> xs : list n A; x\<in> A \<rbrakk> \<Longrightarrow> xs[i := x] : list n A"
-apply (unfold list_def)
-apply simp
-done 
-
-lemma plus_list_Nil [simp]:
-  "[] +[f] xs = []"
-apply (unfold plussub_def map2_def)
-apply simp
-done 
-
-lemma plus_list_Cons [simp]:
-  "(x#xs) +[f] ys = (case ys of [] \<Rightarrow> [] | y#ys \<Rightarrow> (x +_f y)#(xs +[f] ys))"
-  by (simp add: plussub_def map2_def split: list.split)
-
-lemma length_plus_list [rule_format, simp]:
-  "!ys. length(xs +[f] ys) = min(length xs) (length ys)"
-apply (induct xs)
- apply simp
-apply clarify
-apply (simp (no_asm_simp) split: list.split)
-done
-
-lemma nth_plus_list [rule_format, simp]:
-  "!xs ys i. length xs = n \<longrightarrow> length ys = n \<longrightarrow> i<n \<longrightarrow> 
-  (xs +[f] ys)!i = (xs!i) +_f (ys!i)"
-apply (induct n)
- apply simp
-apply clarify
-apply (case_tac xs)
- apply simp
-apply (force simp add: nth_Cons split: list.split nat.split)
-done
-
-
-lemma (in Semilat) plus_list_ub1 [rule_format]:
- "\<lbrakk> set xs <= A; set ys <= A; size xs = size ys \<rbrakk> 
-  \<Longrightarrow> xs <=[r] xs +[f] ys"
-apply (unfold unfold_lesub_list)
-apply (simp add: Listn.le_def list_all2_conv_all_nth)
-done
-
-lemma (in Semilat) plus_list_ub2:
- "\<lbrakk>set xs <= A; set ys <= A; size xs = size ys \<rbrakk>
-  \<Longrightarrow> ys <=[r] xs +[f] ys"
-apply (unfold unfold_lesub_list)
-apply (simp add: Listn.le_def list_all2_conv_all_nth)
-done
-
-lemma (in Semilat) plus_list_lub [rule_format]:
-shows "!xs ys zs. set xs <= A \<longrightarrow> set ys <= A \<longrightarrow> set zs <= A 
-  \<longrightarrow> size xs = n & size ys = n \<longrightarrow> 
-  xs <=[r] zs & ys <=[r] zs \<longrightarrow> xs +[f] ys <=[r] zs"
-apply (unfold unfold_lesub_list)
-apply (simp add: Listn.le_def list_all2_conv_all_nth)
-done
-
-lemma (in Semilat) list_update_incr [rule_format]:
- "x\<in> A \<Longrightarrow> set xs <= A \<longrightarrow> 
-  (!i. i<size xs \<longrightarrow> xs <=[r] xs[i := x +_f xs!i])"
-apply (unfold unfold_lesub_list)
-apply (simp add: Listn.le_def list_all2_conv_all_nth)
-apply (induct xs)
- apply simp
-apply (simp add: in_list_Suc_iff)
-apply clarify
-apply (simp add: nth_Cons split: nat.split)
-done
-
-lemma equals0I_aux:
-  "(\<And>y. A y \<Longrightarrow> False) \<Longrightarrow> A = bot_class.bot"
-  by (rule equals0I) (auto simp add: mem_def)
-
-lemma acc_le_listI [intro!]:
-  "\<lbrakk> order r; acc r \<rbrakk> \<Longrightarrow> acc(Listn.le r)"
-apply (unfold acc_def)
-apply (subgoal_tac
- "wfP (SUP n. (\<lambda>ys xs. size xs = n & size ys = n & xs <_(Listn.le r) ys))")
- apply (erule wfP_subset)
- apply (blast intro: lesssub_list_impl_same_size)
-apply (rule wfP_SUP)
- prefer 2
- apply clarify
- apply (rename_tac m n)
- apply (case_tac "m=n")
-  apply simp
- apply (fast intro!: equals0I_aux dest: not_sym)
-apply clarify
-apply (rename_tac n)
-apply (induct_tac n)
- apply (simp add: lesssub_def cong: conj_cong)
-apply (rename_tac k)
-apply (simp add: wfP_eq_minimal)
-apply (simp (no_asm) add: length_Suc_conv cong: conj_cong)
-apply clarify
-apply (rename_tac M m)
-apply (case_tac "\<exists>x xs. size xs = k & x#xs : M")
- prefer 2
- apply (erule thin_rl)
- apply (erule thin_rl)
- apply blast
-apply (erule_tac x = "{a. \<exists>xs. size xs = k & a#xs:M}" in allE)
-apply (erule impE)
- apply blast
-apply (thin_tac "\<exists>x xs. ?P x xs")
-apply clarify
-apply (rename_tac maxA xs)
-apply (erule_tac x = "{ys. size ys = size xs & maxA#ys : M}" in allE)
-apply (erule impE)
- apply blast
-apply clarify
-apply (thin_tac "m : M")
-apply (thin_tac "maxA#xs : M")
-apply (rule bexI)
- prefer 2
- apply assumption
-apply clarify
-apply simp
-apply blast
-done 
-
-lemma closed_listI:
-  "closed S f \<Longrightarrow> closed (list n S) (map2 f)"
-apply (unfold closed_def)
-apply (induct n)
- apply simp
-apply clarify
-apply (simp add: in_list_Suc_iff)
-apply clarify
-apply simp
-done
-
-
-lemma Listn_sl_aux:
-assumes "semilat (A, r, f)" shows "semilat (Listn.sl n (A,r,f))"
-proof -
-  interpret Semilat A r f using assms by (rule Semilat.intro)
-show ?thesis
-apply (unfold Listn.sl_def)
-apply (simp (no_asm) only: semilat_Def split_conv)
-apply (rule conjI)
- apply simp
-apply (rule conjI)
- apply (simp only: closedI closed_listI)
-apply (simp (no_asm) only: list_def)
-apply (simp (no_asm_simp) add: plus_list_ub1 plus_list_ub2 plus_list_lub)
-done
-qed
-
-lemma Listn_sl: "\<And>L. semilat L \<Longrightarrow> semilat (Listn.sl n L)"
- by(simp add: Listn_sl_aux split_tupled_all)
-
-lemma coalesce_in_err_list [rule_format]:
-  "!xes. xes : list n (err A) \<longrightarrow> coalesce xes : err(list n A)"
-apply (induct n)
- apply simp
-apply clarify
-apply (simp add: in_list_Suc_iff)
-apply clarify
-apply (simp (no_asm) add: plussub_def Err.sup_def lift2_def split: err.split)
-apply force
-done 
-
-lemma lem: "\<And>x xs. x +_(op #) xs = x#xs"
-  by (simp add: plussub_def)
-
-lemma coalesce_eq_OK1_D [rule_format]:
-  "semilat(err A, Err.le r, lift2 f) \<Longrightarrow> 
-  !xs. xs : list n A \<longrightarrow> (!ys. ys : list n A \<longrightarrow> 
-  (!zs. coalesce (xs +[f] ys) = OK zs \<longrightarrow> xs <=[r] zs))"
-apply (induct n)
-  apply simp
-apply clarify
-apply (simp add: in_list_Suc_iff)
-apply clarify
-apply (simp split: err.split_asm add: lem Err.sup_def lift2_def)
-apply (force simp add: semilat_le_err_OK1)
-done
-
-lemma coalesce_eq_OK2_D [rule_format]:
-  "semilat(err A, Err.le r, lift2 f) \<Longrightarrow> 
-  !xs. xs : list n A \<longrightarrow> (!ys. ys : list n A \<longrightarrow> 
-  (!zs. coalesce (xs +[f] ys) = OK zs \<longrightarrow> ys <=[r] zs))"
-apply (induct n)
- apply simp
-apply clarify
-apply (simp add: in_list_Suc_iff)
-apply clarify
-apply (simp split: err.split_asm add: lem Err.sup_def lift2_def)
-apply (force simp add: semilat_le_err_OK2)
-done 
-
-lemma lift2_le_ub:
-  "\<lbrakk> semilat(err A, Err.le r, lift2 f); x\<in> A; y\<in> A; x +_f y = OK z; 
-      u\<in> A; x <=_r u; y <=_r u \<rbrakk> \<Longrightarrow> z <=_r u"
-apply (unfold semilat_Def plussub_def err_def)
-apply (simp add: lift2_def)
-apply clarify
-apply (rotate_tac -3)
-apply (erule thin_rl)
-apply (erule thin_rl)
-apply force
-done
-
-lemma coalesce_eq_OK_ub_D [rule_format]:
-  "semilat(err A, Err.le r, lift2 f) \<Longrightarrow> 
-  !xs. xs : list n A \<longrightarrow> (!ys. ys : list n A \<longrightarrow> 
-  (!zs us. coalesce (xs +[f] ys) = OK zs & xs <=[r] us & ys <=[r] us 
-           & us : list n A \<longrightarrow> zs <=[r] us))"
-apply (induct n)
- apply simp
-apply clarify
-apply (simp add: in_list_Suc_iff)
-apply clarify
-apply (simp (no_asm_use) split: err.split_asm add: lem Err.sup_def lift2_def)
-apply clarify
-apply (rule conjI)
- apply (blast intro: lift2_le_ub)
-apply blast
-done 
-
-lemma lift2_eq_ErrD:
-  "\<lbrakk> x +_f y = Err; semilat(err A, Err.le r, lift2 f); x\<in> A; y\<in> A \<rbrakk> 
-  \<Longrightarrow> ~(\<exists>u\<in> A. x <=_r u & y <=_r u)"
-  by (simp add: OK_plus_OK_eq_Err_conv [THEN iffD1])
-
-
-lemma coalesce_eq_Err_D [rule_format]:
-  "\<lbrakk> semilat(err A, Err.le r, lift2 f) \<rbrakk> 
-  \<Longrightarrow> !xs. xs\<in> list n A \<longrightarrow> (!ys. ys\<in> list n A \<longrightarrow> 
-      coalesce (xs +[f] ys) = Err \<longrightarrow> 
-      ~(\<exists>zs\<in> list n A. xs <=[r] zs & ys <=[r] zs))"
-apply (induct n)
- apply simp
-apply clarify
-apply (simp add: in_list_Suc_iff)
-apply clarify
-apply (simp split: err.split_asm add: lem Err.sup_def lift2_def)
- apply (blast dest: lift2_eq_ErrD)
-done 
-
-lemma closed_err_lift2_conv:
-  "closed (err A) (lift2 f) = (\<forall>x\<in> A. \<forall>y\<in> A. x +_f y : err A)"
-apply (unfold closed_def)
-apply (simp add: err_def)
-done 
-
-lemma closed_map2_list [rule_format]:
-  "closed (err A) (lift2 f) \<Longrightarrow> 
-  \<forall>xs. xs : list n A \<longrightarrow> (\<forall>ys. ys : list n A \<longrightarrow> 
-  map2 f xs ys : list n (err A))"
-apply (unfold map2_def)
-apply (induct n)
- apply simp
-apply clarify
-apply (simp add: in_list_Suc_iff)
-apply clarify
-apply (simp add: plussub_def closed_err_lift2_conv)
-done
-
-lemma closed_lift2_sup:
-  "closed (err A) (lift2 f) \<Longrightarrow> 
-  closed (err (list n A)) (lift2 (sup f))"
-  by (fastsimp  simp add: closed_def plussub_def sup_def lift2_def
-                          coalesce_in_err_list closed_map2_list
-                split: err.split)
-
-lemma err_semilat_sup:
-  "err_semilat (A,r,f) \<Longrightarrow> 
-  err_semilat (list n A, Listn.le r, sup f)"
-apply (unfold Err.sl_def)
-apply (simp only: split_conv)
-apply (simp (no_asm) only: semilat_Def plussub_def)
-apply (simp (no_asm_simp) only: Semilat.closedI [OF Semilat.intro] closed_lift2_sup)
-apply (rule conjI)
- apply (drule Semilat.orderI [OF Semilat.intro])
- apply simp
-apply (simp (no_asm) only: unfold_lesub_err Err.le_def err_def sup_def lift2_def)
-apply (simp (no_asm_simp) add: coalesce_eq_OK1_D coalesce_eq_OK2_D split: err.split)
-apply (blast intro: coalesce_eq_OK_ub_D dest: coalesce_eq_Err_D)
-done 
-
-lemma err_semilat_upto_esl:
-  "\<And>L. err_semilat L \<Longrightarrow> err_semilat(upto_esl m L)"
-apply (unfold Listn.upto_esl_def)
-apply (simp (no_asm_simp) only: split_tupled_all)
-apply simp
-apply (fastsimp intro!: err_semilat_UnionI err_semilat_sup
-                dest: lesub_list_impl_same_size 
-                simp add: plussub_def Listn.sup_def)
-done
-
-end
--- a/src/HOL/MicroJava/BV/Opt.thy	Wed Dec 02 12:04:07 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,295 +0,0 @@
-(*  Title:      HOL/MicroJava/BV/Opt.thy
-    ID:         $Id$
-    Author:     Tobias Nipkow
-    Copyright   2000 TUM
-
-More about options
-*)
-
-header {* \isaheader{More about Options} *}
-
-theory Opt
-imports Err
-begin
-
-constdefs
- le :: "'a ord \<Rightarrow> 'a option ord"
-"le r o1 o2 == case o2 of None \<Rightarrow> o1=None |
-                              Some y \<Rightarrow> (case o1 of None \<Rightarrow> True
-                                                  | Some x \<Rightarrow> x <=_r y)"
-
- opt :: "'a set \<Rightarrow> 'a option set"
-"opt A == insert None {x . ? y:A. x = Some y}"
-
- sup :: "'a ebinop \<Rightarrow> 'a option ebinop"
-"sup f o1 o2 ==  
- case o1 of None \<Rightarrow> OK o2 | Some x \<Rightarrow> (case o2 of None \<Rightarrow> OK o1
-     | Some y \<Rightarrow> (case f x y of Err \<Rightarrow> Err | OK z \<Rightarrow> OK (Some z)))"
-
- esl :: "'a esl \<Rightarrow> 'a option esl"
-"esl == %(A,r,f). (opt A, le r, sup f)"
-
-lemma unfold_le_opt:
-  "o1 <=_(le r) o2 = 
-  (case o2 of None \<Rightarrow> o1=None | 
-              Some y \<Rightarrow> (case o1 of None \<Rightarrow> True | Some x \<Rightarrow> x <=_r y))"
-apply (unfold lesub_def le_def)
-apply (rule refl)
-done
-
-lemma le_opt_refl:
-  "order r \<Longrightarrow> o1 <=_(le r) o1"
-by (simp add: unfold_le_opt split: option.split)
-
-lemma le_opt_trans [rule_format]:
-  "order r \<Longrightarrow> 
-   o1 <=_(le r) o2 \<longrightarrow> o2 <=_(le r) o3 \<longrightarrow> o1 <=_(le r) o3"
-apply (simp add: unfold_le_opt split: option.split)
-apply (blast intro: order_trans)
-done
-
-lemma le_opt_antisym [rule_format]:
-  "order r \<Longrightarrow> o1 <=_(le r) o2 \<longrightarrow> o2 <=_(le r) o1 \<longrightarrow> o1=o2"
-apply (simp add: unfold_le_opt split: option.split)
-apply (blast intro: order_antisym)
-done
-
-lemma order_le_opt [intro!,simp]:
-  "order r \<Longrightarrow> order(le r)"
-apply (subst Semilat.order_def)
-apply (blast intro: le_opt_refl le_opt_trans le_opt_antisym)
-done 
-
-lemma None_bot [iff]: 
-  "None <=_(le r) ox"
-apply (unfold lesub_def le_def)
-apply (simp split: option.split)
-done 
-
-lemma Some_le [iff]:
-  "(Some x <=_(le r) ox) = (? y. ox = Some y & x <=_r y)"
-apply (unfold lesub_def le_def)
-apply (simp split: option.split)
-done 
-
-lemma le_None [iff]:
-  "(ox <=_(le r) None) = (ox = None)";
-apply (unfold lesub_def le_def)
-apply (simp split: option.split)
-done 
-
-
-lemma OK_None_bot [iff]:
-  "OK None <=_(Err.le (le r)) x"
-  by (simp add: lesub_def Err.le_def le_def split: option.split err.split)
-
-lemma sup_None1 [iff]:
-  "x +_(sup f) None = OK x"
-  by (simp add: plussub_def sup_def split: option.split)
-
-lemma sup_None2 [iff]:
-  "None +_(sup f) x = OK x"
-  by (simp add: plussub_def sup_def split: option.split)
-
-
-lemma None_in_opt [iff]:
-  "None : opt A"
-by (simp add: opt_def)
-
-lemma Some_in_opt [iff]:
-  "(Some x : opt A) = (x:A)"
-apply (unfold opt_def)
-apply auto
-done 
-
-
-lemma semilat_opt [intro, simp]:
-  "\<And>L. err_semilat L \<Longrightarrow> err_semilat (Opt.esl L)"
-proof (unfold Opt.esl_def Err.sl_def, simp add: split_tupled_all)
-  
-  fix A r f
-  assume s: "semilat (err A, Err.le r, lift2 f)"
- 
-  let ?A0 = "err A"
-  let ?r0 = "Err.le r"
-  let ?f0 = "lift2 f"
-
-  from s
-  obtain
-    ord: "order ?r0" and
-    clo: "closed ?A0 ?f0" and
-    ub1: "\<forall>x\<in>?A0. \<forall>y\<in>?A0. x <=_?r0 x +_?f0 y" and
-    ub2: "\<forall>x\<in>?A0. \<forall>y\<in>?A0. y <=_?r0 x +_?f0 y" and
-    lub: "\<forall>x\<in>?A0. \<forall>y\<in>?A0. \<forall>z\<in>?A0. x <=_?r0 z \<and> y <=_?r0 z \<longrightarrow> x +_?f0 y <=_?r0 z"
-    by (unfold semilat_def) simp
-
-  let ?A = "err (opt A)" 
-  let ?r = "Err.le (Opt.le r)"
-  let ?f = "lift2 (Opt.sup f)"
-
-  from ord
-  have "order ?r"
-    by simp
-
-  moreover
-
-  have "closed ?A ?f"
-  proof (unfold closed_def, intro strip)
-    fix x y    
-    assume x: "x : ?A" 
-    assume y: "y : ?A" 
-
-    { fix a b
-      assume ab: "x = OK a" "y = OK b"
-      
-      with x 
-      have a: "\<And>c. a = Some c \<Longrightarrow> c : A"
-        by (clarsimp simp add: opt_def)
-
-      from ab y
-      have b: "\<And>d. b = Some d \<Longrightarrow> d : A"
-        by (clarsimp simp add: opt_def)
-      
-      { fix c d assume "a = Some c" "b = Some d"
-        with ab x y
-        have "c:A & d:A"
-          by (simp add: err_def opt_def Bex_def)
-        with clo
-        have "f c d : err A"
-          by (simp add: closed_def plussub_def err_def lift2_def)
-        moreover
-        fix z assume "f c d = OK z"
-        ultimately
-        have "z : A" by simp
-      } note f_closed = this    
-
-      have "sup f a b : ?A"
-      proof (cases a)
-        case None
-        thus ?thesis
-          by (simp add: sup_def opt_def) (cases b, simp, simp add: b Bex_def)
-      next
-        case Some
-        thus ?thesis
-          by (auto simp add: sup_def opt_def Bex_def a b f_closed split: err.split option.split)
-      qed
-    }
-
-    thus "x +_?f y : ?A"
-      by (simp add: plussub_def lift2_def split: err.split)
-  qed
-    
-  moreover
-
-  { fix a b c 
-    assume "a \<in> opt A" "b \<in> opt A" "a +_(sup f) b = OK c" 
-    moreover
-    from ord have "order r" by simp
-    moreover
-    { fix x y z
-      assume "x \<in> A" "y \<in> A" 
-      hence "OK x \<in> err A \<and> OK y \<in> err A" by simp
-      with ub1 ub2
-      have "(OK x) <=_(Err.le r) (OK x) +_(lift2 f) (OK y) \<and>
-            (OK y) <=_(Err.le r) (OK x) +_(lift2 f) (OK y)"
-        by blast
-      moreover
-      assume "x +_f y = OK z"
-      ultimately
-      have "x <=_r z \<and> y <=_r z"
-        by (auto simp add: plussub_def lift2_def Err.le_def lesub_def)
-    }
-    ultimately
-    have "a <=_(le r) c \<and> b <=_(le r) c"
-      by (auto simp add: sup_def le_def lesub_def plussub_def 
-               dest: order_refl split: option.splits err.splits)
-  }
-     
-  hence "(\<forall>x\<in>?A. \<forall>y\<in>?A. x <=_?r x +_?f y) \<and> (\<forall>x\<in>?A. \<forall>y\<in>?A. y <=_?r x +_?f y)"
-    by (auto simp add: lesub_def plussub_def Err.le_def lift2_def split: err.split)
-
-  moreover
-
-  have "\<forall>x\<in>?A. \<forall>y\<in>?A. \<forall>z\<in>?A. x <=_?r z \<and> y <=_?r z \<longrightarrow> x +_?f y <=_?r z"
-  proof (intro strip, elim conjE)
-    fix x y z
-    assume xyz: "x : ?A" "y : ?A" "z : ?A"
-    assume xz: "x <=_?r z"
-    assume yz: "y <=_?r z"
-
-    { fix a b c
-      assume ok: "x = OK a" "y = OK b" "z = OK c"
-
-      { fix d e g
-        assume some: "a = Some d" "b = Some e" "c = Some g"
-        
-        with ok xyz
-        obtain "OK d:err A" "OK e:err A" "OK g:err A"
-          by simp
-        with lub
-        have "\<lbrakk> (OK d) <=_(Err.le r) (OK g); (OK e) <=_(Err.le r) (OK g) \<rbrakk>
-          \<Longrightarrow> (OK d) +_(lift2 f) (OK e) <=_(Err.le r) (OK g)"
-          by blast
-        hence "\<lbrakk> d <=_r g; e <=_r g \<rbrakk> \<Longrightarrow> \<exists>y. d +_f e = OK y \<and> y <=_r g"
-          by simp
-
-        with ok some xyz xz yz
-        have "x +_?f y <=_?r z"
-          by (auto simp add: sup_def le_def lesub_def lift2_def plussub_def Err.le_def)
-      } note this [intro!]
-
-      from ok xyz xz yz
-      have "x +_?f y <=_?r z"
-        by - (cases a, simp, cases b, simp, cases c, simp, blast)
-    }
-    
-    with xyz xz yz
-    show "x +_?f y <=_?r z"
-      by - (cases x, simp, cases y, simp, cases z, simp+)
-  qed
-
-  ultimately
-
-  show "semilat (?A,?r,?f)"
-    by (unfold semilat_def) simp
-qed 
-
-lemma top_le_opt_Some [iff]: 
-  "top (le r) (Some T) = top r T"
-apply (unfold top_def)
-apply (rule iffI)
- apply blast
-apply (rule allI)
-apply (case_tac "x")
-apply simp+
-done 
-
-lemma Top_le_conv:
-  "\<lbrakk> order r; top r T \<rbrakk> \<Longrightarrow> (T <=_r x) = (x = T)"
-apply (unfold top_def)
-apply (blast intro: order_antisym)
-done 
-
-
-lemma acc_le_optI [intro!]:
-  "acc r \<Longrightarrow> acc(le r)"
-apply (unfold acc_def lesub_def le_def lesssub_def)
-apply (simp add: wfP_eq_minimal split: option.split)
-apply clarify
-apply (case_tac "? a. Some a : Q")
- apply (erule_tac x = "{a . Some a : Q}" in allE)
- apply blast
-apply (case_tac "x")
- apply blast
-apply blast
-done 
-
-lemma option_map_in_optionI:
-  "\<lbrakk> ox : opt S; !x:S. ox = Some x \<longrightarrow> f x : S \<rbrakk> 
-  \<Longrightarrow> Option.map f ox : opt S";
-apply (unfold Option.map_def)
-apply (simp split: option.split)
-apply blast
-done 
-
-end
--- a/src/HOL/MicroJava/BV/Product.thy	Wed Dec 02 12:04:07 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,146 +0,0 @@
-(*  Title:      HOL/MicroJava/BV/Product.thy
-    ID:         $Id$
-    Author:     Tobias Nipkow
-    Copyright   2000 TUM
-
-Products as semilattices
-*)
-
-header {* \isaheader{Products as Semilattices} *}
-
-theory Product
-imports Err
-begin
-
-constdefs
- le :: "'a ord \<Rightarrow> 'b ord \<Rightarrow> ('a * 'b) ord"
-"le rA rB == %(a,b) (a',b'). a <=_rA a' & b <=_rB b'"
-
- sup :: "'a ebinop \<Rightarrow> 'b ebinop \<Rightarrow> ('a * 'b)ebinop"
-"sup f g == %(a1,b1)(a2,b2). Err.sup Pair (a1 +_f a2) (b1 +_g b2)"
-
- esl :: "'a esl \<Rightarrow> 'b esl \<Rightarrow> ('a * 'b ) esl"
-"esl == %(A,rA,fA) (B,rB,fB). (A <*> B, le rA rB, sup fA fB)"
-
-syntax "@lesubprod" :: "'a*'b \<Rightarrow> 'a ord \<Rightarrow> 'b ord \<Rightarrow> 'b \<Rightarrow> bool"
-       ("(_ /<='(_,_') _)" [50, 0, 0, 51] 50)
-translations "p <=(rA,rB) q" == "p <=_(Product.le rA rB) q"
-
-lemma unfold_lesub_prod:
-  "p <=(rA,rB) q == le rA rB p q"
-  by (simp add: lesub_def)
-
-lemma le_prod_Pair_conv [iff]:
-  "((a1,b1) <=(rA,rB) (a2,b2)) = (a1 <=_rA a2 & b1 <=_rB b2)"
-  by (simp add: lesub_def le_def)
-
-lemma less_prod_Pair_conv:
-  "((a1,b1) <_(Product.le rA rB) (a2,b2)) = 
-  (a1 <_rA a2 & b1 <=_rB b2 | a1 <=_rA a2 & b1 <_rB b2)"
-apply (unfold lesssub_def)
-apply simp
-apply blast
-done
-
-lemma order_le_prod [iff]:
-  "order(Product.le rA rB) = (order rA & order rB)"
-apply (unfold Semilat.order_def)
-apply simp
-apply blast
-done 
-
-
-lemma acc_le_prodI [intro!]:
-  "\<lbrakk> acc rA; acc rB \<rbrakk> \<Longrightarrow> acc(Product.le rA rB)"
-apply (unfold acc_def)
-apply (rule wfP_subset)
- apply (erule wf_lex_prod [to_pred, THEN wfP_wf_eq [THEN iffD2]])
- apply assumption
-apply (auto simp add: lesssub_def less_prod_Pair_conv lex_prod_def)
-done
-
-
-lemma closed_lift2_sup:
-  "\<lbrakk> closed (err A) (lift2 f); closed (err B) (lift2 g) \<rbrakk> \<Longrightarrow> 
-  closed (err(A<*>B)) (lift2(sup f g))";
-apply (unfold closed_def plussub_def lift2_def err_def sup_def)
-apply (simp split: err.split)
-apply blast
-done 
-
-lemma unfold_plussub_lift2:
-  "e1 +_(lift2 f) e2 == lift2 f e1 e2"
-  by (simp add: plussub_def)
-
-
-lemma plus_eq_Err_conv [simp]:
-  assumes "x:A" and "y:A"
-    and "semilat(err A, Err.le r, lift2 f)"
-  shows "(x +_f y = Err) = (~(? z:A. x <=_r z & y <=_r z))"
-proof -
-  have plus_le_conv2:
-    "\<And>r f z. \<lbrakk> z : err A; semilat (err A, r, f); OK x : err A; OK y : err A;
-                 OK x +_f OK y <=_r z\<rbrakk> \<Longrightarrow> OK x <=_r z \<and> OK y <=_r z"
-    by (rule Semilat.plus_le_conv [OF Semilat.intro, THEN iffD1])
-  from prems show ?thesis
-  apply (rule_tac iffI)
-   apply clarify
-   apply (drule OK_le_err_OK [THEN iffD2])
-   apply (drule OK_le_err_OK [THEN iffD2])
-   apply (drule Semilat.lub [OF Semilat.intro, of _ _ _ "OK x" _ "OK y"])
-        apply assumption
-       apply assumption
-      apply simp
-     apply simp
-    apply simp
-   apply simp
-  apply (case_tac "x +_f y")
-   apply assumption
-  apply (rename_tac "z")
-  apply (subgoal_tac "OK z: err A")
-  apply (frule plus_le_conv2)
-       apply assumption
-      apply simp
-      apply blast
-     apply simp
-    apply (blast dest: Semilat.orderI [OF Semilat.intro] order_refl)
-   apply blast
-  apply (erule subst)
-  apply (unfold semilat_def err_def closed_def)
-  apply simp
-  done
-qed
-
-lemma err_semilat_Product_esl:
-  "\<And>L1 L2. \<lbrakk> err_semilat L1; err_semilat L2 \<rbrakk> \<Longrightarrow> err_semilat(Product.esl L1 L2)"
-apply (unfold esl_def Err.sl_def)
-apply (simp (no_asm_simp) only: split_tupled_all)
-apply simp
-apply (simp (no_asm) only: semilat_Def)
-apply (simp (no_asm_simp) only: Semilat.closedI [OF Semilat.intro] closed_lift2_sup)
-apply (simp (no_asm) only: unfold_lesub_err Err.le_def unfold_plussub_lift2 sup_def)
-apply (auto elim: semilat_le_err_OK1 semilat_le_err_OK2
-            simp add: lift2_def  split: err.split)
-apply (blast dest: Semilat.orderI [OF Semilat.intro])
-apply (blast dest: Semilat.orderI [OF Semilat.intro])
-
-apply (rule OK_le_err_OK [THEN iffD1])
-apply (erule subst, subst OK_lift2_OK [symmetric], rule Semilat.lub [OF Semilat.intro])
-apply simp
-apply simp
-apply simp
-apply simp
-apply simp
-apply simp
-
-apply (rule OK_le_err_OK [THEN iffD1])
-apply (erule subst, subst OK_lift2_OK [symmetric], rule Semilat.lub [OF Semilat.intro])
-apply simp
-apply simp
-apply simp
-apply simp
-apply simp
-apply simp
-done 
-
-end
--- a/src/HOL/MicroJava/BV/Semilat.thy	Wed Dec 02 12:04:07 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,373 +0,0 @@
-(*  Title:      HOL/MicroJava/BV/Semilat.thy
-    ID:         $Id$
-    Author:     Tobias Nipkow
-    Copyright   2000 TUM
-
-Semilattices
-*)
-
-header {* 
-  \chapter{Bytecode Verifier}\label{cha:bv}
-  \isaheader{Semilattices} 
-*}
-
-theory Semilat
-imports Main While_Combinator
-begin
-
-types 'a ord    = "'a \<Rightarrow> 'a \<Rightarrow> bool"
-      'a binop  = "'a \<Rightarrow> 'a \<Rightarrow> 'a"
-      'a sl     = "'a set * 'a ord * 'a binop"
-
-consts
- "@lesub"   :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /<='__ _)" [50, 1000, 51] 50)
- "@lesssub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /<'__ _)" [50, 1000, 51] 50)
-defs
-lesub_def:   "x <=_r y == r x y"
-lesssub_def: "x <_r y  == x <=_r y & x ~= y"
-
-syntax (xsymbols)
- "@lesub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<le>\<^sub>_ _)" [50, 1000, 51] 50)
-
-consts
- "@plussub" :: "'a \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'c" ("(_ /+'__ _)" [65, 1000, 66] 65)
-defs
-plussub_def: "x +_f y == f x y"
-
-syntax (xsymbols)
- "@plussub" :: "'a \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'c" ("(_ /+\<^sub>_ _)" [65, 1000, 66] 65)
-
-syntax (xsymbols)
- "@plussub" :: "'a \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'c" ("(_ /\<squnion>\<^sub>_ _)" [65, 1000, 66] 65)
-
-
-constdefs
- order :: "'a ord \<Rightarrow> bool"
-"order r == (!x. x <=_r x) &
-            (!x y. x <=_r y & y <=_r x \<longrightarrow> x=y) &
-            (!x y z. x <=_r y & y <=_r z \<longrightarrow> x <=_r z)"
-
- acc :: "'a ord \<Rightarrow> bool"
-"acc r == wfP (\<lambda>y x. x <_r y)"
-
- top :: "'a ord \<Rightarrow> 'a \<Rightarrow> bool"
-"top r T == !x. x <=_r T"
-
- closed :: "'a set \<Rightarrow> 'a binop \<Rightarrow> bool"
-"closed A f == !x:A. !y:A. x +_f y : A"
-
- semilat :: "'a sl \<Rightarrow> bool"
-"semilat == %(A,r,f). order r & closed A f &
-                (!x:A. !y:A. x <=_r x +_f y)  &
-                (!x:A. !y:A. y <=_r x +_f y)  &
-                (!x:A. !y:A. !z:A. x <=_r z & y <=_r z \<longrightarrow> x +_f y <=_r z)"
-
- is_ub :: "'a ord \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
-"is_ub r x y u == r x u & r y u"
-
- is_lub :: "'a ord \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
-"is_lub r x y u == is_ub r x y u & (!z. is_ub r x y z \<longrightarrow> r u z)"
-
- some_lub :: "'a ord \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a"
-"some_lub r x y == SOME z. is_lub r x y z";
-
-locale Semilat =
-  fixes A :: "'a set"
-    and r :: "'a ord"
-    and f :: "'a binop"
-  assumes semilat: "semilat(A,r,f)"
-
-lemma order_refl [simp, intro]:
-  "order r \<Longrightarrow> x <=_r x";
-  by (simp add: order_def)
-
-lemma order_antisym:
-  "\<lbrakk> order r; x <=_r y; y <=_r x \<rbrakk> \<Longrightarrow> x = y"
-apply (unfold order_def)
-apply (simp (no_asm_simp))
-done
-
-lemma order_trans:
-   "\<lbrakk> order r; x <=_r y; y <=_r z \<rbrakk> \<Longrightarrow> x <=_r z"
-apply (unfold order_def)
-apply blast
-done 
-
-lemma order_less_irrefl [intro, simp]:
-   "order r \<Longrightarrow> ~ x <_r x"
-apply (unfold order_def lesssub_def)
-apply blast
-done 
-
-lemma order_less_trans:
-  "\<lbrakk> order r; x <_r y; y <_r z \<rbrakk> \<Longrightarrow> x <_r z"
-apply (unfold order_def lesssub_def)
-apply blast
-done 
-
-lemma topD [simp, intro]:
-  "top r T \<Longrightarrow> x <=_r T"
-  by (simp add: top_def)
-
-lemma top_le_conv [simp]:
-  "\<lbrakk> order r; top r T \<rbrakk> \<Longrightarrow> (T <=_r x) = (x = T)"
-  by (blast intro: order_antisym)
-
-lemma semilat_Def:
-"semilat(A,r,f) == order r & closed A f & 
-                 (!x:A. !y:A. x <=_r x +_f y) & 
-                 (!x:A. !y:A. y <=_r x +_f y) & 
-                 (!x:A. !y:A. !z:A. x <=_r z & y <=_r z \<longrightarrow> x +_f y <=_r z)"
-apply (unfold semilat_def split_conv [THEN eq_reflection])
-apply (rule refl [THEN eq_reflection])
-done
-
-lemma (in Semilat) orderI [simp, intro]:
-  "order r"
-  by (insert semilat) (simp add: semilat_Def)
-
-lemma (in Semilat) closedI [simp, intro]:
-  "closed A f"
-  by (insert semilat) (simp add: semilat_Def)
-
-lemma closedD:
-  "\<lbrakk> closed A f; x:A; y:A \<rbrakk> \<Longrightarrow> x +_f y : A"
-  by (unfold closed_def) blast
-
-lemma closed_UNIV [simp]: "closed UNIV f"
-  by (simp add: closed_def)
-
-
-lemma (in Semilat) closed_f [simp, intro]:
-  "\<lbrakk>x:A; y:A\<rbrakk>  \<Longrightarrow> x +_f y : A"
-  by (simp add: closedD [OF closedI])
-
-lemma (in Semilat) refl_r [intro, simp]:
-  "x <=_r x"
-  by simp
-
-lemma (in Semilat) antisym_r [intro?]:
-  "\<lbrakk> x <=_r y; y <=_r x \<rbrakk> \<Longrightarrow> x = y"
-  by (rule order_antisym) auto
-  
-lemma (in Semilat) trans_r [trans, intro?]:
-  "\<lbrakk>x <=_r y; y <=_r z\<rbrakk> \<Longrightarrow> x <=_r z"
-  by (auto intro: order_trans)    
-  
-
-lemma (in Semilat) ub1 [simp, intro?]:
-  "\<lbrakk> x:A; y:A \<rbrakk> \<Longrightarrow> x <=_r x +_f y"
-  by (insert semilat) (unfold semilat_Def, simp)
-
-lemma (in Semilat) ub2 [simp, intro?]:
-  "\<lbrakk> x:A; y:A \<rbrakk> \<Longrightarrow> y <=_r x +_f y"
-  by (insert semilat) (unfold semilat_Def, simp)
-
-lemma (in Semilat) lub [simp, intro?]:
- "\<lbrakk> x <=_r z; y <=_r z; x:A; y:A; z:A \<rbrakk> \<Longrightarrow> x +_f y <=_r z";
-  by (insert semilat) (unfold semilat_Def, simp)
-
-
-lemma (in Semilat) plus_le_conv [simp]:
-  "\<lbrakk> x:A; y:A; z:A \<rbrakk> \<Longrightarrow> (x +_f y <=_r z) = (x <=_r z & y <=_r z)"
-  by (blast intro: ub1 ub2 lub order_trans)
-
-lemma (in Semilat) le_iff_plus_unchanged:
-  "\<lbrakk> x:A; y:A \<rbrakk> \<Longrightarrow> (x <=_r y) = (x +_f y = y)"
-apply (rule iffI)
- apply (blast intro: antisym_r refl_r lub ub2)
-apply (erule subst)
-apply simp
-done
-
-lemma (in Semilat) le_iff_plus_unchanged2:
-  "\<lbrakk> x:A; y:A \<rbrakk> \<Longrightarrow> (x <=_r y) = (y +_f x = y)"
-apply (rule iffI)
- apply (blast intro: order_antisym lub order_refl ub1)
-apply (erule subst)
-apply simp
-done 
-
-
-lemma (in Semilat) plus_assoc [simp]:
-  assumes a: "a \<in> A" and b: "b \<in> A" and c: "c \<in> A"
-  shows "a +_f (b +_f c) = a +_f b +_f c"
-proof -
-  from a b have ab: "a +_f b \<in> A" ..
-  from this c have abc: "(a +_f b) +_f c \<in> A" ..
-  from b c have bc: "b +_f c \<in> A" ..
-  from a this have abc': "a +_f (b +_f c) \<in> A" ..
-
-  show ?thesis
-  proof    
-    show "a +_f (b +_f c) <=_r (a +_f b) +_f c"
-    proof -
-      from a b have "a <=_r a +_f b" .. 
-      also from ab c have "\<dots> <=_r \<dots> +_f c" ..
-      finally have "a<": "a <=_r (a +_f b) +_f c" .
-      from a b have "b <=_r a +_f b" ..
-      also from ab c have "\<dots> <=_r \<dots> +_f c" ..
-      finally have "b<": "b <=_r (a +_f b) +_f c" .
-      from ab c have "c<": "c <=_r (a +_f b) +_f c" ..    
-      from "b<" "c<" b c abc have "b +_f c <=_r (a +_f b) +_f c" ..
-      from "a<" this a bc abc show ?thesis ..
-    qed
-    show "(a +_f b) +_f c <=_r a +_f (b +_f c)" 
-    proof -
-      from b c have "b <=_r b +_f c" .. 
-      also from a bc have "\<dots> <=_r a +_f \<dots>" ..
-      finally have "b<": "b <=_r a +_f (b +_f c)" .
-      from b c have "c <=_r b +_f c" ..
-      also from a bc have "\<dots> <=_r a +_f \<dots>" ..
-      finally have "c<": "c <=_r a +_f (b +_f c)" .
-      from a bc have "a<": "a <=_r a +_f (b +_f c)" ..
-      from "a<" "b<" a b abc' have "a +_f b <=_r a +_f (b +_f c)" ..
-      from this "c<" ab c abc' show ?thesis ..
-    qed
-  qed
-qed
-
-lemma (in Semilat) plus_com_lemma:
-  "\<lbrakk>a \<in> A; b \<in> A\<rbrakk> \<Longrightarrow> a +_f b <=_r b +_f a"
-proof -
-  assume a: "a \<in> A" and b: "b \<in> A"  
-  from b a have "a <=_r b +_f a" .. 
-  moreover from b a have "b <=_r b +_f a" ..
-  moreover note a b
-  moreover from b a have "b +_f a \<in> A" ..
-  ultimately show ?thesis ..
-qed
-
-lemma (in Semilat) plus_commutative:
-  "\<lbrakk>a \<in> A; b \<in> A\<rbrakk> \<Longrightarrow> a +_f b = b +_f a"
-by(blast intro: order_antisym plus_com_lemma)
-
-lemma is_lubD:
-  "is_lub r x y u \<Longrightarrow> is_ub r x y u & (!z. is_ub r x y z \<longrightarrow> r u z)"
-  by (simp add: is_lub_def)
-
-lemma is_ubI:
-  "\<lbrakk> r x u; r y u \<rbrakk> \<Longrightarrow> is_ub r x y u"
-  by (simp add: is_ub_def)
-
-lemma is_ubD:
-  "is_ub r x y u \<Longrightarrow> r x u & r y u"
-  by (simp add: is_ub_def)
-
-
-lemma is_lub_bigger1 [iff]:  
-  "is_lub (r^** ) x y y = r^** x y"
-apply (unfold is_lub_def is_ub_def)
-apply blast
-done
-
-lemma is_lub_bigger2 [iff]:
-  "is_lub (r^** ) x y x = r^** y x"
-apply (unfold is_lub_def is_ub_def)
-apply blast 
-done
-
-lemma extend_lub:
-  "\<lbrakk> single_valuedP r; is_lub (r^** ) x y u; r x' x \<rbrakk> 
-  \<Longrightarrow> EX v. is_lub (r^** ) x' y v"
-apply (unfold is_lub_def is_ub_def)
-apply (case_tac "r^** y x")
- apply (case_tac "r^** y x'")
-  apply blast
- apply (blast elim: converse_rtranclpE dest: single_valuedD)
-apply (rule exI)
-apply (rule conjI)
- apply (blast intro: converse_rtranclp_into_rtranclp dest: single_valuedD)
-apply (blast intro: rtranclp.rtrancl_into_rtrancl converse_rtranclp_into_rtranclp
-             elim: converse_rtranclpE dest: single_valuedD)
-done
-
-lemma single_valued_has_lubs [rule_format]:
-  "\<lbrakk> single_valuedP r; r^** x u \<rbrakk> \<Longrightarrow> (!y. r^** y u \<longrightarrow> 
-  (EX z. is_lub (r^** ) x y z))"
-apply (erule converse_rtranclp_induct)
- apply clarify
- apply (erule converse_rtranclp_induct)
-  apply blast
- apply (blast intro: converse_rtranclp_into_rtranclp)
-apply (blast intro: extend_lub)
-done
-
-lemma some_lub_conv:
-  "\<lbrakk> acyclicP r; is_lub (r^** ) x y u \<rbrakk> \<Longrightarrow> some_lub (r^** ) x y = u"
-apply (unfold some_lub_def is_lub_def)
-apply (rule someI2)
- apply assumption
-apply (blast intro: antisymD dest!: acyclic_impl_antisym_rtrancl [to_pred])
-done
-
-lemma is_lub_some_lub:
-  "\<lbrakk> single_valuedP r; acyclicP r; r^** x u; r^** y u \<rbrakk> 
-  \<Longrightarrow> is_lub (r^** ) x y (some_lub (r^** ) x y)";
-  by (fastsimp dest: single_valued_has_lubs simp add: some_lub_conv)
-
-subsection{*An executable lub-finder*}
-
-constdefs
- exec_lub :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a binop"
-"exec_lub r f x y == while (\<lambda>z. \<not> r\<^sup>*\<^sup>* x z) f y"
-
-
-lemma acyclic_single_valued_finite:
- "\<lbrakk>acyclicP r; single_valuedP r; r\<^sup>*\<^sup>* x y \<rbrakk>
-  \<Longrightarrow> finite ({(x, y). r x y} \<inter> {a. r\<^sup>*\<^sup>* x a} \<times> {b. r\<^sup>*\<^sup>* b y})"
-apply(erule converse_rtranclp_induct)
- apply(rule_tac B = "{}" in finite_subset)
-  apply(simp only:acyclic_def [to_pred])
-  apply(blast intro:rtranclp_into_tranclp2 rtranclp_tranclp_tranclp)
- apply simp
-apply(rename_tac x x')
-apply(subgoal_tac "{(x, y). r x y} \<inter> {a. r\<^sup>*\<^sup>* x a} \<times> {b. r\<^sup>*\<^sup>* b y} =
-                   insert (x,x') ({(x, y). r x y} \<inter> {a. r\<^sup>*\<^sup>* x' a} \<times> {b. r\<^sup>*\<^sup>* b y})")
- apply simp
-apply(blast intro:converse_rtranclp_into_rtranclp
-            elim:converse_rtranclpE dest:single_valuedD)
-done
-
-
-lemma exec_lub_conv:
-  "\<lbrakk> acyclicP r; !x y. r x y \<longrightarrow> f x = y; is_lub (r\<^sup>*\<^sup>*) x y u \<rbrakk> \<Longrightarrow>
-  exec_lub r f x y = u";
-apply(unfold exec_lub_def)
-apply(rule_tac P = "\<lambda>z. r\<^sup>*\<^sup>* y z \<and> r\<^sup>*\<^sup>* z u" and
-               r = "({(x, y). r x y} \<inter> {(a,b). r\<^sup>*\<^sup>* y a \<and> r\<^sup>*\<^sup>* b u})^-1" in while_rule)
-    apply(blast dest: is_lubD is_ubD)
-   apply(erule conjE)
-   apply(erule_tac z = u in converse_rtranclpE)
-    apply(blast dest: is_lubD is_ubD)
-   apply(blast dest: rtranclp.rtrancl_into_rtrancl)
-  apply(rename_tac s)
-  apply(subgoal_tac "is_ub (r\<^sup>*\<^sup>*) x y s")
-   prefer 2; apply(simp add:is_ub_def)
-  apply(subgoal_tac "r\<^sup>*\<^sup>* u s")
-   prefer 2; apply(blast dest:is_lubD)
-  apply(erule converse_rtranclpE)
-   apply blast
-  apply(simp only:acyclic_def [to_pred])
-  apply(blast intro:rtranclp_into_tranclp2 rtranclp_tranclp_tranclp)
- apply(rule finite_acyclic_wf)
-  apply simp
-  apply(erule acyclic_single_valued_finite)
-   apply(blast intro:single_valuedI)
-  apply(simp add:is_lub_def is_ub_def)
- apply simp
- apply(erule acyclic_subset)
- apply blast
-apply simp
-apply(erule conjE)
-apply(erule_tac z = u in converse_rtranclpE)
- apply(blast dest: is_lubD is_ubD)
-apply blast
-done
-
-lemma is_lub_exec_lub:
-  "\<lbrakk> single_valuedP r; acyclicP r; r^** x u; r^** y u; !x y. r x y \<longrightarrow> f x = y \<rbrakk>
-  \<Longrightarrow> is_lub (r^** ) x y (exec_lub r f x y)"
-  by (fastsimp dest: single_valued_has_lubs simp add: exec_lub_conv)
-
-end
--- a/src/HOL/MicroJava/BV/SemilatAlg.thy	Wed Dec 02 12:04:07 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,188 +0,0 @@
-(*  Title:      HOL/MicroJava/BV/SemilatAlg.thy
-    ID:         $Id$
-    Author:     Gerwin Klein
-    Copyright   2002 Technische Universitaet Muenchen
-*)
-
-header {* \isaheader{More on Semilattices} *}
-
-theory SemilatAlg
-imports Typing_Framework Product
-begin
-
-
-constdefs 
-  lesubstep_type :: "(nat \<times> 's) list \<Rightarrow> 's ord \<Rightarrow> (nat \<times> 's) list \<Rightarrow> bool"
-                    ("(_ /<=|_| _)" [50, 0, 51] 50)
-  "x <=|r| y \<equiv> \<forall>(p,s) \<in> set x. \<exists>s'. (p,s') \<in> set y \<and> s <=_r s'"
-
-consts
- "@plusplussub" :: "'a list \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" ("(_ /++'__ _)" [65, 1000, 66] 65)
-primrec
-  "[] ++_f y = y"
-  "(x#xs) ++_f y = xs ++_f (x +_f y)"
-
-constdefs
- bounded :: "'s step_type \<Rightarrow> nat \<Rightarrow> bool"
-"bounded step n == !p<n. !s. !(q,t):set(step p s). q<n"  
-
- pres_type :: "'s step_type \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> bool"
-"pres_type step n A == \<forall>s\<in>A. \<forall>p<n. \<forall>(q,s')\<in>set (step p s). s' \<in> A"
-
- mono :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> bool"
-"mono r step n A ==
- \<forall>s p t. s \<in> A \<and> p < n \<and> s <=_r t \<longrightarrow> step p s <=|r| step p t"
-
-
-lemma pres_typeD:
-  "\<lbrakk> pres_type step n A; s\<in>A; p<n; (q,s')\<in>set (step p s) \<rbrakk> \<Longrightarrow> s' \<in> A"
-  by (unfold pres_type_def, blast)
-
-lemma monoD:
-  "\<lbrakk> mono r step n A; p < n; s\<in>A; s <=_r t \<rbrakk> \<Longrightarrow> step p s <=|r| step p t"
-  by (unfold mono_def, blast)
-
-lemma boundedD: 
-  "\<lbrakk> bounded step n; p < n; (q,t) : set (step p xs) \<rbrakk> \<Longrightarrow> q < n" 
-  by (unfold bounded_def, blast)
-
-lemma lesubstep_type_refl [simp, intro]:
-  "(\<And>x. x <=_r x) \<Longrightarrow> x <=|r| x"
-  by (unfold lesubstep_type_def) auto
-
-lemma lesub_step_typeD:
-  "a <=|r| b \<Longrightarrow> (x,y) \<in> set a \<Longrightarrow> \<exists>y'. (x, y') \<in> set b \<and> y <=_r y'"
-  by (unfold lesubstep_type_def) blast
-
-
-lemma list_update_le_listI [rule_format]:
-  "set xs <= A \<longrightarrow> set ys <= A \<longrightarrow> xs <=[r] ys \<longrightarrow> p < size xs \<longrightarrow>  
-   x <=_r ys!p \<longrightarrow> semilat(A,r,f) \<longrightarrow> x\<in>A \<longrightarrow> 
-   xs[p := x +_f xs!p] <=[r] ys"
-  apply (unfold Listn.le_def lesub_def semilat_def)
-  apply (simp add: list_all2_conv_all_nth nth_list_update)
-  done
-
-
-lemma plusplus_closed: assumes "semilat (A, r, f)" shows
-  "\<And>y. \<lbrakk> set x \<subseteq> A; y \<in> A\<rbrakk> \<Longrightarrow> x ++_f y \<in> A" (is "PROP ?P")
-proof -
-  interpret Semilat A r f using assms by (rule Semilat.intro)
-  show "PROP ?P" proof (induct x)
-    show "\<And>y. y \<in> A \<Longrightarrow> [] ++_f y \<in> A" by simp
-    fix y x xs
-    assume y: "y \<in> A" and xs: "set (x#xs) \<subseteq> A"
-    assume IH: "\<And>y. \<lbrakk> set xs \<subseteq> A; y \<in> A\<rbrakk> \<Longrightarrow> xs ++_f y \<in> A"
-    from xs obtain x: "x \<in> A" and xs': "set xs \<subseteq> A" by simp
-    from x y have "(x +_f y) \<in> A" ..
-    with xs' have "xs ++_f (x +_f y) \<in> A" by (rule IH)
-    thus "(x#xs) ++_f y \<in> A" by simp
-  qed
-qed
-
-lemma (in Semilat) pp_ub2:
- "\<And>y. \<lbrakk> set x \<subseteq> A; y \<in> A\<rbrakk> \<Longrightarrow> y <=_r x ++_f y"
-proof (induct x)
-  from semilat show "\<And>y. y <=_r [] ++_f y" by simp
-  
-  fix y a l
-  assume y:  "y \<in> A"
-  assume "set (a#l) \<subseteq> A"
-  then obtain a: "a \<in> A" and x: "set l \<subseteq> A" by simp
-  assume "\<And>y. \<lbrakk>set l \<subseteq> A; y \<in> A\<rbrakk> \<Longrightarrow> y <=_r l ++_f y"
-  hence IH: "\<And>y. y \<in> A \<Longrightarrow> y <=_r l ++_f y" using x .
-
-  from a y have "y <=_r a +_f y" ..
-  also from a y have "a +_f y \<in> A" ..
-  hence "(a +_f y) <=_r l ++_f (a +_f y)" by (rule IH)
-  finally have "y <=_r l ++_f (a +_f y)" .
-  thus "y <=_r (a#l) ++_f y" by simp
-qed
-
-
-lemma (in Semilat) pp_ub1:
-shows "\<And>y. \<lbrakk>set ls \<subseteq> A; y \<in> A; x \<in> set ls\<rbrakk> \<Longrightarrow> x <=_r ls ++_f y"
-proof (induct ls)
-  show "\<And>y. x \<in> set [] \<Longrightarrow> x <=_r [] ++_f y" by simp
-
-  fix y s ls
-  assume "set (s#ls) \<subseteq> A"
-  then obtain s: "s \<in> A" and ls: "set ls \<subseteq> A" by simp
-  assume y: "y \<in> A" 
-
-  assume 
-    "\<And>y. \<lbrakk>set ls \<subseteq> A; y \<in> A; x \<in> set ls\<rbrakk> \<Longrightarrow> x <=_r ls ++_f y"
-  hence IH: "\<And>y. x \<in> set ls \<Longrightarrow> y \<in> A \<Longrightarrow> x <=_r ls ++_f y" using ls .
-
-  assume "x \<in> set (s#ls)"
-  then obtain xls: "x = s \<or> x \<in> set ls" by simp
-  moreover {
-    assume xs: "x = s"
-    from s y have "s <=_r s +_f y" ..
-    also from s y have "s +_f y \<in> A" ..
-    with ls have "(s +_f y) <=_r ls ++_f (s +_f y)" by (rule pp_ub2)
-    finally have "s <=_r ls ++_f (s +_f y)" .
-    with xs have "x <=_r ls ++_f (s +_f y)" by simp
-  } 
-  moreover {
-    assume "x \<in> set ls"
-    hence "\<And>y. y \<in> A \<Longrightarrow> x <=_r ls ++_f y" by (rule IH)
-    moreover from s y have "s +_f y \<in> A" ..
-    ultimately have "x <=_r ls ++_f (s +_f y)" .
-  }
-  ultimately 
-  have "x <=_r ls ++_f (s +_f y)" by blast
-  thus "x <=_r (s#ls) ++_f y" by simp
-qed
-
-
-lemma (in Semilat) pp_lub:
-  assumes z: "z \<in> A"
-  shows 
-  "\<And>y. y \<in> A \<Longrightarrow> set xs \<subseteq> A \<Longrightarrow> \<forall>x \<in> set xs. x <=_r z \<Longrightarrow> y <=_r z \<Longrightarrow> xs ++_f y <=_r z"
-proof (induct xs)
-  fix y assume "y <=_r z" thus "[] ++_f y <=_r z" by simp
-next
-  fix y l ls assume y: "y \<in> A" and "set (l#ls) \<subseteq> A"
-  then obtain l: "l \<in> A" and ls: "set ls \<subseteq> A" by auto
-  assume "\<forall>x \<in> set (l#ls). x <=_r z"
-  then obtain lz: "l <=_r z" and lsz: "\<forall>x \<in> set ls. x <=_r z" by auto
-  assume "y <=_r z" with lz have "l +_f y <=_r z" using l y z ..
-  moreover
-  from l y have "l +_f y \<in> A" ..
-  moreover
-  assume "\<And>y. y \<in> A \<Longrightarrow> set ls \<subseteq> A \<Longrightarrow> \<forall>x \<in> set ls. x <=_r z \<Longrightarrow> y <=_r z
-          \<Longrightarrow> ls ++_f y <=_r z"
-  ultimately
-  have "ls ++_f (l +_f y) <=_r z" using ls lsz by -
-  thus "(l#ls) ++_f y <=_r z" by simp
-qed
-
-
-lemma ub1':
-  assumes "semilat (A, r, f)"
-  shows "\<lbrakk>\<forall>(p,s) \<in> set S. s \<in> A; y \<in> A; (a,b) \<in> set S\<rbrakk> 
-  \<Longrightarrow> b <=_r map snd [(p', t')\<leftarrow>S. p' = a] ++_f y" 
-proof -
-  interpret Semilat A r f using assms by (rule Semilat.intro)
-
-  let "b <=_r ?map ++_f y" = ?thesis
-
-  assume "y \<in> A"
-  moreover
-  assume "\<forall>(p,s) \<in> set S. s \<in> A"
-  hence "set ?map \<subseteq> A" by auto
-  moreover
-  assume "(a,b) \<in> set S"
-  hence "b \<in> set ?map" by (induct S, auto)
-  ultimately
-  show ?thesis by - (rule pp_ub1)
-qed
-    
-
-lemma plusplus_empty:  
-  "\<forall>s'. (q, s') \<in> set S \<longrightarrow> s' +_f ss ! q = ss ! q \<Longrightarrow>
-   (map snd [(p', t') \<leftarrow> S. p' = q] ++_f ss ! q) = ss ! q"
-  by (induct S) auto 
-
-end
--- a/src/HOL/MicroJava/BV/Typing_Framework.thy	Wed Dec 02 12:04:07 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,37 +0,0 @@
-(*  Title:      HOL/MicroJava/BV/Typing_Framework.thy
-    ID:         $Id$
-    Author:     Tobias Nipkow
-    Copyright   2000 TUM
-*)
-
-header {* \isaheader{Typing and Dataflow Analysis Framework} *}
-
-theory Typing_Framework
-imports Listn
-begin
-
-text {* 
-  The relationship between dataflow analysis and a welltyped-instruction predicate. 
-*}
-types
-  's step_type = "nat \<Rightarrow> 's \<Rightarrow> (nat \<times> 's) list"
-
-constdefs
- stable :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> nat \<Rightarrow> bool"
-"stable r step ss p == !(q,s'):set(step p (ss!p)). s' <=_r ss!q"
-
- stables :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool"
-"stables r step ss == !p<size ss. stable r step ss p"
-
- wt_step ::
-"'s ord \<Rightarrow> 's \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool"
-"wt_step r T step ts ==
- !p<size(ts). ts!p ~= T & stable r step ts p"
-
- is_bcv :: "'s ord \<Rightarrow> 's \<Rightarrow> 's step_type 
-           \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> ('s list \<Rightarrow> 's list) \<Rightarrow> bool"  
-"is_bcv r T step n A bcv == !ss : list n A.
-   (!p<n. (bcv ss)!p ~= T) =
-   (? ts: list n A. ss <=[r] ts & wt_step r T step ts)"
-
-end
--- a/src/HOL/MicroJava/BV/Typing_Framework_JVM.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/BV/Typing_Framework_JVM.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -1,13 +1,13 @@
 (*  Title:      HOL/MicroJava/BV/JVM.thy
-    ID:         $Id$
     Author:     Tobias Nipkow, Gerwin Klein
     Copyright   2000 TUM
 *)
 
 header {* \isaheader{The Typing Framework for the JVM}\label{sec:JVM} *}
 
-theory Typing_Framework_JVM imports Typing_Framework_err JVMType EffectMono BVSpec begin
-
+theory Typing_Framework_JVM
+imports "../DFA/Abstract_BV" JVMType EffectMono BVSpec
+begin
 
 constdefs
   exec :: "jvm_prog \<Rightarrow> nat \<Rightarrow> ty \<Rightarrow> exception_table \<Rightarrow> instr list \<Rightarrow> JVMType.state step_type"
--- a/src/HOL/MicroJava/BV/Typing_Framework_err.thy	Wed Dec 02 12:04:07 2009 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,255 +0,0 @@
-(*  Title:      HOL/MicroJava/BV/Typing_Framework_err.thy
-    ID:         $Id$
-    Author:     Gerwin Klein
-    Copyright   2000 TUM
-
-*)
-
-header {* \isaheader{Lifting the Typing Framework to err, app, and eff} *}
-
-theory Typing_Framework_err
-imports Typing_Framework SemilatAlg
-begin
-
-constdefs
-
-wt_err_step :: "'s ord \<Rightarrow> 's err step_type \<Rightarrow> 's err list \<Rightarrow> bool"
-"wt_err_step r step ts \<equiv> wt_step (Err.le r) Err step ts"
-
-wt_app_eff :: "'s ord \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> bool) \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool"
-"wt_app_eff r app step ts \<equiv>
-  \<forall>p < size ts. app p (ts!p) \<and> (\<forall>(q,t) \<in> set (step p (ts!p)). t <=_r ts!q)"
-
-map_snd :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('a \<times> 'c) list"
-"map_snd f \<equiv> map (\<lambda>(x,y). (x, f y))"
-
-error :: "nat \<Rightarrow> (nat \<times> 'a err) list"
-"error n \<equiv> map (\<lambda>x. (x,Err)) [0..<n]"
-
-err_step :: "nat \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> bool) \<Rightarrow> 's step_type \<Rightarrow> 's err step_type"
-"err_step n app step p t \<equiv> 
-  case t of 
-    Err   \<Rightarrow> error n
-  | OK t' \<Rightarrow> if app p t' then map_snd OK (step p t') else error n"
-
-app_mono :: "'s ord \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> bool"
-"app_mono r app n A \<equiv>
- \<forall>s p t. s \<in> A \<and> p < n \<and> s <=_r t \<longrightarrow> app p t \<longrightarrow> app p s"
-
-
-lemmas err_step_defs = err_step_def map_snd_def error_def
-
-
-lemma bounded_err_stepD:
-  "bounded (err_step n app step) n \<Longrightarrow> 
-  p < n \<Longrightarrow>  app p a \<Longrightarrow> (q,b) \<in> set (step p a) \<Longrightarrow> 
-  q < n"
-  apply (simp add: bounded_def err_step_def)
-  apply (erule allE, erule impE, assumption)
-  apply (erule_tac x = "OK a" in allE, drule bspec)
-   apply (simp add: map_snd_def)
-   apply fast
-  apply simp
-  done
-
-
-lemma in_map_sndD: "(a,b) \<in> set (map_snd f xs) \<Longrightarrow> \<exists>b'. (a,b') \<in> set xs"
-  apply (induct xs)
-  apply (auto simp add: map_snd_def)
-  done
-
-
-lemma bounded_err_stepI:
-  "\<forall>p. p < n \<longrightarrow> (\<forall>s. ap p s \<longrightarrow> (\<forall>(q,s') \<in> set (step p s). q < n))
-  \<Longrightarrow> bounded (err_step n ap step) n"
-apply (clarsimp simp: bounded_def err_step_def split: err.splits)
-apply (simp add: error_def image_def)
-apply (blast dest: in_map_sndD)
-done
-
-
-lemma bounded_lift:
-  "bounded step n \<Longrightarrow> bounded (err_step n app step) n"
-  apply (unfold bounded_def err_step_def error_def)
-  apply clarify
-  apply (erule allE, erule impE, assumption)
-  apply (case_tac s)
-  apply (auto simp add: map_snd_def split: split_if_asm)
-  done
-
-
-lemma le_list_map_OK [simp]:
-  "\<And>b. map OK a <=[Err.le r] map OK b = (a <=[r] b)"
-  apply (induct a)
-   apply simp
-  apply simp
-  apply (case_tac b)
-   apply simp
-  apply simp
-  done
-
-
-lemma map_snd_lessI:
-  "x <=|r| y \<Longrightarrow> map_snd OK x <=|Err.le r| map_snd OK y"
-  apply (induct x)
-  apply (unfold lesubstep_type_def map_snd_def)
-  apply auto
-  done
-
-
-lemma mono_lift:
-  "order r \<Longrightarrow> app_mono r app n A \<Longrightarrow> bounded (err_step n app step) n \<Longrightarrow>
-  \<forall>s p t. s \<in> A \<and> p < n \<and> s <=_r t \<longrightarrow> app p t \<longrightarrow> step p s <=|r| step p t \<Longrightarrow>
-  mono (Err.le r) (err_step n app step) n (err A)"
-apply (unfold app_mono_def mono_def err_step_def)
-apply clarify
-apply (case_tac s)
- apply simp 
-apply simp
-apply (case_tac t)
- apply simp
- apply clarify
- apply (simp add: lesubstep_type_def error_def)
- apply clarify
- apply (drule in_map_sndD)
- apply clarify
- apply (drule bounded_err_stepD, assumption+)
- apply (rule exI [of _ Err])
- apply simp
-apply simp
-apply (erule allE, erule allE, erule allE, erule impE)
- apply (rule conjI, assumption)
- apply (rule conjI, assumption)
- apply assumption
-apply (rule conjI)
-apply clarify
-apply (erule allE, erule allE, erule allE, erule impE)
- apply (rule conjI, assumption)
- apply (rule conjI, assumption)
- apply assumption
-apply (erule impE, assumption)
-apply (rule map_snd_lessI, assumption)
-apply clarify
-apply (simp add: lesubstep_type_def error_def)
-apply clarify
-apply (drule in_map_sndD)
-apply clarify
-apply (drule bounded_err_stepD, assumption+)
-apply (rule exI [of _ Err])
-apply simp
-done
- 
-lemma in_errorD:
-  "(x,y) \<in> set (error n) \<Longrightarrow> y = Err"
-  by (auto simp add: error_def)
-
-lemma pres_type_lift:
-  "\<forall>s\<in>A. \<forall>p. p < n \<longrightarrow> app p s \<longrightarrow> (\<forall>(q, s')\<in>set (step p s). s' \<in> A) 
-  \<Longrightarrow> pres_type (err_step n app step) n (err A)"  
-apply (unfold pres_type_def err_step_def)
-apply clarify
-apply (case_tac b)
- apply simp
-apply (case_tac s)
- apply simp
- apply (drule in_errorD)
- apply simp
-apply (simp add: map_snd_def split: split_if_asm)
- apply fast
-apply (drule in_errorD)
-apply simp
-done
-
-
-
-text {*
-  There used to be a condition here that each instruction must have a
-  successor. This is not needed any more, because the definition of
-  @{term error} trivially ensures that there is a successor for
-  the critical case where @{term app} does not hold. 
-*}
-lemma wt_err_imp_wt_app_eff:
-  assumes wt: "wt_err_step r (err_step (size ts) app step) ts"
-  assumes b:  "bounded (err_step (size ts) app step) (size ts)"
-  shows "wt_app_eff r app step (map ok_val ts)"
-proof (unfold wt_app_eff_def, intro strip, rule conjI)
-  fix p assume "p < size (map ok_val ts)"
-  hence lp: "p < size ts" by simp
-  hence ts: "0 < size ts" by (cases p) auto
-  hence err: "(0,Err) \<in> set (error (size ts))" by (simp add: error_def)
-
-  from wt lp
-  have [intro?]: "\<And>p. p < size ts \<Longrightarrow> ts ! p \<noteq> Err" 
-    by (unfold wt_err_step_def wt_step_def) simp
-
-  show app: "app p (map ok_val ts ! p)"
-  proof (rule ccontr)
-    from wt lp obtain s where
-      OKp:  "ts ! p = OK s" and
-      less: "\<forall>(q,t) \<in> set (err_step (size ts) app step p (ts!p)). t <=_(Err.le r) ts!q"
-      by (unfold wt_err_step_def wt_step_def stable_def) 
-         (auto iff: not_Err_eq)
-    assume "\<not> app p (map ok_val ts ! p)"
-    with OKp lp have "\<not> app p s" by simp
-    with OKp have "err_step (size ts) app step p (ts!p) = error (size ts)" 
-      by (simp add: err_step_def)    
-    with err ts obtain q where 
-      "(q,Err) \<in> set (err_step (size ts) app step p (ts!p))" and
-      q: "q < size ts" by auto    
-    with less have "ts!q = Err" by auto
-    moreover from q have "ts!q \<noteq> Err" ..
-    ultimately show False by blast
-  qed
-  
-  show "\<forall>(q,t)\<in>set(step p (map ok_val ts ! p)). t <=_r map ok_val ts ! q" 
-  proof clarify
-    fix q t assume q: "(q,t) \<in> set (step p (map ok_val ts ! p))"
-
-    from wt lp q
-    obtain s where
-      OKp:  "ts ! p = OK s" and
-      less: "\<forall>(q,t) \<in> set (err_step (size ts) app step p (ts!p)). t <=_(Err.le r) ts!q"
-      by (unfold wt_err_step_def wt_step_def stable_def) 
-         (auto iff: not_Err_eq)
-
-    from b lp app q have lq: "q < size ts" by (rule bounded_err_stepD)
-    hence "ts!q \<noteq> Err" ..
-    then obtain s' where OKq: "ts ! q = OK s'" by (auto iff: not_Err_eq)
-
-    from lp lq OKp OKq app less q
-    show "t <=_r map ok_val ts ! q"
-      by (auto simp add: err_step_def map_snd_def) 
-  qed
-qed
-
-
-lemma wt_app_eff_imp_wt_err:
-  assumes app_eff: "wt_app_eff r app step ts"
-  assumes bounded: "bounded (err_step (size ts) app step) (size ts)"
-  shows "wt_err_step r (err_step (size ts) app step) (map OK ts)"
-proof (unfold wt_err_step_def wt_step_def, intro strip, rule conjI)
-  fix p assume "p < size (map OK ts)" 
-  hence p: "p < size ts" by simp
-  thus "map OK ts ! p \<noteq> Err" by simp
-  { fix q t
-    assume q: "(q,t) \<in> set (err_step (size ts) app step p (map OK ts ! p))" 
-    with p app_eff obtain 
-      "app p (ts ! p)" "\<forall>(q,t) \<in> set (step p (ts!p)). t <=_r ts!q"
-      by (unfold wt_app_eff_def) blast
-    moreover
-    from q p bounded have "q < size ts"
-      by - (rule boundedD)
-    hence "map OK ts ! q = OK (ts!q)" by simp
-    moreover
-    have "p < size ts" by (rule p)
-    moreover note q
-    ultimately     
-    have "t <=_(Err.le r) map OK ts ! q" 
-      by (auto simp add: err_step_def map_snd_def)
-  }
-  thus "stable (Err.le r) (err_step (size ts) app step) (map OK ts) p"
-    by (unfold stable_def) blast
-qed
-
-end
-
--- a/src/HOL/MicroJava/Comp/CorrCompTp.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/Comp/CorrCompTp.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -1122,6 +1122,8 @@
 apply simp+
 done
 
+declare not_Err_eq [iff del]
+
 lemma bc_mt_corresp_Load: "\<lbrakk> i < length LT; LT ! i \<noteq> Err; mxr = length LT \<rbrakk>
   \<Longrightarrow> bc_mt_corresp [Load i] 
          (\<lambda>(ST, LT). pushST [ok_val (LT ! i)] (ST, LT)) (ST, LT) cG rT mxr (Suc 0)"
@@ -1138,7 +1140,7 @@
   apply (frule listE_nth_in) apply assumption
 apply (subgoal_tac "LT ! i \<in> {x. \<exists>y\<in>types cG. x = OK y}")
 apply (drule CollectD) apply (erule bexE)
-apply (simp (no_asm_simp) )
+apply (simp (no_asm_simp))
 apply blast
 apply blast
 done
--- a/src/HOL/MicroJava/Comp/LemmasComp.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/Comp/LemmasComp.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -110,7 +110,7 @@
 by (case_tac "class G C", auto simp: is_class_def dest: comp_class_imp)
 
 lemma comp_subcls1: "subcls1 (comp G) = subcls1 G"
-by (auto simp add: subcls1_def2 comp_classname comp_is_class expand_fun_eq)
+by (auto simp add: subcls1_def2 comp_classname comp_is_class)
 
 lemma comp_widen: "widen (comp G) = widen G"
   apply (simp add: expand_fun_eq)
@@ -183,17 +183,17 @@
 done
 
 
-lemma comp_class_rec: " wfP ((subcls1 G)^--1) \<Longrightarrow> 
+lemma comp_class_rec: " wf ((subcls1 G)^-1) \<Longrightarrow> 
 class_rec (comp G) C t f = 
   class_rec G C t (\<lambda> C' fs' ms' r'. f C' fs' (map (compMethod G C') ms') r')"
-apply (rule_tac a = C in  wfP_induct) apply assumption
-apply (subgoal_tac "wfP ((subcls1 (comp G))\<inverse>\<inverse>)")
+apply (rule_tac a = C in  wf_induct) apply assumption
+apply (subgoal_tac "wf ((subcls1 (comp G))^-1)")
 apply (subgoal_tac "(class G x = None) \<or> (\<exists> D fs ms. (class G x = Some (D, fs, ms)))")
 apply (erule disjE)
 
   (* case class G x = None *)
 apply (simp (no_asm_simp) add: class_rec_def comp_subcls1
-  wfrec [to_pred, where r="(subcls1 G)^--1", simplified])
+  wfrec [where r="(subcls1 G)^-1", simplified])
 apply (simp add: comp_class_None)
 
   (* case \<exists> D fs ms. (class G x = Some (D, fs, ms)) *)
@@ -218,11 +218,11 @@
 apply (simp add: comp_subcls1)
 done
 
-lemma comp_fields: "wfP ((subcls1 G)^--1) \<Longrightarrow> 
+lemma comp_fields: "wf ((subcls1 G)^-1) \<Longrightarrow> 
   fields (comp G,C) = fields (G,C)" 
 by (simp add: fields_def comp_class_rec)
 
-lemma comp_field: "wfP ((subcls1 G)^--1) \<Longrightarrow> 
+lemma comp_field: "wf ((subcls1 G)^-1) \<Longrightarrow> 
   field (comp G,C) = field (G,C)" 
 by (simp add: TypeRel.field_def comp_fields)
 
@@ -234,7 +234,7 @@
   \<Longrightarrow> ((class G C) \<noteq> None) \<longrightarrow> 
   R (class_rec G C t1 f1) (class_rec G C t2 f2)"
 apply (frule wf_subcls1) (* establish wf ((subcls1 G)^-1) *)
-apply (rule_tac a = C in  wfP_induct) apply assumption
+apply (rule_tac a = C in  wf_induct) apply assumption
 apply (intro strip)
 apply (subgoal_tac "(\<exists>D rT mb. class G x = Some (D, rT, mb))")
   apply (erule exE)+
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/MicroJava/DFA/Abstract_BV.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -0,0 +1,14 @@
+(*  Title:      HOL/MicroJava/BV/Semilat.thy
+    Author:     Gerwin Klein
+    Copyright   2003 TUM
+*)
+
+header {* Abstract Bytecode Verifier *}
+
+(*<*)
+theory Abstract_BV
+imports Typing_Framework_err Kildall LBVCorrect LBVComplete
+begin
+
+end
+(*>*)
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/MicroJava/DFA/Err.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -0,0 +1,350 @@
+(*  Title:      HOL/MicroJava/BV/Err.thy
+    Author:     Tobias Nipkow
+    Copyright   2000 TUM
+*)
+
+header {* \isaheader{The Error Type} *}
+
+theory Err
+imports Semilat
+begin
+
+datatype 'a err = Err | OK 'a
+
+types 'a ebinop = "'a \<Rightarrow> 'a \<Rightarrow> 'a err"
+      'a esl =    "'a set * 'a ord * 'a ebinop"
+
+consts
+  ok_val :: "'a err \<Rightarrow> 'a"
+primrec
+  "ok_val (OK x) = x"
+
+constdefs
+ lift :: "('a \<Rightarrow> 'b err) \<Rightarrow> ('a err \<Rightarrow> 'b err)"
+"lift f e == case e of Err \<Rightarrow> Err | OK x \<Rightarrow> f x"
+
+ lift2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c err) \<Rightarrow> 'a err \<Rightarrow> 'b err \<Rightarrow> 'c err"
+"lift2 f e1 e2 ==
+ case e1 of Err  \<Rightarrow> Err
+          | OK x \<Rightarrow> (case e2 of Err \<Rightarrow> Err | OK y \<Rightarrow> f x y)"
+
+ le :: "'a ord \<Rightarrow> 'a err ord"
+"le r e1 e2 ==
+        case e2 of Err \<Rightarrow> True |
+                   OK y \<Rightarrow> (case e1 of Err \<Rightarrow> False | OK x \<Rightarrow> x <=_r y)"
+
+ sup :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> ('a err \<Rightarrow> 'b err \<Rightarrow> 'c err)"
+"sup f == lift2(%x y. OK(x +_f y))"
+
+ err :: "'a set \<Rightarrow> 'a err set"
+"err A == insert Err {x . ? y:A. x = OK y}"
+
+ esl :: "'a sl \<Rightarrow> 'a esl"
+"esl == %(A,r,f). (A,r, %x y. OK(f x y))"
+
+ sl :: "'a esl \<Rightarrow> 'a err sl"
+"sl == %(A,r,f). (err A, le r, lift2 f)"
+
+syntax
+ err_semilat :: "'a esl \<Rightarrow> bool"
+translations
+"err_semilat L" == "semilat(Err.sl L)"
+
+
+consts
+  strict  :: "('a \<Rightarrow> 'b err) \<Rightarrow> ('a err \<Rightarrow> 'b err)"
+primrec
+  "strict f Err    = Err"
+  "strict f (OK x) = f x"
+
+lemma strict_Some [simp]: 
+  "(strict f x = OK y) = (\<exists> z. x = OK z \<and> f z = OK y)"
+  by (cases x, auto)
+
+lemma not_Err_eq:
+  "(x \<noteq> Err) = (\<exists>a. x = OK a)" 
+  by (cases x) auto
+
+lemma not_OK_eq:
+  "(\<forall>y. x \<noteq> OK y) = (x = Err)"
+  by (cases x) auto  
+
+lemma unfold_lesub_err:
+  "e1 <=_(le r) e2 == le r e1 e2"
+  by (simp add: lesub_def)
+
+lemma le_err_refl:
+  "!x. x <=_r x \<Longrightarrow> e <=_(Err.le r) e"
+apply (unfold lesub_def Err.le_def)
+apply (simp split: err.split)
+done 
+
+lemma le_err_trans [rule_format]:
+  "order r \<Longrightarrow> e1 <=_(le r) e2 \<longrightarrow> e2 <=_(le r) e3 \<longrightarrow> e1 <=_(le r) e3"
+apply (unfold unfold_lesub_err le_def)
+apply (simp split: err.split)
+apply (blast intro: order_trans)
+done
+
+lemma le_err_antisym [rule_format]:
+  "order r \<Longrightarrow> e1 <=_(le r) e2 \<longrightarrow> e2 <=_(le r) e1 \<longrightarrow> e1=e2"
+apply (unfold unfold_lesub_err le_def)
+apply (simp split: err.split)
+apply (blast intro: order_antisym)
+done 
+
+lemma OK_le_err_OK:
+  "(OK x <=_(le r) OK y) = (x <=_r y)"
+  by (simp add: unfold_lesub_err le_def)
+
+lemma order_le_err [iff]:
+  "order(le r) = order r"
+apply (rule iffI)
+ apply (subst Semilat.order_def)
+ apply (blast dest: order_antisym OK_le_err_OK [THEN iffD2]
+              intro: order_trans OK_le_err_OK [THEN iffD1])
+apply (subst Semilat.order_def)
+apply (blast intro: le_err_refl le_err_trans le_err_antisym
+             dest: order_refl)
+done 
+
+lemma le_Err [iff]:  "e <=_(le r) Err"
+  by (simp add: unfold_lesub_err le_def)
+
+lemma Err_le_conv [iff]:
+ "Err <=_(le r) e  = (e = Err)"
+  by (simp add: unfold_lesub_err le_def  split: err.split)
+
+lemma le_OK_conv [iff]:
+  "e <=_(le r) OK x  =  (? y. e = OK y & y <=_r x)"
+  by (simp add: unfold_lesub_err le_def split: err.split)
+
+lemma OK_le_conv:
+ "OK x <=_(le r) e  =  (e = Err | (? y. e = OK y & x <=_r y))"
+  by (simp add: unfold_lesub_err le_def split: err.split)
+
+lemma top_Err [iff]: "top (le r) Err";
+  by (simp add: top_def)
+
+lemma OK_less_conv [rule_format, iff]:
+  "OK x <_(le r) e = (e=Err | (? y. e = OK y & x <_r y))"
+  by (simp add: lesssub_def lesub_def le_def split: err.split)
+
+lemma not_Err_less [rule_format, iff]:
+  "~(Err <_(le r) x)"
+  by (simp add: lesssub_def lesub_def le_def split: err.split)
+
+lemma semilat_errI [intro]:
+  assumes semilat: "semilat (A, r, f)"
+  shows "semilat(err A, Err.le r, lift2(%x y. OK(f x y)))"
+  apply(insert semilat)
+  apply (unfold semilat_Def closed_def plussub_def lesub_def 
+    lift2_def Err.le_def err_def)
+  apply (simp split: err.split)
+  done
+
+lemma err_semilat_eslI_aux:
+  assumes semilat: "semilat (A, r, f)"
+  shows "err_semilat(esl(A,r,f))"
+  apply (unfold sl_def esl_def)
+  apply (simp add: semilat_errI[OF semilat])
+  done
+
+lemma err_semilat_eslI [intro, simp]:
+ "\<And>L. semilat L \<Longrightarrow> err_semilat(esl L)"
+by(simp add: err_semilat_eslI_aux split_tupled_all)
+
+lemma acc_err [simp, intro!]:  "acc r \<Longrightarrow> acc(le r)"
+apply (unfold acc_def lesub_def le_def lesssub_def)
+apply (simp add: wf_eq_minimal split: err.split)
+apply clarify
+apply (case_tac "Err : Q")
+ apply blast
+apply (erule_tac x = "{a . OK a : Q}" in allE)
+apply (case_tac "x")
+ apply fast
+apply blast
+done 
+
+lemma Err_in_err [iff]: "Err : err A"
+  by (simp add: err_def)
+
+lemma Ok_in_err [iff]: "(OK x : err A) = (x:A)"
+  by (auto simp add: err_def)
+
+section {* lift *}
+
+lemma lift_in_errI:
+  "\<lbrakk> e : err S; !x:S. e = OK x \<longrightarrow> f x : err S \<rbrakk> \<Longrightarrow> lift f e : err S"
+apply (unfold lift_def)
+apply (simp split: err.split)
+apply blast
+done 
+
+lemma Err_lift2 [simp]: 
+  "Err +_(lift2 f) x = Err"
+  by (simp add: lift2_def plussub_def)
+
+lemma lift2_Err [simp]: 
+  "x +_(lift2 f) Err = Err"
+  by (simp add: lift2_def plussub_def split: err.split)
+
+lemma OK_lift2_OK [simp]:
+  "OK x +_(lift2 f) OK y = x +_f y"
+  by (simp add: lift2_def plussub_def split: err.split)
+
+
+section {* sup *}
+
+lemma Err_sup_Err [simp]:
+  "Err +_(Err.sup f) x = Err"
+  by (simp add: plussub_def Err.sup_def Err.lift2_def)
+
+lemma Err_sup_Err2 [simp]:
+  "x +_(Err.sup f) Err = Err"
+  by (simp add: plussub_def Err.sup_def Err.lift2_def split: err.split)
+
+lemma Err_sup_OK [simp]:
+  "OK x +_(Err.sup f) OK y = OK(x +_f y)"
+  by (simp add: plussub_def Err.sup_def Err.lift2_def)
+
+lemma Err_sup_eq_OK_conv [iff]:
+  "(Err.sup f ex ey = OK z) = (? x y. ex = OK x & ey = OK y & f x y = z)"
+apply (unfold Err.sup_def lift2_def plussub_def)
+apply (rule iffI)
+ apply (simp split: err.split_asm)
+apply clarify
+apply simp
+done
+
+lemma Err_sup_eq_Err [iff]:
+  "(Err.sup f ex ey = Err) = (ex=Err | ey=Err)"
+apply (unfold Err.sup_def lift2_def plussub_def)
+apply (simp split: err.split)
+done 
+
+section {* semilat (err A) (le r) f *}
+
+lemma semilat_le_err_Err_plus [simp]:
+  "\<lbrakk> x: err A; semilat(err A, le r, f) \<rbrakk> \<Longrightarrow> Err +_f x = Err"
+  by (blast intro: Semilat.le_iff_plus_unchanged [OF Semilat.intro, THEN iffD1]
+                   Semilat.le_iff_plus_unchanged2 [OF Semilat.intro, THEN iffD1])
+
+lemma semilat_le_err_plus_Err [simp]:
+  "\<lbrakk> x: err A; semilat(err A, le r, f) \<rbrakk> \<Longrightarrow> x +_f Err = Err"
+  by (blast intro: Semilat.le_iff_plus_unchanged [OF Semilat.intro, THEN iffD1]
+                   Semilat.le_iff_plus_unchanged2 [OF Semilat.intro, THEN iffD1])
+
+lemma semilat_le_err_OK1:
+  "\<lbrakk> x:A; y:A; semilat(err A, le r, f); OK x +_f OK y = OK z \<rbrakk> 
+  \<Longrightarrow> x <=_r z";
+apply (rule OK_le_err_OK [THEN iffD1])
+apply (erule subst)
+apply (simp add: Semilat.ub1 [OF Semilat.intro])
+done
+
+lemma semilat_le_err_OK2:
+  "\<lbrakk> x:A; y:A; semilat(err A, le r, f); OK x +_f OK y = OK z \<rbrakk> 
+  \<Longrightarrow> y <=_r z"
+apply (rule OK_le_err_OK [THEN iffD1])
+apply (erule subst)
+apply (simp add: Semilat.ub2 [OF Semilat.intro])
+done
+
+lemma eq_order_le:
+  "\<lbrakk> x=y; order r \<rbrakk> \<Longrightarrow> x <=_r y"
+apply (unfold Semilat.order_def)
+apply blast
+done
+
+lemma OK_plus_OK_eq_Err_conv [simp]:
+  assumes "x:A" and "y:A" and "semilat(err A, le r, fe)"
+  shows "((OK x) +_fe (OK y) = Err) = (~(? z:A. x <=_r z & y <=_r z))"
+proof -
+  have plus_le_conv3: "\<And>A x y z f r. 
+    \<lbrakk> semilat (A,r,f); x +_f y <=_r z; x:A; y:A; z:A \<rbrakk> 
+    \<Longrightarrow> x <=_r z \<and> y <=_r z"
+    by (rule Semilat.plus_le_conv [OF Semilat.intro, THEN iffD1])
+  from prems show ?thesis
+  apply (rule_tac iffI)
+   apply clarify
+   apply (drule OK_le_err_OK [THEN iffD2])
+   apply (drule OK_le_err_OK [THEN iffD2])
+   apply (drule Semilat.lub [OF Semilat.intro, of _ _ _ "OK x" _ "OK y"])
+        apply assumption
+       apply assumption
+      apply simp
+     apply simp
+    apply simp
+   apply simp
+  apply (case_tac "(OK x) +_fe (OK y)")
+   apply assumption
+  apply (rename_tac z)
+  apply (subgoal_tac "OK z: err A")
+  apply (drule eq_order_le)
+    apply (erule Semilat.orderI [OF Semilat.intro])
+   apply (blast dest: plus_le_conv3) 
+  apply (erule subst)
+  apply (blast intro: Semilat.closedI [OF Semilat.intro] closedD)
+  done 
+qed
+
+section {* semilat (err(Union AS)) *}
+
+(* FIXME? *)
+lemma all_bex_swap_lemma [iff]:
+  "(!x. (? y:A. x = f y) \<longrightarrow> P x) = (!y:A. P(f y))"
+  by blast
+
+lemma closed_err_Union_lift2I: 
+  "\<lbrakk> !A:AS. closed (err A) (lift2 f); AS ~= {}; 
+      !A:AS.!B:AS. A~=B \<longrightarrow> (!a:A.!b:B. a +_f b = Err) \<rbrakk> 
+  \<Longrightarrow> closed (err(Union AS)) (lift2 f)"
+apply (unfold closed_def err_def)
+apply simp
+apply clarify
+apply simp
+apply fast
+done 
+
+text {* 
+  If @{term "AS = {}"} the thm collapses to
+  @{prop "order r & closed {Err} f & Err +_f Err = Err"}
+  which may not hold 
+*}
+lemma err_semilat_UnionI:
+  "\<lbrakk> !A:AS. err_semilat(A, r, f); AS ~= {}; 
+      !A:AS.!B:AS. A~=B \<longrightarrow> (!a:A.!b:B. ~ a <=_r b & a +_f b = Err) \<rbrakk> 
+  \<Longrightarrow> err_semilat(Union AS, r, f)"
+apply (unfold semilat_def sl_def)
+apply (simp add: closed_err_Union_lift2I)
+apply (rule conjI)
+ apply blast
+apply (simp add: err_def)
+apply (rule conjI)
+ apply clarify
+ apply (rename_tac A a u B b)
+ apply (case_tac "A = B")
+  apply simp
+ apply simp
+apply (rule conjI)
+ apply clarify
+ apply (rename_tac A a u B b)
+ apply (case_tac "A = B")
+  apply simp
+ apply simp
+apply clarify
+apply (rename_tac A ya yb B yd z C c a b)
+apply (case_tac "A = B")
+ apply (case_tac "A = C")
+  apply simp
+ apply (rotate_tac -1)
+ apply simp
+apply (rotate_tac -1)
+apply (case_tac "B = C")
+ apply simp
+apply (rotate_tac -1)
+apply simp
+done 
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/MicroJava/DFA/Kildall.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -0,0 +1,495 @@
+(*  Title:      HOL/MicroJava/BV/Kildall.thy
+    Author:     Tobias Nipkow, Gerwin Klein
+    Copyright   2000 TUM
+*)
+
+header {* \isaheader{Kildall's Algorithm}\label{sec:Kildall} *}
+
+theory Kildall
+imports SemilatAlg While_Combinator
+begin
+
+
+consts
+ iter :: "'s binop \<Rightarrow> 's step_type \<Rightarrow>
+          's list \<Rightarrow> nat set \<Rightarrow> 's list \<times> nat set"
+ propa :: "'s binop \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's list \<Rightarrow> nat set \<Rightarrow> 's list * nat set"
+
+primrec
+"propa f []      ss w = (ss,w)"
+"propa f (q'#qs) ss w = (let (q,t) = q';
+                             u = t +_f ss!q;
+                             w' = (if u = ss!q then w else insert q w)
+                         in propa f qs (ss[q := u]) w')"
+
+defs iter_def:
+"iter f step ss w ==
+ while (%(ss,w). w \<noteq> {})
+       (%(ss,w). let p = SOME p. p \<in> w
+                 in propa f (step p (ss!p)) ss (w-{p}))
+       (ss,w)"
+
+constdefs
+ unstables :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> nat set"
+"unstables r step ss == {p. p < size ss \<and> \<not>stable r step ss p}"
+
+ kildall :: "'s ord \<Rightarrow> 's binop \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> 's list"
+"kildall r f step ss == fst(iter f step ss (unstables r step ss))"
+
+consts merges :: "'s binop \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's list \<Rightarrow> 's list"
+primrec
+"merges f []      ss = ss"
+"merges f (p'#ps) ss = (let (p,s) = p' in merges f ps (ss[p := s +_f ss!p]))"
+
+
+lemmas [simp] = Let_def Semilat.le_iff_plus_unchanged [OF Semilat.intro, symmetric]
+
+
+lemma (in Semilat) nth_merges:
+ "\<And>ss. \<lbrakk>p < length ss; ss \<in> list n A; \<forall>(p,t)\<in>set ps. p<n \<and> t\<in>A \<rbrakk> \<Longrightarrow>
+  (merges f ps ss)!p = map snd [(p',t') \<leftarrow> ps. p'=p] ++_f ss!p"
+  (is "\<And>ss. \<lbrakk>_; _; ?steptype ps\<rbrakk> \<Longrightarrow> ?P ss ps")
+proof (induct ps)
+  show "\<And>ss. ?P ss []" by simp
+
+  fix ss p' ps'
+  assume ss: "ss \<in> list n A"
+  assume l:  "p < length ss"
+  assume "?steptype (p'#ps')"
+  then obtain a b where
+    p': "p'=(a,b)" and ab: "a<n" "b\<in>A" and ps': "?steptype ps'"
+    by (cases p') auto
+  assume "\<And>ss. p< length ss \<Longrightarrow> ss \<in> list n A \<Longrightarrow> ?steptype ps' \<Longrightarrow> ?P ss ps'"
+  from this [OF _ _ ps'] have IH: "\<And>ss. ss \<in> list n A \<Longrightarrow> p < length ss \<Longrightarrow> ?P ss ps'" .
+
+  from ss ab
+  have "ss[a := b +_f ss!a] \<in> list n A" by (simp add: closedD)
+  moreover
+  from calculation l
+  have "p < length (ss[a := b +_f ss!a])" by simp
+  ultimately
+  have "?P (ss[a := b +_f ss!a]) ps'" by (rule IH)
+  with p' l
+  show "?P ss (p'#ps')" by simp
+qed
+
+
+(** merges **)
+
+lemma length_merges [rule_format, simp]:
+  "\<forall>ss. size(merges f ps ss) = size ss"
+  by (induct_tac ps, auto)
+
+
+lemma (in Semilat) merges_preserves_type_lemma:
+shows "\<forall>xs. xs \<in> list n A \<longrightarrow> (\<forall>(p,x) \<in> set ps. p<n \<and> x\<in>A)
+          \<longrightarrow> merges f ps xs \<in> list n A"
+apply (insert closedI)
+apply (unfold closed_def)
+apply (induct_tac ps)
+ apply simp
+apply clarsimp
+done
+
+lemma (in Semilat) merges_preserves_type [simp]:
+ "\<lbrakk> xs \<in> list n A; \<forall>(p,x) \<in> set ps. p<n \<and> x\<in>A \<rbrakk>
+  \<Longrightarrow> merges f ps xs \<in> list n A"
+by (simp add: merges_preserves_type_lemma)
+
+lemma (in Semilat) merges_incr_lemma:
+ "\<forall>xs. xs \<in> list n A \<longrightarrow> (\<forall>(p,x)\<in>set ps. p<size xs \<and> x \<in> A) \<longrightarrow> xs <=[r] merges f ps xs"
+apply (induct_tac ps)
+ apply simp
+apply simp
+apply clarify
+apply (rule order_trans)
+  apply simp
+ apply (erule list_update_incr)
+  apply simp
+ apply simp
+apply (blast intro!: listE_set intro: closedD listE_length [THEN nth_in])
+done
+
+lemma (in Semilat) merges_incr:
+ "\<lbrakk> xs \<in> list n A; \<forall>(p,x)\<in>set ps. p<size xs \<and> x \<in> A \<rbrakk> 
+  \<Longrightarrow> xs <=[r] merges f ps xs"
+  by (simp add: merges_incr_lemma)
+
+
+lemma (in Semilat) merges_same_conv [rule_format]:
+ "(\<forall>xs. xs \<in> list n A \<longrightarrow> (\<forall>(p,x)\<in>set ps. p<size xs \<and> x\<in>A) \<longrightarrow> 
+     (merges f ps xs = xs) = (\<forall>(p,x)\<in>set ps. x <=_r xs!p))"
+  apply (induct_tac ps)
+   apply simp
+  apply clarsimp
+  apply (rename_tac p x ps xs)
+  apply (rule iffI)
+   apply (rule context_conjI)
+    apply (subgoal_tac "xs[p := x +_f xs!p] <=[r] xs")
+     apply (drule_tac p = p in le_listD)
+      apply simp
+     apply simp
+    apply (erule subst, rule merges_incr)
+       apply (blast intro!: listE_set intro: closedD listE_length [THEN nth_in])
+      apply clarify
+      apply (rule conjI)
+       apply simp
+       apply (blast dest: boundedD)
+      apply blast
+   apply clarify
+   apply (erule allE)
+   apply (erule impE)
+    apply assumption
+   apply (drule bspec)
+    apply assumption
+   apply (simp add: le_iff_plus_unchanged [THEN iffD1] list_update_same_conv [THEN iffD2])
+   apply blast
+  apply clarify 
+  apply (simp add: le_iff_plus_unchanged [THEN iffD1] list_update_same_conv [THEN iffD2])
+  done
+
+
+lemma (in Semilat) list_update_le_listI [rule_format]:
+  "set xs <= A \<longrightarrow> set ys <= A \<longrightarrow> xs <=[r] ys \<longrightarrow> p < size xs \<longrightarrow>  
+   x <=_r ys!p \<longrightarrow> x\<in>A \<longrightarrow> xs[p := x +_f xs!p] <=[r] ys"
+  apply(insert semilat)
+  apply (unfold Listn.le_def lesub_def semilat_def)
+  apply (simp add: list_all2_conv_all_nth nth_list_update)
+  done
+
+lemma (in Semilat) merges_pres_le_ub:
+  assumes "set ts <= A" and "set ss <= A"
+    and "\<forall>(p,t)\<in>set ps. t <=_r ts!p \<and> t \<in> A \<and> p < size ts" and "ss <=[r] ts"
+  shows "merges f ps ss <=[r] ts"
+proof -
+  { fix t ts ps
+    have
+    "\<And>qs. \<lbrakk>set ts <= A; \<forall>(p,t)\<in>set ps. t <=_r ts!p \<and> t \<in> A \<and> p< size ts \<rbrakk> \<Longrightarrow>
+    set qs <= set ps  \<longrightarrow> 
+    (\<forall>ss. set ss <= A \<longrightarrow> ss <=[r] ts \<longrightarrow> merges f qs ss <=[r] ts)"
+    apply (induct_tac qs)
+     apply simp
+    apply (simp (no_asm_simp))
+    apply clarify
+    apply (rotate_tac -2)
+    apply simp
+    apply (erule allE, erule impE, erule_tac [2] mp)
+     apply (drule bspec, assumption)
+     apply (simp add: closedD)
+    apply (drule bspec, assumption)
+    apply (simp add: list_update_le_listI)
+    done 
+  } note this [dest]
+  
+  from prems show ?thesis by blast
+qed
+
+
+(** propa **)
+
+
+lemma decomp_propa:
+  "\<And>ss w. (\<forall>(q,t)\<in>set qs. q < size ss) \<Longrightarrow> 
+   propa f qs ss w = 
+   (merges f qs ss, {q. \<exists>t. (q,t)\<in>set qs \<and> t +_f ss!q \<noteq> ss!q} Un w)"
+  apply (induct qs)
+   apply simp   
+  apply (simp (no_asm))
+  apply clarify  
+  apply simp
+  apply (rule conjI) 
+   apply blast
+  apply (simp add: nth_list_update)
+  apply blast
+  done 
+
+(** iter **)
+
+lemma (in Semilat) stable_pres_lemma:
+shows "\<lbrakk>pres_type step n A; bounded step n; 
+     ss \<in> list n A; p \<in> w; \<forall>q\<in>w. q < n; 
+     \<forall>q. q < n \<longrightarrow> q \<notin> w \<longrightarrow> stable r step ss q; q < n; 
+     \<forall>s'. (q,s') \<in> set (step p (ss ! p)) \<longrightarrow> s' +_f ss ! q = ss ! q; 
+     q \<notin> w \<or> q = p \<rbrakk> 
+  \<Longrightarrow> stable r step (merges f (step p (ss!p)) ss) q"
+  apply (unfold stable_def)
+  apply (subgoal_tac "\<forall>s'. (q,s') \<in> set (step p (ss!p)) \<longrightarrow> s' : A")
+   prefer 2
+   apply clarify
+   apply (erule pres_typeD)
+    prefer 3 apply assumption
+    apply (rule listE_nth_in)
+     apply assumption
+    apply simp
+   apply simp
+  apply simp
+  apply clarify
+  apply (subst nth_merges)
+       apply simp
+       apply (blast dest: boundedD)
+      apply assumption
+     apply clarify
+     apply (rule conjI)
+      apply (blast dest: boundedD)
+     apply (erule pres_typeD)
+       prefer 3 apply assumption
+      apply simp
+     apply simp
+apply(subgoal_tac "q < length ss")
+prefer 2 apply simp
+  apply (frule nth_merges [of q _ _ "step p (ss!p)"]) (* fixme: why does method subst not work?? *)
+apply assumption
+  apply clarify
+  apply (rule conjI)
+   apply (blast dest: boundedD)
+  apply (erule pres_typeD)
+     prefer 3 apply assumption
+    apply simp
+   apply simp
+  apply (drule_tac P = "\<lambda>x. (a, b) \<in> set (step q x)" in subst)
+   apply assumption
+
+ apply (simp add: plusplus_empty)
+ apply (cases "q \<in> w")
+  apply simp
+  apply (rule ub1')
+     apply (rule semilat)
+    apply clarify
+    apply (rule pres_typeD)
+       apply assumption
+      prefer 3 apply assumption
+     apply (blast intro: listE_nth_in dest: boundedD)
+    apply (blast intro: pres_typeD dest: boundedD)
+   apply (blast intro: listE_nth_in dest: boundedD)
+  apply assumption
+
+ apply simp
+ apply (erule allE, erule impE, assumption, erule impE, assumption)
+ apply (rule order_trans)
+   apply simp
+  defer
+ apply (rule pp_ub2)(*
+    apply assumption*)
+   apply simp
+   apply clarify
+   apply simp
+   apply (rule pres_typeD)
+      apply assumption
+     prefer 3 apply assumption
+    apply (blast intro: listE_nth_in dest: boundedD)
+   apply (blast intro: pres_typeD dest: boundedD)
+  apply (blast intro: listE_nth_in dest: boundedD)
+ apply blast
+ done
+
+
+lemma (in Semilat) merges_bounded_lemma:
+ "\<lbrakk> mono r step n A; bounded step n; 
+    \<forall>(p',s') \<in> set (step p (ss!p)). s' \<in> A; ss \<in> list n A; ts \<in> list n A; p < n; 
+    ss <=[r] ts; \<forall>p. p < n \<longrightarrow> stable r step ts p \<rbrakk> 
+  \<Longrightarrow> merges f (step p (ss!p)) ss <=[r] ts" 
+  apply (unfold stable_def)
+  apply (rule merges_pres_le_ub)
+     apply simp
+    apply simp
+   prefer 2 apply assumption
+
+  apply clarsimp
+  apply (drule boundedD, assumption+)
+  apply (erule allE, erule impE, assumption)
+  apply (drule bspec, assumption)
+  apply simp
+
+  apply (drule monoD [of _ _ _ _ p "ss!p"  "ts!p"])
+     apply assumption
+    apply simp
+   apply (simp add: le_listD)
+  
+  apply (drule lesub_step_typeD, assumption) 
+  apply clarify
+  apply (drule bspec, assumption)
+  apply simp
+  apply (blast intro: order_trans)
+  done
+
+lemma termination_lemma:
+  assumes semilat: "semilat (A, r, f)"
+  shows "\<lbrakk> ss \<in> list n A; \<forall>(q,t)\<in>set qs. q<n \<and> t\<in>A; p\<in>w \<rbrakk> \<Longrightarrow> 
+  ss <[r] merges f qs ss \<or> 
+  merges f qs ss = ss \<and> {q. \<exists>t. (q,t)\<in>set qs \<and> t +_f ss!q \<noteq> ss!q} Un (w-{p}) < w" (is "PROP ?P")
+proof -
+  interpret Semilat A r f using assms by (rule Semilat.intro)
+  show "PROP ?P" apply(insert semilat)
+    apply (unfold lesssub_def)
+    apply (simp (no_asm_simp) add: merges_incr)
+    apply (rule impI)
+    apply (rule merges_same_conv [THEN iffD1, elim_format]) 
+    apply assumption+
+    defer
+    apply (rule sym, assumption)
+    defer apply simp
+    apply (subgoal_tac "\<forall>q t. \<not>((q, t) \<in> set qs \<and> t +_f ss ! q \<noteq> ss ! q)")
+    apply (blast intro!: psubsetI elim: equalityE)
+    apply clarsimp
+    apply (drule bspec, assumption) 
+    apply (drule bspec, assumption)
+    apply clarsimp
+    done
+qed
+
+lemma iter_properties[rule_format]:
+  assumes semilat: "semilat (A, r, f)"
+  shows "\<lbrakk> acc r ; pres_type step n A; mono r step n A;
+     bounded step n; \<forall>p\<in>w0. p < n; ss0 \<in> list n A;
+     \<forall>p<n. p \<notin> w0 \<longrightarrow> stable r step ss0 p \<rbrakk> \<Longrightarrow>
+   iter f step ss0 w0 = (ss',w')
+   \<longrightarrow>
+   ss' \<in> list n A \<and> stables r step ss' \<and> ss0 <=[r] ss' \<and>
+   (\<forall>ts\<in>list n A. ss0 <=[r] ts \<and> stables r step ts \<longrightarrow> ss' <=[r] ts)"
+  (is "PROP ?P")
+proof -
+  interpret Semilat A r f using assms by (rule Semilat.intro)
+  show "PROP ?P" apply(insert semilat)
+apply (unfold iter_def stables_def)
+apply (rule_tac P = "%(ss,w).
+ ss \<in> list n A \<and> (\<forall>p<n. p \<notin> w \<longrightarrow> stable r step ss p) \<and> ss0 <=[r] ss \<and>
+ (\<forall>ts\<in>list n A. ss0 <=[r] ts \<and> stables r step ts \<longrightarrow> ss <=[r] ts) \<and>
+ (\<forall>p\<in>w. p < n)" and
+ r = "{(ss',ss) . ss <[r] ss'} <*lex*> finite_psubset"
+       in while_rule)
+
+-- "Invariant holds initially:"
+apply (simp add:stables_def)
+
+-- "Invariant is preserved:"
+apply(simp add: stables_def split_paired_all)
+apply(rename_tac ss w)
+apply(subgoal_tac "(SOME p. p \<in> w) \<in> w")
+ prefer 2; apply (fast intro: someI)
+apply(subgoal_tac "\<forall>(q,t) \<in> set (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))). q < length ss \<and> t \<in> A")
+ prefer 2
+ apply clarify
+ apply (rule conjI)
+  apply(clarsimp, blast dest!: boundedD)
+ apply (erule pres_typeD)
+  prefer 3
+  apply assumption
+  apply (erule listE_nth_in)
+  apply simp
+ apply simp
+apply (subst decomp_propa)
+ apply fast
+apply simp
+apply (rule conjI)
+ apply (rule merges_preserves_type)
+ apply blast
+ apply clarify
+ apply (rule conjI)
+  apply(clarsimp, fast dest!: boundedD)
+ apply (erule pres_typeD)
+  prefer 3
+  apply assumption
+  apply (erule listE_nth_in)
+  apply blast
+ apply blast
+apply (rule conjI)
+ apply clarify
+ apply (blast intro!: stable_pres_lemma)
+apply (rule conjI)
+ apply (blast intro!: merges_incr intro: le_list_trans)
+apply (rule conjI)
+ apply clarsimp
+ apply (blast intro!: merges_bounded_lemma)
+apply (blast dest!: boundedD)
+
+
+-- "Postcondition holds upon termination:"
+apply(clarsimp simp add: stables_def split_paired_all)
+
+-- "Well-foundedness of the termination relation:"
+apply (rule wf_lex_prod)
+ apply (insert orderI [THEN acc_le_listI])
+ apply (simp add: acc_def lesssub_def wfP_wf_eq [symmetric])
+apply (rule wf_finite_psubset) 
+
+-- "Loop decreases along termination relation:"
+apply(simp add: stables_def split_paired_all)
+apply(rename_tac ss w)
+apply(subgoal_tac "(SOME p. p \<in> w) \<in> w")
+ prefer 2; apply (fast intro: someI)
+apply(subgoal_tac "\<forall>(q,t) \<in> set (step (SOME p. p \<in> w) (ss ! (SOME p. p \<in> w))). q < length ss \<and> t \<in> A")
+ prefer 2
+ apply clarify
+ apply (rule conjI)
+  apply(clarsimp, blast dest!: boundedD)
+ apply (erule pres_typeD)
+  prefer 3
+  apply assumption
+  apply (erule listE_nth_in)
+  apply blast
+ apply blast
+apply (subst decomp_propa)
+ apply blast
+apply clarify
+apply (simp del: listE_length
+    add: lex_prod_def finite_psubset_def 
+         bounded_nat_set_is_finite)
+apply (rule termination_lemma)
+apply assumption+
+defer
+apply assumption
+apply clarsimp
+done
+
+qed
+
+lemma kildall_properties:
+assumes semilat: "semilat (A, r, f)"
+shows "\<lbrakk> acc r; pres_type step n A; mono r step n A;
+     bounded step n; ss0 \<in> list n A \<rbrakk> \<Longrightarrow>
+  kildall r f step ss0 \<in> list n A \<and>
+  stables r step (kildall r f step ss0) \<and>
+  ss0 <=[r] kildall r f step ss0 \<and>
+  (\<forall>ts\<in>list n A. ss0 <=[r] ts \<and> stables r step ts \<longrightarrow>
+                 kildall r f step ss0 <=[r] ts)"
+  (is "PROP ?P")
+proof -
+  interpret Semilat A r f using assms by (rule Semilat.intro)
+  show "PROP ?P"
+apply (unfold kildall_def)
+apply(case_tac "iter f step ss0 (unstables r step ss0)")
+apply(simp)
+apply (rule iter_properties)
+apply (simp_all add: unstables_def stable_def)
+apply (rule semilat)
+done
+qed
+
+lemma is_bcv_kildall:
+assumes semilat: "semilat (A, r, f)"
+shows "\<lbrakk> acc r; top r T; pres_type step n A; bounded step n; mono r step n A \<rbrakk>
+  \<Longrightarrow> is_bcv r T step n A (kildall r f step)"
+  (is "PROP ?P")
+proof -
+  interpret Semilat A r f using assms by (rule Semilat.intro)
+  show "PROP ?P"
+apply(unfold is_bcv_def wt_step_def)
+apply(insert semilat kildall_properties[of A])
+apply(simp add:stables_def)
+apply clarify
+apply(subgoal_tac "kildall r f step ss \<in> list n A")
+ prefer 2 apply (simp(no_asm_simp))
+apply (rule iffI)
+ apply (rule_tac x = "kildall r f step ss" in bexI) 
+  apply (rule conjI)
+   apply (blast)
+  apply (simp  (no_asm_simp))
+ apply(assumption)
+apply clarify
+apply(subgoal_tac "kildall r f step ss!p <=_r ts!p")
+ apply simp
+apply (blast intro!: le_listD less_lengthI)
+done
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/MicroJava/DFA/LBVComplete.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -0,0 +1,378 @@
+(*  Title:      HOL/MicroJava/BV/LBVComplete.thy
+    Author:     Gerwin Klein
+    Copyright   2000 Technische Universitaet Muenchen
+*)
+
+header {* \isaheader{Completeness of the LBV} *}
+
+theory LBVComplete
+imports LBVSpec Typing_Framework
+begin
+
+constdefs
+  is_target :: "['s step_type, 's list, nat] \<Rightarrow> bool" 
+  "is_target step phi pc' \<equiv>
+     \<exists>pc s'. pc' \<noteq> pc+1 \<and> pc < length phi \<and> (pc',s') \<in> set (step pc (phi!pc))"
+
+  make_cert :: "['s step_type, 's list, 's] \<Rightarrow> 's certificate"
+  "make_cert step phi B \<equiv> 
+     map (\<lambda>pc. if is_target step phi pc then phi!pc else B) [0..<length phi] @ [B]"
+
+lemma [code]:
+  "is_target step phi pc' =
+  list_ex (\<lambda>pc. pc' \<noteq> pc+1 \<and> pc' mem (map fst (step pc (phi!pc)))) [0..<length phi]"
+by (force simp: list_ex_iff is_target_def mem_iff)
+
+
+locale lbvc = lbv + 
+  fixes phi :: "'a list" ("\<phi>")
+  fixes c   :: "'a list" 
+  defines cert_def: "c \<equiv> make_cert step \<phi> \<bottom>"
+
+  assumes mono: "mono r step (length \<phi>) A"
+  assumes pres: "pres_type step (length \<phi>) A" 
+  assumes phi:  "\<forall>pc < length \<phi>. \<phi>!pc \<in> A \<and> \<phi>!pc \<noteq> \<top>"
+  assumes bounded: "bounded step (length \<phi>)"
+
+  assumes B_neq_T: "\<bottom> \<noteq> \<top>" 
+
+
+lemma (in lbvc) cert: "cert_ok c (length \<phi>) \<top> \<bottom> A"
+proof (unfold cert_ok_def, intro strip conjI)  
+  note [simp] = make_cert_def cert_def nth_append 
+
+  show "c!length \<phi> = \<bottom>" by simp
+
+  fix pc assume pc: "pc < length \<phi>" 
+  from pc phi B_A show "c!pc \<in> A" by simp
+  from pc phi B_neq_T show "c!pc \<noteq> \<top>" by simp
+qed
+
+lemmas [simp del] = split_paired_Ex
+
+
+lemma (in lbvc) cert_target [intro?]:
+  "\<lbrakk> (pc',s') \<in> set (step pc (\<phi>!pc));
+      pc' \<noteq> pc+1; pc < length \<phi>; pc' < length \<phi> \<rbrakk>
+  \<Longrightarrow> c!pc' = \<phi>!pc'"
+  by (auto simp add: cert_def make_cert_def nth_append is_target_def)
+
+
+lemma (in lbvc) cert_approx [intro?]:
+  "\<lbrakk> pc < length \<phi>; c!pc \<noteq> \<bottom> \<rbrakk>
+  \<Longrightarrow> c!pc = \<phi>!pc"
+  by (auto simp add: cert_def make_cert_def nth_append)
+
+
+lemma (in lbv) le_top [simp, intro]:
+  "x <=_r \<top>"
+  by (insert top) simp
+  
+
+lemma (in lbv) merge_mono:
+  assumes less:  "ss2 <=|r| ss1"
+  assumes x:     "x \<in> A"
+  assumes ss1:   "snd`set ss1 \<subseteq> A"
+  assumes ss2:   "snd`set ss2 \<subseteq> A"
+  shows "merge c pc ss2 x <=_r merge c pc ss1 x" (is "?s2 <=_r ?s1")
+proof-
+  have "?s1 = \<top> \<Longrightarrow> ?thesis" by simp
+  moreover {
+    assume merge: "?s1 \<noteq> T" 
+    from x ss1 have "?s1 =
+      (if \<forall>(pc', s')\<in>set ss1. pc' \<noteq> pc + 1 \<longrightarrow> s' <=_r c!pc'
+      then (map snd [(p', t') \<leftarrow> ss1 . p'=pc+1]) ++_f x
+      else \<top>)" 
+      by (rule merge_def)  
+    with merge obtain
+      app: "\<forall>(pc',s')\<in>set ss1. pc' \<noteq> pc+1 \<longrightarrow> s' <=_r c!pc'" 
+           (is "?app ss1") and
+      sum: "(map snd [(p',t') \<leftarrow> ss1 . p' = pc+1] ++_f x) = ?s1" 
+           (is "?map ss1 ++_f x = _" is "?sum ss1 = _")
+      by (simp split: split_if_asm)
+    from app less 
+    have "?app ss2" by (blast dest: trans_r lesub_step_typeD)
+    moreover {
+      from ss1 have map1: "set (?map ss1) \<subseteq> A" by auto
+      with x have "?sum ss1 \<in> A" by (auto intro!: plusplus_closed semilat)
+      with sum have "?s1 \<in> A" by simp
+      moreover    
+      have mapD: "\<And>x ss. x \<in> set (?map ss) \<Longrightarrow> \<exists>p. (p,x) \<in> set ss \<and> p=pc+1" by auto
+      from x map1 
+      have "\<forall>x \<in> set (?map ss1). x <=_r ?sum ss1"
+        by clarify (rule pp_ub1)
+      with sum have "\<forall>x \<in> set (?map ss1). x <=_r ?s1" by simp
+      with less have "\<forall>x \<in> set (?map ss2). x <=_r ?s1"
+        by (fastsimp dest!: mapD lesub_step_typeD intro: trans_r)
+      moreover 
+      from map1 x have "x <=_r (?sum ss1)" by (rule pp_ub2)
+      with sum have "x <=_r ?s1" by simp
+      moreover 
+      from ss2 have "set (?map ss2) \<subseteq> A" by auto
+      ultimately
+      have "?sum ss2 <=_r ?s1" using x by - (rule pp_lub)
+    }
+    moreover
+    from x ss2 have 
+      "?s2 =
+      (if \<forall>(pc', s')\<in>set ss2. pc' \<noteq> pc + 1 \<longrightarrow> s' <=_r c!pc'
+      then map snd [(p', t') \<leftarrow> ss2 . p' = pc + 1] ++_f x
+      else \<top>)" 
+      by (rule merge_def)
+    ultimately have ?thesis by simp
+  }
+  ultimately show ?thesis by (cases "?s1 = \<top>") auto
+qed
+
+
+lemma (in lbvc) wti_mono:
+  assumes less: "s2 <=_r s1"
+  assumes pc:   "pc < length \<phi>" 
+  assumes s1:   "s1 \<in> A"
+  assumes s2:   "s2 \<in> A"
+  shows "wti c pc s2 <=_r wti c pc s1" (is "?s2' <=_r ?s1'")
+proof -
+  from mono pc s2 less have "step pc s2 <=|r| step pc s1" by (rule monoD)
+  moreover
+  from cert B_A pc have "c!Suc pc \<in> A" by (rule cert_okD3)
+  moreover 
+  from pres s1 pc
+  have "snd`set (step pc s1) \<subseteq> A" by (rule pres_typeD2)
+  moreover
+  from pres s2 pc
+  have "snd`set (step pc s2) \<subseteq> A" by (rule pres_typeD2)
+  ultimately
+  show ?thesis by (simp add: wti merge_mono)
+qed 
+
+lemma (in lbvc) wtc_mono:
+  assumes less: "s2 <=_r s1"
+  assumes pc:   "pc < length \<phi>" 
+  assumes s1:   "s1 \<in> A"
+  assumes s2:   "s2 \<in> A"
+  shows "wtc c pc s2 <=_r wtc c pc s1" (is "?s2' <=_r ?s1'")
+proof (cases "c!pc = \<bottom>")
+  case True 
+  moreover from less pc s1 s2 have "wti c pc s2 <=_r wti c pc s1" by (rule wti_mono)
+  ultimately show ?thesis by (simp add: wtc)
+next
+  case False
+  have "?s1' = \<top> \<Longrightarrow> ?thesis" by simp
+  moreover {
+    assume "?s1' \<noteq> \<top>" 
+    with False have c: "s1 <=_r c!pc" by (simp add: wtc split: split_if_asm)
+    with less have "s2 <=_r c!pc" ..
+    with False c have ?thesis by (simp add: wtc)
+  }
+  ultimately show ?thesis by (cases "?s1' = \<top>") auto
+qed
+
+
+lemma (in lbv) top_le_conv [simp]:
+  "\<top> <=_r x = (x = \<top>)"
+  by (insert semilat) (simp add: top top_le_conv) 
+
+lemma (in lbv) neq_top [simp, elim]:
+  "\<lbrakk> x <=_r y; y \<noteq> \<top> \<rbrakk> \<Longrightarrow> x \<noteq> \<top>"
+  by (cases "x = T") auto
+
+
+lemma (in lbvc) stable_wti:
+  assumes stable:  "stable r step \<phi> pc"
+  assumes pc:      "pc < length \<phi>"
+  shows "wti c pc (\<phi>!pc) \<noteq> \<top>"
+proof -
+  let ?step = "step pc (\<phi>!pc)"
+  from stable 
+  have less: "\<forall>(q,s')\<in>set ?step. s' <=_r \<phi>!q" by (simp add: stable_def)
+  
+  from cert B_A pc 
+  have cert_suc: "c!Suc pc \<in> A" by (rule cert_okD3)
+  moreover  
+  from phi pc have "\<phi>!pc \<in> A" by simp
+  from pres this pc 
+  have stepA: "snd`set ?step \<subseteq> A" by (rule pres_typeD2) 
+  ultimately
+  have "merge c pc ?step (c!Suc pc) =
+    (if \<forall>(pc',s')\<in>set ?step. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'
+    then map snd [(p',t') \<leftarrow> ?step.p'=pc+1] ++_f c!Suc pc
+    else \<top>)" unfolding mrg_def by (rule lbv.merge_def [OF lbvc.axioms(1), OF lbvc_axioms])
+  moreover {
+    fix pc' s' assume s': "(pc', s') \<in> set ?step" and suc_pc: "pc' \<noteq> pc+1"
+    with less have "s' <=_r \<phi>!pc'" by auto
+    also 
+    from bounded pc s' have "pc' < length \<phi>" by (rule boundedD)
+    with s' suc_pc pc have "c!pc' = \<phi>!pc'" ..
+    hence "\<phi>!pc' = c!pc'" ..
+    finally have "s' <=_r c!pc'" .
+  } hence "\<forall>(pc',s')\<in>set ?step. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'" by auto
+  moreover
+  from pc have "Suc pc = length \<phi> \<or> Suc pc < length \<phi>" by auto
+  hence "map snd [(p',t') \<leftarrow> ?step.p'=pc+1] ++_f c!Suc pc \<noteq> \<top>" 
+         (is "?map ++_f _ \<noteq> _")
+  proof (rule disjE)
+    assume pc': "Suc pc = length \<phi>"
+    with cert have "c!Suc pc = \<bottom>" by (simp add: cert_okD2)
+    moreover 
+    from pc' bounded pc 
+    have "\<forall>(p',t')\<in>set ?step. p'\<noteq>pc+1" by clarify (drule boundedD, auto)
+    hence "[(p',t') \<leftarrow> ?step.p'=pc+1] = []" by (blast intro: filter_False) 
+    hence "?map = []" by simp
+    ultimately show ?thesis by (simp add: B_neq_T)  
+  next
+    assume pc': "Suc pc < length \<phi>"
+    from pc' phi have "\<phi>!Suc pc \<in> A" by simp
+    moreover note cert_suc
+    moreover from stepA 
+    have "set ?map \<subseteq> A" by auto
+    moreover
+    have "\<And>s. s \<in> set ?map \<Longrightarrow> \<exists>t. (Suc pc, t) \<in> set ?step" by auto
+    with less have "\<forall>s' \<in> set ?map. s' <=_r \<phi>!Suc pc" by auto
+    moreover
+    from pc' have "c!Suc pc <=_r \<phi>!Suc pc" 
+      by (cases "c!Suc pc = \<bottom>") (auto dest: cert_approx)
+    ultimately
+    have "?map ++_f c!Suc pc <=_r \<phi>!Suc pc" by (rule pp_lub)
+    moreover
+    from pc' phi have "\<phi>!Suc pc \<noteq> \<top>" by simp
+    ultimately
+    show ?thesis by auto
+  qed
+  ultimately
+  have "merge c pc ?step (c!Suc pc) \<noteq> \<top>" by simp
+  thus ?thesis by (simp add: wti)  
+qed
+
+lemma (in lbvc) wti_less:
+  assumes stable:  "stable r step \<phi> pc"
+  assumes suc_pc:   "Suc pc < length \<phi>"
+  shows "wti c pc (\<phi>!pc) <=_r \<phi>!Suc pc" (is "?wti <=_r _")
+proof -
+  let ?step = "step pc (\<phi>!pc)"
+
+  from stable 
+  have less: "\<forall>(q,s')\<in>set ?step. s' <=_r \<phi>!q" by (simp add: stable_def)
+   
+  from suc_pc have pc: "pc < length \<phi>" by simp
+  with cert B_A have cert_suc: "c!Suc pc \<in> A" by (rule cert_okD3)
+  moreover  
+  from phi pc have "\<phi>!pc \<in> A" by simp
+  with pres pc have stepA: "snd`set ?step \<subseteq> A" by - (rule pres_typeD2)
+  moreover
+  from stable pc have "?wti \<noteq> \<top>" by (rule stable_wti)
+  hence "merge c pc ?step (c!Suc pc) \<noteq> \<top>" by (simp add: wti)
+  ultimately
+  have "merge c pc ?step (c!Suc pc) =
+    map snd [(p',t')\<leftarrow> ?step.p'=pc+1] ++_f c!Suc pc" by (rule merge_not_top_s) 
+  hence "?wti = \<dots>" (is "_ = (?map ++_f _)" is "_ = ?sum") by (simp add: wti)
+  also {
+    from suc_pc phi have "\<phi>!Suc pc \<in> A" by simp
+    moreover note cert_suc
+    moreover from stepA have "set ?map \<subseteq> A" by auto
+    moreover
+    have "\<And>s. s \<in> set ?map \<Longrightarrow> \<exists>t. (Suc pc, t) \<in> set ?step" by auto
+    with less have "\<forall>s' \<in> set ?map. s' <=_r \<phi>!Suc pc" by auto
+    moreover
+    from suc_pc have "c!Suc pc <=_r \<phi>!Suc pc"
+      by (cases "c!Suc pc = \<bottom>") (auto dest: cert_approx)
+    ultimately
+    have "?sum <=_r \<phi>!Suc pc" by (rule pp_lub)
+  }
+  finally show ?thesis .
+qed
+
+lemma (in lbvc) stable_wtc:
+  assumes stable:  "stable r step phi pc"
+  assumes pc:      "pc < length \<phi>"
+  shows "wtc c pc (\<phi>!pc) \<noteq> \<top>"
+proof -
+  from stable pc have wti: "wti c pc (\<phi>!pc) \<noteq> \<top>" by (rule stable_wti)
+  show ?thesis
+  proof (cases "c!pc = \<bottom>")
+    case True with wti show ?thesis by (simp add: wtc)
+  next
+    case False
+    with pc have "c!pc = \<phi>!pc" ..    
+    with False wti show ?thesis by (simp add: wtc)
+  qed
+qed
+
+lemma (in lbvc) wtc_less:
+  assumes stable: "stable r step \<phi> pc"
+  assumes suc_pc: "Suc pc < length \<phi>"
+  shows "wtc c pc (\<phi>!pc) <=_r \<phi>!Suc pc" (is "?wtc <=_r _")
+proof (cases "c!pc = \<bottom>")
+  case True
+  moreover from stable suc_pc have "wti c pc (\<phi>!pc) <=_r \<phi>!Suc pc"
+    by (rule wti_less)
+  ultimately show ?thesis by (simp add: wtc)
+next
+  case False
+  from suc_pc have pc: "pc < length \<phi>" by simp
+  with stable have "?wtc \<noteq> \<top>" by (rule stable_wtc)
+  with False have "?wtc = wti c pc (c!pc)" 
+    by (unfold wtc) (simp split: split_if_asm)
+  also from pc False have "c!pc = \<phi>!pc" .. 
+  finally have "?wtc = wti c pc (\<phi>!pc)" .
+  also from stable suc_pc have "wti c pc (\<phi>!pc) <=_r \<phi>!Suc pc" by (rule wti_less)
+  finally show ?thesis .
+qed
+
+
+lemma (in lbvc) wt_step_wtl_lemma:
+  assumes wt_step: "wt_step r \<top> step \<phi>"
+  shows "\<And>pc s. pc+length ls = length \<phi> \<Longrightarrow> s <=_r \<phi>!pc \<Longrightarrow> s \<in> A \<Longrightarrow> s\<noteq>\<top> \<Longrightarrow>
+                wtl ls c pc s \<noteq> \<top>"
+  (is "\<And>pc s. _ \<Longrightarrow> _ \<Longrightarrow> _ \<Longrightarrow> _ \<Longrightarrow> ?wtl ls pc s \<noteq> _")
+proof (induct ls)
+  fix pc s assume "s\<noteq>\<top>" thus "?wtl [] pc s \<noteq> \<top>" by simp
+next
+  fix pc s i ls
+  assume "\<And>pc s. pc+length ls=length \<phi> \<Longrightarrow> s <=_r \<phi>!pc \<Longrightarrow> s \<in> A \<Longrightarrow> s\<noteq>\<top> \<Longrightarrow> 
+                  ?wtl ls pc s \<noteq> \<top>"
+  moreover
+  assume pc_l: "pc + length (i#ls) = length \<phi>"
+  hence suc_pc_l: "Suc pc + length ls = length \<phi>" by simp
+  ultimately
+  have IH: "\<And>s. s <=_r \<phi>!Suc pc \<Longrightarrow> s \<in> A \<Longrightarrow> s \<noteq> \<top> \<Longrightarrow> ?wtl ls (Suc pc) s \<noteq> \<top>" .
+
+  from pc_l obtain pc: "pc < length \<phi>" by simp
+  with wt_step have stable: "stable r step \<phi> pc" by (simp add: wt_step_def)
+  from this pc have wt_phi: "wtc c pc (\<phi>!pc) \<noteq> \<top>" by (rule stable_wtc)
+  assume s_phi: "s <=_r \<phi>!pc"
+  from phi pc have phi_pc: "\<phi>!pc \<in> A" by simp
+  assume s: "s \<in> A"
+  with s_phi pc phi_pc have wt_s_phi: "wtc c pc s <=_r wtc c pc (\<phi>!pc)" by (rule wtc_mono)
+  with wt_phi have wt_s: "wtc c pc s \<noteq> \<top>" by simp
+  moreover
+  assume s': "s \<noteq> \<top>" 
+  ultimately
+  have "ls = [] \<Longrightarrow> ?wtl (i#ls) pc s \<noteq> \<top>" by simp
+  moreover {
+    assume "ls \<noteq> []" 
+    with pc_l have suc_pc: "Suc pc < length \<phi>" by (auto simp add: neq_Nil_conv)
+    with stable have "wtc c pc (phi!pc) <=_r \<phi>!Suc pc" by (rule wtc_less)
+    with wt_s_phi have "wtc c pc s <=_r \<phi>!Suc pc" by (rule trans_r)      
+    moreover
+    from cert suc_pc have "c!pc \<in> A" "c!(pc+1) \<in> A" 
+      by (auto simp add: cert_ok_def)
+    from pres this s pc have "wtc c pc s \<in> A" by (rule wtc_pres)
+    ultimately
+    have "?wtl ls (Suc pc) (wtc c pc s) \<noteq> \<top>" using IH wt_s by blast
+    with s' wt_s have "?wtl (i#ls) pc s \<noteq> \<top>" by simp
+  }
+  ultimately show "?wtl (i#ls) pc s \<noteq> \<top>" by (cases ls) blast+
+qed
+
+  
+theorem (in lbvc) wtl_complete:
+  assumes wt: "wt_step r \<top> step \<phi>"
+    and s: "s <=_r \<phi>!0" "s \<in> A" "s \<noteq> \<top>"
+    and len: "length ins = length phi"
+  shows "wtl ins c 0 s \<noteq> \<top>"
+proof -
+  from len have "0+length ins = length phi" by simp
+  from wt this s show ?thesis by (rule wt_step_wtl_lemma)
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/MicroJava/DFA/LBVCorrect.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -0,0 +1,221 @@
+(*  Author:     Gerwin Klein
+    Copyright   1999 Technische Universitaet Muenchen
+*)
+
+header {* \isaheader{Correctness of the LBV} *}
+
+theory LBVCorrect
+imports LBVSpec Typing_Framework
+begin
+
+locale lbvs = lbv +
+  fixes s0  :: 'a ("s\<^sub>0")
+  fixes c   :: "'a list"
+  fixes ins :: "'b list"
+  fixes phi :: "'a list" ("\<phi>")
+  defines phi_def:
+  "\<phi> \<equiv> map (\<lambda>pc. if c!pc = \<bottom> then wtl (take pc ins) c 0 s0 else c!pc) 
+       [0..<length ins]"
+
+  assumes bounded: "bounded step (length ins)"
+  assumes cert: "cert_ok c (length ins) \<top> \<bottom> A"
+  assumes pres: "pres_type step (length ins) A"
+
+
+lemma (in lbvs) phi_None [intro?]:
+  "\<lbrakk> pc < length ins; c!pc = \<bottom> \<rbrakk> \<Longrightarrow> \<phi> ! pc = wtl (take pc ins) c 0 s0"
+  by (simp add: phi_def)
+
+lemma (in lbvs) phi_Some [intro?]:
+  "\<lbrakk> pc < length ins; c!pc \<noteq> \<bottom> \<rbrakk> \<Longrightarrow> \<phi> ! pc = c ! pc"
+  by (simp add: phi_def)
+
+lemma (in lbvs) phi_len [simp]:
+  "length \<phi> = length ins"
+  by (simp add: phi_def)
+
+
+lemma (in lbvs) wtl_suc_pc:
+  assumes all: "wtl ins c 0 s\<^sub>0 \<noteq> \<top>" 
+  assumes pc:  "pc+1 < length ins"
+  shows "wtl (take (pc+1) ins) c 0 s0 \<sqsubseteq>\<^sub>r \<phi>!(pc+1)"
+proof -
+  from all pc
+  have "wtc c (pc+1) (wtl (take (pc+1) ins) c 0 s0) \<noteq> T" by (rule wtl_all)
+  with pc show ?thesis by (simp add: phi_def wtc split: split_if_asm)
+qed
+
+
+lemma (in lbvs) wtl_stable:
+  assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>" 
+  assumes s0:  "s0 \<in> A" 
+  assumes pc:  "pc < length ins" 
+  shows "stable r step \<phi> pc"
+proof (unfold stable_def, clarify)
+  fix pc' s' assume step: "(pc',s') \<in> set (step pc (\<phi> ! pc))" 
+                      (is "(pc',s') \<in> set (?step pc)")
+  
+  from bounded pc step have pc': "pc' < length ins" by (rule boundedD)
+
+  from wtl have tkpc: "wtl (take pc ins) c 0 s0 \<noteq> \<top>" (is "?s1 \<noteq> _") by (rule wtl_take)
+  from wtl have s2: "wtl (take (pc+1) ins) c 0 s0 \<noteq> \<top>" (is "?s2 \<noteq> _") by (rule wtl_take)
+  
+  from wtl pc have wt_s1: "wtc c pc ?s1 \<noteq> \<top>" by (rule wtl_all)
+
+  have c_Some: "\<forall>pc t. pc < length ins \<longrightarrow> c!pc \<noteq> \<bottom> \<longrightarrow> \<phi>!pc = c!pc" 
+    by (simp add: phi_def)
+  from pc have c_None: "c!pc = \<bottom> \<Longrightarrow> \<phi>!pc = ?s1" ..
+
+  from wt_s1 pc c_None c_Some
+  have inst: "wtc c pc ?s1  = wti c pc (\<phi>!pc)"
+    by (simp add: wtc split: split_if_asm)
+
+  from pres cert s0 wtl pc have "?s1 \<in> A" by (rule wtl_pres)
+  with pc c_Some cert c_None
+  have "\<phi>!pc \<in> A" by (cases "c!pc = \<bottom>") (auto dest: cert_okD1)
+  with pc pres
+  have step_in_A: "snd`set (?step pc) \<subseteq> A" by (auto dest: pres_typeD2)
+
+  show "s' <=_r \<phi>!pc'" 
+  proof (cases "pc' = pc+1")
+    case True
+    with pc' cert
+    have cert_in_A: "c!(pc+1) \<in> A" by (auto dest: cert_okD1)
+    from True pc' have pc1: "pc+1 < length ins" by simp
+    with tkpc have "?s2 = wtc c pc ?s1" by - (rule wtl_Suc)
+    with inst 
+    have merge: "?s2 = merge c pc (?step pc) (c!(pc+1))" by (simp add: wti)
+    also    
+    from s2 merge have "\<dots> \<noteq> \<top>" (is "?merge \<noteq> _") by simp
+    with cert_in_A step_in_A
+    have "?merge = (map snd [(p',t') \<leftarrow> ?step pc. p'=pc+1] ++_f (c!(pc+1)))"
+      by (rule merge_not_top_s) 
+    finally
+    have "s' <=_r ?s2" using step_in_A cert_in_A True step 
+      by (auto intro: pp_ub1')
+    also 
+    from wtl pc1 have "?s2 <=_r \<phi>!(pc+1)" by (rule wtl_suc_pc)
+    also note True [symmetric]
+    finally show ?thesis by simp    
+  next
+    case False
+    from wt_s1 inst
+    have "merge c pc (?step pc) (c!(pc+1)) \<noteq> \<top>" by (simp add: wti)
+    with step_in_A
+    have "\<forall>(pc', s')\<in>set (?step pc). pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'" 
+      by - (rule merge_not_top)
+    with step False 
+    have ok: "s' <=_r c!pc'" by blast
+    moreover
+    from ok
+    have "c!pc' = \<bottom> \<Longrightarrow> s' = \<bottom>" by simp
+    moreover
+    from c_Some pc'
+    have "c!pc' \<noteq> \<bottom> \<Longrightarrow> \<phi>!pc' = c!pc'" by auto
+    ultimately
+    show ?thesis by (cases "c!pc' = \<bottom>") auto 
+  qed
+qed
+
+  
+lemma (in lbvs) phi_not_top:
+  assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>"
+  assumes pc:  "pc < length ins"
+  shows "\<phi>!pc \<noteq> \<top>"
+proof (cases "c!pc = \<bottom>")
+  case False with pc
+  have "\<phi>!pc = c!pc" ..
+  also from cert pc have "\<dots> \<noteq> \<top>" by (rule cert_okD4)
+  finally show ?thesis .
+next
+  case True with pc
+  have "\<phi>!pc = wtl (take pc ins) c 0 s0" ..
+  also from wtl have "\<dots> \<noteq> \<top>" by (rule wtl_take)
+  finally show ?thesis .
+qed
+
+lemma (in lbvs) phi_in_A:
+  assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>"
+  assumes s0:  "s0 \<in> A"
+  shows "\<phi> \<in> list (length ins) A"
+proof -
+  { fix x assume "x \<in> set \<phi>"
+    then obtain xs ys where "\<phi> = xs @ x # ys" 
+      by (auto simp add: in_set_conv_decomp)
+    then obtain pc where pc: "pc < length \<phi>" and x: "\<phi>!pc = x"
+      by (simp add: that [of "length xs"] nth_append)
+    
+    from pres cert wtl s0 pc
+    have "wtl (take pc ins) c 0 s0 \<in> A" by (auto intro!: wtl_pres)
+    moreover
+    from pc have "pc < length ins" by simp
+    with cert have "c!pc \<in> A" ..
+    ultimately
+    have "\<phi>!pc \<in> A" using pc by (simp add: phi_def)
+    hence "x \<in> A" using x by simp
+  } 
+  hence "set \<phi> \<subseteq> A" ..
+  thus ?thesis by (unfold list_def) simp
+qed
+
+
+lemma (in lbvs) phi0:
+  assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>"
+  assumes 0:   "0 < length ins"
+  shows "s0 <=_r \<phi>!0"
+proof (cases "c!0 = \<bottom>")
+  case True
+  with 0 have "\<phi>!0 = wtl (take 0 ins) c 0 s0" ..
+  moreover have "wtl (take 0 ins) c 0 s0 = s0" by simp
+  ultimately have "\<phi>!0 = s0" by simp
+  thus ?thesis by simp
+next
+  case False
+  with 0 have "phi!0 = c!0" ..
+  moreover 
+  from wtl have "wtl (take 1 ins) c 0 s0 \<noteq> \<top>"  by (rule wtl_take)
+  with 0 False 
+  have "s0 <=_r c!0" by (auto simp add: neq_Nil_conv wtc split: split_if_asm)
+  ultimately
+  show ?thesis by simp
+qed
+
+
+theorem (in lbvs) wtl_sound:
+  assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>" 
+  assumes s0: "s0 \<in> A" 
+  shows "\<exists>ts. wt_step r \<top> step ts"
+proof -
+  have "wt_step r \<top> step \<phi>"
+  proof (unfold wt_step_def, intro strip conjI)
+    fix pc assume "pc < length \<phi>"
+    then have pc: "pc < length ins" by simp
+    with wtl show "\<phi>!pc \<noteq> \<top>" by (rule phi_not_top)
+    from wtl s0 pc show "stable r step \<phi> pc" by (rule wtl_stable)
+  qed
+  thus ?thesis ..
+qed
+
+
+theorem (in lbvs) wtl_sound_strong:
+  assumes wtl: "wtl ins c 0 s0 \<noteq> \<top>" 
+  assumes s0: "s0 \<in> A" 
+  assumes nz: "0 < length ins"
+  shows "\<exists>ts \<in> list (length ins) A. wt_step r \<top> step ts \<and> s0 <=_r ts!0"
+proof -
+  from wtl s0 have "\<phi> \<in> list (length ins) A" by (rule phi_in_A)
+  moreover
+  have "wt_step r \<top> step \<phi>"
+  proof (unfold wt_step_def, intro strip conjI)
+    fix pc assume "pc < length \<phi>"
+    then have pc: "pc < length ins" by simp
+    with wtl show "\<phi>!pc \<noteq> \<top>" by (rule phi_not_top)
+    from wtl s0 pc show "stable r step \<phi> pc" by (rule wtl_stable)
+  qed
+  moreover
+  from wtl nz have "s0 <=_r \<phi>!0" by (rule phi0)
+  ultimately
+  show ?thesis by fast
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/MicroJava/DFA/LBVSpec.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -0,0 +1,381 @@
+(*  Title:      HOL/MicroJava/BV/LBVSpec.thy
+    Author:     Gerwin Klein
+    Copyright   1999 Technische Universitaet Muenchen
+*)
+
+header {* \isaheader{The Lightweight Bytecode Verifier} *}
+
+theory LBVSpec
+imports SemilatAlg Opt
+begin
+
+types
+  's certificate = "'s list"   
+
+consts
+merge :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> nat \<Rightarrow> (nat \<times> 's) list \<Rightarrow> 's \<Rightarrow> 's"
+primrec
+"merge cert f r T pc []     x = x"
+"merge cert f r T pc (s#ss) x = merge cert f r T pc ss (let (pc',s') = s in 
+                                  if pc'=pc+1 then s' +_f x
+                                  else if s' <=_r (cert!pc') then x
+                                  else T)"
+
+constdefs
+wtl_inst :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow>
+             's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's"
+"wtl_inst cert f r T step pc s \<equiv> merge cert f r T pc (step pc s) (cert!(pc+1))"
+
+wtl_cert :: "'s certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow>
+             's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's"
+"wtl_cert cert f r T B step pc s \<equiv>
+  if cert!pc = B then 
+    wtl_inst cert f r T step pc s
+  else
+    if s <=_r (cert!pc) then wtl_inst cert f r T step pc (cert!pc) else T"
+
+consts 
+wtl_inst_list :: "'a list \<Rightarrow> 's certificate \<Rightarrow> 's binop \<Rightarrow> 's ord \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow>
+                  's step_type \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's"
+primrec
+"wtl_inst_list []     cert f r T B step pc s = s"
+"wtl_inst_list (i#is) cert f r T B step pc s = 
+    (let s' = wtl_cert cert f r T B step pc s in
+      if s' = T \<or> s = T then T else wtl_inst_list is cert f r T B step (pc+1) s')"
+
+constdefs
+  cert_ok :: "'s certificate \<Rightarrow> nat \<Rightarrow> 's \<Rightarrow> 's \<Rightarrow> 's set \<Rightarrow> bool"
+  "cert_ok cert n T B A \<equiv> (\<forall>i < n. cert!i \<in> A \<and> cert!i \<noteq> T) \<and> (cert!n = B)"
+
+constdefs
+  bottom :: "'a ord \<Rightarrow> 'a \<Rightarrow> bool"
+  "bottom r B \<equiv> \<forall>x. B <=_r x"
+
+
+locale lbv = Semilat +
+  fixes T :: "'a" ("\<top>") 
+  fixes B :: "'a" ("\<bottom>") 
+  fixes step :: "'a step_type" 
+  assumes top: "top r \<top>"
+  assumes T_A: "\<top> \<in> A"
+  assumes bot: "bottom r \<bottom>" 
+  assumes B_A: "\<bottom> \<in> A"
+
+  fixes merge :: "'a certificate \<Rightarrow> nat \<Rightarrow> (nat \<times> 'a) list \<Rightarrow> 'a \<Rightarrow> 'a"
+  defines mrg_def: "merge cert \<equiv> LBVSpec.merge cert f r \<top>"
+
+  fixes wti :: "'a certificate \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a"
+  defines wti_def: "wti cert \<equiv> wtl_inst cert f r \<top> step"
+ 
+  fixes wtc :: "'a certificate \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a"
+  defines wtc_def: "wtc cert \<equiv> wtl_cert cert f r \<top> \<bottom> step"
+
+  fixes wtl :: "'b list \<Rightarrow> 'a certificate \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a"
+  defines wtl_def: "wtl ins cert \<equiv> wtl_inst_list ins cert f r \<top> \<bottom> step"
+
+
+lemma (in lbv) wti:
+  "wti c pc s \<equiv> merge c pc (step pc s) (c!(pc+1))"
+  by (simp add: wti_def mrg_def wtl_inst_def)
+
+lemma (in lbv) wtc: 
+  "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>"
+  by (unfold wtc_def wti_def wtl_cert_def)
+
+
+lemma cert_okD1 [intro?]:
+  "cert_ok c n T B A \<Longrightarrow> pc < n \<Longrightarrow> c!pc \<in> A"
+  by (unfold cert_ok_def) fast
+
+lemma cert_okD2 [intro?]:
+  "cert_ok c n T B A \<Longrightarrow> c!n = B"
+  by (simp add: cert_ok_def)
+
+lemma cert_okD3 [intro?]:
+  "cert_ok c n T B A \<Longrightarrow> B \<in> A \<Longrightarrow> pc < n \<Longrightarrow> c!Suc pc \<in> A"
+  by (drule Suc_leI) (auto simp add: le_eq_less_or_eq dest: cert_okD1 cert_okD2)
+
+lemma cert_okD4 [intro?]:
+  "cert_ok c n T B A \<Longrightarrow> pc < n \<Longrightarrow> c!pc \<noteq> T"
+  by (simp add: cert_ok_def)
+
+declare Let_def [simp]
+
+section "more semilattice lemmas"
+
+
+lemma (in lbv) sup_top [simp, elim]:
+  assumes x: "x \<in> A" 
+  shows "x +_f \<top> = \<top>"
+proof -
+  from top have "x +_f \<top> <=_r \<top>" ..
+  moreover from x T_A have "\<top> <=_r x +_f \<top>" ..
+  ultimately show ?thesis ..
+qed
+  
+lemma (in lbv) plusplussup_top [simp, elim]:
+  "set xs \<subseteq> A \<Longrightarrow> xs ++_f \<top> = \<top>"
+  by (induct xs) auto
+
+
+
+lemma (in Semilat) pp_ub1':
+  assumes S: "snd`set S \<subseteq> A" 
+  assumes y: "y \<in> A" and ab: "(a, b) \<in> set S" 
+  shows "b <=_r map snd [(p', t') \<leftarrow> S . p' = a] ++_f y"
+proof -
+  from S have "\<forall>(x,y) \<in> set S. y \<in> A" by auto
+  with semilat y ab show ?thesis by - (rule ub1')
+qed 
+
+lemma (in lbv) bottom_le [simp, intro]:
+  "\<bottom> <=_r x"
+  by (insert bot) (simp add: bottom_def)
+
+lemma (in lbv) le_bottom [simp]:
+  "x <=_r \<bottom> = (x = \<bottom>)"
+  by (blast intro: antisym_r)
+
+
+
+section "merge"
+
+lemma (in lbv) merge_Nil [simp]:
+  "merge c pc [] x = x" by (simp add: mrg_def)
+
+lemma (in lbv) merge_Cons [simp]:
+  "merge c pc (l#ls) x = merge c pc ls (if fst l=pc+1 then snd l +_f x
+                                        else if snd l <=_r (c!fst l) then x
+                                        else \<top>)"
+  by (simp add: mrg_def split_beta)
+
+lemma (in lbv) merge_Err [simp]:
+  "snd`set ss \<subseteq> A \<Longrightarrow> merge c pc ss \<top> = \<top>"
+  by (induct ss) auto
+
+lemma (in lbv) merge_not_top:
+  "\<And>x. snd`set ss \<subseteq> A \<Longrightarrow> merge c pc ss x \<noteq> \<top> \<Longrightarrow> 
+  \<forall>(pc',s') \<in> set ss. (pc' \<noteq> pc+1 \<longrightarrow> s' <=_r (c!pc'))"
+  (is "\<And>x. ?set ss \<Longrightarrow> ?merge ss x \<Longrightarrow> ?P ss")
+proof (induct ss)
+  show "?P []" by simp
+next
+  fix x ls l
+  assume "?set (l#ls)" then obtain set: "snd`set ls \<subseteq> A" by simp
+  assume merge: "?merge (l#ls) x" 
+  moreover
+  obtain pc' s' where l: "l = (pc',s')" by (cases l)
+  ultimately
+  obtain x' where merge': "?merge ls x'" by simp 
+  assume "\<And>x. ?set ls \<Longrightarrow> ?merge ls x \<Longrightarrow> ?P ls" hence "?P ls" using set merge' .
+  moreover
+  from merge set
+  have "pc' \<noteq> pc+1 \<longrightarrow> s' <=_r (c!pc')" by (simp add: l split: split_if_asm)
+  ultimately
+  show "?P (l#ls)" by (simp add: l)
+qed
+
+
+lemma (in lbv) merge_def:
+  shows 
+  "\<And>x. x \<in> A \<Longrightarrow> snd`set ss \<subseteq> A \<Longrightarrow>
+  merge c pc ss x = 
+  (if \<forall>(pc',s') \<in> set ss. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc' then
+    map snd [(p',t') \<leftarrow> ss. p'=pc+1] ++_f x
+  else \<top>)" 
+  (is "\<And>x. _ \<Longrightarrow> _ \<Longrightarrow> ?merge ss x = ?if ss x" is "\<And>x. _ \<Longrightarrow> _ \<Longrightarrow> ?P ss x")
+proof (induct ss)
+  fix x show "?P [] x" by simp
+next 
+  fix x assume x: "x \<in> A" 
+  fix l::"nat \<times> 'a" and ls  
+  assume "snd`set (l#ls) \<subseteq> A"
+  then obtain l: "snd l \<in> A" and ls: "snd`set ls \<subseteq> A" by auto
+  assume "\<And>x. x \<in> A \<Longrightarrow> snd`set ls \<subseteq> A \<Longrightarrow> ?P ls x" 
+  hence IH: "\<And>x. x \<in> A \<Longrightarrow> ?P ls x" using ls by iprover
+  obtain pc' s' where [simp]: "l = (pc',s')" by (cases l)
+  hence "?merge (l#ls) x = ?merge ls 
+    (if pc'=pc+1 then s' +_f x else if s' <=_r c!pc' then x else \<top>)"
+    (is "?merge (l#ls) x = ?merge ls ?if'")
+    by simp 
+  also have "\<dots> = ?if ls ?if'" 
+  proof -
+    from l have "s' \<in> A" by simp
+    with x have "s' +_f x \<in> A" by simp
+    with x T_A have "?if' \<in> A" by auto
+    hence "?P ls ?if'" by (rule IH) thus ?thesis by simp
+  qed
+  also have "\<dots> = ?if (l#ls) x"
+    proof (cases "\<forall>(pc', s')\<in>set (l#ls). pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'")
+      case True
+      hence "\<forall>(pc', s')\<in>set ls. pc'\<noteq>pc+1 \<longrightarrow> s' <=_r c!pc'" by auto
+      moreover
+      from True have 
+        "map snd [(p',t')\<leftarrow>ls . p'=pc+1] ++_f ?if' = 
+        (map snd [(p',t')\<leftarrow>l#ls . p'=pc+1] ++_f x)"
+        by simp
+      ultimately
+      show ?thesis using True by simp
+    next
+      case False 
+      moreover
+      from ls have "set (map snd [(p', t')\<leftarrow>ls . p' = Suc pc]) \<subseteq> A" by auto
+      ultimately show ?thesis by auto
+    qed
+  finally show "?P (l#ls) x" .
+qed
+
+lemma (in lbv) merge_not_top_s:
+  assumes x:  "x \<in> A" and ss: "snd`set ss \<subseteq> A"
+  assumes m:  "merge c pc ss x \<noteq> \<top>"
+  shows "merge c pc ss x = (map snd [(p',t') \<leftarrow> ss. p'=pc+1] ++_f x)"
+proof -
+  from ss m have "\<forall>(pc',s') \<in> set ss. (pc' \<noteq> pc+1 \<longrightarrow> s' <=_r c!pc')" 
+    by (rule merge_not_top)
+  with x ss m show ?thesis by - (drule merge_def, auto split: split_if_asm)
+qed
+
+section "wtl-inst-list"
+
+lemmas [iff] = not_Err_eq
+
+lemma (in lbv) wtl_Nil [simp]: "wtl [] c pc s = s" 
+  by (simp add: wtl_def)
+
+lemma (in lbv) wtl_Cons [simp]: 
+  "wtl (i#is) c pc s = 
+  (let s' = wtc c pc s in if s' = \<top> \<or> s = \<top> then \<top> else wtl is c (pc+1) s')"
+  by (simp add: wtl_def wtc_def)
+
+lemma (in lbv) wtl_Cons_not_top:
+  "wtl (i#is) c pc s \<noteq> \<top> = 
+  (wtc c pc s \<noteq> \<top> \<and> s \<noteq> T \<and> wtl is c (pc+1) (wtc c pc s) \<noteq> \<top>)"
+  by (auto simp del: split_paired_Ex)
+
+lemma (in lbv) wtl_top [simp]:  "wtl ls c pc \<top> = \<top>"
+  by (cases ls) auto
+
+lemma (in lbv) wtl_not_top:
+  "wtl ls c pc s \<noteq> \<top> \<Longrightarrow> s \<noteq> \<top>"
+  by (cases "s=\<top>") auto
+
+lemma (in lbv) wtl_append [simp]:
+  "\<And>pc s. wtl (a@b) c pc s = wtl b c (pc+length a) (wtl a c pc s)"
+  by (induct a) auto
+
+lemma (in lbv) wtl_take:
+  "wtl is c pc s \<noteq> \<top> \<Longrightarrow> wtl (take pc' is) c pc s \<noteq> \<top>"
+  (is "?wtl is \<noteq> _ \<Longrightarrow> _")
+proof -
+  assume "?wtl is \<noteq> \<top>"
+  hence "?wtl (take pc' is @ drop pc' is) \<noteq> \<top>" by simp  
+  thus ?thesis by (auto dest!: wtl_not_top simp del: append_take_drop_id)
+qed
+
+lemma take_Suc:
+  "\<forall>n. n < length l \<longrightarrow> take (Suc n) l = (take n l)@[l!n]" (is "?P l")
+proof (induct l)
+  show "?P []" by simp
+next
+  fix x xs assume IH: "?P xs"  
+  show "?P (x#xs)"
+  proof (intro strip)
+    fix n assume "n < length (x#xs)"
+    with IH show "take (Suc n) (x # xs) = take n (x # xs) @ [(x # xs) ! n]" 
+      by (cases n, auto)
+  qed
+qed
+
+lemma (in lbv) wtl_Suc:
+  assumes suc: "pc+1 < length is"
+  assumes wtl: "wtl (take pc is) c 0 s \<noteq> \<top>"
+  shows "wtl (take (pc+1) is) c 0 s = wtc c pc (wtl (take pc is) c 0 s)"
+proof -
+  from suc have "take (pc+1) is=(take pc is)@[is!pc]" by (simp add: take_Suc)
+  with suc wtl show ?thesis by (simp add: min_max.inf_absorb2)
+qed
+
+lemma (in lbv) wtl_all:
+  assumes all: "wtl is c 0 s \<noteq> \<top>" (is "?wtl is \<noteq> _") 
+  assumes pc:  "pc < length is"
+  shows  "wtc c pc (wtl (take pc is) c 0 s) \<noteq> \<top>"
+proof -
+  from pc have "0 < length (drop pc is)" by simp
+  then  obtain i r where Cons: "drop pc is = i#r" 
+    by (auto simp add: neq_Nil_conv simp del: length_drop drop_eq_Nil)
+  hence "i#r = drop pc is" ..
+  with all have take: "?wtl (take pc is@i#r) \<noteq> \<top>" by simp 
+  from pc have "is!pc = drop pc is ! 0" by simp
+  with Cons have "is!pc = i" by simp
+  with take pc show ?thesis by (auto simp add: min_max.inf_absorb2)
+qed
+
+section "preserves-type"
+
+lemma (in lbv) merge_pres:
+  assumes s0: "snd`set ss \<subseteq> A" and x: "x \<in> A"
+  shows "merge c pc ss x \<in> A"
+proof -
+  from s0 have "set (map snd [(p', t')\<leftarrow>ss . p'=pc+1]) \<subseteq> A" by auto
+  with x  have "(map snd [(p', t')\<leftarrow>ss . p'=pc+1] ++_f x) \<in> A"
+    by (auto intro!: plusplus_closed semilat)
+  with s0 x show ?thesis by (simp add: merge_def T_A)
+qed
+  
+
+lemma pres_typeD2:
+  "pres_type step n A \<Longrightarrow> s \<in> A \<Longrightarrow> p < n \<Longrightarrow> snd`set (step p s) \<subseteq> A"
+  by auto (drule pres_typeD)
+
+
+lemma (in lbv) wti_pres [intro?]:
+  assumes pres: "pres_type step n A" 
+  assumes cert: "c!(pc+1) \<in> A"
+  assumes s_pc: "s \<in> A" "pc < n"
+  shows "wti c pc s \<in> A"
+proof -
+  from pres s_pc have "snd`set (step pc s) \<subseteq> A" by (rule pres_typeD2)
+  with cert show ?thesis by (simp add: wti merge_pres)
+qed
+
+
+lemma (in lbv) wtc_pres:
+  assumes pres: "pres_type step n A"
+  assumes cert: "c!pc \<in> A" and cert': "c!(pc+1) \<in> A"
+  assumes s: "s \<in> A" and pc: "pc < n"
+  shows "wtc c pc s \<in> A"
+proof -
+  have "wti c pc s \<in> A" using pres cert' s pc ..
+  moreover have "wti c pc (c!pc) \<in> A" using pres cert' cert pc ..
+  ultimately show ?thesis using T_A by (simp add: wtc) 
+qed
+
+
+lemma (in lbv) wtl_pres:
+  assumes pres: "pres_type step (length is) A"
+  assumes cert: "cert_ok c (length is) \<top> \<bottom> A"
+  assumes s:    "s \<in> A" 
+  assumes all:  "wtl is c 0 s \<noteq> \<top>"
+  shows "pc < length is \<Longrightarrow> wtl (take pc is) c 0 s \<in> A"
+  (is "?len pc \<Longrightarrow> ?wtl pc \<in> A")
+proof (induct pc)
+  from s show "?wtl 0 \<in> A" by simp
+next
+  fix n assume IH: "Suc n < length is"
+  then have n: "n < length is" by simp  
+  from IH have n1: "n+1 < length is" by simp
+  assume prem: "n < length is \<Longrightarrow> ?wtl n \<in> A"
+  have "wtc c n (?wtl n) \<in> A"
+  using pres _ _ _ n
+  proof (rule wtc_pres)
+    from prem n show "?wtl n \<in> A" .
+    from cert n show "c!n \<in> A" by (rule cert_okD1)
+    from cert n1 show "c!(n+1) \<in> A" by (rule cert_okD1)
+  qed
+  also
+  from all n have "?wtl n \<noteq> \<top>" by - (rule wtl_take)
+  with n1 have "wtc c n (?wtl n) = ?wtl (n+1)" by (rule wtl_Suc [symmetric])
+  finally  show "?wtl (Suc n) \<in> A" by simp
+qed
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/MicroJava/DFA/Listn.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -0,0 +1,542 @@
+(*  Title:      HOL/MicroJava/BV/Listn.thy
+    Author:     Tobias Nipkow
+    Copyright   2000 TUM
+*)
+
+header {* \isaheader{Fixed Length Lists} *}
+
+theory Listn
+imports Err
+begin
+
+constdefs
+
+ list :: "nat \<Rightarrow> 'a set \<Rightarrow> 'a list set"
+"list n A == {xs. length xs = n & set xs <= A}"
+
+ le :: "'a ord \<Rightarrow> ('a list)ord"
+"le r == list_all2 (%x y. x <=_r y)"
+
+syntax "@lesublist" :: "'a list \<Rightarrow> 'a ord \<Rightarrow> 'a list \<Rightarrow> bool"
+       ("(_ /<=[_] _)" [50, 0, 51] 50)
+syntax "@lesssublist" :: "'a list \<Rightarrow> 'a ord \<Rightarrow> 'a list \<Rightarrow> bool"
+       ("(_ /<[_] _)" [50, 0, 51] 50)
+translations
+ "x <=[r] y" == "x <=_(Listn.le r) y"
+ "x <[r] y"  == "x <_(Listn.le r) y"
+
+constdefs
+ map2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list"
+"map2 f == (%xs ys. map (split f) (zip xs ys))"
+
+syntax "@plussublist" :: "'a list \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b list \<Rightarrow> 'c list"
+       ("(_ /+[_] _)" [65, 0, 66] 65)
+translations  "x +[f] y" == "x +_(map2 f) y"
+
+consts coalesce :: "'a err list \<Rightarrow> 'a list err"
+primrec
+"coalesce [] = OK[]"
+"coalesce (ex#exs) = Err.sup (op #) ex (coalesce exs)"
+
+constdefs
+ sl :: "nat \<Rightarrow> 'a sl \<Rightarrow> 'a list sl"
+"sl n == %(A,r,f). (list n A, le r, map2 f)"
+
+ sup :: "('a \<Rightarrow> 'b \<Rightarrow> 'c err) \<Rightarrow> 'a list \<Rightarrow> 'b list \<Rightarrow> 'c list err"
+"sup f == %xs ys. if size xs = size ys then coalesce(xs +[f] ys) else Err"
+
+ upto_esl :: "nat \<Rightarrow> 'a esl \<Rightarrow> 'a list esl"
+"upto_esl m == %(A,r,f). (Union{list n A |n. n <= m}, le r, sup f)"
+
+lemmas [simp] = set_update_subsetI
+
+lemma unfold_lesub_list:
+  "xs <=[r] ys == Listn.le r xs ys"
+  by (simp add: lesub_def)
+
+lemma Nil_le_conv [iff]:
+  "([] <=[r] ys) = (ys = [])"
+apply (unfold lesub_def Listn.le_def)
+apply simp
+done
+
+lemma Cons_notle_Nil [iff]: 
+  "~ x#xs <=[r] []"
+apply (unfold lesub_def Listn.le_def)
+apply simp
+done
+
+
+lemma Cons_le_Cons [iff]:
+  "x#xs <=[r] y#ys = (x <=_r y & xs <=[r] ys)"
+apply (unfold lesub_def Listn.le_def)
+apply simp
+done
+
+lemma Cons_less_Conss [simp]:
+  "order r \<Longrightarrow> 
+  x#xs <_(Listn.le r) y#ys = 
+  (x <_r y & xs <=[r] ys  |  x = y & xs <_(Listn.le r) ys)"
+apply (unfold lesssub_def)
+apply blast
+done  
+
+lemma list_update_le_cong:
+  "\<lbrakk> i<size xs; xs <=[r] ys; x <=_r y \<rbrakk> \<Longrightarrow> xs[i:=x] <=[r] ys[i:=y]";
+apply (unfold unfold_lesub_list)
+apply (unfold Listn.le_def)
+apply (simp add: list_all2_conv_all_nth nth_list_update)
+done
+
+
+lemma le_listD:
+  "\<lbrakk> xs <=[r] ys; p < size xs \<rbrakk> \<Longrightarrow> xs!p <=_r ys!p"
+apply (unfold Listn.le_def lesub_def)
+apply (simp add: list_all2_conv_all_nth)
+done
+
+lemma le_list_refl:
+  "!x. x <=_r x \<Longrightarrow> xs <=[r] xs"
+apply (unfold unfold_lesub_list)
+apply (simp add: Listn.le_def list_all2_conv_all_nth)
+done
+
+lemma le_list_trans:
+  "\<lbrakk> order r; xs <=[r] ys; ys <=[r] zs \<rbrakk> \<Longrightarrow> xs <=[r] zs"
+apply (unfold unfold_lesub_list)
+apply (simp add: Listn.le_def list_all2_conv_all_nth)
+apply clarify
+apply simp
+apply (blast intro: order_trans)
+done
+
+lemma le_list_antisym:
+  "\<lbrakk> order r; xs <=[r] ys; ys <=[r] xs \<rbrakk> \<Longrightarrow> xs = ys"
+apply (unfold unfold_lesub_list)
+apply (simp add: Listn.le_def list_all2_conv_all_nth)
+apply (rule nth_equalityI)
+ apply blast
+apply clarify
+apply simp
+apply (blast intro: order_antisym)
+done
+
+lemma order_listI [simp, intro!]:
+  "order r \<Longrightarrow> order(Listn.le r)"
+apply (subst Semilat.order_def)
+apply (blast intro: le_list_refl le_list_trans le_list_antisym
+             dest: order_refl)
+done
+
+
+lemma lesub_list_impl_same_size [simp]:
+  "xs <=[r] ys \<Longrightarrow> size ys = size xs"  
+apply (unfold Listn.le_def lesub_def)
+apply (simp add: list_all2_conv_all_nth)
+done 
+
+lemma lesssub_list_impl_same_size:
+  "xs <_(Listn.le r) ys \<Longrightarrow> size ys = size xs"
+apply (unfold lesssub_def)
+apply auto
+done  
+
+lemma le_list_appendI:
+  "\<And>b c d. a <=[r] b \<Longrightarrow> c <=[r] d \<Longrightarrow> a@c <=[r] b@d"
+apply (induct a)
+ apply simp
+apply (case_tac b)
+apply auto
+done
+
+lemma le_listI:
+  "length a = length b \<Longrightarrow> (\<And>n. n < length a \<Longrightarrow> a!n <=_r b!n) \<Longrightarrow> a <=[r] b"
+  apply (unfold lesub_def Listn.le_def)
+  apply (simp add: list_all2_conv_all_nth)
+  done
+
+lemma listI:
+  "\<lbrakk> length xs = n; set xs <= A \<rbrakk> \<Longrightarrow> xs : list n A"
+apply (unfold list_def)
+apply blast
+done
+
+lemma listE_length [simp]:
+   "xs : list n A \<Longrightarrow> length xs = n"
+apply (unfold list_def)
+apply blast
+done 
+
+lemma less_lengthI:
+  "\<lbrakk> xs : list n A; p < n \<rbrakk> \<Longrightarrow> p < length xs"
+  by simp
+
+lemma listE_set [simp]:
+  "xs : list n A \<Longrightarrow> set xs <= A"
+apply (unfold list_def)
+apply blast
+done 
+
+lemma list_0 [simp]:
+  "list 0 A = {[]}"
+apply (unfold list_def)
+apply auto
+done 
+
+lemma in_list_Suc_iff: 
+  "(xs : list (Suc n) A) = (\<exists>y\<in> A. \<exists>ys\<in> list n A. xs = y#ys)"
+apply (unfold list_def)
+apply (case_tac "xs")
+apply auto
+done 
+
+lemma Cons_in_list_Suc [iff]:
+  "(x#xs : list (Suc n) A) = (x\<in> A & xs : list n A)";
+apply (simp add: in_list_Suc_iff)
+done 
+
+lemma list_not_empty:
+  "\<exists>a. a\<in> A \<Longrightarrow> \<exists>xs. xs : list n A";
+apply (induct "n")
+ apply simp
+apply (simp add: in_list_Suc_iff)
+apply blast
+done
+
+
+lemma nth_in [rule_format, simp]:
+  "!i n. length xs = n \<longrightarrow> set xs <= A \<longrightarrow> i < n \<longrightarrow> (xs!i) : A"
+apply (induct "xs")
+ apply simp
+apply (simp add: nth_Cons split: nat.split)
+done
+
+lemma listE_nth_in:
+  "\<lbrakk> xs : list n A; i < n \<rbrakk> \<Longrightarrow> (xs!i) : A"
+  by auto
+
+
+lemma listn_Cons_Suc [elim!]:
+  "l#xs \<in> list n A \<Longrightarrow> (\<And>n'. n = Suc n' \<Longrightarrow> l \<in> A \<Longrightarrow> xs \<in> list n' A \<Longrightarrow> P) \<Longrightarrow> P"
+  by (cases n) auto
+
+lemma listn_appendE [elim!]:
+  "a@b \<in> list n A \<Longrightarrow> (\<And>n1 n2. n=n1+n2 \<Longrightarrow> a \<in> list n1 A \<Longrightarrow> b \<in> list n2 A \<Longrightarrow> P) \<Longrightarrow> P" 
+proof -
+  have "\<And>n. a@b \<in> list n A \<Longrightarrow> \<exists>n1 n2. n=n1+n2 \<and> a \<in> list n1 A \<and> b \<in> list n2 A"
+    (is "\<And>n. ?list a n \<Longrightarrow> \<exists>n1 n2. ?P a n n1 n2")
+  proof (induct a)
+    fix n assume "?list [] n"
+    hence "?P [] n 0 n" by simp
+    thus "\<exists>n1 n2. ?P [] n n1 n2" by fast
+  next
+    fix n l ls
+    assume "?list (l#ls) n"
+    then obtain n' where n: "n = Suc n'" "l \<in> A" and list_n': "ls@b \<in> list n' A" by fastsimp
+    assume "\<And>n. ls @ b \<in> list n A \<Longrightarrow> \<exists>n1 n2. n = n1 + n2 \<and> ls \<in> list n1 A \<and> b \<in> list n2 A"
+    hence "\<exists>n1 n2. n' = n1 + n2 \<and> ls \<in> list n1 A \<and> b \<in> list n2 A" by this (rule list_n')
+    then obtain n1 n2 where "n' = n1 + n2" "ls \<in> list n1 A" "b \<in> list n2 A" by fast
+    with n have "?P (l#ls) n (n1+1) n2" by simp
+    thus "\<exists>n1 n2. ?P (l#ls) n n1 n2" by fastsimp
+  qed
+  moreover
+  assume "a@b \<in> list n A" "\<And>n1 n2. n=n1+n2 \<Longrightarrow> a \<in> list n1 A \<Longrightarrow> b \<in> list n2 A \<Longrightarrow> P"
+  ultimately
+  show ?thesis by blast
+qed
+
+
+lemma listt_update_in_list [simp, intro!]:
+  "\<lbrakk> xs : list n A; x\<in> A \<rbrakk> \<Longrightarrow> xs[i := x] : list n A"
+apply (unfold list_def)
+apply simp
+done 
+
+lemma plus_list_Nil [simp]:
+  "[] +[f] xs = []"
+apply (unfold plussub_def map2_def)
+apply simp
+done 
+
+lemma plus_list_Cons [simp]:
+  "(x#xs) +[f] ys = (case ys of [] \<Rightarrow> [] | y#ys \<Rightarrow> (x +_f y)#(xs +[f] ys))"
+  by (simp add: plussub_def map2_def split: list.split)
+
+lemma length_plus_list [rule_format, simp]:
+  "!ys. length(xs +[f] ys) = min(length xs) (length ys)"
+apply (induct xs)
+ apply simp
+apply clarify
+apply (simp (no_asm_simp) split: list.split)
+done
+
+lemma nth_plus_list [rule_format, simp]:
+  "!xs ys i. length xs = n \<longrightarrow> length ys = n \<longrightarrow> i<n \<longrightarrow> 
+  (xs +[f] ys)!i = (xs!i) +_f (ys!i)"
+apply (induct n)
+ apply simp
+apply clarify
+apply (case_tac xs)
+ apply simp
+apply (force simp add: nth_Cons split: list.split nat.split)
+done
+
+
+lemma (in Semilat) plus_list_ub1 [rule_format]:
+ "\<lbrakk> set xs <= A; set ys <= A; size xs = size ys \<rbrakk> 
+  \<Longrightarrow> xs <=[r] xs +[f] ys"
+apply (unfold unfold_lesub_list)
+apply (simp add: Listn.le_def list_all2_conv_all_nth)
+done
+
+lemma (in Semilat) plus_list_ub2:
+ "\<lbrakk>set xs <= A; set ys <= A; size xs = size ys \<rbrakk>
+  \<Longrightarrow> ys <=[r] xs +[f] ys"
+apply (unfold unfold_lesub_list)
+apply (simp add: Listn.le_def list_all2_conv_all_nth)
+done
+
+lemma (in Semilat) plus_list_lub [rule_format]:
+shows "!xs ys zs. set xs <= A \<longrightarrow> set ys <= A \<longrightarrow> set zs <= A 
+  \<longrightarrow> size xs = n & size ys = n \<longrightarrow> 
+  xs <=[r] zs & ys <=[r] zs \<longrightarrow> xs +[f] ys <=[r] zs"
+apply (unfold unfold_lesub_list)
+apply (simp add: Listn.le_def list_all2_conv_all_nth)
+done
+
+lemma (in Semilat) list_update_incr [rule_format]:
+ "x\<in> A \<Longrightarrow> set xs <= A \<longrightarrow> 
+  (!i. i<size xs \<longrightarrow> xs <=[r] xs[i := x +_f xs!i])"
+apply (unfold unfold_lesub_list)
+apply (simp add: Listn.le_def list_all2_conv_all_nth)
+apply (induct xs)
+ apply simp
+apply (simp add: in_list_Suc_iff)
+apply clarify
+apply (simp add: nth_Cons split: nat.split)
+done
+
+lemma equals0I_aux:
+  "(\<And>y. A y \<Longrightarrow> False) \<Longrightarrow> A = bot_class.bot"
+  by (rule equals0I) (auto simp add: mem_def)
+
+lemma acc_le_listI [intro!]:
+  "\<lbrakk> order r; acc r \<rbrakk> \<Longrightarrow> acc(Listn.le r)"
+apply (unfold acc_def)
+apply (subgoal_tac
+ "wf(UN n. {(ys,xs). size xs = n \<and> size ys = n \<and> xs <_(Listn.le r) ys})")
+ apply (erule wf_subset)
+ apply (blast intro: lesssub_list_impl_same_size)
+apply (rule wf_UN)
+ prefer 2
+ apply clarify
+ apply (rename_tac m n)
+ apply (case_tac "m=n")
+  apply simp
+ apply (fast intro!: equals0I dest: not_sym)
+apply clarify
+apply (rename_tac n)
+apply (induct_tac n)
+ apply (simp add: lesssub_def cong: conj_cong)
+apply (rename_tac k)
+apply (simp add: wf_eq_minimal)
+apply (simp (no_asm) add: length_Suc_conv cong: conj_cong)
+apply clarify
+apply (rename_tac M m)
+apply (case_tac "\<exists>x xs. size xs = k \<and> x#xs \<in> M")
+ prefer 2
+ apply (erule thin_rl)
+ apply (erule thin_rl)
+ apply blast
+apply (erule_tac x = "{a. \<exists>xs. size xs = k \<and> a#xs:M}" in allE)
+apply (erule impE)
+ apply blast
+apply (thin_tac "\<exists>x xs. ?P x xs")
+apply clarify
+apply (rename_tac maxA xs)
+apply (erule_tac x = "{ys. size ys = size xs \<and> maxA#ys \<in> M}" in allE)
+apply (erule impE)
+ apply blast
+apply clarify
+apply (thin_tac "m \<in> M")
+apply (thin_tac "maxA#xs \<in> M")
+apply (rule bexI)
+ prefer 2
+ apply assumption
+apply clarify
+apply simp
+apply blast
+done
+
+lemma closed_listI:
+  "closed S f \<Longrightarrow> closed (list n S) (map2 f)"
+apply (unfold closed_def)
+apply (induct n)
+ apply simp
+apply clarify
+apply (simp add: in_list_Suc_iff)
+apply clarify
+apply simp
+done
+
+
+lemma Listn_sl_aux:
+assumes "semilat (A, r, f)" shows "semilat (Listn.sl n (A,r,f))"
+proof -
+  interpret Semilat A r f using assms by (rule Semilat.intro)
+show ?thesis
+apply (unfold Listn.sl_def)
+apply (simp (no_asm) only: semilat_Def split_conv)
+apply (rule conjI)
+ apply simp
+apply (rule conjI)
+ apply (simp only: closedI closed_listI)
+apply (simp (no_asm) only: list_def)
+apply (simp (no_asm_simp) add: plus_list_ub1 plus_list_ub2 plus_list_lub)
+done
+qed
+
+lemma Listn_sl: "\<And>L. semilat L \<Longrightarrow> semilat (Listn.sl n L)"
+ by(simp add: Listn_sl_aux split_tupled_all)
+
+lemma coalesce_in_err_list [rule_format]:
+  "!xes. xes : list n (err A) \<longrightarrow> coalesce xes : err(list n A)"
+apply (induct n)
+ apply simp
+apply clarify
+apply (simp add: in_list_Suc_iff)
+apply clarify
+apply (simp (no_asm) add: plussub_def Err.sup_def lift2_def split: err.split)
+apply force
+done 
+
+lemma lem: "\<And>x xs. x +_(op #) xs = x#xs"
+  by (simp add: plussub_def)
+
+lemma coalesce_eq_OK1_D [rule_format]:
+  "semilat(err A, Err.le r, lift2 f) \<Longrightarrow> 
+  !xs. xs : list n A \<longrightarrow> (!ys. ys : list n A \<longrightarrow> 
+  (!zs. coalesce (xs +[f] ys) = OK zs \<longrightarrow> xs <=[r] zs))"
+apply (induct n)
+  apply simp
+apply clarify
+apply (simp add: in_list_Suc_iff)
+apply clarify
+apply (simp split: err.split_asm add: lem Err.sup_def lift2_def)
+apply (force simp add: semilat_le_err_OK1)
+done
+
+lemma coalesce_eq_OK2_D [rule_format]:
+  "semilat(err A, Err.le r, lift2 f) \<Longrightarrow> 
+  !xs. xs : list n A \<longrightarrow> (!ys. ys : list n A \<longrightarrow> 
+  (!zs. coalesce (xs +[f] ys) = OK zs \<longrightarrow> ys <=[r] zs))"
+apply (induct n)
+ apply simp
+apply clarify
+apply (simp add: in_list_Suc_iff)
+apply clarify
+apply (simp split: err.split_asm add: lem Err.sup_def lift2_def)
+apply (force simp add: semilat_le_err_OK2)
+done 
+
+lemma lift2_le_ub:
+  "\<lbrakk> semilat(err A, Err.le r, lift2 f); x\<in> A; y\<in> A; x +_f y = OK z; 
+      u\<in> A; x <=_r u; y <=_r u \<rbrakk> \<Longrightarrow> z <=_r u"
+apply (unfold semilat_Def plussub_def err_def)
+apply (simp add: lift2_def)
+apply clarify
+apply (rotate_tac -3)
+apply (erule thin_rl)
+apply (erule thin_rl)
+apply force
+done
+
+lemma coalesce_eq_OK_ub_D [rule_format]:
+  "semilat(err A, Err.le r, lift2 f) \<Longrightarrow> 
+  !xs. xs : list n A \<longrightarrow> (!ys. ys : list n A \<longrightarrow> 
+  (!zs us. coalesce (xs +[f] ys) = OK zs & xs <=[r] us & ys <=[r] us 
+           & us : list n A \<longrightarrow> zs <=[r] us))"
+apply (induct n)
+ apply simp
+apply clarify
+apply (simp add: in_list_Suc_iff)
+apply clarify
+apply (simp (no_asm_use) split: err.split_asm add: lem Err.sup_def lift2_def)
+apply clarify
+apply (rule conjI)
+ apply (blast intro: lift2_le_ub)
+apply blast
+done 
+
+lemma lift2_eq_ErrD:
+  "\<lbrakk> x +_f y = Err; semilat(err A, Err.le r, lift2 f); x\<in> A; y\<in> A \<rbrakk> 
+  \<Longrightarrow> ~(\<exists>u\<in> A. x <=_r u & y <=_r u)"
+  by (simp add: OK_plus_OK_eq_Err_conv [THEN iffD1])
+
+
+lemma coalesce_eq_Err_D [rule_format]:
+  "\<lbrakk> semilat(err A, Err.le r, lift2 f) \<rbrakk> 
+  \<Longrightarrow> !xs. xs\<in> list n A \<longrightarrow> (!ys. ys\<in> list n A \<longrightarrow> 
+      coalesce (xs +[f] ys) = Err \<longrightarrow> 
+      ~(\<exists>zs\<in> list n A. xs <=[r] zs & ys <=[r] zs))"
+apply (induct n)
+ apply simp
+apply clarify
+apply (simp add: in_list_Suc_iff)
+apply clarify
+apply (simp split: err.split_asm add: lem Err.sup_def lift2_def)
+ apply (blast dest: lift2_eq_ErrD)
+done 
+
+lemma closed_err_lift2_conv:
+  "closed (err A) (lift2 f) = (\<forall>x\<in> A. \<forall>y\<in> A. x +_f y : err A)"
+apply (unfold closed_def)
+apply (simp add: err_def)
+done 
+
+lemma closed_map2_list [rule_format]:
+  "closed (err A) (lift2 f) \<Longrightarrow> 
+  \<forall>xs. xs : list n A \<longrightarrow> (\<forall>ys. ys : list n A \<longrightarrow> 
+  map2 f xs ys : list n (err A))"
+apply (unfold map2_def)
+apply (induct n)
+ apply simp
+apply clarify
+apply (simp add: in_list_Suc_iff)
+apply clarify
+apply (simp add: plussub_def closed_err_lift2_conv)
+done
+
+lemma closed_lift2_sup:
+  "closed (err A) (lift2 f) \<Longrightarrow> 
+  closed (err (list n A)) (lift2 (sup f))"
+  by (fastsimp  simp add: closed_def plussub_def sup_def lift2_def
+                          coalesce_in_err_list closed_map2_list
+                split: err.split)
+
+lemma err_semilat_sup:
+  "err_semilat (A,r,f) \<Longrightarrow> 
+  err_semilat (list n A, Listn.le r, sup f)"
+apply (unfold Err.sl_def)
+apply (simp only: split_conv)
+apply (simp (no_asm) only: semilat_Def plussub_def)
+apply (simp (no_asm_simp) only: Semilat.closedI [OF Semilat.intro] closed_lift2_sup)
+apply (rule conjI)
+ apply (drule Semilat.orderI [OF Semilat.intro])
+ apply simp
+apply (simp (no_asm) only: unfold_lesub_err Err.le_def err_def sup_def lift2_def)
+apply (simp (no_asm_simp) add: coalesce_eq_OK1_D coalesce_eq_OK2_D split: err.split)
+apply (blast intro: coalesce_eq_OK_ub_D dest: coalesce_eq_Err_D)
+done 
+
+lemma err_semilat_upto_esl:
+  "\<And>L. err_semilat L \<Longrightarrow> err_semilat(upto_esl m L)"
+apply (unfold Listn.upto_esl_def)
+apply (simp (no_asm_simp) only: split_tupled_all)
+apply simp
+apply (fastsimp intro!: err_semilat_UnionI err_semilat_sup
+                dest: lesub_list_impl_same_size 
+                simp add: plussub_def Listn.sup_def)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/MicroJava/DFA/Opt.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -0,0 +1,292 @@
+(*  Title:      HOL/MicroJava/BV/Opt.thy
+    Author:     Tobias Nipkow
+    Copyright   2000 TUM
+*)
+
+header {* \isaheader{More about Options} *}
+
+theory Opt
+imports Err
+begin
+
+constdefs
+ le :: "'a ord \<Rightarrow> 'a option ord"
+"le r o1 o2 == case o2 of None \<Rightarrow> o1=None |
+                              Some y \<Rightarrow> (case o1 of None \<Rightarrow> True
+                                                  | Some x \<Rightarrow> x <=_r y)"
+
+ opt :: "'a set \<Rightarrow> 'a option set"
+"opt A == insert None {x . ? y:A. x = Some y}"
+
+ sup :: "'a ebinop \<Rightarrow> 'a option ebinop"
+"sup f o1 o2 ==  
+ case o1 of None \<Rightarrow> OK o2 | Some x \<Rightarrow> (case o2 of None \<Rightarrow> OK o1
+     | Some y \<Rightarrow> (case f x y of Err \<Rightarrow> Err | OK z \<Rightarrow> OK (Some z)))"
+
+ esl :: "'a esl \<Rightarrow> 'a option esl"
+"esl == %(A,r,f). (opt A, le r, sup f)"
+
+lemma unfold_le_opt:
+  "o1 <=_(le r) o2 = 
+  (case o2 of None \<Rightarrow> o1=None | 
+              Some y \<Rightarrow> (case o1 of None \<Rightarrow> True | Some x \<Rightarrow> x <=_r y))"
+apply (unfold lesub_def le_def)
+apply (rule refl)
+done
+
+lemma le_opt_refl:
+  "order r \<Longrightarrow> o1 <=_(le r) o1"
+by (simp add: unfold_le_opt split: option.split)
+
+lemma le_opt_trans [rule_format]:
+  "order r \<Longrightarrow> 
+   o1 <=_(le r) o2 \<longrightarrow> o2 <=_(le r) o3 \<longrightarrow> o1 <=_(le r) o3"
+apply (simp add: unfold_le_opt split: option.split)
+apply (blast intro: order_trans)
+done
+
+lemma le_opt_antisym [rule_format]:
+  "order r \<Longrightarrow> o1 <=_(le r) o2 \<longrightarrow> o2 <=_(le r) o1 \<longrightarrow> o1=o2"
+apply (simp add: unfold_le_opt split: option.split)
+apply (blast intro: order_antisym)
+done
+
+lemma order_le_opt [intro!,simp]:
+  "order r \<Longrightarrow> order(le r)"
+apply (subst Semilat.order_def)
+apply (blast intro: le_opt_refl le_opt_trans le_opt_antisym)
+done 
+
+lemma None_bot [iff]: 
+  "None <=_(le r) ox"
+apply (unfold lesub_def le_def)
+apply (simp split: option.split)
+done 
+
+lemma Some_le [iff]:
+  "(Some x <=_(le r) ox) = (? y. ox = Some y & x <=_r y)"
+apply (unfold lesub_def le_def)
+apply (simp split: option.split)
+done 
+
+lemma le_None [iff]:
+  "(ox <=_(le r) None) = (ox = None)";
+apply (unfold lesub_def le_def)
+apply (simp split: option.split)
+done 
+
+
+lemma OK_None_bot [iff]:
+  "OK None <=_(Err.le (le r)) x"
+  by (simp add: lesub_def Err.le_def le_def split: option.split err.split)
+
+lemma sup_None1 [iff]:
+  "x +_(sup f) None = OK x"
+  by (simp add: plussub_def sup_def split: option.split)
+
+lemma sup_None2 [iff]:
+  "None +_(sup f) x = OK x"
+  by (simp add: plussub_def sup_def split: option.split)
+
+
+lemma None_in_opt [iff]:
+  "None : opt A"
+by (simp add: opt_def)
+
+lemma Some_in_opt [iff]:
+  "(Some x : opt A) = (x:A)"
+apply (unfold opt_def)
+apply auto
+done 
+
+
+lemma semilat_opt [intro, simp]:
+  "\<And>L. err_semilat L \<Longrightarrow> err_semilat (Opt.esl L)"
+proof (unfold Opt.esl_def Err.sl_def, simp add: split_tupled_all)
+  
+  fix A r f
+  assume s: "semilat (err A, Err.le r, lift2 f)"
+ 
+  let ?A0 = "err A"
+  let ?r0 = "Err.le r"
+  let ?f0 = "lift2 f"
+
+  from s
+  obtain
+    ord: "order ?r0" and
+    clo: "closed ?A0 ?f0" and
+    ub1: "\<forall>x\<in>?A0. \<forall>y\<in>?A0. x <=_?r0 x +_?f0 y" and
+    ub2: "\<forall>x\<in>?A0. \<forall>y\<in>?A0. y <=_?r0 x +_?f0 y" and
+    lub: "\<forall>x\<in>?A0. \<forall>y\<in>?A0. \<forall>z\<in>?A0. x <=_?r0 z \<and> y <=_?r0 z \<longrightarrow> x +_?f0 y <=_?r0 z"
+    by (unfold semilat_def) simp
+
+  let ?A = "err (opt A)" 
+  let ?r = "Err.le (Opt.le r)"
+  let ?f = "lift2 (Opt.sup f)"
+
+  from ord
+  have "order ?r"
+    by simp
+
+  moreover
+
+  have "closed ?A ?f"
+  proof (unfold closed_def, intro strip)
+    fix x y    
+    assume x: "x : ?A" 
+    assume y: "y : ?A" 
+
+    { fix a b
+      assume ab: "x = OK a" "y = OK b"
+      
+      with x 
+      have a: "\<And>c. a = Some c \<Longrightarrow> c : A"
+        by (clarsimp simp add: opt_def)
+
+      from ab y
+      have b: "\<And>d. b = Some d \<Longrightarrow> d : A"
+        by (clarsimp simp add: opt_def)
+      
+      { fix c d assume "a = Some c" "b = Some d"
+        with ab x y
+        have "c:A & d:A"
+          by (simp add: err_def opt_def Bex_def)
+        with clo
+        have "f c d : err A"
+          by (simp add: closed_def plussub_def err_def lift2_def)
+        moreover
+        fix z assume "f c d = OK z"
+        ultimately
+        have "z : A" by simp
+      } note f_closed = this    
+
+      have "sup f a b : ?A"
+      proof (cases a)
+        case None
+        thus ?thesis
+          by (simp add: sup_def opt_def) (cases b, simp, simp add: b Bex_def)
+      next
+        case Some
+        thus ?thesis
+          by (auto simp add: sup_def opt_def Bex_def a b f_closed split: err.split option.split)
+      qed
+    }
+
+    thus "x +_?f y : ?A"
+      by (simp add: plussub_def lift2_def split: err.split)
+  qed
+    
+  moreover
+
+  { fix a b c 
+    assume "a \<in> opt A" "b \<in> opt A" "a +_(sup f) b = OK c" 
+    moreover
+    from ord have "order r" by simp
+    moreover
+    { fix x y z
+      assume "x \<in> A" "y \<in> A" 
+      hence "OK x \<in> err A \<and> OK y \<in> err A" by simp
+      with ub1 ub2
+      have "(OK x) <=_(Err.le r) (OK x) +_(lift2 f) (OK y) \<and>
+            (OK y) <=_(Err.le r) (OK x) +_(lift2 f) (OK y)"
+        by blast
+      moreover
+      assume "x +_f y = OK z"
+      ultimately
+      have "x <=_r z \<and> y <=_r z"
+        by (auto simp add: plussub_def lift2_def Err.le_def lesub_def)
+    }
+    ultimately
+    have "a <=_(le r) c \<and> b <=_(le r) c"
+      by (auto simp add: sup_def le_def lesub_def plussub_def 
+               dest: order_refl split: option.splits err.splits)
+  }
+     
+  hence "(\<forall>x\<in>?A. \<forall>y\<in>?A. x <=_?r x +_?f y) \<and> (\<forall>x\<in>?A. \<forall>y\<in>?A. y <=_?r x +_?f y)"
+    by (auto simp add: lesub_def plussub_def Err.le_def lift2_def split: err.split)
+
+  moreover
+
+  have "\<forall>x\<in>?A. \<forall>y\<in>?A. \<forall>z\<in>?A. x <=_?r z \<and> y <=_?r z \<longrightarrow> x +_?f y <=_?r z"
+  proof (intro strip, elim conjE)
+    fix x y z
+    assume xyz: "x : ?A" "y : ?A" "z : ?A"
+    assume xz: "x <=_?r z"
+    assume yz: "y <=_?r z"
+
+    { fix a b c
+      assume ok: "x = OK a" "y = OK b" "z = OK c"
+
+      { fix d e g
+        assume some: "a = Some d" "b = Some e" "c = Some g"
+        
+        with ok xyz
+        obtain "OK d:err A" "OK e:err A" "OK g:err A"
+          by simp
+        with lub
+        have "\<lbrakk> (OK d) <=_(Err.le r) (OK g); (OK e) <=_(Err.le r) (OK g) \<rbrakk>
+          \<Longrightarrow> (OK d) +_(lift2 f) (OK e) <=_(Err.le r) (OK g)"
+          by blast
+        hence "\<lbrakk> d <=_r g; e <=_r g \<rbrakk> \<Longrightarrow> \<exists>y. d +_f e = OK y \<and> y <=_r g"
+          by simp
+
+        with ok some xyz xz yz
+        have "x +_?f y <=_?r z"
+          by (auto simp add: sup_def le_def lesub_def lift2_def plussub_def Err.le_def)
+      } note this [intro!]
+
+      from ok xyz xz yz
+      have "x +_?f y <=_?r z"
+        by - (cases a, simp, cases b, simp, cases c, simp, blast)
+    }
+    
+    with xyz xz yz
+    show "x +_?f y <=_?r z"
+      by - (cases x, simp, cases y, simp, cases z, simp+)
+  qed
+
+  ultimately
+
+  show "semilat (?A,?r,?f)"
+    by (unfold semilat_def) simp
+qed 
+
+lemma top_le_opt_Some [iff]: 
+  "top (le r) (Some T) = top r T"
+apply (unfold top_def)
+apply (rule iffI)
+ apply blast
+apply (rule allI)
+apply (case_tac "x")
+apply simp+
+done 
+
+lemma Top_le_conv:
+  "\<lbrakk> order r; top r T \<rbrakk> \<Longrightarrow> (T <=_r x) = (x = T)"
+apply (unfold top_def)
+apply (blast intro: order_antisym)
+done 
+
+
+lemma acc_le_optI [intro!]:
+  "acc r \<Longrightarrow> acc(le r)"
+apply (unfold acc_def lesub_def le_def lesssub_def)
+apply (simp add: wf_eq_minimal split: option.split)
+apply clarify
+apply (case_tac "? a. Some a : Q")
+ apply (erule_tac x = "{a . Some a : Q}" in allE)
+ apply blast
+apply (case_tac "x")
+ apply blast
+apply blast
+done 
+
+lemma option_map_in_optionI:
+  "\<lbrakk> ox : opt S; !x:S. ox = Some x \<longrightarrow> f x : S \<rbrakk> 
+  \<Longrightarrow> Option.map f ox : opt S";
+apply (unfold Option.map_def)
+apply (simp split: option.split)
+apply blast
+done 
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/MicroJava/DFA/Product.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -0,0 +1,141 @@
+(*  Title:      HOL/MicroJava/BV/Product.thy
+    Author:     Tobias Nipkow
+    Copyright   2000 TUM
+*)
+
+header {* \isaheader{Products as Semilattices} *}
+
+theory Product
+imports Err
+begin
+
+constdefs
+ le :: "'a ord \<Rightarrow> 'b ord \<Rightarrow> ('a * 'b) ord"
+"le rA rB == %(a,b) (a',b'). a <=_rA a' & b <=_rB b'"
+
+ sup :: "'a ebinop \<Rightarrow> 'b ebinop \<Rightarrow> ('a * 'b)ebinop"
+"sup f g == %(a1,b1)(a2,b2). Err.sup Pair (a1 +_f a2) (b1 +_g b2)"
+
+ esl :: "'a esl \<Rightarrow> 'b esl \<Rightarrow> ('a * 'b ) esl"
+"esl == %(A,rA,fA) (B,rB,fB). (A <*> B, le rA rB, sup fA fB)"
+
+syntax "@lesubprod" :: "'a*'b \<Rightarrow> 'a ord \<Rightarrow> 'b ord \<Rightarrow> 'b \<Rightarrow> bool"
+       ("(_ /<='(_,_') _)" [50, 0, 0, 51] 50)
+translations "p <=(rA,rB) q" == "p <=_(Product.le rA rB) q"
+
+lemma unfold_lesub_prod:
+  "p <=(rA,rB) q == le rA rB p q"
+  by (simp add: lesub_def)
+
+lemma le_prod_Pair_conv [iff]:
+  "((a1,b1) <=(rA,rB) (a2,b2)) = (a1 <=_rA a2 & b1 <=_rB b2)"
+  by (simp add: lesub_def le_def)
+
+lemma less_prod_Pair_conv:
+  "((a1,b1) <_(Product.le rA rB) (a2,b2)) = 
+  (a1 <_rA a2 & b1 <=_rB b2 | a1 <=_rA a2 & b1 <_rB b2)"
+apply (unfold lesssub_def)
+apply simp
+apply blast
+done
+
+lemma order_le_prod [iff]:
+  "order(Product.le rA rB) = (order rA & order rB)"
+apply (unfold Semilat.order_def)
+apply simp
+apply blast
+done 
+
+lemma acc_le_prodI [intro!]:
+  "\<lbrakk> acc r\<^isub>A; acc r\<^isub>B \<rbrakk> \<Longrightarrow> acc(Product.le r\<^isub>A r\<^isub>B)"
+apply (unfold acc_def)
+apply (rule wf_subset)
+ apply (erule wf_lex_prod)
+ apply assumption
+apply (auto simp add: lesssub_def less_prod_Pair_conv lex_prod_def)
+done
+
+lemma closed_lift2_sup:
+  "\<lbrakk> closed (err A) (lift2 f); closed (err B) (lift2 g) \<rbrakk> \<Longrightarrow> 
+  closed (err(A<*>B)) (lift2(sup f g))";
+apply (unfold closed_def plussub_def lift2_def err_def sup_def)
+apply (simp split: err.split)
+apply blast
+done 
+
+lemma unfold_plussub_lift2:
+  "e1 +_(lift2 f) e2 == lift2 f e1 e2"
+  by (simp add: plussub_def)
+
+
+lemma plus_eq_Err_conv [simp]:
+  assumes "x:A" and "y:A"
+    and "semilat(err A, Err.le r, lift2 f)"
+  shows "(x +_f y = Err) = (~(? z:A. x <=_r z & y <=_r z))"
+proof -
+  have plus_le_conv2:
+    "\<And>r f z. \<lbrakk> z : err A; semilat (err A, r, f); OK x : err A; OK y : err A;
+                 OK x +_f OK y <=_r z\<rbrakk> \<Longrightarrow> OK x <=_r z \<and> OK y <=_r z"
+    by (rule Semilat.plus_le_conv [OF Semilat.intro, THEN iffD1])
+  from prems show ?thesis
+  apply (rule_tac iffI)
+   apply clarify
+   apply (drule OK_le_err_OK [THEN iffD2])
+   apply (drule OK_le_err_OK [THEN iffD2])
+   apply (drule Semilat.lub [OF Semilat.intro, of _ _ _ "OK x" _ "OK y"])
+        apply assumption
+       apply assumption
+      apply simp
+     apply simp
+    apply simp
+   apply simp
+  apply (case_tac "x +_f y")
+   apply assumption
+  apply (rename_tac "z")
+  apply (subgoal_tac "OK z: err A")
+  apply (frule plus_le_conv2)
+       apply assumption
+      apply simp
+      apply blast
+     apply simp
+    apply (blast dest: Semilat.orderI [OF Semilat.intro] order_refl)
+   apply blast
+  apply (erule subst)
+  apply (unfold semilat_def err_def closed_def)
+  apply simp
+  done
+qed
+
+lemma err_semilat_Product_esl:
+  "\<And>L1 L2. \<lbrakk> err_semilat L1; err_semilat L2 \<rbrakk> \<Longrightarrow> err_semilat(Product.esl L1 L2)"
+apply (unfold esl_def Err.sl_def)
+apply (simp (no_asm_simp) only: split_tupled_all)
+apply simp
+apply (simp (no_asm) only: semilat_Def)
+apply (simp (no_asm_simp) only: Semilat.closedI [OF Semilat.intro] closed_lift2_sup)
+apply (simp (no_asm) only: unfold_lesub_err Err.le_def unfold_plussub_lift2 sup_def)
+apply (auto elim: semilat_le_err_OK1 semilat_le_err_OK2
+            simp add: lift2_def  split: err.split)
+apply (blast dest: Semilat.orderI [OF Semilat.intro])
+apply (blast dest: Semilat.orderI [OF Semilat.intro])
+
+apply (rule OK_le_err_OK [THEN iffD1])
+apply (erule subst, subst OK_lift2_OK [symmetric], rule Semilat.lub [OF Semilat.intro])
+apply simp
+apply simp
+apply simp
+apply simp
+apply simp
+apply simp
+
+apply (rule OK_le_err_OK [THEN iffD1])
+apply (erule subst, subst OK_lift2_OK [symmetric], rule Semilat.lub [OF Semilat.intro])
+apply simp
+apply simp
+apply simp
+apply simp
+apply simp
+apply simp
+done 
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/MicroJava/DFA/Semilat.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -0,0 +1,374 @@
+(*  Title:      HOL/MicroJava/BV/Semilat.thy
+    Author:     Tobias Nipkow
+    Copyright   2000 TUM
+*)
+
+header {* 
+  \chapter{Bytecode Verifier}\label{cha:bv}
+  \isaheader{Semilattices} 
+*}
+
+theory Semilat
+imports Main While_Combinator
+begin
+
+types 
+  'a ord    = "'a \<Rightarrow> 'a \<Rightarrow> bool"
+  'a binop  = "'a \<Rightarrow> 'a \<Rightarrow> 'a"
+  'a sl     = "'a set \<times> 'a ord \<times> 'a binop"
+
+consts
+  "lesub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool"
+  "lesssub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool"
+  "plussub" :: "'a \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'c" 
+(*<*)
+syntax 
+  "lesub"   :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /<='__ _)" [50, 1000, 51] 50)
+  "lesssub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /<'__ _)" [50, 1000, 51] 50)
+  "plussub" :: "'a \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'c" ("(_ /+'__ _)" [65, 1000, 66] 65)
+(*>*)
+syntax (xsymbols)
+  "lesub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<sqsubseteq>\<^bsub>_\<^esub> _)" [50, 0, 51] 50)
+  "lesssub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<sqsubset>\<^bsub>_\<^esub> _)" [50, 0, 51] 50)
+  "plussub" :: "'a \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'c" ("(_ /\<squnion>\<^bsub>_\<^esub> _)" [65, 0, 66] 65)
+(*<*)
+ (* allow \<sub> instead of \<bsub>..\<esub> *)  
+  "@lesub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<sqsubseteq>\<^sub>_ _)" [50, 1000, 51] 50)
+  "@lesssub" :: "'a \<Rightarrow> 'a ord \<Rightarrow> 'a \<Rightarrow> bool" ("(_ /\<sqsubset>\<^sub>_ _)" [50, 1000, 51] 50)
+  "@plussub" :: "'a \<Rightarrow> ('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'b \<Rightarrow> 'c" ("(_ /\<squnion>\<^sub>_ _)" [65, 1000, 66] 65)
+
+translations
+  "x \<sqsubseteq>\<^sub>r y" => "x \<sqsubseteq>\<^bsub>r\<^esub> y"
+  "x \<sqsubset>\<^sub>r y" => "x \<sqsubset>\<^bsub>r\<^esub> y" 
+  "x \<squnion>\<^sub>f y" => "x \<squnion>\<^bsub>f\<^esub> y" 
+(*>*)
+
+defs
+  lesub_def:   "x \<sqsubseteq>\<^sub>r y \<equiv> r x y"
+  lesssub_def: "x \<sqsubset>\<^sub>r y \<equiv> x \<sqsubseteq>\<^sub>r y \<and> x \<noteq> y"
+  plussub_def: "x \<squnion>\<^sub>f y \<equiv> f x y"
+
+constdefs
+  ord :: "('a \<times> 'a) set \<Rightarrow> 'a ord"
+  "ord r \<equiv> \<lambda>x y. (x,y) \<in> r"
+
+  order :: "'a ord \<Rightarrow> bool"
+  "order r \<equiv> (\<forall>x. x \<sqsubseteq>\<^sub>r x) \<and> (\<forall>x y. x \<sqsubseteq>\<^sub>r y \<and> y \<sqsubseteq>\<^sub>r x \<longrightarrow> x=y) \<and> (\<forall>x y z. x \<sqsubseteq>\<^sub>r y \<and> y \<sqsubseteq>\<^sub>r z \<longrightarrow> x \<sqsubseteq>\<^sub>r z)"
+
+  top :: "'a ord \<Rightarrow> 'a \<Rightarrow> bool"
+  "top r T \<equiv> \<forall>x. x \<sqsubseteq>\<^sub>r T"
+  
+  acc :: "'a ord \<Rightarrow> bool"
+  "acc r \<equiv> wf {(y,x). x \<sqsubset>\<^sub>r y}"
+
+  closed :: "'a set \<Rightarrow> 'a binop \<Rightarrow> bool"
+  "closed A f \<equiv> \<forall>x\<in>A. \<forall>y\<in>A. x \<squnion>\<^sub>f y \<in> A"
+
+  semilat :: "'a sl \<Rightarrow> bool"
+  "semilat \<equiv> \<lambda>(A,r,f). order r \<and> closed A f \<and> 
+                       (\<forall>x\<in>A. \<forall>y\<in>A. x \<sqsubseteq>\<^sub>r x \<squnion>\<^sub>f y) \<and>
+                       (\<forall>x\<in>A. \<forall>y\<in>A. y \<sqsubseteq>\<^sub>r x \<squnion>\<^sub>f y) \<and>
+                       (\<forall>x\<in>A. \<forall>y\<in>A. \<forall>z\<in>A. x \<sqsubseteq>\<^sub>r z \<and> y \<sqsubseteq>\<^sub>r z \<longrightarrow> x \<squnion>\<^sub>f y \<sqsubseteq>\<^sub>r z)"
+
+
+  is_ub :: "('a \<times> 'a) set \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
+  "is_ub r x y u \<equiv> (x,u)\<in>r \<and> (y,u)\<in>r"
+
+  is_lub :: "('a \<times> 'a) set \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> bool"
+  "is_lub r x y u \<equiv> is_ub r x y u \<and> (\<forall>z. is_ub r x y z \<longrightarrow> (u,z)\<in>r)"
+
+  some_lub :: "('a \<times> 'a) set \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> 'a"
+  "some_lub r x y \<equiv> SOME z. is_lub r x y z"
+
+locale Semilat =
+  fixes A :: "'a set"
+  fixes r :: "'a ord"
+  fixes f :: "'a binop"
+  assumes semilat: "semilat (A, r, f)"
+
+lemma order_refl [simp, intro]: "order r \<Longrightarrow> x \<sqsubseteq>\<^sub>r x"
+  (*<*) by (unfold order_def) (simp (no_asm_simp)) (*>*)
+
+lemma order_antisym: "\<lbrakk> order r; x \<sqsubseteq>\<^sub>r y; y \<sqsubseteq>\<^sub>r x \<rbrakk> \<Longrightarrow> x = y"
+  (*<*) by (unfold order_def) (simp (no_asm_simp)) (*>*)
+
+lemma order_trans: "\<lbrakk> order r; x \<sqsubseteq>\<^sub>r y; y \<sqsubseteq>\<^sub>r z \<rbrakk> \<Longrightarrow> x \<sqsubseteq>\<^sub>r z"
+  (*<*) by (unfold order_def) blast (*>*)
+
+lemma order_less_irrefl [intro, simp]: "order r \<Longrightarrow> \<not> x \<sqsubset>\<^sub>r x"
+  (*<*) by (unfold order_def lesssub_def) blast (*>*)
+
+lemma order_less_trans: "\<lbrakk> order r; x \<sqsubset>\<^sub>r y; y \<sqsubset>\<^sub>r z \<rbrakk> \<Longrightarrow> x \<sqsubset>\<^sub>r z"
+  (*<*) by (unfold order_def lesssub_def) blast (*>*)
+
+lemma topD [simp, intro]: "top r T \<Longrightarrow> x \<sqsubseteq>\<^sub>r T"
+  (*<*) by (simp add: top_def) (*>*)
+
+lemma top_le_conv [simp]: "\<lbrakk> order r; top r T \<rbrakk> \<Longrightarrow> (T \<sqsubseteq>\<^sub>r x) = (x = T)"
+  (*<*) by (blast intro: order_antisym) (*>*)
+
+lemma semilat_Def:
+"semilat(A,r,f) \<equiv> order r \<and> closed A f \<and> 
+                 (\<forall>x\<in>A. \<forall>y\<in>A. x \<sqsubseteq>\<^sub>r x \<squnion>\<^sub>f y) \<and> 
+                 (\<forall>x\<in>A. \<forall>y\<in>A. y \<sqsubseteq>\<^sub>r x \<squnion>\<^sub>f y) \<and> 
+                 (\<forall>x\<in>A. \<forall>y\<in>A. \<forall>z\<in>A. x \<sqsubseteq>\<^sub>r z \<and> y \<sqsubseteq>\<^sub>r z \<longrightarrow> x \<squnion>\<^sub>f y \<sqsubseteq>\<^sub>r z)"
+  (*<*) by (unfold semilat_def) clarsimp (*>*)
+
+lemma (in Semilat) orderI [simp, intro]: "order r"
+  (*<*) using semilat by (simp add: semilat_Def) (*>*)
+
+lemma (in Semilat) closedI [simp, intro]: "closed A f"
+  (*<*) using semilat by (simp add: semilat_Def) (*>*)
+
+lemma closedD: "\<lbrakk> closed A f; x\<in>A; y\<in>A \<rbrakk> \<Longrightarrow> x \<squnion>\<^sub>f y \<in> A"
+  (*<*) by (unfold closed_def) blast (*>*)
+
+lemma closed_UNIV [simp]: "closed UNIV f"
+  (*<*) by (simp add: closed_def) (*>*)
+
+lemma (in Semilat) closed_f [simp, intro]: "\<lbrakk>x \<in> A; y \<in> A\<rbrakk>  \<Longrightarrow> x \<squnion>\<^sub>f y \<in> A"
+  (*<*) by (simp add: closedD [OF closedI]) (*>*)
+
+lemma (in Semilat) refl_r [intro, simp]: "x \<sqsubseteq>\<^sub>r x" by simp
+
+lemma (in Semilat) antisym_r [intro?]: "\<lbrakk> x \<sqsubseteq>\<^sub>r y; y \<sqsubseteq>\<^sub>r x \<rbrakk> \<Longrightarrow> x = y"
+  (*<*) by (rule order_antisym) auto (*>*)
+  
+lemma (in Semilat) trans_r [trans, intro?]: "\<lbrakk>x \<sqsubseteq>\<^sub>r y; y \<sqsubseteq>\<^sub>r z\<rbrakk> \<Longrightarrow> x \<sqsubseteq>\<^sub>r z"
+  (*<*) by (auto intro: order_trans) (*>*)
+  
+lemma (in Semilat) ub1 [simp, intro?]: "\<lbrakk> x \<in> A; y \<in> A \<rbrakk> \<Longrightarrow> x \<sqsubseteq>\<^sub>r x \<squnion>\<^sub>f y"
+  (*<*) by (insert semilat) (unfold semilat_Def, simp) (*>*)
+
+lemma (in Semilat) ub2 [simp, intro?]: "\<lbrakk> x \<in> A; y \<in> A \<rbrakk> \<Longrightarrow> y \<sqsubseteq>\<^sub>r x \<squnion>\<^sub>f y"
+  (*<*) by (insert semilat) (unfold semilat_Def, simp) (*>*)
+
+lemma (in Semilat) lub [simp, intro?]:
+  "\<lbrakk> x \<sqsubseteq>\<^sub>r z; y \<sqsubseteq>\<^sub>r z; x \<in> A; y \<in> A; z \<in> A \<rbrakk> \<Longrightarrow> x \<squnion>\<^sub>f y \<sqsubseteq>\<^sub>r z";
+  (*<*) by (insert semilat) (unfold semilat_Def, simp) (*>*)
+
+lemma (in Semilat) plus_le_conv [simp]:
+  "\<lbrakk> x \<in> A; y \<in> A; z \<in> A \<rbrakk> \<Longrightarrow> (x \<squnion>\<^sub>f y \<sqsubseteq>\<^sub>r z) = (x \<sqsubseteq>\<^sub>r z \<and> y \<sqsubseteq>\<^sub>r z)"
+  (*<*) by (blast intro: ub1 ub2 lub order_trans) (*>*)
+
+lemma (in Semilat) le_iff_plus_unchanged: "\<lbrakk> x \<in> A; y \<in> A \<rbrakk> \<Longrightarrow> (x \<sqsubseteq>\<^sub>r y) = (x \<squnion>\<^sub>f y = y)"
+(*<*)
+apply (rule iffI)
+ apply (blast intro: antisym_r refl_r lub ub2)
+apply (erule subst)
+apply simp
+done
+(*>*)
+
+lemma (in Semilat) le_iff_plus_unchanged2: "\<lbrakk> x \<in> A; y \<in> A \<rbrakk> \<Longrightarrow> (x \<sqsubseteq>\<^sub>r y) = (y \<squnion>\<^sub>f x = y)"
+(*<*)
+apply (rule iffI)
+ apply (blast intro: order_antisym lub order_refl ub1)
+apply (erule subst)
+apply simp
+done 
+(*>*)
+
+
+lemma (in Semilat) plus_assoc [simp]:
+  assumes a: "a \<in> A" and b: "b \<in> A" and c: "c \<in> A"
+  shows "a \<squnion>\<^sub>f (b \<squnion>\<^sub>f c) = a \<squnion>\<^sub>f b \<squnion>\<^sub>f c"
+(*<*)
+proof -
+  from a b have ab: "a \<squnion>\<^sub>f b \<in> A" ..
+  from this c have abc: "(a \<squnion>\<^sub>f b) \<squnion>\<^sub>f c \<in> A" ..
+  from b c have bc: "b \<squnion>\<^sub>f c \<in> A" ..
+  from a this have abc': "a \<squnion>\<^sub>f (b \<squnion>\<^sub>f c) \<in> A" ..
+
+  show ?thesis
+  proof    
+    show "a \<squnion>\<^sub>f (b \<squnion>\<^sub>f c) \<sqsubseteq>\<^sub>r (a \<squnion>\<^sub>f b) \<squnion>\<^sub>f c"
+    proof -
+      from a b have "a \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f b" .. 
+      also from ab c have "\<dots> \<sqsubseteq>\<^sub>r \<dots> \<squnion>\<^sub>f c" ..
+      finally have "a<": "a \<sqsubseteq>\<^sub>r (a \<squnion>\<^sub>f b) \<squnion>\<^sub>f c" .
+      from a b have "b \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f b" ..
+      also from ab c have "\<dots> \<sqsubseteq>\<^sub>r \<dots> \<squnion>\<^sub>f c" ..
+      finally have "b<": "b \<sqsubseteq>\<^sub>r (a \<squnion>\<^sub>f b) \<squnion>\<^sub>f c" .
+      from ab c have "c<": "c \<sqsubseteq>\<^sub>r (a \<squnion>\<^sub>f b) \<squnion>\<^sub>f c" ..    
+      from "b<" "c<" b c abc have "b \<squnion>\<^sub>f c \<sqsubseteq>\<^sub>r (a \<squnion>\<^sub>f b) \<squnion>\<^sub>f c" ..
+      from "a<" this a bc abc show ?thesis ..
+    qed
+    show "(a \<squnion>\<^sub>f b) \<squnion>\<^sub>f c \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f (b \<squnion>\<^sub>f c)" 
+    proof -
+      from b c have "b \<sqsubseteq>\<^sub>r b \<squnion>\<^sub>f c" .. 
+      also from a bc have "\<dots> \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f \<dots>" ..
+      finally have "b<": "b \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f (b \<squnion>\<^sub>f c)" .
+      from b c have "c \<sqsubseteq>\<^sub>r b \<squnion>\<^sub>f c" ..
+      also from a bc have "\<dots> \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f \<dots>" ..
+      finally have "c<": "c \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f (b \<squnion>\<^sub>f c)" .
+      from a bc have "a<": "a \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f (b \<squnion>\<^sub>f c)" ..
+      from "a<" "b<" a b abc' have "a \<squnion>\<^sub>f b \<sqsubseteq>\<^sub>r a \<squnion>\<^sub>f (b \<squnion>\<^sub>f c)" ..
+      from this "c<" ab c abc' show ?thesis ..
+    qed
+  qed
+qed
+(*>*)
+
+lemma (in Semilat) plus_com_lemma:
+  "\<lbrakk>a \<in> A; b \<in> A\<rbrakk> \<Longrightarrow> a \<squnion>\<^sub>f b \<sqsubseteq>\<^sub>r b \<squnion>\<^sub>f a"
+(*<*)
+proof -
+  assume a: "a \<in> A" and b: "b \<in> A"  
+  from b a have "a \<sqsubseteq>\<^sub>r b \<squnion>\<^sub>f a" .. 
+  moreover from b a have "b \<sqsubseteq>\<^sub>r b \<squnion>\<^sub>f a" ..
+  moreover note a b
+  moreover from b a have "b \<squnion>\<^sub>f a \<in> A" ..
+  ultimately show ?thesis ..
+qed
+(*>*)
+
+lemma (in Semilat) plus_commutative:
+  "\<lbrakk>a \<in> A; b \<in> A\<rbrakk> \<Longrightarrow> a \<squnion>\<^sub>f b = b \<squnion>\<^sub>f a"
+  (*<*) by(blast intro: order_antisym plus_com_lemma) (*>*)
+
+lemma is_lubD:
+  "is_lub r x y u \<Longrightarrow> is_ub r x y u \<and> (\<forall>z. is_ub r x y z \<longrightarrow> (u,z) \<in> r)"
+  (*<*) by (simp add: is_lub_def) (*>*)
+
+lemma is_ubI:
+  "\<lbrakk> (x,u) \<in> r; (y,u) \<in> r \<rbrakk> \<Longrightarrow> is_ub r x y u"
+  (*<*) by (simp add: is_ub_def) (*>*)
+
+lemma is_ubD:
+  "is_ub r x y u \<Longrightarrow> (x,u) \<in> r \<and> (y,u) \<in> r"
+  (*<*) by (simp add: is_ub_def) (*>*)
+
+
+lemma is_lub_bigger1 [iff]:  
+  "is_lub (r^* ) x y y = ((x,y)\<in>r^* )"
+(*<*)
+apply (unfold is_lub_def is_ub_def)
+apply blast
+done
+(*>*)
+
+lemma is_lub_bigger2 [iff]:
+  "is_lub (r^* ) x y x = ((y,x)\<in>r^* )"
+(*<*)
+apply (unfold is_lub_def is_ub_def)
+apply blast 
+done
+(*>*)
+
+lemma extend_lub:
+  "\<lbrakk> single_valued r; is_lub (r^* ) x y u; (x',x) \<in> r \<rbrakk> 
+  \<Longrightarrow> EX v. is_lub (r^* ) x' y v"
+(*<*)
+apply (unfold is_lub_def is_ub_def)
+apply (case_tac "(y,x) \<in> r^*")
+ apply (case_tac "(y,x') \<in> r^*")
+  apply blast
+ apply (blast elim: converse_rtranclE dest: single_valuedD)
+apply (rule exI)
+apply (rule conjI)
+ apply (blast intro: converse_rtrancl_into_rtrancl dest: single_valuedD)
+apply (blast intro: rtrancl_into_rtrancl converse_rtrancl_into_rtrancl 
+             elim: converse_rtranclE dest: single_valuedD)
+done
+(*>*)
+
+lemma single_valued_has_lubs [rule_format]:
+  "\<lbrakk> single_valued r; (x,u) \<in> r^* \<rbrakk> \<Longrightarrow> (\<forall>y. (y,u) \<in> r^* \<longrightarrow> 
+  (EX z. is_lub (r^* ) x y z))"
+(*<*)
+apply (erule converse_rtrancl_induct)
+ apply clarify
+ apply (erule converse_rtrancl_induct)
+  apply blast
+ apply (blast intro: converse_rtrancl_into_rtrancl)
+apply (blast intro: extend_lub)
+done
+(*>*)
+
+lemma some_lub_conv:
+  "\<lbrakk> acyclic r; is_lub (r^* ) x y u \<rbrakk> \<Longrightarrow> some_lub (r^* ) x y = u"
+(*<*)
+apply (unfold some_lub_def is_lub_def)
+apply (rule someI2)
+ apply assumption
+apply (blast intro: antisymD dest!: acyclic_impl_antisym_rtrancl)
+done
+(*>*)
+
+lemma is_lub_some_lub:
+  "\<lbrakk> single_valued r; acyclic r; (x,u)\<in>r^*; (y,u)\<in>r^* \<rbrakk> 
+  \<Longrightarrow> is_lub (r^* ) x y (some_lub (r^* ) x y)";
+  (*<*) by (fastsimp dest: single_valued_has_lubs simp add: some_lub_conv) (*>*)
+
+subsection{*An executable lub-finder*}
+
+constdefs
+ exec_lub :: "('a * 'a) set \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> 'a binop"
+"exec_lub r f x y \<equiv> while (\<lambda>z. (x,z) \<notin> r\<^sup>*) f y"
+
+lemma exec_lub_refl: "exec_lub r f T T = T"
+by (simp add: exec_lub_def while_unfold)
+
+lemma acyclic_single_valued_finite:
+ "\<lbrakk>acyclic r; single_valued r; (x,y) \<in> r\<^sup>*\<rbrakk>
+  \<Longrightarrow> finite (r \<inter> {a. (x, a) \<in> r\<^sup>*} \<times> {b. (b, y) \<in> r\<^sup>*})"
+(*<*)
+apply(erule converse_rtrancl_induct)
+ apply(rule_tac B = "{}" in finite_subset)
+  apply(simp only:acyclic_def)
+  apply(blast intro:rtrancl_into_trancl2 rtrancl_trancl_trancl)
+ apply simp
+apply(rename_tac x x')
+apply(subgoal_tac "r \<inter> {a. (x,a) \<in> r\<^sup>*} \<times> {b. (b,y) \<in> r\<^sup>*} =
+                   insert (x,x') (r \<inter> {a. (x', a) \<in> r\<^sup>*} \<times> {b. (b, y) \<in> r\<^sup>*})")
+ apply simp
+apply(blast intro:converse_rtrancl_into_rtrancl
+            elim:converse_rtranclE dest:single_valuedD)
+done
+(*>*)
+
+
+lemma exec_lub_conv:
+  "\<lbrakk> acyclic r; \<forall>x y. (x,y) \<in> r \<longrightarrow> f x = y; is_lub (r\<^sup>*) x y u \<rbrakk> \<Longrightarrow>
+  exec_lub r f x y = u";
+(*<*)
+apply(unfold exec_lub_def)
+apply(rule_tac P = "\<lambda>z. (y,z) \<in> r\<^sup>* \<and> (z,u) \<in> r\<^sup>*" and
+               r = "(r \<inter> {(a,b). (y,a) \<in> r\<^sup>* \<and> (b,u) \<in> r\<^sup>*})^-1" in while_rule)
+    apply(blast dest: is_lubD is_ubD)
+   apply(erule conjE)
+   apply(erule_tac z = u in converse_rtranclE)
+    apply(blast dest: is_lubD is_ubD)
+   apply(blast dest:rtrancl_into_rtrancl)
+  apply(rename_tac s)
+  apply(subgoal_tac "is_ub (r\<^sup>*) x y s")
+   prefer 2; apply(simp add:is_ub_def)
+  apply(subgoal_tac "(u, s) \<in> r\<^sup>*")
+   prefer 2; apply(blast dest:is_lubD)
+  apply(erule converse_rtranclE)
+   apply blast
+  apply(simp only:acyclic_def)
+  apply(blast intro:rtrancl_into_trancl2 rtrancl_trancl_trancl)
+ apply(rule finite_acyclic_wf)
+  apply simp
+  apply(erule acyclic_single_valued_finite)
+   apply(blast intro:single_valuedI)
+  apply(simp add:is_lub_def is_ub_def)
+ apply simp
+ apply(erule acyclic_subset)
+ apply blast
+apply simp
+apply(erule conjE)
+apply(erule_tac z = u in converse_rtranclE)
+ apply(blast dest: is_lubD is_ubD)
+apply(blast dest:rtrancl_into_rtrancl)
+done
+(*>*)
+
+lemma is_lub_exec_lub:
+  "\<lbrakk> single_valued r; acyclic r; (x,u):r^*; (y,u):r^*; \<forall>x y. (x,y) \<in> r \<longrightarrow> f x = y \<rbrakk>
+  \<Longrightarrow> is_lub (r^* ) x y (exec_lub r f x y)"
+  (*<*) by (fastsimp dest: single_valued_has_lubs simp add: exec_lub_conv) (*>*)
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/MicroJava/DFA/SemilatAlg.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -0,0 +1,186 @@
+(*  Title:      HOL/MicroJava/BV/SemilatAlg.thy
+    Author:     Gerwin Klein
+    Copyright   2002 Technische Universitaet Muenchen
+*)
+
+header {* \isaheader{More on Semilattices} *}
+
+theory SemilatAlg
+imports Typing_Framework Product
+begin
+
+constdefs 
+  lesubstep_type :: "(nat \<times> 's) list \<Rightarrow> 's ord \<Rightarrow> (nat \<times> 's) list \<Rightarrow> bool"
+                    ("(_ /<=|_| _)" [50, 0, 51] 50)
+  "x <=|r| y \<equiv> \<forall>(p,s) \<in> set x. \<exists>s'. (p,s') \<in> set y \<and> s <=_r s'"
+
+consts
+ "@plusplussub" :: "'a list \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a" ("(_ /++'__ _)" [65, 1000, 66] 65)
+primrec
+  "[] ++_f y = y"
+  "(x#xs) ++_f y = xs ++_f (x +_f y)"
+
+constdefs
+ bounded :: "'s step_type \<Rightarrow> nat \<Rightarrow> bool"
+"bounded step n == !p<n. !s. !(q,t):set(step p s). q<n"  
+
+ pres_type :: "'s step_type \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> bool"
+"pres_type step n A == \<forall>s\<in>A. \<forall>p<n. \<forall>(q,s')\<in>set (step p s). s' \<in> A"
+
+ mono :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> bool"
+"mono r step n A ==
+ \<forall>s p t. s \<in> A \<and> p < n \<and> s <=_r t \<longrightarrow> step p s <=|r| step p t"
+
+
+lemma pres_typeD:
+  "\<lbrakk> pres_type step n A; s\<in>A; p<n; (q,s')\<in>set (step p s) \<rbrakk> \<Longrightarrow> s' \<in> A"
+  by (unfold pres_type_def, blast)
+
+lemma monoD:
+  "\<lbrakk> mono r step n A; p < n; s\<in>A; s <=_r t \<rbrakk> \<Longrightarrow> step p s <=|r| step p t"
+  by (unfold mono_def, blast)
+
+lemma boundedD: 
+  "\<lbrakk> bounded step n; p < n; (q,t) : set (step p xs) \<rbrakk> \<Longrightarrow> q < n" 
+  by (unfold bounded_def, blast)
+
+lemma lesubstep_type_refl [simp, intro]:
+  "(\<And>x. x <=_r x) \<Longrightarrow> x <=|r| x"
+  by (unfold lesubstep_type_def) auto
+
+lemma lesub_step_typeD:
+  "a <=|r| b \<Longrightarrow> (x,y) \<in> set a \<Longrightarrow> \<exists>y'. (x, y') \<in> set b \<and> y <=_r y'"
+  by (unfold lesubstep_type_def) blast
+
+
+lemma list_update_le_listI [rule_format]:
+  "set xs <= A \<longrightarrow> set ys <= A \<longrightarrow> xs <=[r] ys \<longrightarrow> p < size xs \<longrightarrow>  
+   x <=_r ys!p \<longrightarrow> semilat(A,r,f) \<longrightarrow> x\<in>A \<longrightarrow> 
+   xs[p := x +_f xs!p] <=[r] ys"
+  apply (unfold Listn.le_def lesub_def semilat_def)
+  apply (simp add: list_all2_conv_all_nth nth_list_update)
+  done
+
+
+lemma plusplus_closed: assumes "semilat (A, r, f)" shows
+  "\<And>y. \<lbrakk> set x \<subseteq> A; y \<in> A\<rbrakk> \<Longrightarrow> x ++_f y \<in> A" (is "PROP ?P")
+proof -
+  interpret Semilat A r f using assms by (rule Semilat.intro)
+  show "PROP ?P" proof (induct x)
+    show "\<And>y. y \<in> A \<Longrightarrow> [] ++_f y \<in> A" by simp
+    fix y x xs
+    assume y: "y \<in> A" and xs: "set (x#xs) \<subseteq> A"
+    assume IH: "\<And>y. \<lbrakk> set xs \<subseteq> A; y \<in> A\<rbrakk> \<Longrightarrow> xs ++_f y \<in> A"
+    from xs obtain x: "x \<in> A" and xs': "set xs \<subseteq> A" by simp
+    from x y have "(x +_f y) \<in> A" ..
+    with xs' have "xs ++_f (x +_f y) \<in> A" by (rule IH)
+    thus "(x#xs) ++_f y \<in> A" by simp
+  qed
+qed
+
+lemma (in Semilat) pp_ub2:
+ "\<And>y. \<lbrakk> set x \<subseteq> A; y \<in> A\<rbrakk> \<Longrightarrow> y <=_r x ++_f y"
+proof (induct x)
+  from semilat show "\<And>y. y <=_r [] ++_f y" by simp
+  
+  fix y a l
+  assume y:  "y \<in> A"
+  assume "set (a#l) \<subseteq> A"
+  then obtain a: "a \<in> A" and x: "set l \<subseteq> A" by simp
+  assume "\<And>y. \<lbrakk>set l \<subseteq> A; y \<in> A\<rbrakk> \<Longrightarrow> y <=_r l ++_f y"
+  hence IH: "\<And>y. y \<in> A \<Longrightarrow> y <=_r l ++_f y" using x .
+
+  from a y have "y <=_r a +_f y" ..
+  also from a y have "a +_f y \<in> A" ..
+  hence "(a +_f y) <=_r l ++_f (a +_f y)" by (rule IH)
+  finally have "y <=_r l ++_f (a +_f y)" .
+  thus "y <=_r (a#l) ++_f y" by simp
+qed
+
+
+lemma (in Semilat) pp_ub1:
+shows "\<And>y. \<lbrakk>set ls \<subseteq> A; y \<in> A; x \<in> set ls\<rbrakk> \<Longrightarrow> x <=_r ls ++_f y"
+proof (induct ls)
+  show "\<And>y. x \<in> set [] \<Longrightarrow> x <=_r [] ++_f y" by simp
+
+  fix y s ls
+  assume "set (s#ls) \<subseteq> A"
+  then obtain s: "s \<in> A" and ls: "set ls \<subseteq> A" by simp
+  assume y: "y \<in> A" 
+
+  assume 
+    "\<And>y. \<lbrakk>set ls \<subseteq> A; y \<in> A; x \<in> set ls\<rbrakk> \<Longrightarrow> x <=_r ls ++_f y"
+  hence IH: "\<And>y. x \<in> set ls \<Longrightarrow> y \<in> A \<Longrightarrow> x <=_r ls ++_f y" using ls .
+
+  assume "x \<in> set (s#ls)"
+  then obtain xls: "x = s \<or> x \<in> set ls" by simp
+  moreover {
+    assume xs: "x = s"
+    from s y have "s <=_r s +_f y" ..
+    also from s y have "s +_f y \<in> A" ..
+    with ls have "(s +_f y) <=_r ls ++_f (s +_f y)" by (rule pp_ub2)
+    finally have "s <=_r ls ++_f (s +_f y)" .
+    with xs have "x <=_r ls ++_f (s +_f y)" by simp
+  } 
+  moreover {
+    assume "x \<in> set ls"
+    hence "\<And>y. y \<in> A \<Longrightarrow> x <=_r ls ++_f y" by (rule IH)
+    moreover from s y have "s +_f y \<in> A" ..
+    ultimately have "x <=_r ls ++_f (s +_f y)" .
+  }
+  ultimately 
+  have "x <=_r ls ++_f (s +_f y)" by blast
+  thus "x <=_r (s#ls) ++_f y" by simp
+qed
+
+
+lemma (in Semilat) pp_lub:
+  assumes z: "z \<in> A"
+  shows 
+  "\<And>y. y \<in> A \<Longrightarrow> set xs \<subseteq> A \<Longrightarrow> \<forall>x \<in> set xs. x <=_r z \<Longrightarrow> y <=_r z \<Longrightarrow> xs ++_f y <=_r z"
+proof (induct xs)
+  fix y assume "y <=_r z" thus "[] ++_f y <=_r z" by simp
+next
+  fix y l ls assume y: "y \<in> A" and "set (l#ls) \<subseteq> A"
+  then obtain l: "l \<in> A" and ls: "set ls \<subseteq> A" by auto
+  assume "\<forall>x \<in> set (l#ls). x <=_r z"
+  then obtain lz: "l <=_r z" and lsz: "\<forall>x \<in> set ls. x <=_r z" by auto
+  assume "y <=_r z" with lz have "l +_f y <=_r z" using l y z ..
+  moreover
+  from l y have "l +_f y \<in> A" ..
+  moreover
+  assume "\<And>y. y \<in> A \<Longrightarrow> set ls \<subseteq> A \<Longrightarrow> \<forall>x \<in> set ls. x <=_r z \<Longrightarrow> y <=_r z
+          \<Longrightarrow> ls ++_f y <=_r z"
+  ultimately
+  have "ls ++_f (l +_f y) <=_r z" using ls lsz by -
+  thus "(l#ls) ++_f y <=_r z" by simp
+qed
+
+
+lemma ub1':
+  assumes "semilat (A, r, f)"
+  shows "\<lbrakk>\<forall>(p,s) \<in> set S. s \<in> A; y \<in> A; (a,b) \<in> set S\<rbrakk> 
+  \<Longrightarrow> b <=_r map snd [(p', t')\<leftarrow>S. p' = a] ++_f y" 
+proof -
+  interpret Semilat A r f using assms by (rule Semilat.intro)
+
+  let "b <=_r ?map ++_f y" = ?thesis
+
+  assume "y \<in> A"
+  moreover
+  assume "\<forall>(p,s) \<in> set S. s \<in> A"
+  hence "set ?map \<subseteq> A" by auto
+  moreover
+  assume "(a,b) \<in> set S"
+  hence "b \<in> set ?map" by (induct S, auto)
+  ultimately
+  show ?thesis by - (rule pp_ub1)
+qed
+    
+
+lemma plusplus_empty:  
+  "\<forall>s'. (q, s') \<in> set S \<longrightarrow> s' +_f ss ! q = ss ! q \<Longrightarrow>
+   (map snd [(p', t') \<leftarrow> S. p' = q] ++_f ss ! q) = ss ! q"
+  by (induct S) auto 
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/MicroJava/DFA/Semilattices.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -0,0 +1,14 @@
+(*  Title:      HOL/MicroJava/BV/Semilat.thy
+    Author:     Gerwin Klein
+    Copyright   2003 TUM
+*)
+
+header {* Semilattices *}
+
+(*<*)
+theory Semilattices
+imports Err Opt Product Listn
+begin
+
+end
+(*>*)
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/MicroJava/DFA/Typing_Framework.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -0,0 +1,36 @@
+(*  Title:      HOL/MicroJava/BV/Typing_Framework.thy
+    Author:     Tobias Nipkow
+    Copyright   2000 TUM
+*)
+
+header {* \isaheader{Typing and Dataflow Analysis Framework} *}
+
+theory Typing_Framework
+imports Listn
+begin
+
+text {* 
+  The relationship between dataflow analysis and a welltyped-instruction predicate. 
+*}
+types
+  's step_type = "nat \<Rightarrow> 's \<Rightarrow> (nat \<times> 's) list"
+
+constdefs
+ stable :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> nat \<Rightarrow> bool"
+"stable r step ss p == !(q,s'):set(step p (ss!p)). s' <=_r ss!q"
+
+ stables :: "'s ord \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool"
+"stables r step ss == !p<size ss. stable r step ss p"
+
+ wt_step ::
+"'s ord \<Rightarrow> 's \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool"
+"wt_step r T step ts ==
+ !p<size(ts). ts!p ~= T & stable r step ts p"
+
+ is_bcv :: "'s ord \<Rightarrow> 's \<Rightarrow> 's step_type 
+           \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> ('s list \<Rightarrow> 's list) \<Rightarrow> bool"  
+"is_bcv r T step n A bcv == !ss : list n A.
+   (!p<n. (bcv ss)!p ~= T) =
+   (? ts: list n A. ss <=[r] ts & wt_step r T step ts)"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/MicroJava/DFA/Typing_Framework_err.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -0,0 +1,253 @@
+(*  Title:      HOL/MicroJava/BV/Typing_Framework_err.thy
+    Author:     Gerwin Klein
+    Copyright   2000 TUM
+*)
+
+header {* \isaheader{Lifting the Typing Framework to err, app, and eff} *}
+
+theory Typing_Framework_err
+imports Typing_Framework SemilatAlg
+begin
+
+constdefs
+
+wt_err_step :: "'s ord \<Rightarrow> 's err step_type \<Rightarrow> 's err list \<Rightarrow> bool"
+"wt_err_step r step ts \<equiv> wt_step (Err.le r) Err step ts"
+
+wt_app_eff :: "'s ord \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> bool) \<Rightarrow> 's step_type \<Rightarrow> 's list \<Rightarrow> bool"
+"wt_app_eff r app step ts \<equiv>
+  \<forall>p < size ts. app p (ts!p) \<and> (\<forall>(q,t) \<in> set (step p (ts!p)). t <=_r ts!q)"
+
+map_snd :: "('b \<Rightarrow> 'c) \<Rightarrow> ('a \<times> 'b) list \<Rightarrow> ('a \<times> 'c) list"
+"map_snd f \<equiv> map (\<lambda>(x,y). (x, f y))"
+
+error :: "nat \<Rightarrow> (nat \<times> 'a err) list"
+"error n \<equiv> map (\<lambda>x. (x,Err)) [0..<n]"
+
+err_step :: "nat \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> bool) \<Rightarrow> 's step_type \<Rightarrow> 's err step_type"
+"err_step n app step p t \<equiv> 
+  case t of 
+    Err   \<Rightarrow> error n
+  | OK t' \<Rightarrow> if app p t' then map_snd OK (step p t') else error n"
+
+app_mono :: "'s ord \<Rightarrow> (nat \<Rightarrow> 's \<Rightarrow> bool) \<Rightarrow> nat \<Rightarrow> 's set \<Rightarrow> bool"
+"app_mono r app n A \<equiv>
+ \<forall>s p t. s \<in> A \<and> p < n \<and> s <=_r t \<longrightarrow> app p t \<longrightarrow> app p s"
+
+
+lemmas err_step_defs = err_step_def map_snd_def error_def
+
+
+lemma bounded_err_stepD:
+  "bounded (err_step n app step) n \<Longrightarrow> 
+  p < n \<Longrightarrow>  app p a \<Longrightarrow> (q,b) \<in> set (step p a) \<Longrightarrow> 
+  q < n"
+  apply (simp add: bounded_def err_step_def)
+  apply (erule allE, erule impE, assumption)
+  apply (erule_tac x = "OK a" in allE, drule bspec)
+   apply (simp add: map_snd_def)
+   apply fast
+  apply simp
+  done
+
+
+lemma in_map_sndD: "(a,b) \<in> set (map_snd f xs) \<Longrightarrow> \<exists>b'. (a,b') \<in> set xs"
+  apply (induct xs)
+  apply (auto simp add: map_snd_def)
+  done
+
+
+lemma bounded_err_stepI:
+  "\<forall>p. p < n \<longrightarrow> (\<forall>s. ap p s \<longrightarrow> (\<forall>(q,s') \<in> set (step p s). q < n))
+  \<Longrightarrow> bounded (err_step n ap step) n"
+apply (clarsimp simp: bounded_def err_step_def split: err.splits)
+apply (simp add: error_def image_def)
+apply (blast dest: in_map_sndD)
+done
+
+
+lemma bounded_lift:
+  "bounded step n \<Longrightarrow> bounded (err_step n app step) n"
+  apply (unfold bounded_def err_step_def error_def)
+  apply clarify
+  apply (erule allE, erule impE, assumption)
+  apply (case_tac s)
+  apply (auto simp add: map_snd_def split: split_if_asm)
+  done
+
+
+lemma le_list_map_OK [simp]:
+  "\<And>b. map OK a <=[Err.le r] map OK b = (a <=[r] b)"
+  apply (induct a)
+   apply simp
+  apply simp
+  apply (case_tac b)
+   apply simp
+  apply simp
+  done
+
+
+lemma map_snd_lessI:
+  "x <=|r| y \<Longrightarrow> map_snd OK x <=|Err.le r| map_snd OK y"
+  apply (induct x)
+  apply (unfold lesubstep_type_def map_snd_def)
+  apply auto
+  done
+
+
+lemma mono_lift:
+  "order r \<Longrightarrow> app_mono r app n A \<Longrightarrow> bounded (err_step n app step) n \<Longrightarrow>
+  \<forall>s p t. s \<in> A \<and> p < n \<and> s <=_r t \<longrightarrow> app p t \<longrightarrow> step p s <=|r| step p t \<Longrightarrow>
+  mono (Err.le r) (err_step n app step) n (err A)"
+apply (unfold app_mono_def mono_def err_step_def)
+apply clarify
+apply (case_tac s)
+ apply simp 
+apply simp
+apply (case_tac t)
+ apply simp
+ apply clarify
+ apply (simp add: lesubstep_type_def error_def)
+ apply clarify
+ apply (drule in_map_sndD)
+ apply clarify
+ apply (drule bounded_err_stepD, assumption+)
+ apply (rule exI [of _ Err])
+ apply simp
+apply simp
+apply (erule allE, erule allE, erule allE, erule impE)
+ apply (rule conjI, assumption)
+ apply (rule conjI, assumption)
+ apply assumption
+apply (rule conjI)
+apply clarify
+apply (erule allE, erule allE, erule allE, erule impE)
+ apply (rule conjI, assumption)
+ apply (rule conjI, assumption)
+ apply assumption
+apply (erule impE, assumption)
+apply (rule map_snd_lessI, assumption)
+apply clarify
+apply (simp add: lesubstep_type_def error_def)
+apply clarify
+apply (drule in_map_sndD)
+apply clarify
+apply (drule bounded_err_stepD, assumption+)
+apply (rule exI [of _ Err])
+apply simp
+done
+ 
+lemma in_errorD:
+  "(x,y) \<in> set (error n) \<Longrightarrow> y = Err"
+  by (auto simp add: error_def)
+
+lemma pres_type_lift:
+  "\<forall>s\<in>A. \<forall>p. p < n \<longrightarrow> app p s \<longrightarrow> (\<forall>(q, s')\<in>set (step p s). s' \<in> A) 
+  \<Longrightarrow> pres_type (err_step n app step) n (err A)"  
+apply (unfold pres_type_def err_step_def)
+apply clarify
+apply (case_tac b)
+ apply simp
+apply (case_tac s)
+ apply simp
+ apply (drule in_errorD)
+ apply simp
+apply (simp add: map_snd_def split: split_if_asm)
+ apply fast
+apply (drule in_errorD)
+apply simp
+done
+
+
+
+text {*
+  There used to be a condition here that each instruction must have a
+  successor. This is not needed any more, because the definition of
+  @{term error} trivially ensures that there is a successor for
+  the critical case where @{term app} does not hold. 
+*}
+lemma wt_err_imp_wt_app_eff:
+  assumes wt: "wt_err_step r (err_step (size ts) app step) ts"
+  assumes b:  "bounded (err_step (size ts) app step) (size ts)"
+  shows "wt_app_eff r app step (map ok_val ts)"
+proof (unfold wt_app_eff_def, intro strip, rule conjI)
+  fix p assume "p < size (map ok_val ts)"
+  hence lp: "p < size ts" by simp
+  hence ts: "0 < size ts" by (cases p) auto
+  hence err: "(0,Err) \<in> set (error (size ts))" by (simp add: error_def)
+
+  from wt lp
+  have [intro?]: "\<And>p. p < size ts \<Longrightarrow> ts ! p \<noteq> Err" 
+    by (unfold wt_err_step_def wt_step_def) simp
+
+  show app: "app p (map ok_val ts ! p)"
+  proof (rule ccontr)
+    from wt lp obtain s where
+      OKp:  "ts ! p = OK s" and
+      less: "\<forall>(q,t) \<in> set (err_step (size ts) app step p (ts!p)). t <=_(Err.le r) ts!q"
+      by (unfold wt_err_step_def wt_step_def stable_def) 
+         (auto iff: not_Err_eq)
+    assume "\<not> app p (map ok_val ts ! p)"
+    with OKp lp have "\<not> app p s" by simp
+    with OKp have "err_step (size ts) app step p (ts!p) = error (size ts)" 
+      by (simp add: err_step_def)    
+    with err ts obtain q where 
+      "(q,Err) \<in> set (err_step (size ts) app step p (ts!p))" and
+      q: "q < size ts" by auto    
+    with less have "ts!q = Err" by auto
+    moreover from q have "ts!q \<noteq> Err" ..
+    ultimately show False by blast
+  qed
+  
+  show "\<forall>(q,t)\<in>set(step p (map ok_val ts ! p)). t <=_r map ok_val ts ! q" 
+  proof clarify
+    fix q t assume q: "(q,t) \<in> set (step p (map ok_val ts ! p))"
+
+    from wt lp q
+    obtain s where
+      OKp:  "ts ! p = OK s" and
+      less: "\<forall>(q,t) \<in> set (err_step (size ts) app step p (ts!p)). t <=_(Err.le r) ts!q"
+      by (unfold wt_err_step_def wt_step_def stable_def) 
+         (auto iff: not_Err_eq)
+
+    from b lp app q have lq: "q < size ts" by (rule bounded_err_stepD)
+    hence "ts!q \<noteq> Err" ..
+    then obtain s' where OKq: "ts ! q = OK s'" by (auto iff: not_Err_eq)
+
+    from lp lq OKp OKq app less q
+    show "t <=_r map ok_val ts ! q"
+      by (auto simp add: err_step_def map_snd_def) 
+  qed
+qed
+
+
+lemma wt_app_eff_imp_wt_err:
+  assumes app_eff: "wt_app_eff r app step ts"
+  assumes bounded: "bounded (err_step (size ts) app step) (size ts)"
+  shows "wt_err_step r (err_step (size ts) app step) (map OK ts)"
+proof (unfold wt_err_step_def wt_step_def, intro strip, rule conjI)
+  fix p assume "p < size (map OK ts)" 
+  hence p: "p < size ts" by simp
+  thus "map OK ts ! p \<noteq> Err" by simp
+  { fix q t
+    assume q: "(q,t) \<in> set (err_step (size ts) app step p (map OK ts ! p))" 
+    with p app_eff obtain 
+      "app p (ts ! p)" "\<forall>(q,t) \<in> set (step p (ts!p)). t <=_r ts!q"
+      by (unfold wt_app_eff_def) blast
+    moreover
+    from q p bounded have "q < size ts"
+      by - (rule boundedD)
+    hence "map OK ts ! q = OK (ts!q)" by simp
+    moreover
+    have "p < size ts" by (rule p)
+    moreover note q
+    ultimately     
+    have "t <=_(Err.le r) map OK ts ! q" 
+      by (auto simp add: err_step_def map_snd_def)
+  }
+  thus "stable (Err.le r) (err_step (size ts) app step) (map OK ts) p"
+    by (unfold stable_def) blast
+qed
+
+end
+
--- a/src/HOL/MicroJava/J/Example.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/J/Example.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -173,19 +173,19 @@
 apply (simp (no_asm))
 done
 
-lemma not_Object_subcls [elim!]: "(subcls1 tprg)^++ Object C ==> R"
-apply (auto dest!: tranclpD subcls1D)
+lemma not_Object_subcls [elim!]: "(Object, C) \<in> (subcls1 tprg)^+ ==> R"
+apply (auto dest!: tranclD subcls1D)
 done
 
 lemma subcls_ObjectD [dest!]: "tprg\<turnstile>Object\<preceq>C C ==> C = Object"
-apply (erule rtranclp_induct)
+apply (erule rtrancl_induct)
 apply  auto
 apply (drule subcls1D)
 apply auto
 done
 
-lemma not_Base_subcls_Ext [elim!]: "(subcls1 tprg)^++ Base Ext ==> R"
-apply (auto dest!: tranclpD subcls1D)
+lemma not_Base_subcls_Ext [elim!]: "(Base, Ext) \<in> (subcls1 tprg)^+  ==> R"
+apply (auto dest!: tranclD subcls1D)
 done
 
 lemma class_tprgD: 
@@ -194,11 +194,11 @@
 apply (auto split add: split_if_asm simp add: map_of_Cons)
 done
 
-lemma not_class_subcls_class [elim!]: "(subcls1 tprg)^++ C C ==> R"
-apply (auto dest!: tranclpD subcls1D)
+lemma not_class_subcls_class [elim!]: "(C, C) \<in> (subcls1 tprg)^+ ==> R"
+apply (auto dest!: tranclD subcls1D)
 apply (frule class_tprgD)
 apply (auto dest!:)
-apply (drule rtranclpD)
+apply (drule rtranclD)
 apply auto
 done
 
@@ -206,7 +206,7 @@
 apply (simp (no_asm) add: ObjectC_def BaseC_def ExtC_def NullPointerC_def ClassCastC_def OutOfMemoryC_def)
 done
 
-lemmas subcls_direct = subcls1I [THEN r_into_rtranclp [where r="subcls1 G"], standard]
+lemmas subcls_direct = subcls1I [THEN r_into_rtrancl [where r="subcls1 G"], standard]
 
 lemma Ext_subcls_Base [simp]: "tprg\<turnstile>Ext\<preceq>C Base"
 apply (rule subcls_direct)
@@ -220,12 +220,12 @@
 
 declare ty_expr_ty_exprs_wt_stmt.intros [intro!]
 
-lemma acyclic_subcls1': "acyclicP (subcls1 tprg)"
-apply (rule acyclicI [to_pred])
+lemma acyclic_subcls1': "acyclic (subcls1 tprg)"
+apply (rule acyclicI)
 apply safe
 done
 
-lemmas wf_subcls1' = acyclic_subcls1' [THEN finite_subcls1 [THEN finite_acyclic_wf_converse [to_pred]]]
+lemmas wf_subcls1' = acyclic_subcls1' [THEN finite_subcls1 [THEN finite_acyclic_wf_converse]]
 
 lemmas fields_rec' = wf_subcls1' [THEN [2] fields_rec_lemma]
 
@@ -346,7 +346,7 @@
 apply (fold ExtC_def)
 apply (rule mp) defer apply (rule wf_foo_Ext)
 apply (auto simp add: wf_mdecl_def)
-apply (drule rtranclpD)
+apply (drule rtranclD)
 apply auto
 done
 
--- a/src/HOL/MicroJava/J/JTypeSafe.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/J/JTypeSafe.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -218,7 +218,7 @@
 apply( rule conjI)
 apply(  force elim!: NewC_conforms)
 apply( rule conf_obj_AddrI)
-apply(  rule_tac [2] rtranclp.rtrancl_refl)
+apply(  rule_tac [2] rtrancl.rtrancl_refl)
 apply( simp (no_asm))
 
 -- "for Cast"
--- a/src/HOL/MicroJava/J/TypeRel.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/J/TypeRel.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -9,44 +9,47 @@
 theory TypeRel imports Decl begin
 
 -- "direct subclass, cf. 8.1.3"
-inductive
-  subcls1 :: "'c prog => [cname, cname] => bool" ("_ \<turnstile> _ \<prec>C1 _" [71,71,71] 70)
+
+inductive_set
+  subcls1 :: "'c prog => (cname \<times> cname) set"
+  and subcls1' :: "'c prog => cname \<Rightarrow> cname => bool" ("_ \<turnstile> _ \<prec>C1 _" [71,71,71] 70)
   for G :: "'c prog"
 where
-  subcls1I: "\<lbrakk>class G C = Some (D,rest); C \<noteq> Object\<rbrakk> \<Longrightarrow> G\<turnstile>C\<prec>C1D"
+  "G \<turnstile> C \<prec>C1 D \<equiv> (C, D) \<in> subcls1 G"
+  | subcls1I: "\<lbrakk>class G C = Some (D,rest); C \<noteq> Object\<rbrakk> \<Longrightarrow> G \<turnstile> C \<prec>C1 D"
 
 abbreviation
-  subcls  :: "'c prog => [cname, cname] => bool" ("_ \<turnstile> _ \<preceq>C _"  [71,71,71] 70)
-  where "G\<turnstile>C \<preceq>C  D \<equiv> (subcls1 G)^** C D"
-  
+  subcls  :: "'c prog => cname \<Rightarrow> cname => bool" ("_ \<turnstile> _ \<preceq>C _"  [71,71,71] 70)
+  where "G \<turnstile> C \<preceq>C D \<equiv> (C, D) \<in> (subcls1 G)^*"
+
 lemma subcls1D: 
   "G\<turnstile>C\<prec>C1D \<Longrightarrow> C \<noteq> Object \<and> (\<exists>fs ms. class G C = Some (D,fs,ms))"
 apply (erule subcls1.cases)
 apply auto
 done
 
-lemma subcls1_def2: 
-  "subcls1 G = (\<lambda>C D. (C, D) \<in>
-     (SIGMA C: {C. is_class G C} . {D. C\<noteq>Object \<and> fst (the (class G C))=D}))"
-  by (auto simp add: is_class_def expand_fun_eq dest: subcls1D intro: subcls1I)
+lemma subcls1_def2:
+  "subcls1 P =
+     (SIGMA C:{C. is_class P C}. {D. C\<noteq>Object \<and> fst (the (class P C))=D})"
+  by (auto simp add: is_class_def dest: subcls1D intro: subcls1I)
 
-lemma finite_subcls1: "finite {(C, D). subcls1 G C D}"
+lemma finite_subcls1: "finite (subcls1 G)"
 apply(simp add: subcls1_def2 del: mem_Sigma_iff)
 apply(rule finite_SigmaI [OF finite_is_class])
 apply(rule_tac B = "{fst (the (class G C))}" in finite_subset)
 apply  auto
 done
 
-lemma subcls_is_class: "(subcls1 G)^++ C D ==> is_class G C"
+lemma subcls_is_class: "(C, D) \<in> (subcls1 G)^+  ==> is_class G C"
 apply (unfold is_class_def)
-apply(erule tranclp_trans_induct)
+apply(erule trancl_trans_induct)
 apply (auto dest!: subcls1D)
 done
 
 lemma subcls_is_class2 [rule_format (no_asm)]: 
   "G\<turnstile>C\<preceq>C D \<Longrightarrow> is_class G D \<longrightarrow> is_class G C"
 apply (unfold is_class_def)
-apply (erule rtranclp_induct)
+apply (erule rtrancl_induct)
 apply  (drule_tac [2] subcls1D)
 apply  auto
 done
@@ -54,48 +57,28 @@
 constdefs
   class_rec :: "'c prog \<Rightarrow> cname \<Rightarrow> 'a \<Rightarrow>
     (cname \<Rightarrow> fdecl list \<Rightarrow> 'c mdecl list \<Rightarrow> 'a \<Rightarrow> 'a) \<Rightarrow> 'a"
-  "class_rec G == wfrec {(C, D). (subcls1 G)^--1 C D}
+  "class_rec G == wfrec ((subcls1 G)^-1)
     (\<lambda>r C t f. case class G C of
          None \<Rightarrow> undefined
        | Some (D,fs,ms) \<Rightarrow> 
            f C fs ms (if C = Object then t else r D t f))"
 
-lemma class_rec_lemma: "wfP ((subcls1 G)^--1) \<Longrightarrow> class G C = Some (D,fs,ms) \<Longrightarrow>
- class_rec G C t f = f C fs ms (if C=Object then t else class_rec G D t f)"
-  by (simp add: class_rec_def wfrec [to_pred, where r="(subcls1 G)^--1", simplified]
-    cut_apply [where r="{(C, D). subcls1 G D C}", simplified, OF subcls1I])
+lemma class_rec_lemma:
+  assumes wf: "wf ((subcls1 G)^-1)"
+    and cls: "class G C = Some (D, fs, ms)"
+  shows "class_rec G C t f = f C fs ms (if C=Object then t else class_rec G D t f)"
+proof -
+  from wf have step: "\<And>H a. wfrec ((subcls1 G)\<inverse>) H a =
+    H (cut (wfrec ((subcls1 G)\<inverse>) H) ((subcls1 G)\<inverse>) a) a"
+    by (rule wfrec)
+  have cut: "\<And>f. C \<noteq> Object \<Longrightarrow> cut f ((subcls1 G)\<inverse>) C D = f D"
+    by (rule cut_apply [where r="(subcls1 G)^-1", simplified, OF subcls1I, OF cls])
+  from cls show ?thesis by (simp add: step cut class_rec_def)
+qed
 
 definition
-  "wf_class G = wfP ((subcls1 G)^--1)"
+  "wf_class G = wf ((subcls1 G)^-1)"
 
-lemma class_rec_func (*[code]*):
-  "class_rec G C t f = (if wf_class G then
-    (case class G C
-      of None \<Rightarrow> undefined
-       | Some (D, fs, ms) \<Rightarrow> f C fs ms (if C = Object then t else class_rec G D t f))
-    else class_rec G C t f)"
-proof (cases "wf_class G")
-  case False then show ?thesis by auto
-next
-  case True
-  from `wf_class G` have wf: "wfP ((subcls1 G)^--1)"
-    unfolding wf_class_def .
-  show ?thesis
-  proof (cases "class G C")
-    case None
-    with wf show ?thesis
-      by (simp add: class_rec_def wfrec [to_pred, where r="(subcls1 G)^--1", simplified]
-        cut_apply [where r="{(C, D).subcls1 G D C}", simplified, OF subcls1I])
-  next
-    case (Some x) show ?thesis
-    proof (cases x)
-      case (fields D fs ms)
-      then have is_some: "class G C = Some (D, fs, ms)" using Some by simp
-      note class_rec = class_rec_lemma [OF wf is_some]
-      show ?thesis unfolding class_rec by (simp add: is_some)
-    qed
-  qed
-qed
 
 text {* Code generator setup (FIXME!) *}
 
@@ -115,7 +98,7 @@
 defs method_def: "method \<equiv> \<lambda>(G,C). class_rec G C empty (\<lambda>C fs ms ts.
                            ts ++ map_of (map (\<lambda>(s,m). (s,(C,m))) ms))"
 
-lemma method_rec_lemma: "[|class G C = Some (D,fs,ms); wfP ((subcls1 G)^--1)|] ==>
+lemma method_rec_lemma: "[|class G C = Some (D,fs,ms); wf ((subcls1 G)^-1)|] ==>
   method (G,C) = (if C = Object then empty else method (G,D)) ++  
   map_of (map (\<lambda>(s,m). (s,(C,m))) ms)"
 apply (unfold method_def)
@@ -129,7 +112,7 @@
 defs fields_def: "fields \<equiv> \<lambda>(G,C). class_rec G C []    (\<lambda>C fs ms ts.
                            map (\<lambda>(fn,ft). ((fn,C),ft)) fs @ ts)"
 
-lemma fields_rec_lemma: "[|class G C = Some (D,fs,ms); wfP ((subcls1 G)^--1)|] ==>
+lemma fields_rec_lemma: "[|class G C = Some (D,fs,ms); wf ((subcls1 G)^-1)|] ==>
  fields (G,C) = 
   map (\<lambda>(fn,ft). ((fn,C),ft)) fs @ (if C = Object then [] else fields (G,D))"
 apply (unfold fields_def)
--- a/src/HOL/MicroJava/J/WellForm.thy	Wed Dec 02 12:04:07 2009 +0100
+++ b/src/HOL/MicroJava/J/WellForm.thy	Tue Nov 24 14:37:23 2009 +0100
@@ -5,7 +5,9 @@
 
 header {* \isaheader{Well-formedness of Java programs} *}
 
-theory WellForm imports TypeRel SystemClasses begin
+theory WellForm
+imports TypeRel SystemClasses
+begin
 
 text {*
 for static checks on expressions and statements, see WellType.
@@ -133,13 +135,13 @@
   apply (auto intro!: map_of_SomeI)
   done
 
-lemma subcls1_wfD: "[|G\<turnstile>C\<prec>C1D; ws_prog G|] ==> D \<noteq> C \<and> \<not> (subcls1 G)^++ D C"
-apply( frule tranclp.r_into_trancl [where r="subcls1 G"])
+lemma subcls1_wfD: "[|G\<turnstile>C\<prec>C1D; ws_prog G|] ==> D \<noteq> C \<and> (D, C) \<notin> (subcls1 G)^+"
+apply( frule trancl.r_into_trancl [where r="subcls1 G"])
 apply( drule subcls1D)
 apply(clarify)
 apply( drule (1) class_wf_struct)
 apply( unfold ws_cdecl_def)
-apply(force simp add: reflcl_tranclp [THEN sym] simp del: reflcl_tranclp)
+apply(force simp add: reflcl_trancl [THEN sym] simp del: reflcl_trancl)
 done
 
 lemma wf_cdecl_supD: 
@@ -148,47 +150,50 @@
 apply (auto split add: option.split_asm)
 done
 
-lemma subcls_asym: "[|ws_prog G; (subcls1 G)^++ C D|] ==> \<not> (subcls1 G)^++ D C"
-apply(erule tranclp.cases)
+lemma subcls_asym: "[|ws_prog G; (C, D) \<in> (subcls1 G)^+|] ==> (D, C) \<notin> (subcls1 G)^+"
+apply(erule trancl.cases)
 apply(fast dest!: subcls1_wfD )
-apply(fast dest!: subcls1_wfD intro: tranclp_trans)
+apply(fast dest!: subcls1_wfD intro: trancl_trans)
 done
 
-lemma subcls_irrefl: "[|ws_prog G; (subcls1 G)^++ C D|] ==> C \<noteq> D"
-apply (erule tranclp_trans_induct)
+lemma subcls_irrefl: "[|ws_prog G; (C, D) \<in> (subcls1 G)^+|] ==> C \<noteq> D"
+apply (erule trancl_trans_induct)
 apply  (auto dest: subcls1_wfD subcls_asym)
 done
 
-lemma acyclic_subcls1: "ws_prog G ==> acyclicP (subcls1 G)"
-apply (simp add: acyclic_def [to_pred])
+lemma acyclic_subcls1: "ws_prog G ==> acyclic (subcls1 G)"
+apply (simp add: acyclic_def)
 apply (fast dest: subcls_irrefl)
 done
 
-lemma wf_subcls1: "ws_prog G ==> wfP ((subcls1 G)^--1)"
-apply (rule finite_acyclic_wf [to_pred])
-apply ( subst finite_converse [to_pred])
+lemma wf_subcls1: "ws_prog G ==> wf ((subcls1 G)^-1)"
+apply (rule finite_acyclic_wf)
+apply ( subst finite_converse)
 apply ( rule finite_subcls1)
-apply (subst acyclic_converse [to_pred])
+apply (subst acyclic_converse)
 apply (erule acyclic_subcls1)
 done
 
-
-lemma subcls_induct: 
-"[|wf_prog wf_mb G; !!C. \<forall>D. (subcls1 G)^++ C D --> P D ==> P C|] ==> P C"
+lemma subcls_induct_struct: 
+"[|ws_prog G; !!C. \<forall>D. (C, D) \<in> (subcls1 G)^+ --> P D ==> P C|] ==> P C"
 (is "?A \<Longrightarrow> PROP ?P \<Longrightarrow> _")
 proof -
   assume p: "PROP ?P"
   assume ?A thus ?thesis apply -
-apply (drule wf_prog_ws_prog)
 apply(drule wf_subcls1)
-apply(drule wfP_trancl)
-apply(simp only: tranclp_converse)
-apply(erule_tac a = C in wfP_induct)
+apply(drule wf_trancl)
+apply(simp only: trancl_converse)
+apply(erule_tac a = C in wf_induct)
 apply(rule p)
 apply(auto)
 done
 qed
 
+lemma subcls_induct: 
+"[|wf_prog wf_mb G; !!C. \<forall>D. (C, D) \<in> (subcls1 G)^+ --> P D ==> P C|] ==> P C"
+(is "?A \<Longrightarrow> PROP ?P \<Longrightarrow> _")
+by (fact subcls_induct_struct [OF wf_prog_ws_prog])
+
 lemma subcls1_induct:
 "[|is_class G C; wf_prog wf_mb G; P Object;  
    !!C D fs ms. [|C \<noteq> Object; is_class G C; class G C = Some (D,fs,ms) \<and>  
@@ -223,22 +228,6 @@
 done
 qed
 
-lemma subcls_induct_struct: 
-"[|ws_prog G; !!C. \<forall>D. (subcls1 G)^++ C D --> P D ==> P C|] ==> P C"
-(is "?A \<Longrightarrow> PROP ?P \<Longrightarrow> _")
-proof -
-  assume p: "PROP ?P"
-  assume ?A thus ?thesis apply -
-apply(drule wf_subcls1)
-apply(drule wfP_trancl)
-apply(simp only: tranclp_converse)
-apply(erule_tac a = C in wfP_induct)
-apply(rule p)
-apply(auto)
-done
-qed
-
-
 lemma subcls1_induct_struct:
 "[|is_class G C; ws_prog G; P Object;  
    !!C D fs ms. [|C \<noteq> Object; is_class G C; class G C = Some (D,fs,ms) \<and>  
@@ -337,7 +326,7 @@
 apply( simp (no_asm))
 apply( erule UnE)
 apply(  force)
-apply( erule r_into_rtranclp [THEN rtranclp_trans])
+apply( erule r_into_rtrancl [THEN rtrancl_trans])
 apply auto
 done
 
@@ -379,9 +368,9 @@
 done
 
 lemma fields_mono_lemma [rule_format (no_asm)]: 
-  "[|ws_prog G; (subcls1 G)^** C' C|] ==>  
+  "[|ws_prog G; (C', C) \<in> (subcls1 G)^*|] ==>  
   x \<in> set (fields (G,C)) --> x \<in> set (fields (G,C'))"
-apply(erule converse_rtranclp_induct)
+apply(erule converse_rtrancl_induct)
 apply( safe dest!: subcls1D)
 apply(subst fields_rec)
 apply(  auto)
@@ -400,10 +389,10 @@
 "[|field (G,C) fn = Some (fd, fT); G\<turnstile>D\<preceq>C C; ws_prog G|]==>  
   map_of (fields (G,D)) (fn, fd) = Some fT"
 apply (drule field_fields)
-apply (drule rtranclpD)
+apply (drule rtranclD)
 apply safe
 apply (frule subcls_is_class)
-apply (drule tranclp_into_rtranclp)
+apply (drule trancl_into_rtrancl)
 apply (fast dest: fields_mono)
 done
 
@@ -435,10 +424,10 @@
 apply (frule map_of_SomeD)
 apply (clarsimp simp add: wf_cdecl_def)
 apply( clarify)
-apply( rule rtranclp_trans)
+apply( rule rtrancl_trans)
 prefer 2
 apply(  assumption)
-apply( rule r_into_rtranclp)
+apply( rule r_into_rtrancl)
 apply( fast intro: subcls1I)
 done
 
@@ -471,10 +460,10 @@
 apply (clarsimp simp add: ws_cdecl_def)
 apply blast
 apply clarify
-apply( rule rtranclp_trans)
+apply( rule rtrancl_trans)
 prefer 2
 apply(  assumption)
-apply( rule r_into_rtranclp)
+apply( rule r_into_rtrancl)
 apply( fast intro: subcls1I)
 done
 
@@ -482,15 +471,15 @@
   "[|G\<turnstile>T'\<preceq>C T; wf_prog wf_mb G|] ==>  
    \<forall>D rT b. method (G,T) sig = Some (D,rT ,b) --> 
   (\<exists>D' rT' b'. method (G,T') sig = Some (D',rT',b') \<and> G\<turnstile>D'\<preceq>C D \<and> G\<turnstile>rT'\<preceq>rT)"
-apply( drule rtranclpD)
+apply( drule rtranclD)
 apply( erule disjE)
 apply(  fast)
 apply( erule conjE)
-apply( erule tranclp_trans_induct)
+apply( erule trancl_trans_induct)
 prefer 2
 apply(  clarify)
 apply(  drule spec, drule spec, drule spec, erule (1) impE)
-apply(  fast elim: widen_trans rtranclp_trans)
+apply(  fast elim: widen_trans rtrancl_trans)
 apply( clarify)
 apply( drule subcls1D)
 apply( clarify)
@@ -510,7 +499,7 @@
 apply( unfold wf_cdecl_def)
 apply( drule map_of_SomeD)
 apply (auto simp add: wf_mrT_def)
-apply (rule rtranclp_trans)
+apply (rule rtrancl_trans)
 defer
 apply (rule method_wf_mhead [THEN conjunct1])
 apply (simp only: wf_prog_def)