# HG changeset patch # User haftmann # Date 1259069843 -3600 # Node ID 1bc3b688548c814b81d6be8ab9740e0cb46abb33 # Parent 6a973bd4394996c31f638e5c59ea6bb953335c9a backported parts of abstract byte code verifier from AFP/Jinja diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/IsaMakefile --- 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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/Altern.thy --- 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 \ nat \ nat \ JVMType.state \ bool" diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/BVExample.thy --- 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 = (\C D. (C, D) \ {(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 \ a \ b \ (\y. \ r a y) \ False" - by (auto elim: converse_rtranclpE) + "(a, b) \ r\<^sup>* \ a \ b \ (\y. (a, y) \ r) \ 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)\\)" - apply (rule finite_acyclic_wf_converse [to_pred]) +lemma wf_subcls1_E: "wf ((subcls1 E)\)" + apply (rule finite_acyclic_wf_converse) apply (simp add: subcls1 del: insert_iff) apply (rule acyclic_subcls1_E) done diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/BVNoTypeError.thy --- 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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/BVSpec.thy --- 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} *} diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/BVSpecTypeSafe.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/Correct.thy --- 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] \ bool" diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/Effect.thy --- 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 \ state_type option) list" diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/EffectMono.thy --- 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 \ xb \ PrimT p) = (xb = PrimT p)" by (auto elim: widen.cases) diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/Err.thy --- 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 \ 'a \ 'a err" - 'a esl = "'a set * 'a ord * 'a ebinop" - -consts - ok_val :: "'a err \ 'a" -primrec - "ok_val (OK x) = x" - -constdefs - lift :: "('a \ 'b err) \ ('a err \ 'b err)" -"lift f e == case e of Err \ Err | OK x \ f x" - - lift2 :: "('a \ 'b \ 'c err) \ 'a err \ 'b err \ 'c err" -"lift2 f e1 e2 == - case e1 of Err \ Err - | OK x \ (case e2 of Err \ Err | OK y \ f x y)" - - le :: "'a ord \ 'a err ord" -"le r e1 e2 == - case e2 of Err \ True | - OK y \ (case e1 of Err \ False | OK x \ x <=_r y)" - - sup :: "('a \ 'b \ 'c) \ ('a err \ 'b err \ 'c err)" -"sup f == lift2(%x y. OK(x +_f y))" - - err :: "'a set \ 'a err set" -"err A == insert Err {x . ? y:A. x = OK y}" - - esl :: "'a sl \ 'a esl" -"esl == %(A,r,f). (A,r, %x y. OK(f x y))" - - sl :: "'a esl \ 'a err sl" -"sl == %(A,r,f). (err A, le r, lift2 f)" - -syntax - err_semilat :: "'a esl \ bool" -translations -"err_semilat L" == "semilat(Err.sl L)" - - -consts - strict :: "('a \ 'b err) \ ('a err \ 'b err)" -primrec - "strict f Err = Err" - "strict f (OK x) = f x" - -lemma strict_Some [simp]: - "(strict f x = OK y) = (\ z. x = OK z \ f z = OK y)" - by (cases x, auto) - -lemma not_Err_eq: - "(x \ Err) = (\a. x = OK a)" - by (cases x) auto - -lemma not_OK_eq: - "(\y. x \ 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 \ 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 \ e1 <=_(le r) e2 \ e2 <=_(le r) e3 \ 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 \ e1 <=_(le r) e2 \ e2 <=_(le r) e1 \ 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]: - "\L. semilat L \ err_semilat(esl L)" -by(simp add: err_semilat_eslI_aux split_tupled_all) - -lemma acc_err [simp, intro!]: "acc r \ 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: - "\ e : err S; !x:S. e = OK x \ f x : err S \ \ 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]: - "\ x: err A; semilat(err A, le r, f) \ \ 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]: - "\ x: err A; semilat(err A, le r, f) \ \ 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: - "\ x:A; y:A; semilat(err A, le r, f); OK x +_f OK y = OK z \ - \ 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: - "\ x:A; y:A; semilat(err A, le r, f); OK x +_f OK y = OK z \ - \ 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: - "\ x=y; order r \ \ 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: "\A x y z f r. - \ semilat (A,r,f); x +_f y <=_r z; x:A; y:A; z:A \ - \ x <=_r z \ 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) \ P x) = (!y:A. P(f y))" - by blast - -lemma closed_err_Union_lift2I: - "\ !A:AS. closed (err A) (lift2 f); AS ~= {}; - !A:AS.!B:AS. A~=B \ (!a:A.!b:B. a +_f b = Err) \ - \ 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: - "\ !A:AS. err_semilat(A, r, f); AS ~= {}; - !A:AS.!B:AS. A~=B \ (!a:A.!b:B. ~ a <=_r b & a +_f b = Err) \ - \ 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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/JType.thy --- 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 \ cname \ cname" @@ -34,7 +35,7 @@ is_ty :: "'c prog \ ty \ bool" "is_ty G T == case T of PrimT P \ True | RefT R \ - (case R of NullT \ True | ClassT C \ (subcls1 G)^** C Object)" + (case R of NullT \ True | ClassT C \ (C, Object) \ (subcls1 G)^*)" translations @@ -78,7 +79,7 @@ have "R \ ClassT Object \ ?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) \ order (subtype G)" + "acyclic (subcls1 G) \ 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) \ acc (subtype G)" + "wf ((subcls1 G)^-1) \ acc (subtype G)" apply (unfold Semilat.acc_def lesssub_def) -apply (drule_tac p = "inf ((subcls1 G)^--1) op \" 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 \)^--1 = (inf ((subcls1 G)^--1) op \)") +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: - "\ ws_prog G; single_valuedP (subcls1 G); acyclicP (subcls1 G) \ + "\ ws_prog G; single_valued (subcls1 G); acyclic (subcls1 G) \ \ 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: - "\ ws_prog G; single_valuedP (subcls1 G); acyclicP (subcls1 G); + "\ ws_prog G; single_valued (subcls1 G); acyclic (subcls1 G); is_type G t1; is_type G t2; sup G t1 t2 = OK s \ \ subtype G t1 s \ 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: - "\ ws_prog G; single_valuedP (subcls1 G); acyclicP (subcls1 G); + "\ 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 \ \ 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: - "\ ws_prog G; single_valuedP (subcls1 G); acyclicP (subcls1 G) \ + "\ ws_prog G; single_valued (subcls1 G); acyclic (subcls1 G) \ \ 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 - "(\x\err (types G). \y\err (types G). x <=_(le (subtype G)) x +_(lift2 (sup G)) y) \ - (\x\err (types G). \y\err (types G). y <=_(le (subtype G)) x +_(lift2 (sup G)) y)" + "(\x\err (types G). \y\err (types G). x <=_(Err.le (subtype G)) x +_(lift2 (sup G)) y) \ + (\x\err (types G). \y\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 "\x\err (types G). \y\err (types G). \z\err (types G). - x <=_(le (subtype G)) z \ y <=_(le (subtype G)) z \ x +_(lift2 (sup G)) y <=_(le (subtype G)) z" + x <=_(Err.le (subtype G)) z \ y <=_(Err.le (subtype G)) z \ 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 \ single_valuedP (subcls1 G)" + "ws_prog G \ single_valued (subcls1 G)" by (auto simp add: ws_prog_def unique_def single_valued_def intro: subcls1I elim!: subcls1.cases) diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/JVM.thy --- 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 \ nat \ nat \ ty \ exception_table \ @@ -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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/JVMType.thy --- 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" diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/Kildall.thy --- 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 \ 's step_type \ - 's list \ nat set \ 's list \ nat set" - propa :: "'s binop \ (nat \ 's) list \ 's list \ nat set \ '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 \ {}) - (%(ss,w). let p = SOME p. p \ w - in propa f (step p (ss!p)) ss (w-{p})) - (ss,w)" - -constdefs - unstables :: "'s ord \ 's step_type \ 's list \ nat set" -"unstables r step ss == {p. p < size ss \ \stable r step ss p}" - - kildall :: "'s ord \ 's binop \ 's step_type \ 's list \ 's list" -"kildall r f step ss == fst(iter f step ss (unstables r step ss))" - -consts merges :: "'s binop \ (nat \ 's) list \ 's list \ '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: - "\ss. \p < length ss; ss \ list n A; \(p,t)\set ps. p t\A \ \ - (merges f ps ss)!p = map snd [(p',t') \ ps. p'=p] ++_f ss!p" - (is "\ss. \_; _; ?steptype ps\ \ ?P ss ps") -proof (induct ps) - show "\ss. ?P ss []" by simp - - fix ss p' ps' - assume ss: "ss \ list n A" - assume l: "p < length ss" - assume "?steptype (p'#ps')" - then obtain a b where - p': "p'=(a,b)" and ab: "aA" and ps': "?steptype ps'" - by (cases p') auto - assume "\ss. p< length ss \ ss \ list n A \ ?steptype ps' \ ?P ss ps'" - from this [OF _ _ ps'] have IH: "\ss. ss \ list n A \ p < length ss \ ?P ss ps'" . - - from ss ab - have "ss[a := b +_f ss!a] \ 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]: - "\ss. size(merges f ps ss) = size ss" - by (induct_tac ps, auto) - - -lemma (in Semilat) merges_preserves_type_lemma: -shows "\xs. xs \ list n A \ (\(p,x) \ set ps. p x\A) - \ merges f ps xs \ 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]: - "\ xs \ list n A; \(p,x) \ set ps. p x\A \ - \ merges f ps xs \ list n A" -by (simp add: merges_preserves_type_lemma) - -lemma (in Semilat) merges_incr_lemma: - "\xs. xs \ list n A \ (\(p,x)\set ps. p x \ A) \ 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: - "\ xs \ list n A; \(p,x)\set ps. p x \ A \ - \ xs <=[r] merges f ps xs" - by (simp add: merges_incr_lemma) - - -lemma (in Semilat) merges_same_conv [rule_format]: - "(\xs. xs \ list n A \ (\(p,x)\set ps. p x\A) \ - (merges f ps xs = xs) = (\(p,x)\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 \ set ys <= A \ xs <=[r] ys \ p < size xs \ - x <=_r ys!p \ x\A \ 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 "\(p,t)\set ps. t <=_r ts!p \ t \ A \ p < size ts" and "ss <=[r] ts" - shows "merges f ps ss <=[r] ts" -proof - - { fix t ts ps - have - "\qs. \set ts <= A; \(p,t)\set ps. t <=_r ts!p \ t \ A \ p< size ts \ \ - set qs <= set ps \ - (\ss. set ss <= A \ ss <=[r] ts \ 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: - "\ss w. (\(q,t)\set qs. q < size ss) \ - propa f qs ss w = - (merges f qs ss, {q. \t. (q,t)\set qs \ t +_f ss!q \ 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 "\pres_type step n A; bounded step n; - ss \ list n A; p \ w; \q\w. q < n; - \q. q < n \ q \ w \ stable r step ss q; q < n; - \s'. (q,s') \ set (step p (ss ! p)) \ s' +_f ss ! q = ss ! q; - q \ w \ q = p \ - \ stable r step (merges f (step p (ss!p)) ss) q" - apply (unfold stable_def) - apply (subgoal_tac "\s'. (q,s') \ set (step p (ss!p)) \ 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 = "\x. (a, b) \ set (step q x)" in subst) - apply assumption - - apply (simp add: plusplus_empty) - apply (cases "q \ 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: - "\ mono r step n A; bounded step n; - \(p',s') \ set (step p (ss!p)). s' \ A; ss \ list n A; ts \ list n A; p < n; - ss <=[r] ts; \p. p < n \ stable r step ts p \ - \ 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 "\ ss \ list n A; \(q,t)\set qs. q t\A; p\w \ \ - ss <[r] merges f qs ss \ - merges f qs ss = ss \ {q. \t. (q,t)\set qs \ t +_f ss!q \ 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 "\q t. \((q, t) \ set qs \ t +_f ss ! q \ 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 "\ acc r ; pres_type step n A; mono r step n A; - bounded step n; \p\w0. p < n; ss0 \ list n A; - \p w0 \ stable r step ss0 p \ \ - iter f step ss0 w0 = (ss',w') - \ - ss' \ list n A \ stables r step ss' \ ss0 <=[r] ss' \ - (\ts\list n A. ss0 <=[r] ts \ stables r step ts \ 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 \ list n A \ (\p w \ stable r step ss p) \ ss0 <=[r] ss \ - (\ts\list n A. ss0 <=[r] ts \ stables r step ts \ ss <=[r] ts) \ - (\p\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 \ w) \ w") - prefer 2; apply (fast intro: someI) -apply(subgoal_tac "\(q,t) \ set (step (SOME p. p \ w) (ss ! (SOME p. p \ w))). q < length ss \ t \ 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 \ w) \ w") - prefer 2; apply (fast intro: someI) -apply(subgoal_tac "\(q,t) \ set (step (SOME p. p \ w) (ss ! (SOME p. p \ w))). q < length ss \ t \ 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 "\ acc r; pres_type step n A; mono r step n A; - bounded step n; ss0 \ list n A \ \ - kildall r f step ss0 \ list n A \ - stables r step (kildall r f step ss0) \ - ss0 <=[r] kildall r f step ss0 \ - (\ts\list n A. ss0 <=[r] ts \ stables r step ts \ - 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 "\ acc r; top r T; pres_type step n A; bounded step n; mono r step n A \ - \ 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 \ 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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/LBVComplete.thy --- 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] \ bool" - "is_target step phi pc' \ - \pc s'. pc' \ pc+1 \ pc < length phi \ (pc',s') \ set (step pc (phi!pc))" - - make_cert :: "['s step_type, 's list, 's] \ 's certificate" - "make_cert step phi B \ - map (\pc. if is_target step phi pc then phi!pc else B) [0..pc. pc' \ pc+1 \ pc' mem (map fst (step pc (phi!pc)))) [0..") - fixes c :: "'a list" - defines cert_def: "c \ make_cert step \ \" - - assumes mono: "mono r step (length \) A" - assumes pres: "pres_type step (length \) A" - assumes phi: "\pc < length \. \!pc \ A \ \!pc \ \" - assumes bounded: "bounded step (length \)" - - assumes B_neq_T: "\ \ \" - - -lemma (in lbvc) cert: "cert_ok c (length \) \ \ A" -proof (unfold cert_ok_def, intro strip conjI) - note [simp] = make_cert_def cert_def nth_append - - show "c!length \ = \" by simp - - fix pc assume pc: "pc < length \" - from pc phi B_A show "c!pc \ A" by simp - from pc phi B_neq_T show "c!pc \ \" by simp -qed - -lemmas [simp del] = split_paired_Ex - - -lemma (in lbvc) cert_target [intro?]: - "\ (pc',s') \ set (step pc (\!pc)); - pc' \ pc+1; pc < length \; pc' < length \ \ - \ c!pc' = \!pc'" - by (auto simp add: cert_def make_cert_def nth_append is_target_def) - - -lemma (in lbvc) cert_approx [intro?]: - "\ pc < length \; c!pc \ \ \ - \ c!pc = \!pc" - by (auto simp add: cert_def make_cert_def nth_append) - - -lemma (in lbv) le_top [simp, intro]: - "x <=_r \" - by (insert top) simp - - -lemma (in lbv) merge_mono: - assumes less: "ss2 <=|r| ss1" - assumes x: "x \ A" - assumes ss1: "snd`set ss1 \ A" - assumes ss2: "snd`set ss2 \ A" - shows "merge c pc ss2 x <=_r merge c pc ss1 x" (is "?s2 <=_r ?s1") -proof- - have "?s1 = \ \ ?thesis" by simp - moreover { - assume merge: "?s1 \ T" - from x ss1 have "?s1 = - (if \(pc', s')\set ss1. pc' \ pc + 1 \ s' <=_r c!pc' - then (map snd [(p', t') \ ss1 . p'=pc+1]) ++_f x - else \)" - by (rule merge_def) - with merge obtain - app: "\(pc',s')\set ss1. pc' \ pc+1 \ s' <=_r c!pc'" - (is "?app ss1") and - sum: "(map snd [(p',t') \ 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) \ A" by auto - with x have "?sum ss1 \ A" by (auto intro!: plusplus_closed semilat) - with sum have "?s1 \ A" by simp - moreover - have mapD: "\x ss. x \ set (?map ss) \ \p. (p,x) \ set ss \ p=pc+1" by auto - from x map1 - have "\x \ set (?map ss1). x <=_r ?sum ss1" - by clarify (rule pp_ub1) - with sum have "\x \ set (?map ss1). x <=_r ?s1" by simp - with less have "\x \ 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) \ A" by auto - ultimately - have "?sum ss2 <=_r ?s1" using x by - (rule pp_lub) - } - moreover - from x ss2 have - "?s2 = - (if \(pc', s')\set ss2. pc' \ pc + 1 \ s' <=_r c!pc' - then map snd [(p', t') \ ss2 . p' = pc + 1] ++_f x - else \)" - by (rule merge_def) - ultimately have ?thesis by simp - } - ultimately show ?thesis by (cases "?s1 = \") auto -qed - - -lemma (in lbvc) wti_mono: - assumes less: "s2 <=_r s1" - assumes pc: "pc < length \" - assumes s1: "s1 \ A" - assumes s2: "s2 \ 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 \ A" by (rule cert_okD3) - moreover - from pres s1 pc - have "snd`set (step pc s1) \ A" by (rule pres_typeD2) - moreover - from pres s2 pc - have "snd`set (step pc s2) \ 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 \" - assumes s1: "s1 \ A" - assumes s2: "s2 \ A" - shows "wtc c pc s2 <=_r wtc c pc s1" (is "?s2' <=_r ?s1'") -proof (cases "c!pc = \") - 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' = \ \ ?thesis" by simp - moreover { - assume "?s1' \ \" - 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' = \") auto -qed - - -lemma (in lbv) top_le_conv [simp]: - "\ <=_r x = (x = \)" - by (insert semilat) (simp add: top top_le_conv) - -lemma (in lbv) neq_top [simp, elim]: - "\ x <=_r y; y \ \ \ \ x \ \" - by (cases "x = T") auto - - -lemma (in lbvc) stable_wti: - assumes stable: "stable r step \ pc" - assumes pc: "pc < length \" - shows "wti c pc (\!pc) \ \" -proof - - let ?step = "step pc (\!pc)" - from stable - have less: "\(q,s')\set ?step. s' <=_r \!q" by (simp add: stable_def) - - from cert B_A pc - have cert_suc: "c!Suc pc \ A" by (rule cert_okD3) - moreover - from phi pc have "\!pc \ A" by simp - from pres this pc - have stepA: "snd`set ?step \ A" by (rule pres_typeD2) - ultimately - have "merge c pc ?step (c!Suc pc) = - (if \(pc',s')\set ?step. pc'\pc+1 \ s' <=_r c!pc' - then map snd [(p',t') \ ?step.p'=pc+1] ++_f c!Suc pc - else \)" unfolding mrg_def by (rule lbv.merge_def [OF lbvc.axioms(1), OF lbvc_axioms]) - moreover { - fix pc' s' assume s': "(pc', s') \ set ?step" and suc_pc: "pc' \ pc+1" - with less have "s' <=_r \!pc'" by auto - also - from bounded pc s' have "pc' < length \" by (rule boundedD) - with s' suc_pc pc have "c!pc' = \!pc'" .. - hence "\!pc' = c!pc'" .. - finally have "s' <=_r c!pc'" . - } hence "\(pc',s')\set ?step. pc'\pc+1 \ s' <=_r c!pc'" by auto - moreover - from pc have "Suc pc = length \ \ Suc pc < length \" by auto - hence "map snd [(p',t') \ ?step.p'=pc+1] ++_f c!Suc pc \ \" - (is "?map ++_f _ \ _") - proof (rule disjE) - assume pc': "Suc pc = length \" - with cert have "c!Suc pc = \" by (simp add: cert_okD2) - moreover - from pc' bounded pc - have "\(p',t')\set ?step. p'\pc+1" by clarify (drule boundedD, auto) - hence "[(p',t') \ ?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 \" - from pc' phi have "\!Suc pc \ A" by simp - moreover note cert_suc - moreover from stepA - have "set ?map \ A" by auto - moreover - have "\s. s \ set ?map \ \t. (Suc pc, t) \ set ?step" by auto - with less have "\s' \ set ?map. s' <=_r \!Suc pc" by auto - moreover - from pc' have "c!Suc pc <=_r \!Suc pc" - by (cases "c!Suc pc = \") (auto dest: cert_approx) - ultimately - have "?map ++_f c!Suc pc <=_r \!Suc pc" by (rule pp_lub) - moreover - from pc' phi have "\!Suc pc \ \" by simp - ultimately - show ?thesis by auto - qed - ultimately - have "merge c pc ?step (c!Suc pc) \ \" by simp - thus ?thesis by (simp add: wti) -qed - -lemma (in lbvc) wti_less: - assumes stable: "stable r step \ pc" - assumes suc_pc: "Suc pc < length \" - shows "wti c pc (\!pc) <=_r \!Suc pc" (is "?wti <=_r _") -proof - - let ?step = "step pc (\!pc)" - - from stable - have less: "\(q,s')\set ?step. s' <=_r \!q" by (simp add: stable_def) - - from suc_pc have pc: "pc < length \" by simp - with cert B_A have cert_suc: "c!Suc pc \ A" by (rule cert_okD3) - moreover - from phi pc have "\!pc \ A" by simp - with pres pc have stepA: "snd`set ?step \ A" by - (rule pres_typeD2) - moreover - from stable pc have "?wti \ \" by (rule stable_wti) - hence "merge c pc ?step (c!Suc pc) \ \" by (simp add: wti) - ultimately - have "merge c pc ?step (c!Suc pc) = - map snd [(p',t')\ ?step.p'=pc+1] ++_f c!Suc pc" by (rule merge_not_top_s) - hence "?wti = \" (is "_ = (?map ++_f _)" is "_ = ?sum") by (simp add: wti) - also { - from suc_pc phi have "\!Suc pc \ A" by simp - moreover note cert_suc - moreover from stepA have "set ?map \ A" by auto - moreover - have "\s. s \ set ?map \ \t. (Suc pc, t) \ set ?step" by auto - with less have "\s' \ set ?map. s' <=_r \!Suc pc" by auto - moreover - from suc_pc have "c!Suc pc <=_r \!Suc pc" - by (cases "c!Suc pc = \") (auto dest: cert_approx) - ultimately - have "?sum <=_r \!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 \" - shows "wtc c pc (\!pc) \ \" -proof - - from stable pc have wti: "wti c pc (\!pc) \ \" by (rule stable_wti) - show ?thesis - proof (cases "c!pc = \") - case True with wti show ?thesis by (simp add: wtc) - next - case False - with pc have "c!pc = \!pc" .. - with False wti show ?thesis by (simp add: wtc) - qed -qed - -lemma (in lbvc) wtc_less: - assumes stable: "stable r step \ pc" - assumes suc_pc: "Suc pc < length \" - shows "wtc c pc (\!pc) <=_r \!Suc pc" (is "?wtc <=_r _") -proof (cases "c!pc = \") - case True - moreover from stable suc_pc have "wti c pc (\!pc) <=_r \!Suc pc" - by (rule wti_less) - ultimately show ?thesis by (simp add: wtc) -next - case False - from suc_pc have pc: "pc < length \" by simp - with stable have "?wtc \ \" 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 = \!pc" .. - finally have "?wtc = wti c pc (\!pc)" . - also from stable suc_pc have "wti c pc (\!pc) <=_r \!Suc pc" by (rule wti_less) - finally show ?thesis . -qed - - -lemma (in lbvc) wt_step_wtl_lemma: - assumes wt_step: "wt_step r \ step \" - shows "\pc s. pc+length ls = length \ \ s <=_r \!pc \ s \ A \ s\\ \ - wtl ls c pc s \ \" - (is "\pc s. _ \ _ \ _ \ _ \ ?wtl ls pc s \ _") -proof (induct ls) - fix pc s assume "s\\" thus "?wtl [] pc s \ \" by simp -next - fix pc s i ls - assume "\pc s. pc+length ls=length \ \ s <=_r \!pc \ s \ A \ s\\ \ - ?wtl ls pc s \ \" - moreover - assume pc_l: "pc + length (i#ls) = length \" - hence suc_pc_l: "Suc pc + length ls = length \" by simp - ultimately - have IH: "\s. s <=_r \!Suc pc \ s \ A \ s \ \ \ ?wtl ls (Suc pc) s \ \" . - - from pc_l obtain pc: "pc < length \" by simp - with wt_step have stable: "stable r step \ pc" by (simp add: wt_step_def) - from this pc have wt_phi: "wtc c pc (\!pc) \ \" by (rule stable_wtc) - assume s_phi: "s <=_r \!pc" - from phi pc have phi_pc: "\!pc \ A" by simp - assume s: "s \ A" - with s_phi pc phi_pc have wt_s_phi: "wtc c pc s <=_r wtc c pc (\!pc)" by (rule wtc_mono) - with wt_phi have wt_s: "wtc c pc s \ \" by simp - moreover - assume s': "s \ \" - ultimately - have "ls = [] \ ?wtl (i#ls) pc s \ \" by simp - moreover { - assume "ls \ []" - with pc_l have suc_pc: "Suc pc < length \" by (auto simp add: neq_Nil_conv) - with stable have "wtc c pc (phi!pc) <=_r \!Suc pc" by (rule wtc_less) - with wt_s_phi have "wtc c pc s <=_r \!Suc pc" by (rule trans_r) - moreover - from cert suc_pc have "c!pc \ A" "c!(pc+1) \ A" - by (auto simp add: cert_ok_def) - from pres this s pc have "wtc c pc s \ A" by (rule wtc_pres) - ultimately - have "?wtl ls (Suc pc) (wtc c pc s) \ \" using IH wt_s by blast - with s' wt_s have "?wtl (i#ls) pc s \ \" by simp - } - ultimately show "?wtl (i#ls) pc s \ \" by (cases ls) blast+ -qed - - -theorem (in lbvc) wtl_complete: - assumes wt: "wt_step r \ step \" - and s: "s <=_r \!0" "s \ A" "s \ \" - and len: "length ins = length phi" - shows "wtl ins c 0 s \ \" -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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/LBVCorrect.thy --- 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" ("\") - defines phi_def: - "\ \ map (\pc. if c!pc = \ then wtl (take pc ins) c 0 s0 else c!pc) - [0.. \ A" - assumes pres: "pres_type step (length ins) A" - - -lemma (in lbvs) phi_None [intro?]: - "\ pc < length ins; c!pc = \ \ \ \ ! pc = wtl (take pc ins) c 0 s0" - by (simp add: phi_def) - -lemma (in lbvs) phi_Some [intro?]: - "\ pc < length ins; c!pc \ \ \ \ \ ! pc = c ! pc" - by (simp add: phi_def) - -lemma (in lbvs) phi_len [simp]: - "length \ = length ins" - by (simp add: phi_def) - - -lemma (in lbvs) wtl_suc_pc: - assumes all: "wtl ins c 0 s\<^sub>0 \ \" - assumes pc: "pc+1 < length ins" - shows "wtl (take (pc+1) ins) c 0 s0 \\<^sub>r \!(pc+1)" -proof - - from all pc - have "wtc c (pc+1) (wtl (take (pc+1) ins) c 0 s0) \ 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 \ \" - assumes s0: "s0 \ A" - assumes pc: "pc < length ins" - shows "stable r step \ pc" -proof (unfold stable_def, clarify) - fix pc' s' assume step: "(pc',s') \ set (step pc (\ ! pc))" - (is "(pc',s') \ 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 \ \" (is "?s1 \ _") by (rule wtl_take) - from wtl have s2: "wtl (take (pc+1) ins) c 0 s0 \ \" (is "?s2 \ _") by (rule wtl_take) - - from wtl pc have wt_s1: "wtc c pc ?s1 \ \" by (rule wtl_all) - - have c_Some: "\pc t. pc < length ins \ c!pc \ \ \ \!pc = c!pc" - by (simp add: phi_def) - from pc have c_None: "c!pc = \ \ \!pc = ?s1" .. - - from wt_s1 pc c_None c_Some - have inst: "wtc c pc ?s1 = wti c pc (\!pc)" - by (simp add: wtc split: split_if_asm) - - from pres cert s0 wtl pc have "?s1 \ A" by (rule wtl_pres) - with pc c_Some cert c_None - have "\!pc \ A" by (cases "c!pc = \") (auto dest: cert_okD1) - with pc pres - have step_in_A: "snd`set (?step pc) \ A" by (auto dest: pres_typeD2) - - show "s' <=_r \!pc'" - proof (cases "pc' = pc+1") - case True - with pc' cert - have cert_in_A: "c!(pc+1) \ 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 "\ \ \" (is "?merge \ _") by simp - with cert_in_A step_in_A - have "?merge = (map snd [(p',t') \ ?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 \!(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)) \ \" by (simp add: wti) - with step_in_A - have "\(pc', s')\set (?step pc). pc'\pc+1 \ 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' = \ \ s' = \" by simp - moreover - from c_Some pc' - have "c!pc' \ \ \ \!pc' = c!pc'" by auto - ultimately - show ?thesis by (cases "c!pc' = \") auto - qed -qed - - -lemma (in lbvs) phi_not_top: - assumes wtl: "wtl ins c 0 s0 \ \" - assumes pc: "pc < length ins" - shows "\!pc \ \" -proof (cases "c!pc = \") - case False with pc - have "\!pc = c!pc" .. - also from cert pc have "\ \ \" by (rule cert_okD4) - finally show ?thesis . -next - case True with pc - have "\!pc = wtl (take pc ins) c 0 s0" .. - also from wtl have "\ \ \" by (rule wtl_take) - finally show ?thesis . -qed - -lemma (in lbvs) phi_in_A: - assumes wtl: "wtl ins c 0 s0 \ \" - assumes s0: "s0 \ A" - shows "\ \ list (length ins) A" -proof - - { fix x assume "x \ set \" - then obtain xs ys where "\ = xs @ x # ys" - by (auto simp add: in_set_conv_decomp) - then obtain pc where pc: "pc < length \" and x: "\!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 \ A" by (auto intro!: wtl_pres) - moreover - from pc have "pc < length ins" by simp - with cert have "c!pc \ A" .. - ultimately - have "\!pc \ A" using pc by (simp add: phi_def) - hence "x \ A" using x by simp - } - hence "set \ \ A" .. - thus ?thesis by (unfold list_def) simp -qed - - -lemma (in lbvs) phi0: - assumes wtl: "wtl ins c 0 s0 \ \" - assumes 0: "0 < length ins" - shows "s0 <=_r \!0" -proof (cases "c!0 = \") - case True - with 0 have "\!0 = wtl (take 0 ins) c 0 s0" .. - moreover have "wtl (take 0 ins) c 0 s0 = s0" by simp - ultimately have "\!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 \ \" 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 \ \" - assumes s0: "s0 \ A" - shows "\ts. wt_step r \ step ts" -proof - - have "wt_step r \ step \" - proof (unfold wt_step_def, intro strip conjI) - fix pc assume "pc < length \" - then have pc: "pc < length ins" by simp - with wtl show "\!pc \ \" by (rule phi_not_top) - from wtl s0 pc show "stable r step \ pc" by (rule wtl_stable) - qed - thus ?thesis .. -qed - - -theorem (in lbvs) wtl_sound_strong: - assumes wtl: "wtl ins c 0 s0 \ \" - assumes s0: "s0 \ A" - assumes nz: "0 < length ins" - shows "\ts \ list (length ins) A. wt_step r \ step ts \ s0 <=_r ts!0" -proof - - from wtl s0 have "\ \ list (length ins) A" by (rule phi_in_A) - moreover - have "wt_step r \ step \" - proof (unfold wt_step_def, intro strip conjI) - fix pc assume "pc < length \" - then have pc: "pc < length ins" by simp - with wtl show "\!pc \ \" by (rule phi_not_top) - from wtl s0 pc show "stable r step \ pc" by (rule wtl_stable) - qed - moreover - from wtl nz have "s0 <=_r \!0" by (rule phi0) - ultimately - show ?thesis by fast -qed - -end diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/LBVJVM.thy --- 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 \ sig \ JVMType.state list" diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/LBVSpec.thy --- 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 \ 's binop \ 's ord \ 's \ nat \ (nat \ 's) list \ 's \ '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 \ 's binop \ 's ord \ 's \ - 's step_type \ nat \ 's \ 's" -"wtl_inst cert f r T step pc s \ merge cert f r T pc (step pc s) (cert!(pc+1))" - -wtl_cert :: "'s certificate \ 's binop \ 's ord \ 's \ 's \ - 's step_type \ nat \ 's \ 's" -"wtl_cert cert f r T B step pc s \ - 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 \ 's certificate \ 's binop \ 's ord \ 's \ 's \ - 's step_type \ nat \ 's \ '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 \ s = T then T else wtl_inst_list is cert f r T B step (pc+1) s')" - -constdefs - cert_ok :: "'s certificate \ nat \ 's \ 's \ 's set \ bool" - "cert_ok cert n T B A \ (\i < n. cert!i \ A \ cert!i \ T) \ (cert!n = B)" - -constdefs - bottom :: "'a ord \ 'a \ bool" - "bottom r B \ \x. B <=_r x" - - -locale lbv = Semilat + - fixes T :: "'a" ("\") - fixes B :: "'a" ("\") - fixes step :: "'a step_type" - assumes top: "top r \" - assumes T_A: "\ \ A" - assumes bot: "bottom r \" - assumes B_A: "\ \ A" - - fixes merge :: "'a certificate \ nat \ (nat \ 'a) list \ 'a \ 'a" - defines mrg_def: "merge cert \ LBVSpec.merge cert f r \" - - fixes wti :: "'a certificate \ nat \ 'a \ 'a" - defines wti_def: "wti cert \ wtl_inst cert f r \ step" - - fixes wtc :: "'a certificate \ nat \ 'a \ 'a" - defines wtc_def: "wtc cert \ wtl_cert cert f r \ \ step" - - fixes wtl :: "'b list \ 'a certificate \ nat \ 'a \ 'a" - defines wtl_def: "wtl ins cert \ wtl_inst_list ins cert f r \ \ step" - - -lemma (in lbv) wti: - "wti c pc s \ 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 \ if c!pc = \ then wti c pc s else if s <=_r c!pc then wti c pc (c!pc) else \" - by (unfold wtc_def wti_def wtl_cert_def) - - -lemma cert_okD1 [intro?]: - "cert_ok c n T B A \ pc < n \ c!pc \ A" - by (unfold cert_ok_def) fast - -lemma cert_okD2 [intro?]: - "cert_ok c n T B A \ c!n = B" - by (simp add: cert_ok_def) - -lemma cert_okD3 [intro?]: - "cert_ok c n T B A \ B \ A \ pc < n \ c!Suc pc \ 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 \ pc < n \ c!pc \ 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 \ A" - shows "x +_f \ = \" -proof - - from top have "x +_f \ <=_r \" .. - moreover from x T_A have "\ <=_r x +_f \" .. - ultimately show ?thesis .. -qed - -lemma (in lbv) plusplussup_top [simp, elim]: - "set xs \ A \ xs ++_f \ = \" - by (induct xs) auto - - - -lemma (in Semilat) pp_ub1': - assumes S: "snd`set S \ A" - assumes y: "y \ A" and ab: "(a, b) \ set S" - shows "b <=_r map snd [(p', t') \ S . p' = a] ++_f y" -proof - - from S have "\(x,y) \ set S. y \ A" by auto - with semilat y ab show ?thesis by - (rule ub1') -qed - -lemma (in lbv) bottom_le [simp, intro]: - "\ <=_r x" - by (insert bot) (simp add: bottom_def) - -lemma (in lbv) le_bottom [simp]: - "x <=_r \ = (x = \)" - 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 \)" - by (simp add: mrg_def split_beta) - -lemma (in lbv) merge_Err [simp]: - "snd`set ss \ A \ merge c pc ss \ = \" - by (induct ss) auto - -lemma (in lbv) merge_not_top: - "\x. snd`set ss \ A \ merge c pc ss x \ \ \ - \(pc',s') \ set ss. (pc' \ pc+1 \ s' <=_r (c!pc'))" - (is "\x. ?set ss \ ?merge ss x \ ?P ss") -proof (induct ss) - show "?P []" by simp -next - fix x ls l - assume "?set (l#ls)" then obtain set: "snd`set ls \ 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 "\x. ?set ls \ ?merge ls x \ ?P ls" hence "?P ls" using set merge' . - moreover - from merge set - have "pc' \ pc+1 \ 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 - "\x. x \ A \ snd`set ss \ A \ - merge c pc ss x = - (if \(pc',s') \ set ss. pc'\pc+1 \ s' <=_r c!pc' then - map snd [(p',t') \ ss. p'=pc+1] ++_f x - else \)" - (is "\x. _ \ _ \ ?merge ss x = ?if ss x" is "\x. _ \ _ \ ?P ss x") -proof (induct ss) - fix x show "?P [] x" by simp -next - fix x assume x: "x \ A" - fix l::"nat \ 'a" and ls - assume "snd`set (l#ls) \ A" - then obtain l: "snd l \ A" and ls: "snd`set ls \ A" by auto - assume "\x. x \ A \ snd`set ls \ A \ ?P ls x" - hence IH: "\x. x \ A \ ?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 \)" - (is "?merge (l#ls) x = ?merge ls ?if'") - by simp - also have "\ = ?if ls ?if'" - proof - - from l have "s' \ A" by simp - with x have "s' +_f x \ A" by simp - with x T_A have "?if' \ A" by auto - hence "?P ls ?if'" by (rule IH) thus ?thesis by simp - qed - also have "\ = ?if (l#ls) x" - proof (cases "\(pc', s')\set (l#ls). pc'\pc+1 \ s' <=_r c!pc'") - case True - hence "\(pc', s')\set ls. pc'\pc+1 \ s' <=_r c!pc'" by auto - moreover - from True have - "map snd [(p',t')\ls . p'=pc+1] ++_f ?if' = - (map snd [(p',t')\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')\ls . p' = Suc pc]) \ 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 \ A" and ss: "snd`set ss \ A" - assumes m: "merge c pc ss x \ \" - shows "merge c pc ss x = (map snd [(p',t') \ ss. p'=pc+1] ++_f x)" -proof - - from ss m have "\(pc',s') \ set ss. (pc' \ pc+1 \ 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' = \ \ s = \ then \ 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 \ \ = - (wtc c pc s \ \ \ s \ T \ wtl is c (pc+1) (wtc c pc s) \ \)" - by (auto simp del: split_paired_Ex) - -lemma (in lbv) wtl_top [simp]: "wtl ls c pc \ = \" - by (cases ls) auto - -lemma (in lbv) wtl_not_top: - "wtl ls c pc s \ \ \ s \ \" - by (cases "s=\") auto - -lemma (in lbv) wtl_append [simp]: - "\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 \ \ \ wtl (take pc' is) c pc s \ \" - (is "?wtl is \ _ \ _") -proof - - assume "?wtl is \ \" - hence "?wtl (take pc' is @ drop pc' is) \ \" by simp - thus ?thesis by (auto dest!: wtl_not_top simp del: append_take_drop_id) -qed - -lemma take_Suc: - "\n. n < length l \ 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 \ \" - 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 \ \" (is "?wtl is \ _") - assumes pc: "pc < length is" - shows "wtc c pc (wtl (take pc is) c 0 s) \ \" -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) \ \" 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 \ A" and x: "x \ A" - shows "merge c pc ss x \ A" -proof - - from s0 have "set (map snd [(p', t')\ss . p'=pc+1]) \ A" by auto - with x have "(map snd [(p', t')\ss . p'=pc+1] ++_f x) \ 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 \ s \ A \ p < n \ snd`set (step p s) \ A" - by auto (drule pres_typeD) - - -lemma (in lbv) wti_pres [intro?]: - assumes pres: "pres_type step n A" - assumes cert: "c!(pc+1) \ A" - assumes s_pc: "s \ A" "pc < n" - shows "wti c pc s \ A" -proof - - from pres s_pc have "snd`set (step pc s) \ 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 \ A" and cert': "c!(pc+1) \ A" - assumes s: "s \ A" and pc: "pc < n" - shows "wtc c pc s \ A" -proof - - have "wti c pc s \ A" using pres cert' s pc .. - moreover have "wti c pc (c!pc) \ 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) \ \ A" - assumes s: "s \ A" - assumes all: "wtl is c 0 s \ \" - shows "pc < length is \ wtl (take pc is) c 0 s \ A" - (is "?len pc \ ?wtl pc \ A") -proof (induct pc) - from s show "?wtl 0 \ 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 \ ?wtl n \ A" - have "wtc c n (?wtl n) \ A" - using pres _ _ _ n - proof (rule wtc_pres) - from prem n show "?wtl n \ A" . - from cert n show "c!n \ A" by (rule cert_okD1) - from cert n1 show "c!(n+1) \ A" by (rule cert_okD1) - qed - also - from all n have "?wtl n \ \" 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) \ A" by simp -qed - -end diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/Listn.thy --- 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 \ 'a set \ 'a list set" -"list n A == {xs. length xs = n & set xs <= A}" - - le :: "'a ord \ ('a list)ord" -"le r == list_all2 (%x y. x <=_r y)" - -syntax "@lesublist" :: "'a list \ 'a ord \ 'a list \ bool" - ("(_ /<=[_] _)" [50, 0, 51] 50) -syntax "@lesssublist" :: "'a list \ 'a ord \ 'a list \ 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 \ 'b \ 'c) \ 'a list \ 'b list \ 'c list" -"map2 f == (%xs ys. map (split f) (zip xs ys))" - -syntax "@plussublist" :: "'a list \ ('a \ 'b \ 'c) \ 'b list \ 'c list" - ("(_ /+[_] _)" [65, 0, 66] 65) -translations "x +[f] y" == "x +_(map2 f) y" - -consts coalesce :: "'a err list \ 'a list err" -primrec -"coalesce [] = OK[]" -"coalesce (ex#exs) = Err.sup (op #) ex (coalesce exs)" - -constdefs - sl :: "nat \ 'a sl \ 'a list sl" -"sl n == %(A,r,f). (list n A, le r, map2 f)" - - sup :: "('a \ 'b \ 'c err) \ 'a list \ 'b list \ 'c list err" -"sup f == %xs ys. if size xs = size ys then coalesce(xs +[f] ys) else Err" - - upto_esl :: "nat \ 'a esl \ '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 \ - 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: - "\ i \ 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: - "\ xs <=[r] ys; p < size xs \ \ 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 \ xs <=[r] xs" -apply (unfold unfold_lesub_list) -apply (simp add: Listn.le_def list_all2_conv_all_nth) -done - -lemma le_list_trans: - "\ order r; xs <=[r] ys; ys <=[r] zs \ \ 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: - "\ order r; xs <=[r] ys; ys <=[r] xs \ \ 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 \ 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 \ 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 \ size ys = size xs" -apply (unfold lesssub_def) -apply auto -done - -lemma le_list_appendI: - "\b c d. a <=[r] b \ c <=[r] d \ a@c <=[r] b@d" -apply (induct a) - apply simp -apply (case_tac b) -apply auto -done - -lemma le_listI: - "length a = length b \ (\n. n < length a \ a!n <=_r b!n) \ a <=[r] b" - apply (unfold lesub_def Listn.le_def) - apply (simp add: list_all2_conv_all_nth) - done - -lemma listI: - "\ length xs = n; set xs <= A \ \ xs : list n A" -apply (unfold list_def) -apply blast -done - -lemma listE_length [simp]: - "xs : list n A \ length xs = n" -apply (unfold list_def) -apply blast -done - -lemma less_lengthI: - "\ xs : list n A; p < n \ \ p < length xs" - by simp - -lemma listE_set [simp]: - "xs : list n A \ 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) = (\y\ A. \ys\ 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\ A & xs : list n A)"; -apply (simp add: in_list_Suc_iff) -done - -lemma list_not_empty: - "\a. a\ A \ \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 \ set xs <= A \ i < n \ (xs!i) : A" -apply (induct "xs") - apply simp -apply (simp add: nth_Cons split: nat.split) -done - -lemma listE_nth_in: - "\ xs : list n A; i < n \ \ (xs!i) : A" - by auto - - -lemma listn_Cons_Suc [elim!]: - "l#xs \ list n A \ (\n'. n = Suc n' \ l \ A \ xs \ list n' A \ P) \ P" - by (cases n) auto - -lemma listn_appendE [elim!]: - "a@b \ list n A \ (\n1 n2. n=n1+n2 \ a \ list n1 A \ b \ list n2 A \ P) \ P" -proof - - have "\n. a@b \ list n A \ \n1 n2. n=n1+n2 \ a \ list n1 A \ b \ list n2 A" - (is "\n. ?list a n \ \n1 n2. ?P a n n1 n2") - proof (induct a) - fix n assume "?list [] n" - hence "?P [] n 0 n" by simp - thus "\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 \ A" and list_n': "ls@b \ list n' A" by fastsimp - assume "\n. ls @ b \ list n A \ \n1 n2. n = n1 + n2 \ ls \ list n1 A \ b \ list n2 A" - hence "\n1 n2. n' = n1 + n2 \ ls \ list n1 A \ b \ list n2 A" by this (rule list_n') - then obtain n1 n2 where "n' = n1 + n2" "ls \ list n1 A" "b \ list n2 A" by fast - with n have "?P (l#ls) n (n1+1) n2" by simp - thus "\n1 n2. ?P (l#ls) n n1 n2" by fastsimp - qed - moreover - assume "a@b \ list n A" "\n1 n2. n=n1+n2 \ a \ list n1 A \ b \ list n2 A \ P" - ultimately - show ?thesis by blast -qed - - -lemma listt_update_in_list [simp, intro!]: - "\ xs : list n A; x\ A \ \ 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 [] \ [] | y#ys \ (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 \ length ys = n \ i - (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]: - "\ set xs <= A; set ys <= A; size xs = size ys \ - \ 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: - "\set xs <= A; set ys <= A; size xs = size ys \ - \ 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 \ set ys <= A \ set zs <= A - \ size xs = n & size ys = n \ - xs <=[r] zs & ys <=[r] zs \ 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\ A \ set xs <= A \ - (!i. i 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: - "(\y. A y \ False) \ A = bot_class.bot" - by (rule equals0I) (auto simp add: mem_def) - -lemma acc_le_listI [intro!]: - "\ order r; acc r \ \ acc(Listn.le r)" -apply (unfold acc_def) -apply (subgoal_tac - "wfP (SUP n. (\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 "\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. \xs. size xs = k & a#xs:M}" in allE) -apply (erule impE) - apply blast -apply (thin_tac "\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 \ 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: "\L. semilat L \ 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) \ 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: "\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) \ - !xs. xs : list n A \ (!ys. ys : list n A \ - (!zs. coalesce (xs +[f] ys) = OK zs \ 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) \ - !xs. xs : list n A \ (!ys. ys : list n A \ - (!zs. coalesce (xs +[f] ys) = OK 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 (force simp add: semilat_le_err_OK2) -done - -lemma lift2_le_ub: - "\ semilat(err A, Err.le r, lift2 f); x\ A; y\ A; x +_f y = OK z; - u\ A; x <=_r u; y <=_r u \ \ 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) \ - !xs. xs : list n A \ (!ys. ys : list n A \ - (!zs us. coalesce (xs +[f] ys) = OK zs & xs <=[r] us & ys <=[r] us - & us : list n A \ 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: - "\ x +_f y = Err; semilat(err A, Err.le r, lift2 f); x\ A; y\ A \ - \ ~(\u\ 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]: - "\ semilat(err A, Err.le r, lift2 f) \ - \ !xs. xs\ list n A \ (!ys. ys\ list n A \ - coalesce (xs +[f] ys) = Err \ - ~(\zs\ 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) = (\x\ A. \y\ 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) \ - \xs. xs : list n A \ (\ys. ys : list n A \ - 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) \ - 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) \ - 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: - "\L. err_semilat L \ 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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/Opt.thy --- 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 \ 'a option ord" -"le r o1 o2 == case o2 of None \ o1=None | - Some y \ (case o1 of None \ True - | Some x \ x <=_r y)" - - opt :: "'a set \ 'a option set" -"opt A == insert None {x . ? y:A. x = Some y}" - - sup :: "'a ebinop \ 'a option ebinop" -"sup f o1 o2 == - case o1 of None \ OK o2 | Some x \ (case o2 of None \ OK o1 - | Some y \ (case f x y of Err \ Err | OK z \ OK (Some z)))" - - esl :: "'a esl \ '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 \ o1=None | - Some y \ (case o1 of None \ True | Some x \ x <=_r y))" -apply (unfold lesub_def le_def) -apply (rule refl) -done - -lemma le_opt_refl: - "order r \ o1 <=_(le r) o1" -by (simp add: unfold_le_opt split: option.split) - -lemma le_opt_trans [rule_format]: - "order r \ - o1 <=_(le r) o2 \ o2 <=_(le r) o3 \ 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 \ o1 <=_(le r) o2 \ o2 <=_(le r) o1 \ 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 \ 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]: - "\L. err_semilat L \ 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: "\x\?A0. \y\?A0. x <=_?r0 x +_?f0 y" and - ub2: "\x\?A0. \y\?A0. y <=_?r0 x +_?f0 y" and - lub: "\x\?A0. \y\?A0. \z\?A0. x <=_?r0 z \ y <=_?r0 z \ 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: "\c. a = Some c \ c : A" - by (clarsimp simp add: opt_def) - - from ab y - have b: "\d. b = Some d \ 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 \ opt A" "b \ opt A" "a +_(sup f) b = OK c" - moreover - from ord have "order r" by simp - moreover - { fix x y z - assume "x \ A" "y \ A" - hence "OK x \ err A \ OK y \ err A" by simp - with ub1 ub2 - have "(OK x) <=_(Err.le r) (OK x) +_(lift2 f) (OK y) \ - (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 \ y <=_r z" - by (auto simp add: plussub_def lift2_def Err.le_def lesub_def) - } - ultimately - have "a <=_(le r) c \ 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 "(\x\?A. \y\?A. x <=_?r x +_?f y) \ (\x\?A. \y\?A. y <=_?r x +_?f y)" - by (auto simp add: lesub_def plussub_def Err.le_def lift2_def split: err.split) - - moreover - - have "\x\?A. \y\?A. \z\?A. x <=_?r z \ y <=_?r z \ 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 "\ (OK d) <=_(Err.le r) (OK g); (OK e) <=_(Err.le r) (OK g) \ - \ (OK d) +_(lift2 f) (OK e) <=_(Err.le r) (OK g)" - by blast - hence "\ d <=_r g; e <=_r g \ \ \y. d +_f e = OK y \ 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: - "\ order r; top r T \ \ (T <=_r x) = (x = T)" -apply (unfold top_def) -apply (blast intro: order_antisym) -done - - -lemma acc_le_optI [intro!]: - "acc r \ 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: - "\ ox : opt S; !x:S. ox = Some x \ f x : S \ - \ Option.map f ox : opt S"; -apply (unfold Option.map_def) -apply (simp split: option.split) -apply blast -done - -end diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/Product.thy --- 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 \ 'b ord \ ('a * 'b) ord" -"le rA rB == %(a,b) (a',b'). a <=_rA a' & b <=_rB b'" - - sup :: "'a ebinop \ 'b ebinop \ ('a * 'b)ebinop" -"sup f g == %(a1,b1)(a2,b2). Err.sup Pair (a1 +_f a2) (b1 +_g b2)" - - esl :: "'a esl \ 'b esl \ ('a * 'b ) esl" -"esl == %(A,rA,fA) (B,rB,fB). (A <*> B, le rA rB, sup fA fB)" - -syntax "@lesubprod" :: "'a*'b \ 'a ord \ 'b ord \ 'b \ 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!]: - "\ acc rA; acc rB \ \ 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: - "\ closed (err A) (lift2 f); closed (err B) (lift2 g) \ \ - 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: - "\r f z. \ z : err A; semilat (err A, r, f); OK x : err A; OK y : err A; - OK x +_f OK y <=_r z\ \ OK x <=_r z \ 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: - "\L1 L2. \ err_semilat L1; err_semilat L2 \ \ 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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/Semilat.thy --- 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 \ 'a \ bool" - 'a binop = "'a \ 'a \ 'a" - 'a sl = "'a set * 'a ord * 'a binop" - -consts - "@lesub" :: "'a \ 'a ord \ 'a \ bool" ("(_ /<='__ _)" [50, 1000, 51] 50) - "@lesssub" :: "'a \ 'a ord \ 'a \ 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 \ 'a ord \ 'a \ bool" ("(_ /\\<^sub>_ _)" [50, 1000, 51] 50) - -consts - "@plussub" :: "'a \ ('a \ 'b \ 'c) \ 'b \ 'c" ("(_ /+'__ _)" [65, 1000, 66] 65) -defs -plussub_def: "x +_f y == f x y" - -syntax (xsymbols) - "@plussub" :: "'a \ ('a \ 'b \ 'c) \ 'b \ 'c" ("(_ /+\<^sub>_ _)" [65, 1000, 66] 65) - -syntax (xsymbols) - "@plussub" :: "'a \ ('a \ 'b \ 'c) \ 'b \ 'c" ("(_ /\\<^sub>_ _)" [65, 1000, 66] 65) - - -constdefs - order :: "'a ord \ bool" -"order r == (!x. x <=_r x) & - (!x y. x <=_r y & y <=_r x \ x=y) & - (!x y z. x <=_r y & y <=_r z \ x <=_r z)" - - acc :: "'a ord \ bool" -"acc r == wfP (\y x. x <_r y)" - - top :: "'a ord \ 'a \ bool" -"top r T == !x. x <=_r T" - - closed :: "'a set \ 'a binop \ bool" -"closed A f == !x:A. !y:A. x +_f y : A" - - semilat :: "'a sl \ 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 \ x +_f y <=_r z)" - - is_ub :: "'a ord \ 'a \ 'a \ 'a \ bool" -"is_ub r x y u == r x u & r y u" - - is_lub :: "'a ord \ 'a \ 'a \ 'a \ bool" -"is_lub r x y u == is_ub r x y u & (!z. is_ub r x y z \ r u z)" - - some_lub :: "'a ord \ 'a \ 'a \ '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 \ x <=_r x"; - by (simp add: order_def) - -lemma order_antisym: - "\ order r; x <=_r y; y <=_r x \ \ x = y" -apply (unfold order_def) -apply (simp (no_asm_simp)) -done - -lemma order_trans: - "\ order r; x <=_r y; y <=_r z \ \ x <=_r z" -apply (unfold order_def) -apply blast -done - -lemma order_less_irrefl [intro, simp]: - "order r \ ~ x <_r x" -apply (unfold order_def lesssub_def) -apply blast -done - -lemma order_less_trans: - "\ order r; x <_r y; y <_r z \ \ x <_r z" -apply (unfold order_def lesssub_def) -apply blast -done - -lemma topD [simp, intro]: - "top r T \ x <=_r T" - by (simp add: top_def) - -lemma top_le_conv [simp]: - "\ order r; top r T \ \ (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 \ 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: - "\ closed A f; x:A; y:A \ \ 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]: - "\x:A; y:A\ \ 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?]: - "\ x <=_r y; y <=_r x \ \ x = y" - by (rule order_antisym) auto - -lemma (in Semilat) trans_r [trans, intro?]: - "\x <=_r y; y <=_r z\ \ x <=_r z" - by (auto intro: order_trans) - - -lemma (in Semilat) ub1 [simp, intro?]: - "\ x:A; y:A \ \ x <=_r x +_f y" - by (insert semilat) (unfold semilat_Def, simp) - -lemma (in Semilat) ub2 [simp, intro?]: - "\ x:A; y:A \ \ y <=_r x +_f y" - by (insert semilat) (unfold semilat_Def, simp) - -lemma (in Semilat) lub [simp, intro?]: - "\ x <=_r z; y <=_r z; x:A; y:A; z:A \ \ x +_f y <=_r z"; - by (insert semilat) (unfold semilat_Def, simp) - - -lemma (in Semilat) plus_le_conv [simp]: - "\ x:A; y:A; z:A \ \ (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: - "\ x:A; y:A \ \ (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: - "\ x:A; y:A \ \ (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 \ A" and b: "b \ A" and c: "c \ A" - shows "a +_f (b +_f c) = a +_f b +_f c" -proof - - from a b have ab: "a +_f b \ A" .. - from this c have abc: "(a +_f b) +_f c \ A" .. - from b c have bc: "b +_f c \ A" .. - from a this have abc': "a +_f (b +_f c) \ 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 "\ <=_r \ +_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 "\ <=_r \ +_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 "\ <=_r a +_f \" .. - finally have "b<": "b <=_r a +_f (b +_f c)" . - from b c have "c <=_r b +_f c" .. - also from a bc have "\ <=_r a +_f \" .. - 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: - "\a \ A; b \ A\ \ a +_f b <=_r b +_f a" -proof - - assume a: "a \ A" and b: "b \ 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 \ A" .. - ultimately show ?thesis .. -qed - -lemma (in Semilat) plus_commutative: - "\a \ A; b \ A\ \ a +_f b = b +_f a" -by(blast intro: order_antisym plus_com_lemma) - -lemma is_lubD: - "is_lub r x y u \ is_ub r x y u & (!z. is_ub r x y z \ r u z)" - by (simp add: is_lub_def) - -lemma is_ubI: - "\ r x u; r y u \ \ is_ub r x y u" - by (simp add: is_ub_def) - -lemma is_ubD: - "is_ub r x y u \ 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: - "\ single_valuedP r; is_lub (r^** ) x y u; r x' x \ - \ 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]: - "\ single_valuedP r; r^** x u \ \ (!y. r^** y u \ - (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: - "\ acyclicP r; is_lub (r^** ) x y u \ \ 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: - "\ single_valuedP r; acyclicP r; r^** x u; r^** y u \ - \ 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 \ bool) \ ('a \ 'a) \ 'a binop" -"exec_lub r f x y == while (\z. \ r\<^sup>*\<^sup>* x z) f y" - - -lemma acyclic_single_valued_finite: - "\acyclicP r; single_valuedP r; r\<^sup>*\<^sup>* x y \ - \ finite ({(x, y). r x y} \ {a. r\<^sup>*\<^sup>* x a} \ {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} \ {a. r\<^sup>*\<^sup>* x a} \ {b. r\<^sup>*\<^sup>* b y} = - insert (x,x') ({(x, y). r x y} \ {a. r\<^sup>*\<^sup>* x' a} \ {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: - "\ acyclicP r; !x y. r x y \ f x = y; is_lub (r\<^sup>*\<^sup>*) x y u \ \ - exec_lub r f x y = u"; -apply(unfold exec_lub_def) -apply(rule_tac P = "\z. r\<^sup>*\<^sup>* y z \ r\<^sup>*\<^sup>* z u" and - r = "({(x, y). r x y} \ {(a,b). r\<^sup>*\<^sup>* y a \ 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: - "\ single_valuedP r; acyclicP r; r^** x u; r^** y u; !x y. r x y \ f x = y \ - \ is_lub (r^** ) x y (exec_lub r f x y)" - by (fastsimp dest: single_valued_has_lubs simp add: exec_lub_conv) - -end diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/SemilatAlg.thy --- 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 \ 's) list \ 's ord \ (nat \ 's) list \ bool" - ("(_ /<=|_| _)" [50, 0, 51] 50) - "x <=|r| y \ \(p,s) \ set x. \s'. (p,s') \ set y \ s <=_r s'" - -consts - "@plusplussub" :: "'a list \ ('a \ 'a \ 'a) \ 'a \ 'a" ("(_ /++'__ _)" [65, 1000, 66] 65) -primrec - "[] ++_f y = y" - "(x#xs) ++_f y = xs ++_f (x +_f y)" - -constdefs - bounded :: "'s step_type \ nat \ bool" -"bounded step n == !p nat \ 's set \ bool" -"pres_type step n A == \s\A. \p(q,s')\set (step p s). s' \ A" - - mono :: "'s ord \ 's step_type \ nat \ 's set \ bool" -"mono r step n A == - \s p t. s \ A \ p < n \ s <=_r t \ step p s <=|r| step p t" - - -lemma pres_typeD: - "\ pres_type step n A; s\A; pset (step p s) \ \ s' \ A" - by (unfold pres_type_def, blast) - -lemma monoD: - "\ mono r step n A; p < n; s\A; s <=_r t \ \ step p s <=|r| step p t" - by (unfold mono_def, blast) - -lemma boundedD: - "\ bounded step n; p < n; (q,t) : set (step p xs) \ \ q < n" - by (unfold bounded_def, blast) - -lemma lesubstep_type_refl [simp, intro]: - "(\x. x <=_r x) \ x <=|r| x" - by (unfold lesubstep_type_def) auto - -lemma lesub_step_typeD: - "a <=|r| b \ (x,y) \ set a \ \y'. (x, y') \ set b \ y <=_r y'" - by (unfold lesubstep_type_def) blast - - -lemma list_update_le_listI [rule_format]: - "set xs <= A \ set ys <= A \ xs <=[r] ys \ p < size xs \ - x <=_r ys!p \ semilat(A,r,f) \ x\A \ - 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 - "\y. \ set x \ A; y \ A\ \ x ++_f y \ A" (is "PROP ?P") -proof - - interpret Semilat A r f using assms by (rule Semilat.intro) - show "PROP ?P" proof (induct x) - show "\y. y \ A \ [] ++_f y \ A" by simp - fix y x xs - assume y: "y \ A" and xs: "set (x#xs) \ A" - assume IH: "\y. \ set xs \ A; y \ A\ \ xs ++_f y \ A" - from xs obtain x: "x \ A" and xs': "set xs \ A" by simp - from x y have "(x +_f y) \ A" .. - with xs' have "xs ++_f (x +_f y) \ A" by (rule IH) - thus "(x#xs) ++_f y \ A" by simp - qed -qed - -lemma (in Semilat) pp_ub2: - "\y. \ set x \ A; y \ A\ \ y <=_r x ++_f y" -proof (induct x) - from semilat show "\y. y <=_r [] ++_f y" by simp - - fix y a l - assume y: "y \ A" - assume "set (a#l) \ A" - then obtain a: "a \ A" and x: "set l \ A" by simp - assume "\y. \set l \ A; y \ A\ \ y <=_r l ++_f y" - hence IH: "\y. y \ A \ y <=_r l ++_f y" using x . - - from a y have "y <=_r a +_f y" .. - also from a y have "a +_f y \ 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 "\y. \set ls \ A; y \ A; x \ set ls\ \ x <=_r ls ++_f y" -proof (induct ls) - show "\y. x \ set [] \ x <=_r [] ++_f y" by simp - - fix y s ls - assume "set (s#ls) \ A" - then obtain s: "s \ A" and ls: "set ls \ A" by simp - assume y: "y \ A" - - assume - "\y. \set ls \ A; y \ A; x \ set ls\ \ x <=_r ls ++_f y" - hence IH: "\y. x \ set ls \ y \ A \ x <=_r ls ++_f y" using ls . - - assume "x \ set (s#ls)" - then obtain xls: "x = s \ x \ 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 \ 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 \ set ls" - hence "\y. y \ A \ x <=_r ls ++_f y" by (rule IH) - moreover from s y have "s +_f y \ 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 \ A" - shows - "\y. y \ A \ set xs \ A \ \x \ set xs. x <=_r z \ y <=_r z \ 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 \ A" and "set (l#ls) \ A" - then obtain l: "l \ A" and ls: "set ls \ A" by auto - assume "\x \ set (l#ls). x <=_r z" - then obtain lz: "l <=_r z" and lsz: "\x \ 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 \ A" .. - moreover - assume "\y. y \ A \ set ls \ A \ \x \ set ls. x <=_r z \ y <=_r z - \ 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 "\\(p,s) \ set S. s \ A; y \ A; (a,b) \ set S\ - \ b <=_r map snd [(p', t')\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 \ A" - moreover - assume "\(p,s) \ set S. s \ A" - hence "set ?map \ A" by auto - moreover - assume "(a,b) \ set S" - hence "b \ set ?map" by (induct S, auto) - ultimately - show ?thesis by - (rule pp_ub1) -qed - - -lemma plusplus_empty: - "\s'. (q, s') \ set S \ s' +_f ss ! q = ss ! q \ - (map snd [(p', t') \ S. p' = q] ++_f ss ! q) = ss ! q" - by (induct S) auto - -end diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/Typing_Framework.thy --- 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 \ 's \ (nat \ 's) list" - -constdefs - stable :: "'s ord \ 's step_type \ 's list \ nat \ bool" -"stable r step ss p == !(q,s'):set(step p (ss!p)). s' <=_r ss!q" - - stables :: "'s ord \ 's step_type \ 's list \ bool" -"stables r step ss == !p 's \ 's step_type \ 's list \ bool" -"wt_step r T step ts == - !p 's \ 's step_type - \ nat \ 's set \ ('s list \ 's list) \ bool" -"is_bcv r T step n A bcv == !ss : list n A. - (!p nat \ ty \ exception_table \ instr list \ JVMType.state step_type" diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/BV/Typing_Framework_err.thy --- 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 \ 's err step_type \ 's err list \ bool" -"wt_err_step r step ts \ wt_step (Err.le r) Err step ts" - -wt_app_eff :: "'s ord \ (nat \ 's \ bool) \ 's step_type \ 's list \ bool" -"wt_app_eff r app step ts \ - \p < size ts. app p (ts!p) \ (\(q,t) \ set (step p (ts!p)). t <=_r ts!q)" - -map_snd :: "('b \ 'c) \ ('a \ 'b) list \ ('a \ 'c) list" -"map_snd f \ map (\(x,y). (x, f y))" - -error :: "nat \ (nat \ 'a err) list" -"error n \ map (\x. (x,Err)) [0.. (nat \ 's \ bool) \ 's step_type \ 's err step_type" -"err_step n app step p t \ - case t of - Err \ error n - | OK t' \ if app p t' then map_snd OK (step p t') else error n" - -app_mono :: "'s ord \ (nat \ 's \ bool) \ nat \ 's set \ bool" -"app_mono r app n A \ - \s p t. s \ A \ p < n \ s <=_r t \ app p t \ 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 \ - p < n \ app p a \ (q,b) \ set (step p a) \ - 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) \ set (map_snd f xs) \ \b'. (a,b') \ set xs" - apply (induct xs) - apply (auto simp add: map_snd_def) - done - - -lemma bounded_err_stepI: - "\p. p < n \ (\s. ap p s \ (\(q,s') \ set (step p s). q < n)) - \ 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 \ 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]: - "\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 \ 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 \ app_mono r app n A \ bounded (err_step n app step) n \ - \s p t. s \ A \ p < n \ s <=_r t \ app p t \ step p s <=|r| step p t \ - 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) \ set (error n) \ y = Err" - by (auto simp add: error_def) - -lemma pres_type_lift: - "\s\A. \p. p < n \ app p s \ (\(q, s')\set (step p s). s' \ A) - \ 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) \ set (error (size ts))" by (simp add: error_def) - - from wt lp - have [intro?]: "\p. p < size ts \ ts ! p \ 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: "\(q,t) \ 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 "\ app p (map ok_val ts ! p)" - with OKp lp have "\ 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) \ 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 \ Err" .. - ultimately show False by blast - qed - - show "\(q,t)\set(step p (map ok_val ts ! p)). t <=_r map ok_val ts ! q" - proof clarify - fix q t assume q: "(q,t) \ set (step p (map ok_val ts ! p))" - - from wt lp q - obtain s where - OKp: "ts ! p = OK s" and - less: "\(q,t) \ 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 \ 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 \ Err" by simp - { fix q t - assume q: "(q,t) \ set (err_step (size ts) app step p (map OK ts ! p))" - with p app_eff obtain - "app p (ts ! p)" "\(q,t) \ 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 - diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/Comp/AuxLemmas.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/Comp/CorrComp.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/Comp/CorrCompTp.thy --- 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: "\ i < length LT; LT ! i \ Err; mxr = length LT \ \ bc_mt_corresp [Load i] (\(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 \ {x. \y\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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/Comp/DefsComp.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/Comp/Index.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/Comp/LemmasComp.thy --- 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) \ +lemma comp_class_rec: " wf ((subcls1 G)^-1) \ class_rec (comp G) C t f = class_rec G C t (\ 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))\\)") +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) \ (\ 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 \ 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) \ +lemma comp_fields: "wf ((subcls1 G)^-1) \ fields (comp G,C) = fields (G,C)" by (simp add: fields_def comp_class_rec) -lemma comp_field: "wfP ((subcls1 G)^--1) \ +lemma comp_field: "wf ((subcls1 G)^-1) \ field (comp G,C) = field (G,C)" by (simp add: TypeRel.field_def comp_fields) @@ -234,7 +234,7 @@ \ ((class G C) \ None) \ 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 "(\D rT mb. class G x = Some (D, rT, mb))") apply (erule exE)+ diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/Comp/NatCanonify.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/Comp/TranslComp.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/Comp/TranslCompTp.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/Comp/TypeInf.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/DFA/Abstract_BV.thy --- /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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/DFA/Err.thy --- /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 \ 'a \ 'a err" + 'a esl = "'a set * 'a ord * 'a ebinop" + +consts + ok_val :: "'a err \ 'a" +primrec + "ok_val (OK x) = x" + +constdefs + lift :: "('a \ 'b err) \ ('a err \ 'b err)" +"lift f e == case e of Err \ Err | OK x \ f x" + + lift2 :: "('a \ 'b \ 'c err) \ 'a err \ 'b err \ 'c err" +"lift2 f e1 e2 == + case e1 of Err \ Err + | OK x \ (case e2 of Err \ Err | OK y \ f x y)" + + le :: "'a ord \ 'a err ord" +"le r e1 e2 == + case e2 of Err \ True | + OK y \ (case e1 of Err \ False | OK x \ x <=_r y)" + + sup :: "('a \ 'b \ 'c) \ ('a err \ 'b err \ 'c err)" +"sup f == lift2(%x y. OK(x +_f y))" + + err :: "'a set \ 'a err set" +"err A == insert Err {x . ? y:A. x = OK y}" + + esl :: "'a sl \ 'a esl" +"esl == %(A,r,f). (A,r, %x y. OK(f x y))" + + sl :: "'a esl \ 'a err sl" +"sl == %(A,r,f). (err A, le r, lift2 f)" + +syntax + err_semilat :: "'a esl \ bool" +translations +"err_semilat L" == "semilat(Err.sl L)" + + +consts + strict :: "('a \ 'b err) \ ('a err \ 'b err)" +primrec + "strict f Err = Err" + "strict f (OK x) = f x" + +lemma strict_Some [simp]: + "(strict f x = OK y) = (\ z. x = OK z \ f z = OK y)" + by (cases x, auto) + +lemma not_Err_eq: + "(x \ Err) = (\a. x = OK a)" + by (cases x) auto + +lemma not_OK_eq: + "(\y. x \ 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 \ 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 \ e1 <=_(le r) e2 \ e2 <=_(le r) e3 \ 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 \ e1 <=_(le r) e2 \ e2 <=_(le r) e1 \ 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]: + "\L. semilat L \ err_semilat(esl L)" +by(simp add: err_semilat_eslI_aux split_tupled_all) + +lemma acc_err [simp, intro!]: "acc r \ 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: + "\ e : err S; !x:S. e = OK x \ f x : err S \ \ 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]: + "\ x: err A; semilat(err A, le r, f) \ \ 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]: + "\ x: err A; semilat(err A, le r, f) \ \ 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: + "\ x:A; y:A; semilat(err A, le r, f); OK x +_f OK y = OK z \ + \ 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: + "\ x:A; y:A; semilat(err A, le r, f); OK x +_f OK y = OK z \ + \ 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: + "\ x=y; order r \ \ 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: "\A x y z f r. + \ semilat (A,r,f); x +_f y <=_r z; x:A; y:A; z:A \ + \ x <=_r z \ 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) \ P x) = (!y:A. P(f y))" + by blast + +lemma closed_err_Union_lift2I: + "\ !A:AS. closed (err A) (lift2 f); AS ~= {}; + !A:AS.!B:AS. A~=B \ (!a:A.!b:B. a +_f b = Err) \ + \ 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: + "\ !A:AS. err_semilat(A, r, f); AS ~= {}; + !A:AS.!B:AS. A~=B \ (!a:A.!b:B. ~ a <=_r b & a +_f b = Err) \ + \ 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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/DFA/Kildall.thy --- /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 \ 's step_type \ + 's list \ nat set \ 's list \ nat set" + propa :: "'s binop \ (nat \ 's) list \ 's list \ nat set \ '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 \ {}) + (%(ss,w). let p = SOME p. p \ w + in propa f (step p (ss!p)) ss (w-{p})) + (ss,w)" + +constdefs + unstables :: "'s ord \ 's step_type \ 's list \ nat set" +"unstables r step ss == {p. p < size ss \ \stable r step ss p}" + + kildall :: "'s ord \ 's binop \ 's step_type \ 's list \ 's list" +"kildall r f step ss == fst(iter f step ss (unstables r step ss))" + +consts merges :: "'s binop \ (nat \ 's) list \ 's list \ '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: + "\ss. \p < length ss; ss \ list n A; \(p,t)\set ps. p t\A \ \ + (merges f ps ss)!p = map snd [(p',t') \ ps. p'=p] ++_f ss!p" + (is "\ss. \_; _; ?steptype ps\ \ ?P ss ps") +proof (induct ps) + show "\ss. ?P ss []" by simp + + fix ss p' ps' + assume ss: "ss \ list n A" + assume l: "p < length ss" + assume "?steptype (p'#ps')" + then obtain a b where + p': "p'=(a,b)" and ab: "aA" and ps': "?steptype ps'" + by (cases p') auto + assume "\ss. p< length ss \ ss \ list n A \ ?steptype ps' \ ?P ss ps'" + from this [OF _ _ ps'] have IH: "\ss. ss \ list n A \ p < length ss \ ?P ss ps'" . + + from ss ab + have "ss[a := b +_f ss!a] \ 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]: + "\ss. size(merges f ps ss) = size ss" + by (induct_tac ps, auto) + + +lemma (in Semilat) merges_preserves_type_lemma: +shows "\xs. xs \ list n A \ (\(p,x) \ set ps. p x\A) + \ merges f ps xs \ 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]: + "\ xs \ list n A; \(p,x) \ set ps. p x\A \ + \ merges f ps xs \ list n A" +by (simp add: merges_preserves_type_lemma) + +lemma (in Semilat) merges_incr_lemma: + "\xs. xs \ list n A \ (\(p,x)\set ps. p x \ A) \ 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: + "\ xs \ list n A; \(p,x)\set ps. p x \ A \ + \ xs <=[r] merges f ps xs" + by (simp add: merges_incr_lemma) + + +lemma (in Semilat) merges_same_conv [rule_format]: + "(\xs. xs \ list n A \ (\(p,x)\set ps. p x\A) \ + (merges f ps xs = xs) = (\(p,x)\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 \ set ys <= A \ xs <=[r] ys \ p < size xs \ + x <=_r ys!p \ x\A \ 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 "\(p,t)\set ps. t <=_r ts!p \ t \ A \ p < size ts" and "ss <=[r] ts" + shows "merges f ps ss <=[r] ts" +proof - + { fix t ts ps + have + "\qs. \set ts <= A; \(p,t)\set ps. t <=_r ts!p \ t \ A \ p< size ts \ \ + set qs <= set ps \ + (\ss. set ss <= A \ ss <=[r] ts \ 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: + "\ss w. (\(q,t)\set qs. q < size ss) \ + propa f qs ss w = + (merges f qs ss, {q. \t. (q,t)\set qs \ t +_f ss!q \ 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 "\pres_type step n A; bounded step n; + ss \ list n A; p \ w; \q\w. q < n; + \q. q < n \ q \ w \ stable r step ss q; q < n; + \s'. (q,s') \ set (step p (ss ! p)) \ s' +_f ss ! q = ss ! q; + q \ w \ q = p \ + \ stable r step (merges f (step p (ss!p)) ss) q" + apply (unfold stable_def) + apply (subgoal_tac "\s'. (q,s') \ set (step p (ss!p)) \ 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 = "\x. (a, b) \ set (step q x)" in subst) + apply assumption + + apply (simp add: plusplus_empty) + apply (cases "q \ 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: + "\ mono r step n A; bounded step n; + \(p',s') \ set (step p (ss!p)). s' \ A; ss \ list n A; ts \ list n A; p < n; + ss <=[r] ts; \p. p < n \ stable r step ts p \ + \ 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 "\ ss \ list n A; \(q,t)\set qs. q t\A; p\w \ \ + ss <[r] merges f qs ss \ + merges f qs ss = ss \ {q. \t. (q,t)\set qs \ t +_f ss!q \ 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 "\q t. \((q, t) \ set qs \ t +_f ss ! q \ 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 "\ acc r ; pres_type step n A; mono r step n A; + bounded step n; \p\w0. p < n; ss0 \ list n A; + \p w0 \ stable r step ss0 p \ \ + iter f step ss0 w0 = (ss',w') + \ + ss' \ list n A \ stables r step ss' \ ss0 <=[r] ss' \ + (\ts\list n A. ss0 <=[r] ts \ stables r step ts \ 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 \ list n A \ (\p w \ stable r step ss p) \ ss0 <=[r] ss \ + (\ts\list n A. ss0 <=[r] ts \ stables r step ts \ ss <=[r] ts) \ + (\p\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 \ w) \ w") + prefer 2; apply (fast intro: someI) +apply(subgoal_tac "\(q,t) \ set (step (SOME p. p \ w) (ss ! (SOME p. p \ w))). q < length ss \ t \ 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 \ w) \ w") + prefer 2; apply (fast intro: someI) +apply(subgoal_tac "\(q,t) \ set (step (SOME p. p \ w) (ss ! (SOME p. p \ w))). q < length ss \ t \ 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 "\ acc r; pres_type step n A; mono r step n A; + bounded step n; ss0 \ list n A \ \ + kildall r f step ss0 \ list n A \ + stables r step (kildall r f step ss0) \ + ss0 <=[r] kildall r f step ss0 \ + (\ts\list n A. ss0 <=[r] ts \ stables r step ts \ + 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 "\ acc r; top r T; pres_type step n A; bounded step n; mono r step n A \ + \ 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 \ 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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/DFA/LBVComplete.thy --- /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] \ bool" + "is_target step phi pc' \ + \pc s'. pc' \ pc+1 \ pc < length phi \ (pc',s') \ set (step pc (phi!pc))" + + make_cert :: "['s step_type, 's list, 's] \ 's certificate" + "make_cert step phi B \ + map (\pc. if is_target step phi pc then phi!pc else B) [0..pc. pc' \ pc+1 \ pc' mem (map fst (step pc (phi!pc)))) [0..") + fixes c :: "'a list" + defines cert_def: "c \ make_cert step \ \" + + assumes mono: "mono r step (length \) A" + assumes pres: "pres_type step (length \) A" + assumes phi: "\pc < length \. \!pc \ A \ \!pc \ \" + assumes bounded: "bounded step (length \)" + + assumes B_neq_T: "\ \ \" + + +lemma (in lbvc) cert: "cert_ok c (length \) \ \ A" +proof (unfold cert_ok_def, intro strip conjI) + note [simp] = make_cert_def cert_def nth_append + + show "c!length \ = \" by simp + + fix pc assume pc: "pc < length \" + from pc phi B_A show "c!pc \ A" by simp + from pc phi B_neq_T show "c!pc \ \" by simp +qed + +lemmas [simp del] = split_paired_Ex + + +lemma (in lbvc) cert_target [intro?]: + "\ (pc',s') \ set (step pc (\!pc)); + pc' \ pc+1; pc < length \; pc' < length \ \ + \ c!pc' = \!pc'" + by (auto simp add: cert_def make_cert_def nth_append is_target_def) + + +lemma (in lbvc) cert_approx [intro?]: + "\ pc < length \; c!pc \ \ \ + \ c!pc = \!pc" + by (auto simp add: cert_def make_cert_def nth_append) + + +lemma (in lbv) le_top [simp, intro]: + "x <=_r \" + by (insert top) simp + + +lemma (in lbv) merge_mono: + assumes less: "ss2 <=|r| ss1" + assumes x: "x \ A" + assumes ss1: "snd`set ss1 \ A" + assumes ss2: "snd`set ss2 \ A" + shows "merge c pc ss2 x <=_r merge c pc ss1 x" (is "?s2 <=_r ?s1") +proof- + have "?s1 = \ \ ?thesis" by simp + moreover { + assume merge: "?s1 \ T" + from x ss1 have "?s1 = + (if \(pc', s')\set ss1. pc' \ pc + 1 \ s' <=_r c!pc' + then (map snd [(p', t') \ ss1 . p'=pc+1]) ++_f x + else \)" + by (rule merge_def) + with merge obtain + app: "\(pc',s')\set ss1. pc' \ pc+1 \ s' <=_r c!pc'" + (is "?app ss1") and + sum: "(map snd [(p',t') \ 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) \ A" by auto + with x have "?sum ss1 \ A" by (auto intro!: plusplus_closed semilat) + with sum have "?s1 \ A" by simp + moreover + have mapD: "\x ss. x \ set (?map ss) \ \p. (p,x) \ set ss \ p=pc+1" by auto + from x map1 + have "\x \ set (?map ss1). x <=_r ?sum ss1" + by clarify (rule pp_ub1) + with sum have "\x \ set (?map ss1). x <=_r ?s1" by simp + with less have "\x \ 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) \ A" by auto + ultimately + have "?sum ss2 <=_r ?s1" using x by - (rule pp_lub) + } + moreover + from x ss2 have + "?s2 = + (if \(pc', s')\set ss2. pc' \ pc + 1 \ s' <=_r c!pc' + then map snd [(p', t') \ ss2 . p' = pc + 1] ++_f x + else \)" + by (rule merge_def) + ultimately have ?thesis by simp + } + ultimately show ?thesis by (cases "?s1 = \") auto +qed + + +lemma (in lbvc) wti_mono: + assumes less: "s2 <=_r s1" + assumes pc: "pc < length \" + assumes s1: "s1 \ A" + assumes s2: "s2 \ 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 \ A" by (rule cert_okD3) + moreover + from pres s1 pc + have "snd`set (step pc s1) \ A" by (rule pres_typeD2) + moreover + from pres s2 pc + have "snd`set (step pc s2) \ 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 \" + assumes s1: "s1 \ A" + assumes s2: "s2 \ A" + shows "wtc c pc s2 <=_r wtc c pc s1" (is "?s2' <=_r ?s1'") +proof (cases "c!pc = \") + 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' = \ \ ?thesis" by simp + moreover { + assume "?s1' \ \" + 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' = \") auto +qed + + +lemma (in lbv) top_le_conv [simp]: + "\ <=_r x = (x = \)" + by (insert semilat) (simp add: top top_le_conv) + +lemma (in lbv) neq_top [simp, elim]: + "\ x <=_r y; y \ \ \ \ x \ \" + by (cases "x = T") auto + + +lemma (in lbvc) stable_wti: + assumes stable: "stable r step \ pc" + assumes pc: "pc < length \" + shows "wti c pc (\!pc) \ \" +proof - + let ?step = "step pc (\!pc)" + from stable + have less: "\(q,s')\set ?step. s' <=_r \!q" by (simp add: stable_def) + + from cert B_A pc + have cert_suc: "c!Suc pc \ A" by (rule cert_okD3) + moreover + from phi pc have "\!pc \ A" by simp + from pres this pc + have stepA: "snd`set ?step \ A" by (rule pres_typeD2) + ultimately + have "merge c pc ?step (c!Suc pc) = + (if \(pc',s')\set ?step. pc'\pc+1 \ s' <=_r c!pc' + then map snd [(p',t') \ ?step.p'=pc+1] ++_f c!Suc pc + else \)" unfolding mrg_def by (rule lbv.merge_def [OF lbvc.axioms(1), OF lbvc_axioms]) + moreover { + fix pc' s' assume s': "(pc', s') \ set ?step" and suc_pc: "pc' \ pc+1" + with less have "s' <=_r \!pc'" by auto + also + from bounded pc s' have "pc' < length \" by (rule boundedD) + with s' suc_pc pc have "c!pc' = \!pc'" .. + hence "\!pc' = c!pc'" .. + finally have "s' <=_r c!pc'" . + } hence "\(pc',s')\set ?step. pc'\pc+1 \ s' <=_r c!pc'" by auto + moreover + from pc have "Suc pc = length \ \ Suc pc < length \" by auto + hence "map snd [(p',t') \ ?step.p'=pc+1] ++_f c!Suc pc \ \" + (is "?map ++_f _ \ _") + proof (rule disjE) + assume pc': "Suc pc = length \" + with cert have "c!Suc pc = \" by (simp add: cert_okD2) + moreover + from pc' bounded pc + have "\(p',t')\set ?step. p'\pc+1" by clarify (drule boundedD, auto) + hence "[(p',t') \ ?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 \" + from pc' phi have "\!Suc pc \ A" by simp + moreover note cert_suc + moreover from stepA + have "set ?map \ A" by auto + moreover + have "\s. s \ set ?map \ \t. (Suc pc, t) \ set ?step" by auto + with less have "\s' \ set ?map. s' <=_r \!Suc pc" by auto + moreover + from pc' have "c!Suc pc <=_r \!Suc pc" + by (cases "c!Suc pc = \") (auto dest: cert_approx) + ultimately + have "?map ++_f c!Suc pc <=_r \!Suc pc" by (rule pp_lub) + moreover + from pc' phi have "\!Suc pc \ \" by simp + ultimately + show ?thesis by auto + qed + ultimately + have "merge c pc ?step (c!Suc pc) \ \" by simp + thus ?thesis by (simp add: wti) +qed + +lemma (in lbvc) wti_less: + assumes stable: "stable r step \ pc" + assumes suc_pc: "Suc pc < length \" + shows "wti c pc (\!pc) <=_r \!Suc pc" (is "?wti <=_r _") +proof - + let ?step = "step pc (\!pc)" + + from stable + have less: "\(q,s')\set ?step. s' <=_r \!q" by (simp add: stable_def) + + from suc_pc have pc: "pc < length \" by simp + with cert B_A have cert_suc: "c!Suc pc \ A" by (rule cert_okD3) + moreover + from phi pc have "\!pc \ A" by simp + with pres pc have stepA: "snd`set ?step \ A" by - (rule pres_typeD2) + moreover + from stable pc have "?wti \ \" by (rule stable_wti) + hence "merge c pc ?step (c!Suc pc) \ \" by (simp add: wti) + ultimately + have "merge c pc ?step (c!Suc pc) = + map snd [(p',t')\ ?step.p'=pc+1] ++_f c!Suc pc" by (rule merge_not_top_s) + hence "?wti = \" (is "_ = (?map ++_f _)" is "_ = ?sum") by (simp add: wti) + also { + from suc_pc phi have "\!Suc pc \ A" by simp + moreover note cert_suc + moreover from stepA have "set ?map \ A" by auto + moreover + have "\s. s \ set ?map \ \t. (Suc pc, t) \ set ?step" by auto + with less have "\s' \ set ?map. s' <=_r \!Suc pc" by auto + moreover + from suc_pc have "c!Suc pc <=_r \!Suc pc" + by (cases "c!Suc pc = \") (auto dest: cert_approx) + ultimately + have "?sum <=_r \!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 \" + shows "wtc c pc (\!pc) \ \" +proof - + from stable pc have wti: "wti c pc (\!pc) \ \" by (rule stable_wti) + show ?thesis + proof (cases "c!pc = \") + case True with wti show ?thesis by (simp add: wtc) + next + case False + with pc have "c!pc = \!pc" .. + with False wti show ?thesis by (simp add: wtc) + qed +qed + +lemma (in lbvc) wtc_less: + assumes stable: "stable r step \ pc" + assumes suc_pc: "Suc pc < length \" + shows "wtc c pc (\!pc) <=_r \!Suc pc" (is "?wtc <=_r _") +proof (cases "c!pc = \") + case True + moreover from stable suc_pc have "wti c pc (\!pc) <=_r \!Suc pc" + by (rule wti_less) + ultimately show ?thesis by (simp add: wtc) +next + case False + from suc_pc have pc: "pc < length \" by simp + with stable have "?wtc \ \" 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 = \!pc" .. + finally have "?wtc = wti c pc (\!pc)" . + also from stable suc_pc have "wti c pc (\!pc) <=_r \!Suc pc" by (rule wti_less) + finally show ?thesis . +qed + + +lemma (in lbvc) wt_step_wtl_lemma: + assumes wt_step: "wt_step r \ step \" + shows "\pc s. pc+length ls = length \ \ s <=_r \!pc \ s \ A \ s\\ \ + wtl ls c pc s \ \" + (is "\pc s. _ \ _ \ _ \ _ \ ?wtl ls pc s \ _") +proof (induct ls) + fix pc s assume "s\\" thus "?wtl [] pc s \ \" by simp +next + fix pc s i ls + assume "\pc s. pc+length ls=length \ \ s <=_r \!pc \ s \ A \ s\\ \ + ?wtl ls pc s \ \" + moreover + assume pc_l: "pc + length (i#ls) = length \" + hence suc_pc_l: "Suc pc + length ls = length \" by simp + ultimately + have IH: "\s. s <=_r \!Suc pc \ s \ A \ s \ \ \ ?wtl ls (Suc pc) s \ \" . + + from pc_l obtain pc: "pc < length \" by simp + with wt_step have stable: "stable r step \ pc" by (simp add: wt_step_def) + from this pc have wt_phi: "wtc c pc (\!pc) \ \" by (rule stable_wtc) + assume s_phi: "s <=_r \!pc" + from phi pc have phi_pc: "\!pc \ A" by simp + assume s: "s \ A" + with s_phi pc phi_pc have wt_s_phi: "wtc c pc s <=_r wtc c pc (\!pc)" by (rule wtc_mono) + with wt_phi have wt_s: "wtc c pc s \ \" by simp + moreover + assume s': "s \ \" + ultimately + have "ls = [] \ ?wtl (i#ls) pc s \ \" by simp + moreover { + assume "ls \ []" + with pc_l have suc_pc: "Suc pc < length \" by (auto simp add: neq_Nil_conv) + with stable have "wtc c pc (phi!pc) <=_r \!Suc pc" by (rule wtc_less) + with wt_s_phi have "wtc c pc s <=_r \!Suc pc" by (rule trans_r) + moreover + from cert suc_pc have "c!pc \ A" "c!(pc+1) \ A" + by (auto simp add: cert_ok_def) + from pres this s pc have "wtc c pc s \ A" by (rule wtc_pres) + ultimately + have "?wtl ls (Suc pc) (wtc c pc s) \ \" using IH wt_s by blast + with s' wt_s have "?wtl (i#ls) pc s \ \" by simp + } + ultimately show "?wtl (i#ls) pc s \ \" by (cases ls) blast+ +qed + + +theorem (in lbvc) wtl_complete: + assumes wt: "wt_step r \ step \" + and s: "s <=_r \!0" "s \ A" "s \ \" + and len: "length ins = length phi" + shows "wtl ins c 0 s \ \" +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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/DFA/LBVCorrect.thy --- /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" ("\") + defines phi_def: + "\ \ map (\pc. if c!pc = \ then wtl (take pc ins) c 0 s0 else c!pc) + [0.. \ A" + assumes pres: "pres_type step (length ins) A" + + +lemma (in lbvs) phi_None [intro?]: + "\ pc < length ins; c!pc = \ \ \ \ ! pc = wtl (take pc ins) c 0 s0" + by (simp add: phi_def) + +lemma (in lbvs) phi_Some [intro?]: + "\ pc < length ins; c!pc \ \ \ \ \ ! pc = c ! pc" + by (simp add: phi_def) + +lemma (in lbvs) phi_len [simp]: + "length \ = length ins" + by (simp add: phi_def) + + +lemma (in lbvs) wtl_suc_pc: + assumes all: "wtl ins c 0 s\<^sub>0 \ \" + assumes pc: "pc+1 < length ins" + shows "wtl (take (pc+1) ins) c 0 s0 \\<^sub>r \!(pc+1)" +proof - + from all pc + have "wtc c (pc+1) (wtl (take (pc+1) ins) c 0 s0) \ 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 \ \" + assumes s0: "s0 \ A" + assumes pc: "pc < length ins" + shows "stable r step \ pc" +proof (unfold stable_def, clarify) + fix pc' s' assume step: "(pc',s') \ set (step pc (\ ! pc))" + (is "(pc',s') \ 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 \ \" (is "?s1 \ _") by (rule wtl_take) + from wtl have s2: "wtl (take (pc+1) ins) c 0 s0 \ \" (is "?s2 \ _") by (rule wtl_take) + + from wtl pc have wt_s1: "wtc c pc ?s1 \ \" by (rule wtl_all) + + have c_Some: "\pc t. pc < length ins \ c!pc \ \ \ \!pc = c!pc" + by (simp add: phi_def) + from pc have c_None: "c!pc = \ \ \!pc = ?s1" .. + + from wt_s1 pc c_None c_Some + have inst: "wtc c pc ?s1 = wti c pc (\!pc)" + by (simp add: wtc split: split_if_asm) + + from pres cert s0 wtl pc have "?s1 \ A" by (rule wtl_pres) + with pc c_Some cert c_None + have "\!pc \ A" by (cases "c!pc = \") (auto dest: cert_okD1) + with pc pres + have step_in_A: "snd`set (?step pc) \ A" by (auto dest: pres_typeD2) + + show "s' <=_r \!pc'" + proof (cases "pc' = pc+1") + case True + with pc' cert + have cert_in_A: "c!(pc+1) \ 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 "\ \ \" (is "?merge \ _") by simp + with cert_in_A step_in_A + have "?merge = (map snd [(p',t') \ ?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 \!(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)) \ \" by (simp add: wti) + with step_in_A + have "\(pc', s')\set (?step pc). pc'\pc+1 \ 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' = \ \ s' = \" by simp + moreover + from c_Some pc' + have "c!pc' \ \ \ \!pc' = c!pc'" by auto + ultimately + show ?thesis by (cases "c!pc' = \") auto + qed +qed + + +lemma (in lbvs) phi_not_top: + assumes wtl: "wtl ins c 0 s0 \ \" + assumes pc: "pc < length ins" + shows "\!pc \ \" +proof (cases "c!pc = \") + case False with pc + have "\!pc = c!pc" .. + also from cert pc have "\ \ \" by (rule cert_okD4) + finally show ?thesis . +next + case True with pc + have "\!pc = wtl (take pc ins) c 0 s0" .. + also from wtl have "\ \ \" by (rule wtl_take) + finally show ?thesis . +qed + +lemma (in lbvs) phi_in_A: + assumes wtl: "wtl ins c 0 s0 \ \" + assumes s0: "s0 \ A" + shows "\ \ list (length ins) A" +proof - + { fix x assume "x \ set \" + then obtain xs ys where "\ = xs @ x # ys" + by (auto simp add: in_set_conv_decomp) + then obtain pc where pc: "pc < length \" and x: "\!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 \ A" by (auto intro!: wtl_pres) + moreover + from pc have "pc < length ins" by simp + with cert have "c!pc \ A" .. + ultimately + have "\!pc \ A" using pc by (simp add: phi_def) + hence "x \ A" using x by simp + } + hence "set \ \ A" .. + thus ?thesis by (unfold list_def) simp +qed + + +lemma (in lbvs) phi0: + assumes wtl: "wtl ins c 0 s0 \ \" + assumes 0: "0 < length ins" + shows "s0 <=_r \!0" +proof (cases "c!0 = \") + case True + with 0 have "\!0 = wtl (take 0 ins) c 0 s0" .. + moreover have "wtl (take 0 ins) c 0 s0 = s0" by simp + ultimately have "\!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 \ \" 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 \ \" + assumes s0: "s0 \ A" + shows "\ts. wt_step r \ step ts" +proof - + have "wt_step r \ step \" + proof (unfold wt_step_def, intro strip conjI) + fix pc assume "pc < length \" + then have pc: "pc < length ins" by simp + with wtl show "\!pc \ \" by (rule phi_not_top) + from wtl s0 pc show "stable r step \ pc" by (rule wtl_stable) + qed + thus ?thesis .. +qed + + +theorem (in lbvs) wtl_sound_strong: + assumes wtl: "wtl ins c 0 s0 \ \" + assumes s0: "s0 \ A" + assumes nz: "0 < length ins" + shows "\ts \ list (length ins) A. wt_step r \ step ts \ s0 <=_r ts!0" +proof - + from wtl s0 have "\ \ list (length ins) A" by (rule phi_in_A) + moreover + have "wt_step r \ step \" + proof (unfold wt_step_def, intro strip conjI) + fix pc assume "pc < length \" + then have pc: "pc < length ins" by simp + with wtl show "\!pc \ \" by (rule phi_not_top) + from wtl s0 pc show "stable r step \ pc" by (rule wtl_stable) + qed + moreover + from wtl nz have "s0 <=_r \!0" by (rule phi0) + ultimately + show ?thesis by fast +qed + +end diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/DFA/LBVSpec.thy --- /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 \ 's binop \ 's ord \ 's \ nat \ (nat \ 's) list \ 's \ '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 \ 's binop \ 's ord \ 's \ + 's step_type \ nat \ 's \ 's" +"wtl_inst cert f r T step pc s \ merge cert f r T pc (step pc s) (cert!(pc+1))" + +wtl_cert :: "'s certificate \ 's binop \ 's ord \ 's \ 's \ + 's step_type \ nat \ 's \ 's" +"wtl_cert cert f r T B step pc s \ + 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 \ 's certificate \ 's binop \ 's ord \ 's \ 's \ + 's step_type \ nat \ 's \ '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 \ s = T then T else wtl_inst_list is cert f r T B step (pc+1) s')" + +constdefs + cert_ok :: "'s certificate \ nat \ 's \ 's \ 's set \ bool" + "cert_ok cert n T B A \ (\i < n. cert!i \ A \ cert!i \ T) \ (cert!n = B)" + +constdefs + bottom :: "'a ord \ 'a \ bool" + "bottom r B \ \x. B <=_r x" + + +locale lbv = Semilat + + fixes T :: "'a" ("\") + fixes B :: "'a" ("\") + fixes step :: "'a step_type" + assumes top: "top r \" + assumes T_A: "\ \ A" + assumes bot: "bottom r \" + assumes B_A: "\ \ A" + + fixes merge :: "'a certificate \ nat \ (nat \ 'a) list \ 'a \ 'a" + defines mrg_def: "merge cert \ LBVSpec.merge cert f r \" + + fixes wti :: "'a certificate \ nat \ 'a \ 'a" + defines wti_def: "wti cert \ wtl_inst cert f r \ step" + + fixes wtc :: "'a certificate \ nat \ 'a \ 'a" + defines wtc_def: "wtc cert \ wtl_cert cert f r \ \ step" + + fixes wtl :: "'b list \ 'a certificate \ nat \ 'a \ 'a" + defines wtl_def: "wtl ins cert \ wtl_inst_list ins cert f r \ \ step" + + +lemma (in lbv) wti: + "wti c pc s \ 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 \ if c!pc = \ then wti c pc s else if s <=_r c!pc then wti c pc (c!pc) else \" + by (unfold wtc_def wti_def wtl_cert_def) + + +lemma cert_okD1 [intro?]: + "cert_ok c n T B A \ pc < n \ c!pc \ A" + by (unfold cert_ok_def) fast + +lemma cert_okD2 [intro?]: + "cert_ok c n T B A \ c!n = B" + by (simp add: cert_ok_def) + +lemma cert_okD3 [intro?]: + "cert_ok c n T B A \ B \ A \ pc < n \ c!Suc pc \ 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 \ pc < n \ c!pc \ 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 \ A" + shows "x +_f \ = \" +proof - + from top have "x +_f \ <=_r \" .. + moreover from x T_A have "\ <=_r x +_f \" .. + ultimately show ?thesis .. +qed + +lemma (in lbv) plusplussup_top [simp, elim]: + "set xs \ A \ xs ++_f \ = \" + by (induct xs) auto + + + +lemma (in Semilat) pp_ub1': + assumes S: "snd`set S \ A" + assumes y: "y \ A" and ab: "(a, b) \ set S" + shows "b <=_r map snd [(p', t') \ S . p' = a] ++_f y" +proof - + from S have "\(x,y) \ set S. y \ A" by auto + with semilat y ab show ?thesis by - (rule ub1') +qed + +lemma (in lbv) bottom_le [simp, intro]: + "\ <=_r x" + by (insert bot) (simp add: bottom_def) + +lemma (in lbv) le_bottom [simp]: + "x <=_r \ = (x = \)" + 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 \)" + by (simp add: mrg_def split_beta) + +lemma (in lbv) merge_Err [simp]: + "snd`set ss \ A \ merge c pc ss \ = \" + by (induct ss) auto + +lemma (in lbv) merge_not_top: + "\x. snd`set ss \ A \ merge c pc ss x \ \ \ + \(pc',s') \ set ss. (pc' \ pc+1 \ s' <=_r (c!pc'))" + (is "\x. ?set ss \ ?merge ss x \ ?P ss") +proof (induct ss) + show "?P []" by simp +next + fix x ls l + assume "?set (l#ls)" then obtain set: "snd`set ls \ 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 "\x. ?set ls \ ?merge ls x \ ?P ls" hence "?P ls" using set merge' . + moreover + from merge set + have "pc' \ pc+1 \ 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 + "\x. x \ A \ snd`set ss \ A \ + merge c pc ss x = + (if \(pc',s') \ set ss. pc'\pc+1 \ s' <=_r c!pc' then + map snd [(p',t') \ ss. p'=pc+1] ++_f x + else \)" + (is "\x. _ \ _ \ ?merge ss x = ?if ss x" is "\x. _ \ _ \ ?P ss x") +proof (induct ss) + fix x show "?P [] x" by simp +next + fix x assume x: "x \ A" + fix l::"nat \ 'a" and ls + assume "snd`set (l#ls) \ A" + then obtain l: "snd l \ A" and ls: "snd`set ls \ A" by auto + assume "\x. x \ A \ snd`set ls \ A \ ?P ls x" + hence IH: "\x. x \ A \ ?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 \)" + (is "?merge (l#ls) x = ?merge ls ?if'") + by simp + also have "\ = ?if ls ?if'" + proof - + from l have "s' \ A" by simp + with x have "s' +_f x \ A" by simp + with x T_A have "?if' \ A" by auto + hence "?P ls ?if'" by (rule IH) thus ?thesis by simp + qed + also have "\ = ?if (l#ls) x" + proof (cases "\(pc', s')\set (l#ls). pc'\pc+1 \ s' <=_r c!pc'") + case True + hence "\(pc', s')\set ls. pc'\pc+1 \ s' <=_r c!pc'" by auto + moreover + from True have + "map snd [(p',t')\ls . p'=pc+1] ++_f ?if' = + (map snd [(p',t')\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')\ls . p' = Suc pc]) \ 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 \ A" and ss: "snd`set ss \ A" + assumes m: "merge c pc ss x \ \" + shows "merge c pc ss x = (map snd [(p',t') \ ss. p'=pc+1] ++_f x)" +proof - + from ss m have "\(pc',s') \ set ss. (pc' \ pc+1 \ 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' = \ \ s = \ then \ 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 \ \ = + (wtc c pc s \ \ \ s \ T \ wtl is c (pc+1) (wtc c pc s) \ \)" + by (auto simp del: split_paired_Ex) + +lemma (in lbv) wtl_top [simp]: "wtl ls c pc \ = \" + by (cases ls) auto + +lemma (in lbv) wtl_not_top: + "wtl ls c pc s \ \ \ s \ \" + by (cases "s=\") auto + +lemma (in lbv) wtl_append [simp]: + "\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 \ \ \ wtl (take pc' is) c pc s \ \" + (is "?wtl is \ _ \ _") +proof - + assume "?wtl is \ \" + hence "?wtl (take pc' is @ drop pc' is) \ \" by simp + thus ?thesis by (auto dest!: wtl_not_top simp del: append_take_drop_id) +qed + +lemma take_Suc: + "\n. n < length l \ 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 \ \" + 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 \ \" (is "?wtl is \ _") + assumes pc: "pc < length is" + shows "wtc c pc (wtl (take pc is) c 0 s) \ \" +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) \ \" 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 \ A" and x: "x \ A" + shows "merge c pc ss x \ A" +proof - + from s0 have "set (map snd [(p', t')\ss . p'=pc+1]) \ A" by auto + with x have "(map snd [(p', t')\ss . p'=pc+1] ++_f x) \ 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 \ s \ A \ p < n \ snd`set (step p s) \ A" + by auto (drule pres_typeD) + + +lemma (in lbv) wti_pres [intro?]: + assumes pres: "pres_type step n A" + assumes cert: "c!(pc+1) \ A" + assumes s_pc: "s \ A" "pc < n" + shows "wti c pc s \ A" +proof - + from pres s_pc have "snd`set (step pc s) \ 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 \ A" and cert': "c!(pc+1) \ A" + assumes s: "s \ A" and pc: "pc < n" + shows "wtc c pc s \ A" +proof - + have "wti c pc s \ A" using pres cert' s pc .. + moreover have "wti c pc (c!pc) \ 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) \ \ A" + assumes s: "s \ A" + assumes all: "wtl is c 0 s \ \" + shows "pc < length is \ wtl (take pc is) c 0 s \ A" + (is "?len pc \ ?wtl pc \ A") +proof (induct pc) + from s show "?wtl 0 \ 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 \ ?wtl n \ A" + have "wtc c n (?wtl n) \ A" + using pres _ _ _ n + proof (rule wtc_pres) + from prem n show "?wtl n \ A" . + from cert n show "c!n \ A" by (rule cert_okD1) + from cert n1 show "c!(n+1) \ A" by (rule cert_okD1) + qed + also + from all n have "?wtl n \ \" 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) \ A" by simp +qed + +end diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/DFA/Listn.thy --- /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 \ 'a set \ 'a list set" +"list n A == {xs. length xs = n & set xs <= A}" + + le :: "'a ord \ ('a list)ord" +"le r == list_all2 (%x y. x <=_r y)" + +syntax "@lesublist" :: "'a list \ 'a ord \ 'a list \ bool" + ("(_ /<=[_] _)" [50, 0, 51] 50) +syntax "@lesssublist" :: "'a list \ 'a ord \ 'a list \ 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 \ 'b \ 'c) \ 'a list \ 'b list \ 'c list" +"map2 f == (%xs ys. map (split f) (zip xs ys))" + +syntax "@plussublist" :: "'a list \ ('a \ 'b \ 'c) \ 'b list \ 'c list" + ("(_ /+[_] _)" [65, 0, 66] 65) +translations "x +[f] y" == "x +_(map2 f) y" + +consts coalesce :: "'a err list \ 'a list err" +primrec +"coalesce [] = OK[]" +"coalesce (ex#exs) = Err.sup (op #) ex (coalesce exs)" + +constdefs + sl :: "nat \ 'a sl \ 'a list sl" +"sl n == %(A,r,f). (list n A, le r, map2 f)" + + sup :: "('a \ 'b \ 'c err) \ 'a list \ 'b list \ 'c list err" +"sup f == %xs ys. if size xs = size ys then coalesce(xs +[f] ys) else Err" + + upto_esl :: "nat \ 'a esl \ '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 \ + 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: + "\ i \ 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: + "\ xs <=[r] ys; p < size xs \ \ 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 \ xs <=[r] xs" +apply (unfold unfold_lesub_list) +apply (simp add: Listn.le_def list_all2_conv_all_nth) +done + +lemma le_list_trans: + "\ order r; xs <=[r] ys; ys <=[r] zs \ \ 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: + "\ order r; xs <=[r] ys; ys <=[r] xs \ \ 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 \ 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 \ 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 \ size ys = size xs" +apply (unfold lesssub_def) +apply auto +done + +lemma le_list_appendI: + "\b c d. a <=[r] b \ c <=[r] d \ a@c <=[r] b@d" +apply (induct a) + apply simp +apply (case_tac b) +apply auto +done + +lemma le_listI: + "length a = length b \ (\n. n < length a \ a!n <=_r b!n) \ a <=[r] b" + apply (unfold lesub_def Listn.le_def) + apply (simp add: list_all2_conv_all_nth) + done + +lemma listI: + "\ length xs = n; set xs <= A \ \ xs : list n A" +apply (unfold list_def) +apply blast +done + +lemma listE_length [simp]: + "xs : list n A \ length xs = n" +apply (unfold list_def) +apply blast +done + +lemma less_lengthI: + "\ xs : list n A; p < n \ \ p < length xs" + by simp + +lemma listE_set [simp]: + "xs : list n A \ 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) = (\y\ A. \ys\ 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\ A & xs : list n A)"; +apply (simp add: in_list_Suc_iff) +done + +lemma list_not_empty: + "\a. a\ A \ \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 \ set xs <= A \ i < n \ (xs!i) : A" +apply (induct "xs") + apply simp +apply (simp add: nth_Cons split: nat.split) +done + +lemma listE_nth_in: + "\ xs : list n A; i < n \ \ (xs!i) : A" + by auto + + +lemma listn_Cons_Suc [elim!]: + "l#xs \ list n A \ (\n'. n = Suc n' \ l \ A \ xs \ list n' A \ P) \ P" + by (cases n) auto + +lemma listn_appendE [elim!]: + "a@b \ list n A \ (\n1 n2. n=n1+n2 \ a \ list n1 A \ b \ list n2 A \ P) \ P" +proof - + have "\n. a@b \ list n A \ \n1 n2. n=n1+n2 \ a \ list n1 A \ b \ list n2 A" + (is "\n. ?list a n \ \n1 n2. ?P a n n1 n2") + proof (induct a) + fix n assume "?list [] n" + hence "?P [] n 0 n" by simp + thus "\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 \ A" and list_n': "ls@b \ list n' A" by fastsimp + assume "\n. ls @ b \ list n A \ \n1 n2. n = n1 + n2 \ ls \ list n1 A \ b \ list n2 A" + hence "\n1 n2. n' = n1 + n2 \ ls \ list n1 A \ b \ list n2 A" by this (rule list_n') + then obtain n1 n2 where "n' = n1 + n2" "ls \ list n1 A" "b \ list n2 A" by fast + with n have "?P (l#ls) n (n1+1) n2" by simp + thus "\n1 n2. ?P (l#ls) n n1 n2" by fastsimp + qed + moreover + assume "a@b \ list n A" "\n1 n2. n=n1+n2 \ a \ list n1 A \ b \ list n2 A \ P" + ultimately + show ?thesis by blast +qed + + +lemma listt_update_in_list [simp, intro!]: + "\ xs : list n A; x\ A \ \ 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 [] \ [] | y#ys \ (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 \ length ys = n \ i + (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]: + "\ set xs <= A; set ys <= A; size xs = size ys \ + \ 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: + "\set xs <= A; set ys <= A; size xs = size ys \ + \ 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 \ set ys <= A \ set zs <= A + \ size xs = n & size ys = n \ + xs <=[r] zs & ys <=[r] zs \ 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\ A \ set xs <= A \ + (!i. i 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: + "(\y. A y \ False) \ A = bot_class.bot" + by (rule equals0I) (auto simp add: mem_def) + +lemma acc_le_listI [intro!]: + "\ order r; acc r \ \ acc(Listn.le r)" +apply (unfold acc_def) +apply (subgoal_tac + "wf(UN n. {(ys,xs). size xs = n \ size ys = n \ 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 "\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. \xs. size xs = k \ a#xs:M}" in allE) +apply (erule impE) + apply blast +apply (thin_tac "\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 \ 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: "\L. semilat L \ 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) \ 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: "\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) \ + !xs. xs : list n A \ (!ys. ys : list n A \ + (!zs. coalesce (xs +[f] ys) = OK zs \ 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) \ + !xs. xs : list n A \ (!ys. ys : list n A \ + (!zs. coalesce (xs +[f] ys) = OK 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 (force simp add: semilat_le_err_OK2) +done + +lemma lift2_le_ub: + "\ semilat(err A, Err.le r, lift2 f); x\ A; y\ A; x +_f y = OK z; + u\ A; x <=_r u; y <=_r u \ \ 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) \ + !xs. xs : list n A \ (!ys. ys : list n A \ + (!zs us. coalesce (xs +[f] ys) = OK zs & xs <=[r] us & ys <=[r] us + & us : list n A \ 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: + "\ x +_f y = Err; semilat(err A, Err.le r, lift2 f); x\ A; y\ A \ + \ ~(\u\ 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]: + "\ semilat(err A, Err.le r, lift2 f) \ + \ !xs. xs\ list n A \ (!ys. ys\ list n A \ + coalesce (xs +[f] ys) = Err \ + ~(\zs\ 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) = (\x\ A. \y\ 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) \ + \xs. xs : list n A \ (\ys. ys : list n A \ + 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) \ + 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) \ + 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: + "\L. err_semilat L \ 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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/DFA/Opt.thy --- /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 \ 'a option ord" +"le r o1 o2 == case o2 of None \ o1=None | + Some y \ (case o1 of None \ True + | Some x \ x <=_r y)" + + opt :: "'a set \ 'a option set" +"opt A == insert None {x . ? y:A. x = Some y}" + + sup :: "'a ebinop \ 'a option ebinop" +"sup f o1 o2 == + case o1 of None \ OK o2 | Some x \ (case o2 of None \ OK o1 + | Some y \ (case f x y of Err \ Err | OK z \ OK (Some z)))" + + esl :: "'a esl \ '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 \ o1=None | + Some y \ (case o1 of None \ True | Some x \ x <=_r y))" +apply (unfold lesub_def le_def) +apply (rule refl) +done + +lemma le_opt_refl: + "order r \ o1 <=_(le r) o1" +by (simp add: unfold_le_opt split: option.split) + +lemma le_opt_trans [rule_format]: + "order r \ + o1 <=_(le r) o2 \ o2 <=_(le r) o3 \ 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 \ o1 <=_(le r) o2 \ o2 <=_(le r) o1 \ 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 \ 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]: + "\L. err_semilat L \ 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: "\x\?A0. \y\?A0. x <=_?r0 x +_?f0 y" and + ub2: "\x\?A0. \y\?A0. y <=_?r0 x +_?f0 y" and + lub: "\x\?A0. \y\?A0. \z\?A0. x <=_?r0 z \ y <=_?r0 z \ 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: "\c. a = Some c \ c : A" + by (clarsimp simp add: opt_def) + + from ab y + have b: "\d. b = Some d \ 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 \ opt A" "b \ opt A" "a +_(sup f) b = OK c" + moreover + from ord have "order r" by simp + moreover + { fix x y z + assume "x \ A" "y \ A" + hence "OK x \ err A \ OK y \ err A" by simp + with ub1 ub2 + have "(OK x) <=_(Err.le r) (OK x) +_(lift2 f) (OK y) \ + (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 \ y <=_r z" + by (auto simp add: plussub_def lift2_def Err.le_def lesub_def) + } + ultimately + have "a <=_(le r) c \ 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 "(\x\?A. \y\?A. x <=_?r x +_?f y) \ (\x\?A. \y\?A. y <=_?r x +_?f y)" + by (auto simp add: lesub_def plussub_def Err.le_def lift2_def split: err.split) + + moreover + + have "\x\?A. \y\?A. \z\?A. x <=_?r z \ y <=_?r z \ 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 "\ (OK d) <=_(Err.le r) (OK g); (OK e) <=_(Err.le r) (OK g) \ + \ (OK d) +_(lift2 f) (OK e) <=_(Err.le r) (OK g)" + by blast + hence "\ d <=_r g; e <=_r g \ \ \y. d +_f e = OK y \ 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: + "\ order r; top r T \ \ (T <=_r x) = (x = T)" +apply (unfold top_def) +apply (blast intro: order_antisym) +done + + +lemma acc_le_optI [intro!]: + "acc r \ 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: + "\ ox : opt S; !x:S. ox = Some x \ f x : S \ + \ Option.map f ox : opt S"; +apply (unfold Option.map_def) +apply (simp split: option.split) +apply blast +done + +end diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/DFA/Product.thy --- /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 \ 'b ord \ ('a * 'b) ord" +"le rA rB == %(a,b) (a',b'). a <=_rA a' & b <=_rB b'" + + sup :: "'a ebinop \ 'b ebinop \ ('a * 'b)ebinop" +"sup f g == %(a1,b1)(a2,b2). Err.sup Pair (a1 +_f a2) (b1 +_g b2)" + + esl :: "'a esl \ 'b esl \ ('a * 'b ) esl" +"esl == %(A,rA,fA) (B,rB,fB). (A <*> B, le rA rB, sup fA fB)" + +syntax "@lesubprod" :: "'a*'b \ 'a ord \ 'b ord \ 'b \ 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!]: + "\ acc r\<^isub>A; acc r\<^isub>B \ \ 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: + "\ closed (err A) (lift2 f); closed (err B) (lift2 g) \ \ + 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: + "\r f z. \ z : err A; semilat (err A, r, f); OK x : err A; OK y : err A; + OK x +_f OK y <=_r z\ \ OK x <=_r z \ 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: + "\L1 L2. \ err_semilat L1; err_semilat L2 \ \ 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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/DFA/Semilat.thy --- /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 \ 'a \ bool" + 'a binop = "'a \ 'a \ 'a" + 'a sl = "'a set \ 'a ord \ 'a binop" + +consts + "lesub" :: "'a \ 'a ord \ 'a \ bool" + "lesssub" :: "'a \ 'a ord \ 'a \ bool" + "plussub" :: "'a \ ('a \ 'b \ 'c) \ 'b \ 'c" +(*<*) +syntax + "lesub" :: "'a \ 'a ord \ 'a \ bool" ("(_ /<='__ _)" [50, 1000, 51] 50) + "lesssub" :: "'a \ 'a ord \ 'a \ bool" ("(_ /<'__ _)" [50, 1000, 51] 50) + "plussub" :: "'a \ ('a \ 'b \ 'c) \ 'b \ 'c" ("(_ /+'__ _)" [65, 1000, 66] 65) +(*>*) +syntax (xsymbols) + "lesub" :: "'a \ 'a ord \ 'a \ bool" ("(_ /\\<^bsub>_\<^esub> _)" [50, 0, 51] 50) + "lesssub" :: "'a \ 'a ord \ 'a \ bool" ("(_ /\\<^bsub>_\<^esub> _)" [50, 0, 51] 50) + "plussub" :: "'a \ ('a \ 'b \ 'c) \ 'b \ 'c" ("(_ /\\<^bsub>_\<^esub> _)" [65, 0, 66] 65) +(*<*) + (* allow \ instead of \..\ *) + "@lesub" :: "'a \ 'a ord \ 'a \ bool" ("(_ /\\<^sub>_ _)" [50, 1000, 51] 50) + "@lesssub" :: "'a \ 'a ord \ 'a \ bool" ("(_ /\\<^sub>_ _)" [50, 1000, 51] 50) + "@plussub" :: "'a \ ('a \ 'b \ 'c) \ 'b \ 'c" ("(_ /\\<^sub>_ _)" [65, 1000, 66] 65) + +translations + "x \\<^sub>r y" => "x \\<^bsub>r\<^esub> y" + "x \\<^sub>r y" => "x \\<^bsub>r\<^esub> y" + "x \\<^sub>f y" => "x \\<^bsub>f\<^esub> y" +(*>*) + +defs + lesub_def: "x \\<^sub>r y \ r x y" + lesssub_def: "x \\<^sub>r y \ x \\<^sub>r y \ x \ y" + plussub_def: "x \\<^sub>f y \ f x y" + +constdefs + ord :: "('a \ 'a) set \ 'a ord" + "ord r \ \x y. (x,y) \ r" + + order :: "'a ord \ bool" + "order r \ (\x. x \\<^sub>r x) \ (\x y. x \\<^sub>r y \ y \\<^sub>r x \ x=y) \ (\x y z. x \\<^sub>r y \ y \\<^sub>r z \ x \\<^sub>r z)" + + top :: "'a ord \ 'a \ bool" + "top r T \ \x. x \\<^sub>r T" + + acc :: "'a ord \ bool" + "acc r \ wf {(y,x). x \\<^sub>r y}" + + closed :: "'a set \ 'a binop \ bool" + "closed A f \ \x\A. \y\A. x \\<^sub>f y \ A" + + semilat :: "'a sl \ bool" + "semilat \ \(A,r,f). order r \ closed A f \ + (\x\A. \y\A. x \\<^sub>r x \\<^sub>f y) \ + (\x\A. \y\A. y \\<^sub>r x \\<^sub>f y) \ + (\x\A. \y\A. \z\A. x \\<^sub>r z \ y \\<^sub>r z \ x \\<^sub>f y \\<^sub>r z)" + + + is_ub :: "('a \ 'a) set \ 'a \ 'a \ 'a \ bool" + "is_ub r x y u \ (x,u)\r \ (y,u)\r" + + is_lub :: "('a \ 'a) set \ 'a \ 'a \ 'a \ bool" + "is_lub r x y u \ is_ub r x y u \ (\z. is_ub r x y z \ (u,z)\r)" + + some_lub :: "('a \ 'a) set \ 'a \ 'a \ 'a" + "some_lub r x y \ 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 \ x \\<^sub>r x" + (*<*) by (unfold order_def) (simp (no_asm_simp)) (*>*) + +lemma order_antisym: "\ order r; x \\<^sub>r y; y \\<^sub>r x \ \ x = y" + (*<*) by (unfold order_def) (simp (no_asm_simp)) (*>*) + +lemma order_trans: "\ order r; x \\<^sub>r y; y \\<^sub>r z \ \ x \\<^sub>r z" + (*<*) by (unfold order_def) blast (*>*) + +lemma order_less_irrefl [intro, simp]: "order r \ \ x \\<^sub>r x" + (*<*) by (unfold order_def lesssub_def) blast (*>*) + +lemma order_less_trans: "\ order r; x \\<^sub>r y; y \\<^sub>r z \ \ x \\<^sub>r z" + (*<*) by (unfold order_def lesssub_def) blast (*>*) + +lemma topD [simp, intro]: "top r T \ x \\<^sub>r T" + (*<*) by (simp add: top_def) (*>*) + +lemma top_le_conv [simp]: "\ order r; top r T \ \ (T \\<^sub>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 \\<^sub>r x \\<^sub>f y) \ + (\x\A. \y\A. y \\<^sub>r x \\<^sub>f y) \ + (\x\A. \y\A. \z\A. x \\<^sub>r z \ y \\<^sub>r z \ x \\<^sub>f y \\<^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: "\ closed A f; x\A; y\A \ \ x \\<^sub>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]: "\x \ A; y \ A\ \ x \\<^sub>f y \ A" + (*<*) by (simp add: closedD [OF closedI]) (*>*) + +lemma (in Semilat) refl_r [intro, simp]: "x \\<^sub>r x" by simp + +lemma (in Semilat) antisym_r [intro?]: "\ x \\<^sub>r y; y \\<^sub>r x \ \ x = y" + (*<*) by (rule order_antisym) auto (*>*) + +lemma (in Semilat) trans_r [trans, intro?]: "\x \\<^sub>r y; y \\<^sub>r z\ \ x \\<^sub>r z" + (*<*) by (auto intro: order_trans) (*>*) + +lemma (in Semilat) ub1 [simp, intro?]: "\ x \ A; y \ A \ \ x \\<^sub>r x \\<^sub>f y" + (*<*) by (insert semilat) (unfold semilat_Def, simp) (*>*) + +lemma (in Semilat) ub2 [simp, intro?]: "\ x \ A; y \ A \ \ y \\<^sub>r x \\<^sub>f y" + (*<*) by (insert semilat) (unfold semilat_Def, simp) (*>*) + +lemma (in Semilat) lub [simp, intro?]: + "\ x \\<^sub>r z; y \\<^sub>r z; x \ A; y \ A; z \ A \ \ x \\<^sub>f y \\<^sub>r z"; + (*<*) by (insert semilat) (unfold semilat_Def, simp) (*>*) + +lemma (in Semilat) plus_le_conv [simp]: + "\ x \ A; y \ A; z \ A \ \ (x \\<^sub>f y \\<^sub>r z) = (x \\<^sub>r z \ y \\<^sub>r z)" + (*<*) by (blast intro: ub1 ub2 lub order_trans) (*>*) + +lemma (in Semilat) le_iff_plus_unchanged: "\ x \ A; y \ A \ \ (x \\<^sub>r y) = (x \\<^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: "\ x \ A; y \ A \ \ (x \\<^sub>r y) = (y \\<^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 \ A" and b: "b \ A" and c: "c \ A" + shows "a \\<^sub>f (b \\<^sub>f c) = a \\<^sub>f b \\<^sub>f c" +(*<*) +proof - + from a b have ab: "a \\<^sub>f b \ A" .. + from this c have abc: "(a \\<^sub>f b) \\<^sub>f c \ A" .. + from b c have bc: "b \\<^sub>f c \ A" .. + from a this have abc': "a \\<^sub>f (b \\<^sub>f c) \ A" .. + + show ?thesis + proof + show "a \\<^sub>f (b \\<^sub>f c) \\<^sub>r (a \\<^sub>f b) \\<^sub>f c" + proof - + from a b have "a \\<^sub>r a \\<^sub>f b" .. + also from ab c have "\ \\<^sub>r \ \\<^sub>f c" .. + finally have "a<": "a \\<^sub>r (a \\<^sub>f b) \\<^sub>f c" . + from a b have "b \\<^sub>r a \\<^sub>f b" .. + also from ab c have "\ \\<^sub>r \ \\<^sub>f c" .. + finally have "b<": "b \\<^sub>r (a \\<^sub>f b) \\<^sub>f c" . + from ab c have "c<": "c \\<^sub>r (a \\<^sub>f b) \\<^sub>f c" .. + from "b<" "c<" b c abc have "b \\<^sub>f c \\<^sub>r (a \\<^sub>f b) \\<^sub>f c" .. + from "a<" this a bc abc show ?thesis .. + qed + show "(a \\<^sub>f b) \\<^sub>f c \\<^sub>r a \\<^sub>f (b \\<^sub>f c)" + proof - + from b c have "b \\<^sub>r b \\<^sub>f c" .. + also from a bc have "\ \\<^sub>r a \\<^sub>f \" .. + finally have "b<": "b \\<^sub>r a \\<^sub>f (b \\<^sub>f c)" . + from b c have "c \\<^sub>r b \\<^sub>f c" .. + also from a bc have "\ \\<^sub>r a \\<^sub>f \" .. + finally have "c<": "c \\<^sub>r a \\<^sub>f (b \\<^sub>f c)" . + from a bc have "a<": "a \\<^sub>r a \\<^sub>f (b \\<^sub>f c)" .. + from "a<" "b<" a b abc' have "a \\<^sub>f b \\<^sub>r a \\<^sub>f (b \\<^sub>f c)" .. + from this "c<" ab c abc' show ?thesis .. + qed + qed +qed +(*>*) + +lemma (in Semilat) plus_com_lemma: + "\a \ A; b \ A\ \ a \\<^sub>f b \\<^sub>r b \\<^sub>f a" +(*<*) +proof - + assume a: "a \ A" and b: "b \ A" + from b a have "a \\<^sub>r b \\<^sub>f a" .. + moreover from b a have "b \\<^sub>r b \\<^sub>f a" .. + moreover note a b + moreover from b a have "b \\<^sub>f a \ A" .. + ultimately show ?thesis .. +qed +(*>*) + +lemma (in Semilat) plus_commutative: + "\a \ A; b \ A\ \ a \\<^sub>f b = b \\<^sub>f a" + (*<*) by(blast intro: order_antisym plus_com_lemma) (*>*) + +lemma is_lubD: + "is_lub r x y u \ is_ub r x y u \ (\z. is_ub r x y z \ (u,z) \ r)" + (*<*) by (simp add: is_lub_def) (*>*) + +lemma is_ubI: + "\ (x,u) \ r; (y,u) \ r \ \ is_ub r x y u" + (*<*) by (simp add: is_ub_def) (*>*) + +lemma is_ubD: + "is_ub r x y u \ (x,u) \ r \ (y,u) \ r" + (*<*) by (simp add: is_ub_def) (*>*) + + +lemma is_lub_bigger1 [iff]: + "is_lub (r^* ) x y y = ((x,y)\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)\r^* )" +(*<*) +apply (unfold is_lub_def is_ub_def) +apply blast +done +(*>*) + +lemma extend_lub: + "\ single_valued r; is_lub (r^* ) x y u; (x',x) \ r \ + \ EX v. is_lub (r^* ) x' y v" +(*<*) +apply (unfold is_lub_def is_ub_def) +apply (case_tac "(y,x) \ r^*") + apply (case_tac "(y,x') \ 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]: + "\ single_valued r; (x,u) \ r^* \ \ (\y. (y,u) \ r^* \ + (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: + "\ acyclic r; is_lub (r^* ) x y u \ \ 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: + "\ single_valued r; acyclic r; (x,u)\r^*; (y,u)\r^* \ + \ 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 \ ('a \ 'a) \ 'a binop" +"exec_lub r f x y \ while (\z. (x,z) \ 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: + "\acyclic r; single_valued r; (x,y) \ r\<^sup>*\ + \ finite (r \ {a. (x, a) \ r\<^sup>*} \ {b. (b, y) \ 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 \ {a. (x,a) \ r\<^sup>*} \ {b. (b,y) \ r\<^sup>*} = + insert (x,x') (r \ {a. (x', a) \ r\<^sup>*} \ {b. (b, y) \ r\<^sup>*})") + apply simp +apply(blast intro:converse_rtrancl_into_rtrancl + elim:converse_rtranclE dest:single_valuedD) +done +(*>*) + + +lemma exec_lub_conv: + "\ acyclic r; \x y. (x,y) \ r \ f x = y; is_lub (r\<^sup>*) x y u \ \ + exec_lub r f x y = u"; +(*<*) +apply(unfold exec_lub_def) +apply(rule_tac P = "\z. (y,z) \ r\<^sup>* \ (z,u) \ r\<^sup>*" and + r = "(r \ {(a,b). (y,a) \ r\<^sup>* \ (b,u) \ 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) \ 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: + "\ single_valued r; acyclic r; (x,u):r^*; (y,u):r^*; \x y. (x,y) \ r \ f x = y \ + \ is_lub (r^* ) x y (exec_lub r f x y)" + (*<*) by (fastsimp dest: single_valued_has_lubs simp add: exec_lub_conv) (*>*) + +end diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/DFA/SemilatAlg.thy --- /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 \ 's) list \ 's ord \ (nat \ 's) list \ bool" + ("(_ /<=|_| _)" [50, 0, 51] 50) + "x <=|r| y \ \(p,s) \ set x. \s'. (p,s') \ set y \ s <=_r s'" + +consts + "@plusplussub" :: "'a list \ ('a \ 'a \ 'a) \ 'a \ 'a" ("(_ /++'__ _)" [65, 1000, 66] 65) +primrec + "[] ++_f y = y" + "(x#xs) ++_f y = xs ++_f (x +_f y)" + +constdefs + bounded :: "'s step_type \ nat \ bool" +"bounded step n == !p nat \ 's set \ bool" +"pres_type step n A == \s\A. \p(q,s')\set (step p s). s' \ A" + + mono :: "'s ord \ 's step_type \ nat \ 's set \ bool" +"mono r step n A == + \s p t. s \ A \ p < n \ s <=_r t \ step p s <=|r| step p t" + + +lemma pres_typeD: + "\ pres_type step n A; s\A; pset (step p s) \ \ s' \ A" + by (unfold pres_type_def, blast) + +lemma monoD: + "\ mono r step n A; p < n; s\A; s <=_r t \ \ step p s <=|r| step p t" + by (unfold mono_def, blast) + +lemma boundedD: + "\ bounded step n; p < n; (q,t) : set (step p xs) \ \ q < n" + by (unfold bounded_def, blast) + +lemma lesubstep_type_refl [simp, intro]: + "(\x. x <=_r x) \ x <=|r| x" + by (unfold lesubstep_type_def) auto + +lemma lesub_step_typeD: + "a <=|r| b \ (x,y) \ set a \ \y'. (x, y') \ set b \ y <=_r y'" + by (unfold lesubstep_type_def) blast + + +lemma list_update_le_listI [rule_format]: + "set xs <= A \ set ys <= A \ xs <=[r] ys \ p < size xs \ + x <=_r ys!p \ semilat(A,r,f) \ x\A \ + 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 + "\y. \ set x \ A; y \ A\ \ x ++_f y \ A" (is "PROP ?P") +proof - + interpret Semilat A r f using assms by (rule Semilat.intro) + show "PROP ?P" proof (induct x) + show "\y. y \ A \ [] ++_f y \ A" by simp + fix y x xs + assume y: "y \ A" and xs: "set (x#xs) \ A" + assume IH: "\y. \ set xs \ A; y \ A\ \ xs ++_f y \ A" + from xs obtain x: "x \ A" and xs': "set xs \ A" by simp + from x y have "(x +_f y) \ A" .. + with xs' have "xs ++_f (x +_f y) \ A" by (rule IH) + thus "(x#xs) ++_f y \ A" by simp + qed +qed + +lemma (in Semilat) pp_ub2: + "\y. \ set x \ A; y \ A\ \ y <=_r x ++_f y" +proof (induct x) + from semilat show "\y. y <=_r [] ++_f y" by simp + + fix y a l + assume y: "y \ A" + assume "set (a#l) \ A" + then obtain a: "a \ A" and x: "set l \ A" by simp + assume "\y. \set l \ A; y \ A\ \ y <=_r l ++_f y" + hence IH: "\y. y \ A \ y <=_r l ++_f y" using x . + + from a y have "y <=_r a +_f y" .. + also from a y have "a +_f y \ 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 "\y. \set ls \ A; y \ A; x \ set ls\ \ x <=_r ls ++_f y" +proof (induct ls) + show "\y. x \ set [] \ x <=_r [] ++_f y" by simp + + fix y s ls + assume "set (s#ls) \ A" + then obtain s: "s \ A" and ls: "set ls \ A" by simp + assume y: "y \ A" + + assume + "\y. \set ls \ A; y \ A; x \ set ls\ \ x <=_r ls ++_f y" + hence IH: "\y. x \ set ls \ y \ A \ x <=_r ls ++_f y" using ls . + + assume "x \ set (s#ls)" + then obtain xls: "x = s \ x \ 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 \ 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 \ set ls" + hence "\y. y \ A \ x <=_r ls ++_f y" by (rule IH) + moreover from s y have "s +_f y \ 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 \ A" + shows + "\y. y \ A \ set xs \ A \ \x \ set xs. x <=_r z \ y <=_r z \ 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 \ A" and "set (l#ls) \ A" + then obtain l: "l \ A" and ls: "set ls \ A" by auto + assume "\x \ set (l#ls). x <=_r z" + then obtain lz: "l <=_r z" and lsz: "\x \ 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 \ A" .. + moreover + assume "\y. y \ A \ set ls \ A \ \x \ set ls. x <=_r z \ y <=_r z + \ 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 "\\(p,s) \ set S. s \ A; y \ A; (a,b) \ set S\ + \ b <=_r map snd [(p', t')\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 \ A" + moreover + assume "\(p,s) \ set S. s \ A" + hence "set ?map \ A" by auto + moreover + assume "(a,b) \ set S" + hence "b \ set ?map" by (induct S, auto) + ultimately + show ?thesis by - (rule pp_ub1) +qed + + +lemma plusplus_empty: + "\s'. (q, s') \ set S \ s' +_f ss ! q = ss ! q \ + (map snd [(p', t') \ S. p' = q] ++_f ss ! q) = ss ! q" + by (induct S) auto + +end diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/DFA/Semilattices.thy --- /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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/DFA/Typing_Framework.thy --- /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 \ 's \ (nat \ 's) list" + +constdefs + stable :: "'s ord \ 's step_type \ 's list \ nat \ bool" +"stable r step ss p == !(q,s'):set(step p (ss!p)). s' <=_r ss!q" + + stables :: "'s ord \ 's step_type \ 's list \ bool" +"stables r step ss == !p 's \ 's step_type \ 's list \ bool" +"wt_step r T step ts == + !p 's \ 's step_type + \ nat \ 's set \ ('s list \ 's list) \ bool" +"is_bcv r T step n A bcv == !ss : list n A. + (!p 's err step_type \ 's err list \ bool" +"wt_err_step r step ts \ wt_step (Err.le r) Err step ts" + +wt_app_eff :: "'s ord \ (nat \ 's \ bool) \ 's step_type \ 's list \ bool" +"wt_app_eff r app step ts \ + \p < size ts. app p (ts!p) \ (\(q,t) \ set (step p (ts!p)). t <=_r ts!q)" + +map_snd :: "('b \ 'c) \ ('a \ 'b) list \ ('a \ 'c) list" +"map_snd f \ map (\(x,y). (x, f y))" + +error :: "nat \ (nat \ 'a err) list" +"error n \ map (\x. (x,Err)) [0.. (nat \ 's \ bool) \ 's step_type \ 's err step_type" +"err_step n app step p t \ + case t of + Err \ error n + | OK t' \ if app p t' then map_snd OK (step p t') else error n" + +app_mono :: "'s ord \ (nat \ 's \ bool) \ nat \ 's set \ bool" +"app_mono r app n A \ + \s p t. s \ A \ p < n \ s <=_r t \ app p t \ 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 \ + p < n \ app p a \ (q,b) \ set (step p a) \ + 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) \ set (map_snd f xs) \ \b'. (a,b') \ set xs" + apply (induct xs) + apply (auto simp add: map_snd_def) + done + + +lemma bounded_err_stepI: + "\p. p < n \ (\s. ap p s \ (\(q,s') \ set (step p s). q < n)) + \ 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 \ 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]: + "\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 \ 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 \ app_mono r app n A \ bounded (err_step n app step) n \ + \s p t. s \ A \ p < n \ s <=_r t \ app p t \ step p s <=|r| step p t \ + 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) \ set (error n) \ y = Err" + by (auto simp add: error_def) + +lemma pres_type_lift: + "\s\A. \p. p < n \ app p s \ (\(q, s')\set (step p s). s' \ A) + \ 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) \ set (error (size ts))" by (simp add: error_def) + + from wt lp + have [intro?]: "\p. p < size ts \ ts ! p \ 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: "\(q,t) \ 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 "\ app p (map ok_val ts ! p)" + with OKp lp have "\ 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) \ 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 \ Err" .. + ultimately show False by blast + qed + + show "\(q,t)\set(step p (map ok_val ts ! p)). t <=_r map ok_val ts ! q" + proof clarify + fix q t assume q: "(q,t) \ set (step p (map ok_val ts ! p))" + + from wt lp q + obtain s where + OKp: "ts ! p = OK s" and + less: "\(q,t) \ 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 \ 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 \ Err" by simp + { fix q t + assume q: "(q,t) \ set (err_step (size ts) app step p (map OK ts ! p))" + with p app_eff obtain + "app p (ts ! p)" "\(q,t) \ 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 + diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/J/Conform.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/J/Decl.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/J/Eval.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/J/Example.thy --- 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) \ (subcls1 tprg)^+ ==> R" +apply (auto dest!: tranclD subcls1D) done lemma subcls_ObjectD [dest!]: "tprg\Object\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) \ (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) \ (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\Ext\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 diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/J/Exceptions.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/J/JBasis.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/J/JListExample.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/J/JTypeSafe.thy --- 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" diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/J/State.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/J/SystemClasses.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/J/Term.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/J/Type.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/J/TypeRel.thy --- 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" ("_ \ _ \C1 _" [71,71,71] 70) + +inductive_set + subcls1 :: "'c prog => (cname \ cname) set" + and subcls1' :: "'c prog => cname \ cname => bool" ("_ \ _ \C1 _" [71,71,71] 70) for G :: "'c prog" where - subcls1I: "\class G C = Some (D,rest); C \ Object\ \ G\C\C1D" + "G \ C \C1 D \ (C, D) \ subcls1 G" + | subcls1I: "\class G C = Some (D,rest); C \ Object\ \ G \ C \C1 D" abbreviation - subcls :: "'c prog => [cname, cname] => bool" ("_ \ _ \C _" [71,71,71] 70) - where "G\C \C D \ (subcls1 G)^** C D" - + subcls :: "'c prog => cname \ cname => bool" ("_ \ _ \C _" [71,71,71] 70) + where "G \ C \C D \ (C, D) \ (subcls1 G)^*" + lemma subcls1D: "G\C\C1D \ C \ Object \ (\fs ms. class G C = Some (D,fs,ms))" apply (erule subcls1.cases) apply auto done -lemma subcls1_def2: - "subcls1 G = (\C D. (C, D) \ - (SIGMA C: {C. is_class G C} . {D. C\Object \ 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\Object \ 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) \ (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\C\C D \ is_class G D \ 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 \ cname \ 'a \ (cname \ fdecl list \ 'c mdecl list \ 'a \ 'a) \ 'a" - "class_rec G == wfrec {(C, D). (subcls1 G)^--1 C D} + "class_rec G == wfrec ((subcls1 G)^-1) (\r C t f. case class G C of None \ undefined | Some (D,fs,ms) \ f C fs ms (if C = Object then t else r D t f))" -lemma class_rec_lemma: "wfP ((subcls1 G)^--1) \ class G C = Some (D,fs,ms) \ - 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: "\H a. wfrec ((subcls1 G)\) H a = + H (cut (wfrec ((subcls1 G)\) H) ((subcls1 G)\) a) a" + by (rule wfrec) + have cut: "\f. C \ Object \ cut f ((subcls1 G)\) 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 \ undefined - | Some (D, fs, ms) \ 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 \ \(G,C). class_rec G C empty (\C fs ms ts. ts ++ map_of (map (\(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 (\(s,m). (s,(C,m))) ms)" apply (unfold method_def) @@ -129,7 +112,7 @@ defs fields_def: "fields \ \(G,C). class_rec G C [] (\C fs ms ts. map (\(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 (\(fn,ft). ((fn,C),ft)) fs @ (if C = Object then [] else fields (G,D))" apply (unfold fields_def) diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/J/Value.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/J/WellForm.thy --- 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\C\C1D; ws_prog G|] ==> D \ C \ \ (subcls1 G)^++ D C" -apply( frule tranclp.r_into_trancl [where r="subcls1 G"]) +lemma subcls1_wfD: "[|G\C\C1D; ws_prog G|] ==> D \ C \ (D, C) \ (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|] ==> \ (subcls1 G)^++ D C" -apply(erule tranclp.cases) +lemma subcls_asym: "[|ws_prog G; (C, D) \ (subcls1 G)^+|] ==> (D, C) \ (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 \ D" -apply (erule tranclp_trans_induct) +lemma subcls_irrefl: "[|ws_prog G; (C, D) \ (subcls1 G)^+|] ==> C \ 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. \D. (subcls1 G)^++ C D --> P D ==> P C|] ==> P C" +lemma subcls_induct_struct: +"[|ws_prog G; !!C. \D. (C, D) \ (subcls1 G)^+ --> P D ==> P C|] ==> P C" (is "?A \ PROP ?P \ _") 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. \D. (C, D) \ (subcls1 G)^+ --> P D ==> P C|] ==> P C" +(is "?A \ PROP ?P \ _") +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 \ Object; is_class G C; class G C = Some (D,fs,ms) \ @@ -223,22 +228,6 @@ done qed -lemma subcls_induct_struct: -"[|ws_prog G; !!C. \D. (subcls1 G)^++ C D --> P D ==> P C|] ==> P C" -(is "?A \ PROP ?P \ _") -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 \ Object; is_class G C; class G C = Some (D,fs,ms) \ @@ -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) \ (subcls1 G)^*|] ==> x \ set (fields (G,C)) --> x \ 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\D\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\T'\C T; wf_prog wf_mb G|] ==> \D rT b. method (G,T) sig = Some (D,rT ,b) --> (\D' rT' b'. method (G,T') sig = Some (D',rT',b') \ G\D'\C D \ G\rT'\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) diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/J/WellType.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/JVM/JVMDefensive.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/JVM/JVMExceptions.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/JVM/JVMExec.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/JVM/JVMExecInstr.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/JVM/JVMInstructions.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/JVM/JVMListExample.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/JVM/JVMState.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/MicroJava.thy diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/ROOT.ML diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/document/introduction.tex diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/document/root.bib diff -r 6a973bd43949 -r 1bc3b688548c src/HOL/MicroJava/document/root.tex