# HG changeset patch # User wenzelm # Date 1259936424 -3600 # Node ID 01dcd9b926bf22b54c1b436194c97a234cfc3b7f # Parent 651028e34b5d82762f2bd9212d3b84f117d292df# Parent 9c7fa7f76950260f547b33a4840898c0c6a2f8f3 merged, resolving minor conflicts; diff -r 651028e34b5d -r 01dcd9b926bf src/FOLP/simp.ML --- a/src/FOLP/simp.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/FOLP/simp.ML Fri Dec 04 15:20:24 2009 +0100 @@ -419,11 +419,11 @@ if lhs_is_NORM(thm,i) then (ss, res1(thm,trans_norms,i), anet,ats,cs) else case Seq.pull(cong_tac i thm) of SOME(thm',_) => - let val ps = prems_of thm and ps' = prems_of thm'; + let val ps = prems_of thm + and ps' = prems_of thm'; val n = length(ps')-length(ps); val a = length(Logic.strip_assums_hyp(List.nth(ps,i-1))) - val l = map (fn p => length(Logic.strip_assums_hyp(p))) - (Library.take(n,Library.drop(i-1,ps'))); + val l = map (length o Logic.strip_assums_hyp) (take n (drop (i-1) ps')); in (simp_refl(rev(l),a,REW::ss),thm',anet,ats,cs) end | NONE => (REW::ss,thm,anet,ats,cs); @@ -535,7 +535,7 @@ fun Pinst(f,fT,(eq,eqT),k,i,T,yik,Ts) = let fun xn_list(x,n) = let val ixs = map_range (fn i => (x^(radixstring(26,"a",i)),0)) (n - 1); - in ListPair.map eta_Var (ixs, Library.take(n+1,Ts)) end + in ListPair.map eta_Var (ixs, take (n+1) Ts) end val lhs = list_comb(f,xn_list("X",k-1)) val rhs = list_comb(f,xn_list("X",i-1) @ [Bound 0] @ yik) in Abs("", T, Const(eq,[fT,fT]--->eqT) $ lhs $ rhs) end; diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Bali/AxExample.thy --- a/src/HOL/Bali/AxExample.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Bali/AxExample.thy Fri Dec 04 15:20:24 2009 +0100 @@ -1,11 +1,12 @@ (* Title: HOL/Bali/AxExample.thy - ID: $Id$ Author: David von Oheimb *) header {* Example of a proof based on the Bali axiomatic semantics *} -theory AxExample imports AxSem Example begin +theory AxExample +imports AxSem Example +begin constdefs arr_inv :: "st \ bool" diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Bali/Basis.thy --- a/src/HOL/Bali/Basis.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Bali/Basis.thy Fri Dec 04 15:20:24 2009 +0100 @@ -216,8 +216,8 @@ In1l :: "'al \ ('al + 'ar, 'b, 'c) sum3" In1r :: "'ar \ ('al + 'ar, 'b, 'c) sum3" translations - "In1l e" == "In1 (Inl e)" - "In1r c" == "In1 (Inr c)" + "In1l e" == "In1 (CONST Inl e)" + "In1r c" == "In1 (CONST Inr c)" syntax the_In1l :: "('al + 'ar, 'b, 'c) sum3 \ 'al" the_In1r :: "('al + 'ar, 'b, 'c) sum3 \ 'ar" @@ -233,7 +233,7 @@ (* e.g. lemmas is_stmt_rews = is_stmt_def [of "In1l x", simplified] *) translations - "option"<= (type) "Datatype.option" + "option"<= (type) "Option.option" "list" <= (type) "List.list" "sum3" <= (type) "Basis.sum3" diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Bali/Example.thy --- a/src/HOL/Bali/Example.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Bali/Example.thy Fri Dec 04 15:20:24 2009 +0100 @@ -1,10 +1,11 @@ (* Title: HOL/Bali/Example.thy - ID: $Id$ Author: David von Oheimb *) header {* Example Bali program *} -theory Example imports Eval WellForm begin +theory Example +imports Eval WellForm +begin text {* The following example Bali program includes: @@ -1235,24 +1236,24 @@ translations "obj_a" <= "\tag=Arr (PrimT Boolean) (CONST two) - ,values=CONST empty(Inr 0\Bool False)(Inr (CONST one)\Bool False)\" + ,values=CONST empty(CONST Inr 0\Bool False)(CONST Inr (CONST one)\Bool False)\" "obj_b" <= "\tag=CInst (CONST Ext) - ,values=(CONST empty(Inl (CONST vee, CONST Base)\Null ) - (Inl (CONST vee, CONST Ext )\Intg 0))\" + ,values=(CONST empty(CONST Inl (CONST vee, CONST Base)\Null ) + (CONST Inl (CONST vee, CONST Ext )\Intg 0))\" "obj_c" == "\tag=CInst (SXcpt NullPointer),values=CONST empty\" - "arr_N" == "CONST empty(Inl (CONST arr, CONST Base)\Null)" - "arr_a" == "CONST empty(Inl (CONST arr, CONST Base)\Addr a)" - "globs1" == "CONST empty(Inr (CONST Ext) \\tag=CONST undefined, values=CONST empty\) - (Inr (CONST Base) \\tag=CONST undefined, values=arr_N\) - (Inr Object\\tag=CONST undefined, values=CONST empty\)" - "globs2" == "CONST empty(Inr (CONST Ext) \\tag=CONST undefined, values=CONST empty\) - (Inr Object\\tag=CONST undefined, values=CONST empty\) - (Inl a\obj_a) - (Inr (CONST Base) \\tag=CONST undefined, values=arr_a\)" - "globs3" == "globs2(Inl b\obj_b)" - "globs8" == "globs3(Inl c\obj_c)" + "arr_N" == "CONST empty(CONST Inl (CONST arr, CONST Base)\Null)" + "arr_a" == "CONST empty(CONST Inl (CONST arr, CONST Base)\Addr a)" + "globs1" == "CONST empty(CONST Inr (CONST Ext) \\tag=CONST undefined, values=CONST empty\) + (CONST Inr (CONST Base) \\tag=CONST undefined, values=arr_N\) + (CONST Inr Object\\tag=CONST undefined, values=CONST empty\)" + "globs2" == "CONST empty(CONST Inr (CONST Ext) \\tag=CONST undefined, values=CONST empty\) + (CONST Inr Object\\tag=CONST undefined, values=CONST empty\) + (CONST Inl a\obj_a) + (CONST Inr (CONST Base) \\tag=CONST undefined, values=arr_a\)" + "globs3" == "globs2(CONST Inl b\obj_b)" + "globs8" == "globs3(CONST Inl c\obj_c)" "locs3" == "CONST empty(VName (CONST e)\Addr b)" - "locs4" == "CONST empty(VName (CONST z)\Null)(Inr()\Addr b)" + "locs4" == "CONST empty(VName (CONST z)\Null)(CONST Inr()\Addr b)" "locs8" == "locs3(VName (CONST z)\Addr c)" "s0" == " st (CONST empty) (CONST empty)" "s0'" == " Norm s0" diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Bali/State.thy --- a/src/HOL/Bali/State.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Bali/State.thy Fri Dec 04 15:20:24 2009 +0100 @@ -1,10 +1,11 @@ (* Title: HOL/Bali/State.thy - ID: $Id$ Author: David von Oheimb *) header {* State for evaluation of Java expressions and statements *} -theory State imports DeclConcepts begin +theory State +imports DeclConcepts +begin text {* design issues: @@ -138,8 +139,8 @@ Stat :: "qtname \ oref" translations - "Heap" => "Inl" - "Stat" => "Inr" + "Heap" => "CONST Inl" + "Stat" => "CONST Inr" "oref" <= (type) "loc + qtname" constdefs @@ -328,7 +329,7 @@ init_class_obj :: "prog \ qtname \ st \ st" translations - "init_class_obj G C" == "init_obj G CONST undefined (Inr C)" + "init_class_obj G C" == "init_obj G CONST undefined (CONST Inr C)" lemma gupd_def2 [simp]: "gupd(r\obj) (st g l) = st (g(r\obj)) l" apply (unfold gupd_def) diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Bali/WellType.thy --- a/src/HOL/Bali/WellType.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Bali/WellType.thy Fri Dec 04 15:20:24 2009 +0100 @@ -1,10 +1,11 @@ (* Title: HOL/Bali/WellType.thy - ID: $Id$ Author: David von Oheimb *) header {* Well-typedness of Java programs *} -theory WellType imports DeclConcepts begin +theory WellType +imports DeclConcepts +begin text {* improvements over Java Specification 1.0: @@ -443,10 +444,10 @@ translations "E\t\ T" == "E,empty_dt\t\ T" - "E\s\\" == "E\In1r s\Inl (PrimT Void)" - "E\e\-T" == "E\In1l e\Inl T" - "E\e\=T" == "E\In2 e\Inl T" - "E\e\\T" == "E\In3 e\Inr T" + "E\s\\" == "E\In1r s\CONST Inl (PrimT Void)" + "E\e\-T" == "E\In1l e\CONST Inl T" + "E\e\=T" == "E\In2 e\CONST Inl T" + "E\e\\T" == "E\In3 e\CONST Inr T" declare not_None_eq [simp del] diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Datatype.thy --- a/src/HOL/Datatype.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Datatype.thy Fri Dec 04 15:20:24 2009 +0100 @@ -1,16 +1,20 @@ (* Title: HOL/Datatype.thy Author: Lawrence C Paulson, Cambridge University Computer Laboratory Author: Stefan Berghofer and Markus Wenzel, TU Muenchen - -Could <*> be generalized to a general summation (Sigma)? *) -header {* Analogues of the Cartesian Product and Disjoint Sum for Datatypes *} +header {* Datatype package: constructing datatypes from Cartesian Products and Disjoint Sums *} theory Datatype -imports Nat Product_Type +imports Product_Type Sum_Type Nat +uses + ("Tools/Datatype/datatype.ML") + ("Tools/inductive_realizer.ML") + ("Tools/Datatype/datatype_realizer.ML") begin +subsection {* The datatype universe *} + typedef (Node) ('a,'b) node = "{p. EX f x k. p = (f::nat=>'b+nat, x::'a+nat) & f k = Inr 0}" --{*it is a subtype of @{text "(nat=>'b+nat) * ('a+nat)"}*} @@ -513,75 +517,12 @@ hide (open) type node item hide (open) const Push Node Atom Leaf Numb Lim Split Case - -section {* Datatypes *} - -subsection {* Representing sums *} - -rep_datatype (sum) Inl Inr -proof - - fix P - fix s :: "'a + 'b" - assume x: "\x\'a. P (Inl x)" and y: "\y\'b. P (Inr y)" - then show "P s" by (auto intro: sumE [of s]) -qed simp_all - -lemma sum_case_KK[simp]: "sum_case (%x. a) (%x. a) = (%x. a)" - by (rule ext) (simp split: sum.split) - -lemma surjective_sum: "sum_case (%x::'a. f (Inl x)) (%y::'b. f (Inr y)) s = f(s)" - apply (rule_tac s = s in sumE) - apply (erule ssubst) - apply (rule sum.cases(1)) - apply (erule ssubst) - apply (rule sum.cases(2)) - done - -lemma sum_case_weak_cong: "s = t ==> sum_case f g s = sum_case f g t" - -- {* Prevents simplification of @{text f} and @{text g}: much faster. *} - by simp +use "Tools/Datatype/datatype.ML" -lemma sum_case_inject: - "sum_case f1 f2 = sum_case g1 g2 ==> (f1 = g1 ==> f2 = g2 ==> P) ==> P" -proof - - assume a: "sum_case f1 f2 = sum_case g1 g2" - assume r: "f1 = g1 ==> f2 = g2 ==> P" - show P - apply (rule r) - apply (rule ext) - apply (cut_tac x = "Inl x" in a [THEN fun_cong], simp) - apply (rule ext) - apply (cut_tac x = "Inr x" in a [THEN fun_cong], simp) - done -qed - -constdefs - Suml :: "('a => 'c) => 'a + 'b => 'c" - "Suml == (%f. sum_case f undefined)" - - Sumr :: "('b => 'c) => 'a + 'b => 'c" - "Sumr == sum_case undefined" +use "Tools/inductive_realizer.ML" +setup InductiveRealizer.setup -lemma [code]: - "Suml f (Inl x) = f x" - by (simp add: Suml_def) - -lemma [code]: - "Sumr f (Inr x) = f x" - by (simp add: Sumr_def) - -lemma Suml_inject: "Suml f = Suml g ==> f = g" - by (unfold Suml_def) (erule sum_case_inject) - -lemma Sumr_inject: "Sumr f = Sumr g ==> f = g" - by (unfold Sumr_def) (erule sum_case_inject) - -primrec Projl :: "'a + 'b => 'a" -where Projl_Inl: "Projl (Inl x) = x" - -primrec Projr :: "'a + 'b => 'b" -where Projr_Inr: "Projr (Inr x) = x" - -hide (open) const Suml Sumr Projl Projr +use "Tools/Datatype/datatype_realizer.ML" +setup Datatype_Realizer.setup end diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Finite_Set.thy --- a/src/HOL/Finite_Set.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Finite_Set.thy Fri Dec 04 15:20:24 2009 +0100 @@ -6,7 +6,7 @@ header {* Finite sets *} theory Finite_Set -imports Nat Product_Type Power +imports Power Product_Type Sum_Type begin subsection {* Definition and basic properties *} @@ -1157,8 +1157,8 @@ ==> setsum (%x. f x) A = setsum (%x. g x) B" by(fastsimp simp: simp_implies_def setsum_def intro: comm_monoid_add.fold_image_cong) -lemma setsum_cong2: "\\x. x \ A \ f x = g x\ \ setsum f A = setsum g A"; -by (rule setsum_cong[OF refl], auto); +lemma setsum_cong2: "\\x. x \ A \ f x = g x\ \ setsum f A = setsum g A" +by (rule setsum_cong[OF refl], auto) lemma setsum_reindex_cong: "[|inj_on f A; B = f ` A; !!a. a:A \ g a = h (f a)|] diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Import/hol4rews.ML --- a/src/HOL/Import/hol4rews.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Import/hol4rews.ML Fri Dec 04 15:20:24 2009 +0100 @@ -480,7 +480,7 @@ then let val paths = Long_Name.explode isa - val i = Library.drop(length paths - 2,paths) + val i = drop (length paths - 2) paths in case i of [seg,con] => output_thy ^ "." ^ seg ^ "." ^ con diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Inductive.thy --- a/src/HOL/Inductive.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Inductive.thy Fri Dec 04 15:20:24 2009 +0100 @@ -5,17 +5,16 @@ header {* Knaster-Tarski Fixpoint Theorem and inductive definitions *} theory Inductive -imports Lattices Sum_Type +imports Complete_Lattice uses ("Tools/inductive.ML") "Tools/dseq.ML" ("Tools/inductive_codegen.ML") - ("Tools/Datatype/datatype_aux.ML") - ("Tools/Datatype/datatype_prop.ML") - ("Tools/Datatype/datatype_rep_proofs.ML") + "Tools/Datatype/datatype_aux.ML" + "Tools/Datatype/datatype_prop.ML" + "Tools/Datatype/datatype_case.ML" ("Tools/Datatype/datatype_abs_proofs.ML") - ("Tools/Datatype/datatype_case.ML") - ("Tools/Datatype/datatype.ML") + ("Tools/Datatype/datatype_data.ML") ("Tools/old_primrec.ML") ("Tools/primrec.ML") ("Tools/Datatype/datatype_codegen.ML") @@ -277,20 +276,16 @@ text {* Package setup. *} -use "Tools/Datatype/datatype_aux.ML" -use "Tools/Datatype/datatype_prop.ML" -use "Tools/Datatype/datatype_rep_proofs.ML" use "Tools/Datatype/datatype_abs_proofs.ML" -use "Tools/Datatype/datatype_case.ML" -use "Tools/Datatype/datatype.ML" -setup Datatype.setup +use "Tools/Datatype/datatype_data.ML" +setup Datatype_Data.setup + +use "Tools/Datatype/datatype_codegen.ML" +setup Datatype_Codegen.setup use "Tools/old_primrec.ML" use "Tools/primrec.ML" -use "Tools/Datatype/datatype_codegen.ML" -setup DatatypeCodegen.setup - use "Tools/inductive_codegen.ML" setup InductiveCodegen.setup @@ -306,7 +301,7 @@ fun fun_tr ctxt [cs] = let val x = Free (Name.variant (Term.add_free_names cs []) "x", dummyT); - val ft = DatatypeCase.case_tr true Datatype.info_of_constr + val ft = Datatype_Case.case_tr true Datatype_Data.info_of_constr ctxt [x, cs] in lambda x ft end in [("_lam_pats_syntax", fun_tr)] end diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/IsaMakefile Fri Dec 04 15:20:24 2009 +0100 @@ -168,10 +168,10 @@ Tools/Datatype/datatype_aux.ML \ Tools/Datatype/datatype_case.ML \ Tools/Datatype/datatype_codegen.ML \ - Tools/Datatype/datatype.ML \ + Tools/Datatype/datatype_data.ML \ Tools/Datatype/datatype_prop.ML \ Tools/Datatype/datatype_realizer.ML \ - Tools/Datatype/datatype_rep_proofs.ML \ + Tools/Datatype/datatype.ML \ Tools/dseq.ML \ Tools/Function/context_tree.ML \ Tools/Function/decompose.ML \ @@ -849,20 +849,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 651028e34b5d -r 01dcd9b926bf src/HOL/List.thy --- a/src/HOL/List.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/List.thy Fri Dec 04 15:20:24 2009 +0100 @@ -366,7 +366,7 @@ val case2 = Syntax.const "_case1" $ Syntax.const Term.dummy_patternN $ NilC; val cs = Syntax.const "_case2" $ case1 $ case2 - val ft = DatatypeCase.case_tr false Datatype.info_of_constr + val ft = Datatype_Case.case_tr false Datatype.info_of_constr ctxt [x, cs] in lambda x ft end; diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/Altern.thy --- a/src/HOL/MicroJava/BV/Altern.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/BV/Altern.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/BVExample.thy --- a/src/HOL/MicroJava/BV/BVExample.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/BV/BVExample.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/BVNoTypeError.thy --- a/src/HOL/MicroJava/BV/BVNoTypeError.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/BV/BVNoTypeError.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/BVSpec.thy --- a/src/HOL/MicroJava/BV/BVSpec.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/BV/BVSpec.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/BVSpecTypeSafe.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/Correct.thy --- a/src/HOL/MicroJava/BV/Correct.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/BV/Correct.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/Effect.thy --- a/src/HOL/MicroJava/BV/Effect.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/BV/Effect.thy Fri Dec 04 15:20:24 2009 +0100 @@ -9,7 +9,6 @@ imports JVMType "../JVM/JVMExceptions" begin - types succ_type = "(p_count \ state_type option) list" diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/EffectMono.thy --- a/src/HOL/MicroJava/BV/EffectMono.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/BV/EffectMono.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/Err.thy --- a/src/HOL/MicroJava/BV/Err.thy Fri Dec 04 11:44:57 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/JType.thy --- a/src/HOL/MicroJava/BV/JType.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/BV/JType.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/JVM.thy --- a/src/HOL/MicroJava/BV/JVM.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/BV/JVM.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/JVMType.thy --- a/src/HOL/MicroJava/BV/JVMType.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/BV/JVMType.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/Kildall.thy --- a/src/HOL/MicroJava/BV/Kildall.thy Fri Dec 04 11:44:57 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/LBVComplete.thy --- a/src/HOL/MicroJava/BV/LBVComplete.thy Fri Dec 04 11:44:57 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/LBVCorrect.thy --- a/src/HOL/MicroJava/BV/LBVCorrect.thy Fri Dec 04 11:44:57 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/LBVJVM.thy --- a/src/HOL/MicroJava/BV/LBVJVM.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/BV/LBVJVM.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/LBVSpec.thy --- a/src/HOL/MicroJava/BV/LBVSpec.thy Fri Dec 04 11:44:57 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/Listn.thy --- a/src/HOL/MicroJava/BV/Listn.thy Fri Dec 04 11:44:57 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/Opt.thy --- a/src/HOL/MicroJava/BV/Opt.thy Fri Dec 04 11:44:57 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/Product.thy --- a/src/HOL/MicroJava/BV/Product.thy Fri Dec 04 11:44:57 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/Semilat.thy --- a/src/HOL/MicroJava/BV/Semilat.thy Fri Dec 04 11:44:57 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/SemilatAlg.thy --- a/src/HOL/MicroJava/BV/SemilatAlg.thy Fri Dec 04 11:44:57 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/Typing_Framework.thy --- a/src/HOL/MicroJava/BV/Typing_Framework.thy Fri Dec 04 11:44:57 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/BV/Typing_Framework_err.thy --- a/src/HOL/MicroJava/BV/Typing_Framework_err.thy Fri Dec 04 11:44:57 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/Comp/AuxLemmas.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/Comp/CorrComp.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/Comp/CorrCompTp.thy --- a/src/HOL/MicroJava/Comp/CorrCompTp.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/Comp/CorrCompTp.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/Comp/DefsComp.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/Comp/Index.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/Comp/LemmasComp.thy --- a/src/HOL/MicroJava/Comp/LemmasComp.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/Comp/LemmasComp.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/Comp/NatCanonify.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/Comp/TranslComp.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/Comp/TranslCompTp.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/Comp/TypeInf.thy diff -r 651028e34b5d -r 01dcd9b926bf 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 Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/DFA/Err.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MicroJava/DFA/Err.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/DFA/Kildall.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MicroJava/DFA/Kildall.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/DFA/LBVComplete.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MicroJava/DFA/LBVComplete.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/DFA/LBVCorrect.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MicroJava/DFA/LBVCorrect.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/DFA/LBVSpec.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MicroJava/DFA/LBVSpec.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/DFA/Listn.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MicroJava/DFA/Listn.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/DFA/Opt.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MicroJava/DFA/Opt.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/DFA/Product.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MicroJava/DFA/Product.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/DFA/Semilat.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MicroJava/DFA/Semilat.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/DFA/SemilatAlg.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MicroJava/DFA/SemilatAlg.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/DFA/Semilattices.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/MicroJava/DFA/Semilattices.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf 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 Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/J/Conform.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/J/Decl.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/J/Eval.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/J/Example.thy --- a/src/HOL/MicroJava/J/Example.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/J/Example.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/J/Exceptions.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/J/JBasis.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/J/JListExample.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/J/JTypeSafe.thy --- a/src/HOL/MicroJava/J/JTypeSafe.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/J/JTypeSafe.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/J/State.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/J/SystemClasses.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/J/Term.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/J/Type.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/J/TypeRel.thy --- a/src/HOL/MicroJava/J/TypeRel.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/J/TypeRel.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/J/Value.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/J/WellForm.thy --- a/src/HOL/MicroJava/J/WellForm.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/MicroJava/J/WellForm.thy Fri Dec 04 15:20:24 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 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/J/WellType.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/JVM/JVMDefensive.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/JVM/JVMExceptions.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/JVM/JVMExec.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/JVM/JVMExecInstr.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/JVM/JVMInstructions.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/JVM/JVMListExample.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/JVM/JVMState.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/MicroJava.thy diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/ROOT.ML diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/document/introduction.tex diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/document/root.bib diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/MicroJava/document/root.tex diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Modelcheck/MuckeSyn.thy --- a/src/HOL/Modelcheck/MuckeSyn.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Modelcheck/MuckeSyn.thy Fri Dec 04 15:20:24 2009 +0100 @@ -130,7 +130,7 @@ fun move_mus i state = let val sign = Thm.theory_of_thm state; - val (subgoal::_) = Library.drop(i-1,prems_of state); + val subgoal = nth (prems_of state) i; val concl = Logic.strip_imp_concl subgoal; (* recursive mu's in prems? *) val redex = search_mu concl; val idx = let val t = #maxidx (rep_thm state) in @@ -222,9 +222,9 @@ (* the interface *) fun mc_mucke_tac defs i state = - (case Library.drop (i - 1, Thm.prems_of state) of - [] => no_tac state - | subgoal :: _ => + (case try (nth (Thm.prems_of state)) i of + NONE => no_tac state + | SOME subgoal => EVERY [ REPEAT (etac thin_rl i), cut_facts_tac (mk_lam_defs defs) i, diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Modelcheck/mucke_oracle.ML --- a/src/HOL/Modelcheck/mucke_oracle.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Modelcheck/mucke_oracle.ML Fri Dec 04 15:20:24 2009 +0100 @@ -95,7 +95,7 @@ fun if_full_simp_tac sset i state = let val sign = Thm.theory_of_thm state; - val (subgoal::_) = Library.drop(i-1,prems_of state); + val subgoal = nth (prems_of state) i; val prems = Logic.strip_imp_prems subgoal; val concl = Logic.strip_imp_concl subgoal; val prems = prems @ [concl]; diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Nominal/nominal_datatype.ML --- a/src/HOL/Nominal/nominal_datatype.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Nominal/nominal_datatype.ML Fri Dec 04 15:20:24 2009 +0100 @@ -37,7 +37,7 @@ val Collect_False_empty = @{thm empty_def [THEN sym, THEN eq_reflection]}; val empty_iff = thm "empty_iff"; -open DatatypeAux; +open Datatype_Aux; open NominalAtoms; (** FIXME: Datatype should export this function **) @@ -56,7 +56,7 @@ fun induct_cases descr = - DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr)); + Datatype_Prop.indexify_names (maps (dt_cases descr) (map #2 descr)); fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i)); @@ -258,7 +258,7 @@ val perm_types = map (fn (i, _) => let val T = nth_dtyp i in permT --> T --> T end) descr; - val perm_names' = DatatypeProp.indexify_names (map (fn (i, _) => + val perm_names' = Datatype_Prop.indexify_names (map (fn (i, _) => "perm_" ^ name_of_typ (nth_dtyp i)) descr); val perm_names = replicate (length new_type_names) "Nominal.perm" @ map (Sign.full_bname thy1) (List.drop (perm_names', length new_type_names)); @@ -270,7 +270,7 @@ in map (fn (cname, dts) => let val Ts = map (typ_of_dtyp descr sorts) dts; - val names = Name.variant_list ["pi"] (DatatypeProp.make_tnames Ts); + val names = Name.variant_list ["pi"] (Datatype_Prop.make_tnames Ts); val args = map Free (names ~~ Ts); val c = Const (cname, Ts ---> T); fun perm_arg (dt, x) = @@ -307,7 +307,7 @@ val _ = warning ("length descr: " ^ string_of_int (length descr)); val _ = warning ("length new_type_names: " ^ string_of_int (length new_type_names)); - val perm_indnames = DatatypeProp.make_tnames (map body_type perm_types); + val perm_indnames = Datatype_Prop.make_tnames (map body_type perm_types); val perm_fun_def = PureThy.get_thm thy2 "perm_fun_def"; val unfolded_perm_eq_thms = @@ -502,10 +502,10 @@ val _ = warning "representing sets"; - val rep_set_names = DatatypeProp.indexify_names + val rep_set_names = Datatype_Prop.indexify_names (map (fn (i, _) => name_of_typ (nth_dtyp i) ^ "_set") descr); val big_rep_name = - space_implode "_" (DatatypeProp.indexify_names (map_filter + space_implode "_" (Datatype_Prop.indexify_names (map_filter (fn (i, ("Nominal.noption", _, _)) => NONE | (i, _) => SOME (name_of_typ (nth_dtyp i))) descr)) ^ "_set"; val _ = warning ("big_rep_name: " ^ big_rep_name); @@ -766,8 +766,8 @@ Sign.intern_const thy7 ("Abs_" ^ s)) new_type_names; val recTs = get_rec_types descr'' sorts; - val newTs' = Library.take (length new_type_names, recTs'); - val newTs = Library.take (length new_type_names, recTs); + val newTs' = take (length new_type_names) recTs'; + val newTs = take (length new_type_names) recTs; val full_new_type_names = map (Sign.full_bname thy) new_type_names; @@ -867,7 +867,7 @@ (* prove distinctness theorems *) - val distinct_props = DatatypeProp.make_distincts descr' sorts; + val distinct_props = Datatype_Prop.make_distincts descr' sorts; val dist_rewrites = map2 (fn rep_thms => fn dist_lemma => dist_lemma :: rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0]) constr_rep_thmss dist_lemmas; @@ -1067,7 +1067,7 @@ val Abs_inverse_thms' = map (fn r => r RS subst) Abs_inverse_thms; - val dt_induct_prop = DatatypeProp.make_ind descr' sorts; + val dt_induct_prop = Datatype_Prop.make_ind descr' sorts; val dt_induct = Goal.prove_global thy8 [] (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop) (fn {prems, ...} => EVERY @@ -1085,7 +1085,7 @@ val _ = warning "proving finite support for the new datatype"; - val indnames = DatatypeProp.make_tnames recTs; + val indnames = Datatype_Prop.make_tnames recTs; val abs_supp = PureThy.get_thms thy8 "abs_supp"; val supp_atm = PureThy.get_thms thy8 "supp_atm"; @@ -1120,12 +1120,12 @@ PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] ||>> PureThy.add_thmss [((Binding.name "inducts", projections dt_induct), [case_names_induct])] ||> Sign.parent_path ||>> - DatatypeAux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>> - DatatypeAux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>> - DatatypeAux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>> - DatatypeAux.store_thmss "inject" new_type_names inject_thms ||>> - DatatypeAux.store_thmss "supp" new_type_names supp_thms ||>> - DatatypeAux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||> + Datatype_Aux.store_thmss_atts "distinct" new_type_names simp_atts distinct_thms ||>> + Datatype_Aux.store_thmss "constr_rep" new_type_names constr_rep_thmss ||>> + Datatype_Aux.store_thmss_atts "perm" new_type_names simp_eqvt_atts perm_simps' ||>> + Datatype_Aux.store_thmss "inject" new_type_names inject_thms ||>> + Datatype_Aux.store_thmss "supp" new_type_names supp_thms ||>> + Datatype_Aux.store_thmss_atts "fresh" new_type_names simp_atts fresh_thms ||> fold (fn (atom, ths) => fn thy => let val class = fs_class_of thy atom; @@ -1145,7 +1145,7 @@ val fsT' = TFree ("'n", HOLogic.typeS); val fresh_fs = map (fn (s, T) => (T, Free (s, fsT' --> HOLogic.mk_setT T))) - (DatatypeProp.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs); + (Datatype_Prop.indexify_names (replicate (length dt_atomTs) "f") ~~ dt_atomTs); fun make_pred fsT i T = Free (List.nth (pnames, i), fsT --> T --> HOLogic.boolT); @@ -1167,7 +1167,7 @@ val recs = filter is_rec_type cargs; val Ts = map (typ_of_dtyp descr'' sorts) cargs; val recTs' = map (typ_of_dtyp descr'' sorts) recs; - val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts); + val tnames = Name.variant_list pnames (Datatype_Prop.make_tnames Ts); val rec_tnames = map fst (filter (is_rec_type o snd) (tnames ~~ cargs)); val frees = tnames ~~ Ts; val frees' = partition_cargs idxs frees; @@ -1196,7 +1196,7 @@ map (make_ind_prem fsT (fn T => fn t => fn u => fresh_const T fsT $ t $ u) i T) (constrs ~~ idxss)) (descr'' ~~ ndescr ~~ recTs); - val tnames = DatatypeProp.make_tnames recTs; + val tnames = Datatype_Prop.make_tnames recTs; val zs = Name.variant_list tnames (replicate (length descr'') "z"); val ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") (map (fn ((((i, _), T), tname), z) => @@ -1220,7 +1220,7 @@ val induct' = Logic.list_implies (ind_prems', ind_concl'); val aux_ind_vars = - (DatatypeProp.indexify_names (replicate (length dt_atomTs) "pi") ~~ + (Datatype_Prop.indexify_names (replicate (length dt_atomTs) "pi") ~~ map mk_permT dt_atomTs) @ [("z", fsT')]; val aux_ind_Ts = rev (map snd aux_ind_vars); val aux_ind_concl = HOLogic.mk_Trueprop (foldr1 (HOLogic.mk_binop "op &") @@ -1418,7 +1418,7 @@ val used = List.foldr OldTerm.add_typ_tfree_names [] recTs; - val (rec_result_Ts', rec_fn_Ts') = DatatypeProp.make_primrec_Ts descr' sorts used; + val (rec_result_Ts', rec_fn_Ts') = Datatype_Prop.make_primrec_Ts descr' sorts used; val rec_sort = if null dt_atomTs then HOLogic.typeS else Sign.certify_sort thy10 pt_cp_sort; @@ -1664,10 +1664,10 @@ val fun_tuple = foldr1 HOLogic.mk_prod (rec_ctxt :: rec_fns); val fun_tupleT = fastype_of fun_tuple; val rec_unique_frees = - DatatypeProp.indexify_names (replicate (length recTs) "x") ~~ recTs; + Datatype_Prop.indexify_names (replicate (length recTs) "x") ~~ recTs; val rec_unique_frees'' = map (fn (s, T) => (s ^ "'", T)) rec_unique_frees; val rec_unique_frees' = - DatatypeProp.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts; + Datatype_Prop.indexify_names (replicate (length recTs) "y") ~~ rec_result_Ts; val rec_unique_concls = map (fn ((x, U), R) => Const ("Ex1", (U --> HOLogic.boolT) --> HOLogic.boolT) $ Abs ("y", U, R $ Free x $ Bound 0)) @@ -2048,7 +2048,7 @@ resolve_tac rec_intrs 1, REPEAT (solve (prems @ rec_total_thms) prems 1)]) end) (rec_eq_prems ~~ - DatatypeProp.make_primrecs new_type_names descr' sorts thy12); + Datatype_Prop.make_primrecs new_type_names descr' sorts thy12); val dt_infos = map_index (make_dt_info pdescr sorts induct reccomb_names rec_thms) (descr1 ~~ distinct_thms ~~ inject_thms); diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Nominal/nominal_induct.ML --- a/src/HOL/Nominal/nominal_induct.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Nominal/nominal_induct.ML Fri Dec 04 15:20:24 2009 +0100 @@ -46,7 +46,7 @@ val m = length vars and n = length inst; val _ = if m >= n + 2 then () else error "Too few variables in conclusion of rule"; val P :: x :: ys = vars; - val zs = Library.drop (m - n - 2, ys); + val zs = drop (m - n - 2) ys; in (P, tuple_fun (map #2 avoiding) (Term.dest_Var P)) :: (x, tuple (map Free avoiding)) :: @@ -71,7 +71,7 @@ val p = length ps; val ys = if p < n then [] - else map (tune o #1) (Library.take (p - n, ps)) @ xs; + else map (tune o #1) (take (p - n) ps) @ xs; in Logic.list_rename_params (ys, prem) end; fun rename_prems prop = let val (As, C) = Logic.strip_horn (Thm.prop_of rule) diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Nominal/nominal_inductive.ML --- a/src/HOL/Nominal/nominal_inductive.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Nominal/nominal_inductive.ML Fri Dec 04 15:20:24 2009 +0100 @@ -247,7 +247,7 @@ end) prems); val ind_vars = - (DatatypeProp.indexify_names (replicate (length atomTs) "pi") ~~ + (Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~ map NominalAtoms.mk_permT atomTs) @ [("z", fsT)]; val ind_Ts = rev (map snd ind_vars); @@ -647,7 +647,7 @@ val thss = map (fn atom => let val pi' = Free (pi, NominalAtoms.mk_permT (Type (atom, []))) in map (fn th => zero_var_indexes (th RS mp)) - (DatatypeAux.split_conj_thm (Goal.prove ctxt' [] [] + (Datatype_Aux.split_conj_thm (Goal.prove ctxt' [] [] (HOLogic.mk_Trueprop (foldr1 HOLogic.mk_conj (map (fn p => let val (h, ts) = strip_comb p; diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Nominal/nominal_inductive2.ML --- a/src/HOL/Nominal/nominal_inductive2.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Nominal/nominal_inductive2.ML Fri Dec 04 15:20:24 2009 +0100 @@ -263,7 +263,7 @@ in abs_params params' prem end) prems); val ind_vars = - (DatatypeProp.indexify_names (replicate (length atomTs) "pi") ~~ + (Datatype_Prop.indexify_names (replicate (length atomTs) "pi") ~~ map NominalAtoms.mk_permT atomTs) @ [("z", fsT)]; val ind_Ts = rev (map snd ind_vars); diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Nominal/nominal_primrec.ML --- a/src/HOL/Nominal/nominal_primrec.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Nominal/nominal_primrec.ML Fri Dec 04 15:20:24 2009 +0100 @@ -21,7 +21,7 @@ structure NominalPrimrec : NOMINAL_PRIMREC = struct -open DatatypeAux; +open Datatype_Aux; exception RecError of string; diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Product_Type.thy --- a/src/HOL/Product_Type.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Product_Type.thy Fri Dec 04 15:20:24 2009 +0100 @@ -6,12 +6,10 @@ header {* Cartesian products *} theory Product_Type -imports Inductive +imports Typedef Inductive Fun uses ("Tools/split_rule.ML") ("Tools/inductive_set.ML") - ("Tools/inductive_realizer.ML") - ("Tools/Datatype/datatype_realizer.ML") begin subsection {* @{typ bool} is a datatype *} @@ -1195,16 +1193,7 @@ val unit_eq = thm "unit_eq"; *} - -subsection {* Further inductive packages *} - -use "Tools/inductive_realizer.ML" -setup InductiveRealizer.setup - use "Tools/inductive_set.ML" setup Inductive_Set.setup -use "Tools/Datatype/datatype_realizer.ML" -setup DatatypeRealizer.setup - end diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Sum_Type.thy --- a/src/HOL/Sum_Type.thy Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Sum_Type.thy Fri Dec 04 15:20:24 2009 +0100 @@ -1,5 +1,4 @@ (* Title: HOL/Sum_Type.thy - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1992 University of Cambridge *) @@ -7,224 +6,186 @@ header{*The Disjoint Sum of Two Types*} theory Sum_Type -imports Typedef Fun +imports Typedef Inductive Fun begin -text{*The representations of the two injections*} +subsection {* Construction of the sum type and its basic abstract operations *} -constdefs - Inl_Rep :: "['a, 'a, 'b, bool] => bool" - "Inl_Rep == (%a. %x y p. x=a & p)" +definition Inl_Rep :: "'a \ 'a \ 'b \ bool \ bool" where + "Inl_Rep a x y p \ x = a \ p" - Inr_Rep :: "['b, 'a, 'b, bool] => bool" - "Inr_Rep == (%b. %x y p. y=b & ~p)" - +definition Inr_Rep :: "'b \ 'a \ 'b \ bool \ bool" where + "Inr_Rep b x y p \ y = b \ \ p" global -typedef (Sum) - ('a, 'b) "+" (infixr "+" 10) - = "{f. (? a. f = Inl_Rep(a::'a)) | (? b. f = Inr_Rep(b::'b))}" +typedef (Sum) ('a, 'b) "+" (infixr "+" 10) = "{f. (\a. f = Inl_Rep (a::'a)) \ (\b. f = Inr_Rep (b::'b))}" by auto local +lemma Inl_RepI: "Inl_Rep a \ Sum" + by (auto simp add: Sum_def) -text{*abstract constants and syntax*} +lemma Inr_RepI: "Inr_Rep b \ Sum" + by (auto simp add: Sum_def) + +lemma inj_on_Abs_Sum: "A \ Sum \ inj_on Abs_Sum A" + by (rule inj_on_inverseI, rule Abs_Sum_inverse) auto + +lemma Inl_Rep_inject: "inj_on Inl_Rep A" +proof (rule inj_onI) + show "\a c. Inl_Rep a = Inl_Rep c \ a = c" + by (auto simp add: Inl_Rep_def expand_fun_eq) +qed -constdefs - Inl :: "'a => 'a + 'b" - "Inl == (%a. Abs_Sum(Inl_Rep(a)))" +lemma Inr_Rep_inject: "inj_on Inr_Rep A" +proof (rule inj_onI) + show "\b d. Inr_Rep b = Inr_Rep d \ b = d" + by (auto simp add: Inr_Rep_def expand_fun_eq) +qed + +lemma Inl_Rep_not_Inr_Rep: "Inl_Rep a \ Inr_Rep b" + by (auto simp add: Inl_Rep_def Inr_Rep_def expand_fun_eq) + +definition Inl :: "'a \ 'a + 'b" where + "Inl = Abs_Sum \ Inl_Rep" + +definition Inr :: "'b \ 'a + 'b" where + "Inr = Abs_Sum \ Inr_Rep" + +lemma inj_Inl [simp]: "inj_on Inl A" +by (auto simp add: Inl_def intro!: comp_inj_on Inl_Rep_inject inj_on_Abs_Sum Inl_RepI) - Inr :: "'b => 'a + 'b" - "Inr == (%b. Abs_Sum(Inr_Rep(b)))" +lemma Inl_inject: "Inl x = Inl y \ x = y" +using inj_Inl by (rule injD) + +lemma inj_Inr [simp]: "inj_on Inr A" +by (auto simp add: Inr_def intro!: comp_inj_on Inr_Rep_inject inj_on_Abs_Sum Inr_RepI) + +lemma Inr_inject: "Inr x = Inr y \ x = y" +using inj_Inr by (rule injD) + +lemma Inl_not_Inr: "Inl a \ Inr b" +proof - + from Inl_RepI [of a] Inr_RepI [of b] have "{Inl_Rep a, Inr_Rep b} \ Sum" by auto + with inj_on_Abs_Sum have "inj_on Abs_Sum {Inl_Rep a, Inr_Rep b}" . + with Inl_Rep_not_Inr_Rep [of a b] inj_on_contraD have "Abs_Sum (Inl_Rep a) \ Abs_Sum (Inr_Rep b)" by auto + then show ?thesis by (simp add: Inl_def Inr_def) +qed - Plus :: "['a set, 'b set] => ('a + 'b) set" (infixr "<+>" 65) - "A <+> B == (Inl`A) Un (Inr`B)" - --{*disjoint sum for sets; the operator + is overloaded with wrong type!*} +lemma Inr_not_Inl: "Inr b \ Inl a" + using Inl_not_Inr by (rule not_sym) - Part :: "['a set, 'b => 'a] => 'a set" - "Part A h == A Int {x. ? z. x = h(z)}" - --{*for selecting out the components of a mutually recursive definition*} +lemma sumE: + assumes "\x::'a. s = Inl x \ P" + and "\y::'b. s = Inr y \ P" + shows P +proof (rule Abs_Sum_cases [of s]) + fix f + assume "s = Abs_Sum f" and "f \ Sum" + with assms show P by (auto simp add: Sum_def Inl_def Inr_def) +qed +rep_datatype (sum) Inl Inr +proof - + fix P + fix s :: "'a + 'b" + assume x: "\x\'a. P (Inl x)" and y: "\y\'b. P (Inr y)" + then show "P s" by (auto intro: sumE [of s]) +qed (auto dest: Inl_inject Inr_inject simp add: Inl_not_Inr) -(** Inl_Rep and Inr_Rep: Representations of the constructors **) +subsection {* Projections *} -(*This counts as a non-emptiness result for admitting 'a+'b as a type*) -lemma Inl_RepI: "Inl_Rep(a) : Sum" -by (auto simp add: Sum_def) - -lemma Inr_RepI: "Inr_Rep(b) : Sum" -by (auto simp add: Sum_def) +lemma sum_case_KK [simp]: "sum_case (\x. a) (\x. a) = (\x. a)" + by (rule ext) (simp split: sum.split) -lemma inj_on_Abs_Sum: "inj_on Abs_Sum Sum" -apply (rule inj_on_inverseI) -apply (erule Abs_Sum_inverse) -done - -subsection{*Freeness Properties for @{term Inl} and @{term Inr}*} - -text{*Distinctness*} +lemma surjective_sum: "sum_case (\x::'a. f (Inl x)) (\y::'b. f (Inr y)) = f" +proof + fix s :: "'a + 'b" + show "(case s of Inl (x\'a) \ f (Inl x) | Inr (y\'b) \ f (Inr y)) = f s" + by (cases s) simp_all +qed -lemma Inl_Rep_not_Inr_Rep: "Inl_Rep(a) ~= Inr_Rep(b)" -by (auto simp add: Inl_Rep_def Inr_Rep_def expand_fun_eq) +lemma sum_case_inject: + assumes a: "sum_case f1 f2 = sum_case g1 g2" + assumes r: "f1 = g1 \ f2 = g2 \ P" + shows P +proof (rule r) + show "f1 = g1" proof + fix x :: 'a + from a have "sum_case f1 f2 (Inl x) = sum_case g1 g2 (Inl x)" by simp + then show "f1 x = g1 x" by simp + qed + show "f2 = g2" proof + fix y :: 'b + from a have "sum_case f1 f2 (Inr y) = sum_case g1 g2 (Inr y)" by simp + then show "f2 y = g2 y" by simp + qed +qed -lemma Inl_not_Inr [iff]: "Inl(a) ~= Inr(b)" -apply (simp add: Inl_def Inr_def) -apply (rule inj_on_Abs_Sum [THEN inj_on_contraD]) -apply (rule Inl_Rep_not_Inr_Rep) -apply (rule Inl_RepI) -apply (rule Inr_RepI) -done - -lemmas Inr_not_Inl = Inl_not_Inr [THEN not_sym, standard] -declare Inr_not_Inl [iff] - -lemmas Inl_neq_Inr = Inl_not_Inr [THEN notE, standard] -lemmas Inr_neq_Inl = sym [THEN Inl_neq_Inr, standard] - - -text{*Injectiveness*} +lemma sum_case_weak_cong: + "s = t \ sum_case f g s = sum_case f g t" + -- {* Prevents simplification of @{text f} and @{text g}: much faster. *} + by simp -lemma Inl_Rep_inject: "Inl_Rep(a) = Inl_Rep(c) ==> a=c" -by (auto simp add: Inl_Rep_def expand_fun_eq) +primrec Projl :: "'a + 'b \ 'a" where + Projl_Inl: "Projl (Inl x) = x" -lemma Inr_Rep_inject: "Inr_Rep(b) = Inr_Rep(d) ==> b=d" -by (auto simp add: Inr_Rep_def expand_fun_eq) +primrec Projr :: "'a + 'b \ 'b" where + Projr_Inr: "Projr (Inr x) = x" -lemma inj_Inl [simp]: "inj_on Inl A" -apply (simp add: Inl_def) -apply (rule inj_onI) -apply (erule inj_on_Abs_Sum [THEN inj_onD, THEN Inl_Rep_inject]) -apply (rule Inl_RepI) -apply (rule Inl_RepI) -done +primrec Suml :: "('a \ 'c) \ 'a + 'b \ 'c" where + "Suml f (Inl x) = f x" -lemmas Inl_inject = inj_Inl [THEN injD, standard] +primrec Sumr :: "('b \ 'c) \ 'a + 'b \ 'c" where + "Sumr f (Inr x) = f x" -lemma inj_Inr [simp]: "inj_on Inr A" -apply (simp add: Inr_def) -apply (rule inj_onI) -apply (erule inj_on_Abs_Sum [THEN inj_onD, THEN Inr_Rep_inject]) -apply (rule Inr_RepI) -apply (rule Inr_RepI) -done - -lemmas Inr_inject = inj_Inr [THEN injD, standard] +lemma Suml_inject: + assumes "Suml f = Suml g" shows "f = g" +proof + fix x :: 'a + let ?s = "Inl x \ 'a + 'b" + from assms have "Suml f ?s = Suml g ?s" by simp + then show "f x = g x" by simp +qed -lemma Inl_eq [iff]: "(Inl(x)=Inl(y)) = (x=y)" -by (blast dest!: Inl_inject) - -lemma Inr_eq [iff]: "(Inr(x)=Inr(y)) = (x=y)" -by (blast dest!: Inr_inject) +lemma Sumr_inject: + assumes "Sumr f = Sumr g" shows "f = g" +proof + fix x :: 'b + let ?s = "Inr x \ 'a + 'b" + from assms have "Sumr f ?s = Sumr g ?s" by simp + then show "f x = g x" by simp +qed - -subsection{*The Disjoint Sum of Sets*} +subsection {* The Disjoint Sum of Sets *} -(** Introduction rules for the injections **) +definition Plus :: "'a set \ 'b set \ ('a + 'b) set" (infixr "<+>" 65) where + "A <+> B = Inl ` A \ Inr ` B" -lemma InlI [intro!]: "a : A ==> Inl(a) : A <+> B" +lemma InlI [intro!]: "a \ A \ Inl a \ A <+> B" by (simp add: Plus_def) -lemma InrI [intro!]: "b : B ==> Inr(b) : A <+> B" +lemma InrI [intro!]: "b \ B \ Inr b \ A <+> B" by (simp add: Plus_def) -(** Elimination rules **) +text {* Exhaustion rule for sums, a degenerate form of induction *} lemma PlusE [elim!]: - "[| u: A <+> B; - !!x. [| x:A; u=Inl(x) |] ==> P; - !!y. [| y:B; u=Inr(y) |] ==> P - |] ==> P" + "u \ A <+> B \ (\x. x \ A \ u = Inl x \ P) \ (\y. y \ B \ u = Inr y \ P) \ P" by (auto simp add: Plus_def) - - -text{*Exhaustion rule for sums, a degenerate form of induction*} -lemma sumE: - "[| !!x::'a. s = Inl(x) ==> P; !!y::'b. s = Inr(y) ==> P - |] ==> P" -apply (rule Abs_Sum_cases [of s]) -apply (auto simp add: Sum_def Inl_def Inr_def) -done - +lemma Plus_eq_empty_conv [simp]: "A <+> B = {} \ A = {} \ B = {}" +by auto lemma UNIV_Plus_UNIV [simp]: "UNIV <+> UNIV = UNIV" -apply (rule set_ext) -apply(rename_tac s) -apply(rule_tac s=s in sumE) -apply auto -done - -lemma Plus_eq_empty_conv[simp]: "A <+> B = {} \ A = {} \ B = {}" -by(auto) - -subsection{*The @{term Part} Primitive*} - -lemma Part_eqI [intro]: "[| a : A; a=h(b) |] ==> a : Part A h" -by (auto simp add: Part_def) - -lemmas PartI = Part_eqI [OF _ refl, standard] - -lemma PartE [elim!]: "[| a : Part A h; !!z. [| a : A; a=h(z) |] ==> P |] ==> P" -by (auto simp add: Part_def) - - -lemma Part_subset: "Part A h <= A" -by (auto simp add: Part_def) - -lemma Part_mono: "A<=B ==> Part A h <= Part B h" -by blast - -lemmas basic_monos = basic_monos Part_mono - -lemma PartD1: "a : Part A h ==> a : A" -by (simp add: Part_def) - -lemma Part_id: "Part A (%x. x) = A" -by blast - -lemma Part_Int: "Part (A Int B) h = (Part A h) Int (Part B h)" -by blast +proof (rule set_ext) + fix u :: "'a + 'b" + show "u \ UNIV <+> UNIV \ u \ UNIV" by (cases u) auto +qed -lemma Part_Collect: "Part (A Int {x. P x}) h = (Part A h) Int {x. P x}" -by blast - - -ML -{* -val Inl_RepI = thm "Inl_RepI"; -val Inr_RepI = thm "Inr_RepI"; -val inj_on_Abs_Sum = thm "inj_on_Abs_Sum"; -val Inl_Rep_not_Inr_Rep = thm "Inl_Rep_not_Inr_Rep"; -val Inl_not_Inr = thm "Inl_not_Inr"; -val Inr_not_Inl = thm "Inr_not_Inl"; -val Inl_neq_Inr = thm "Inl_neq_Inr"; -val Inr_neq_Inl = thm "Inr_neq_Inl"; -val Inl_Rep_inject = thm "Inl_Rep_inject"; -val Inr_Rep_inject = thm "Inr_Rep_inject"; -val inj_Inl = thm "inj_Inl"; -val Inl_inject = thm "Inl_inject"; -val inj_Inr = thm "inj_Inr"; -val Inr_inject = thm "Inr_inject"; -val Inl_eq = thm "Inl_eq"; -val Inr_eq = thm "Inr_eq"; -val InlI = thm "InlI"; -val InrI = thm "InrI"; -val PlusE = thm "PlusE"; -val sumE = thm "sumE"; -val Part_eqI = thm "Part_eqI"; -val PartI = thm "PartI"; -val PartE = thm "PartE"; -val Part_subset = thm "Part_subset"; -val Part_mono = thm "Part_mono"; -val PartD1 = thm "PartD1"; -val Part_id = thm "Part_id"; -val Part_Int = thm "Part_Int"; -val Part_Collect = thm "Part_Collect"; - -val basic_monos = thms "basic_monos"; -*} +hide (open) const Suml Sumr Projl Projr end diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Datatype/datatype.ML --- a/src/HOL/Tools/Datatype/datatype.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype.ML Fri Dec 04 15:20:24 2009 +0100 @@ -1,445 +1,638 @@ -(* Title: HOL/Tools/datatype.ML +(* Title: HOL/Tools/Datatype/datatype.ML Author: Stefan Berghofer, TU Muenchen -Datatype package for Isabelle/HOL. +Datatype package: definitional introduction of datatypes +with proof of characteristic theorems: injectivity / distinctness +of constructors and induction. Main interface to datatypes +after full bootstrap of datatype package. *) signature DATATYPE = sig - include DATATYPE_COMMON + include DATATYPE_DATA val add_datatype : config -> string list -> (string list * binding * mixfix * (binding * typ list * mixfix) list) list -> theory -> string list * theory val datatype_cmd : string list -> (string list * binding * mixfix * (binding * string list * mixfix) list) list -> theory -> theory - val rep_datatype : config -> (string list -> Proof.context -> Proof.context) - -> string list option -> term list -> theory -> Proof.state - val rep_datatype_cmd : string list option -> string list -> theory -> Proof.state - val get_info : theory -> string -> info option - val the_info : theory -> string -> info - val the_descr : theory -> string list - -> descr * (string * sort) list * string list - * string * (string list * string list) * (typ list * typ list) - val the_spec : theory -> string -> (string * sort) list * (string * typ list) list - val all_distincts : theory -> typ list -> thm list list - val get_constrs : theory -> string -> (string * typ) list option - val get_all : theory -> info Symtab.table - val info_of_constr : theory -> string * typ -> info option - val info_of_case : theory -> string -> info option - val interpretation : (config -> string list -> theory -> theory) -> theory -> theory - val make_case : Proof.context -> DatatypeCase.config -> string list -> term -> - (term * term) list -> term * (term * (int * bool)) list - val strip_case : Proof.context -> bool -> term -> (term * (term * term) list) option - val read_typ: theory -> string -> (string * sort) list -> typ * (string * sort) list - val setup: theory -> theory end; structure Datatype : DATATYPE = struct -open DatatypeAux; - -(** theory data **) - -(* data management *) - -structure DatatypesData = Theory_Data -( - type T = - {types: info Symtab.table, - constrs: (string * info) list Symtab.table, - cases: info Symtab.table}; +(** auxiliary **) - val empty = - {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty}; - val extend = I; - fun merge - ({types = types1, constrs = constrs1, cases = cases1}, - {types = types2, constrs = constrs2, cases = cases2}) : T = - {types = Symtab.merge (K true) (types1, types2), - constrs = Symtab.join (K (AList.merge (op =) (K true))) (constrs1, constrs2), - cases = Symtab.merge (K true) (cases1, cases2)}; -); - -val get_all = #types o DatatypesData.get; -val get_info = Symtab.lookup o get_all; - -fun the_info thy name = - (case get_info thy name of - SOME info => info - | NONE => error ("Unknown datatype " ^ quote name)); - -fun info_of_constr thy (c, T) = - let - val tab = Symtab.lookup_list ((#constrs o DatatypesData.get) thy) c; - val hint = case strip_type T of (_, Type (tyco, _)) => SOME tyco | _ => NONE; - val default = if null tab then NONE - else SOME (snd (Library.last_elem tab)) - (*conservative wrt. overloaded constructors*); - in case hint - of NONE => default - | SOME tyco => case AList.lookup (op =) tab tyco - of NONE => default (*permissive*) - | SOME info => SOME info - end; +open Datatype_Aux; +open Datatype_Data; -val info_of_case = Symtab.lookup o #cases o DatatypesData.get; - -fun register (dt_infos : (string * info) list) = - DatatypesData.map (fn {types, constrs, cases} => - {types = types |> fold Symtab.update dt_infos, - constrs = constrs |> fold (fn (constr, dtname_info) => - Symtab.map_default (constr, []) (cons dtname_info)) - (maps (fn (dtname, info as {descr, index, ...}) => - map (rpair (dtname, info) o fst) - (#3 (the (AList.lookup op = descr index)))) dt_infos), - cases = cases |> fold Symtab.update - (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)}); - - -(* complex queries *) +val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma); -fun the_spec thy dtco = - let - val info as { descr, index, sorts = raw_sorts, ... } = the_info thy dtco; - val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr index; - val sorts = map ((fn v => (v, (the o AList.lookup (op =) raw_sorts) v)) - o DatatypeAux.dest_DtTFree) dtys; - val cos = map - (fn (co, tys) => (co, map (DatatypeAux.typ_of_dtyp descr sorts) tys)) raw_cos; - in (sorts, cos) end; - -fun the_descr thy (raw_tycos as raw_tyco :: _) = - let - val info = the_info thy raw_tyco; - val descr = #descr info; +val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq]; - val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr (#index info); - val vs = map ((fn v => (v, (the o AList.lookup (op =) (#sorts info)) v)) - o dest_DtTFree) dtys; - - fun is_DtTFree (DtTFree _) = true - | is_DtTFree _ = false - val k = find_index (fn (_, (_, dTs, _)) => not (forall is_DtTFree dTs)) descr; - val protoTs as (dataTs, _) = chop k descr - |> (pairself o map) (fn (_, (tyco, dTs, _)) => (tyco, map (typ_of_dtyp descr vs) dTs)); - - val tycos = map fst dataTs; - val _ = if eq_set (op =) (tycos, raw_tycos) then () - else error ("Type constructors " ^ commas (map quote raw_tycos) - ^ " do not belong exhaustively to one mutual recursive datatype"); - - val (Ts, Us) = (pairself o map) Type protoTs; +fun exh_thm_of (dt_info : info Symtab.table) tname = + #exhaust (the (Symtab.lookup dt_info tname)); - val names = map Long_Name.base_name (the_default tycos (#alt_names info)); - val (auxnames, _) = Name.make_context names - |> fold_map (yield_singleton Name.variants o name_of_typ) Us; - val prefix = space_implode "_" names; - - in (descr, vs, tycos, prefix, (names, auxnames), (Ts, Us)) end; - -fun all_distincts thy Ts = - let - fun add_tycos (Type (tyco, Ts)) = insert (op =) tyco #> fold add_tycos Ts - | add_tycos _ = I; - val tycos = fold add_tycos Ts []; - in map_filter (Option.map #distinct o get_info thy) tycos end; - -fun get_constrs thy dtco = - case try (the_spec thy) dtco - of SOME (sorts, cos) => - let - fun subst (v, sort) = TVar ((v, 0), sort); - fun subst_ty (TFree v) = subst v - | subst_ty ty = ty; - val dty = Type (dtco, map subst sorts); - fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty); - in SOME (map mk_co cos) end - | NONE => NONE; - - - -(** various auxiliary **) - -(* prepare datatype specifications *) - -fun read_typ thy str sorts = - let - val ctxt = ProofContext.init thy - |> fold (Variable.declare_typ o TFree) sorts; - val T = Syntax.read_typ ctxt str; - in (T, Term.add_tfreesT T sorts) end; - -fun cert_typ sign raw_T sorts = - let - val T = Type.no_tvars (Sign.certify_typ sign raw_T) - handle TYPE (msg, _, _) => error msg; - val sorts' = Term.add_tfreesT T sorts; - val _ = - case duplicates (op =) (map fst sorts') of - [] => () - | dups => error ("Inconsistent sort constraints for " ^ commas dups) - in (T, sorts') end; +val node_name = @{type_name "Datatype.node"}; +val In0_name = @{const_name "Datatype.In0"}; +val In1_name = @{const_name "Datatype.In1"}; +val Scons_name = @{const_name "Datatype.Scons"}; +val Leaf_name = @{const_name "Datatype.Leaf"}; +val Lim_name = @{const_name "Datatype.Lim"}; +val Suml_name = @{const_name "Sum_Type.Suml"}; +val Sumr_name = @{const_name "Sum_Type.Sumr"}; - -(* case names *) - -local - -fun dt_recs (DtTFree _) = [] - | dt_recs (DtType (_, dts)) = maps dt_recs dts - | dt_recs (DtRec i) = [i]; - -fun dt_cases (descr: descr) (_, args, constrs) = - let - fun the_bname i = Long_Name.base_name (#1 (the (AList.lookup (op =) descr i))); - val bnames = map the_bname (distinct (op =) (maps dt_recs args)); - in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end; - -fun induct_cases descr = - DatatypeProp.indexify_names (maps (dt_cases descr) (map #2 descr)); - -fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i)); - -in - -fun mk_case_names_induct descr = Rule_Cases.case_names (induct_cases descr); - -fun mk_case_names_exhausts descr new = - map (Rule_Cases.case_names o exhaust_cases descr o #1) - (filter (fn ((_, (name, _, _))) => member (op =) new name) descr); - -end; - - -(* translation rules for case *) - -fun make_case ctxt = DatatypeCase.make_case - (info_of_constr (ProofContext.theory_of ctxt)) ctxt; - -fun strip_case ctxt = DatatypeCase.strip_case - (info_of_case (ProofContext.theory_of ctxt)); - -fun add_case_tr' case_names thy = - Sign.add_advanced_trfuns ([], [], - map (fn case_name => - let val case_name' = Sign.const_syntax_name thy case_name - in (case_name', DatatypeCase.case_tr' info_of_case case_name') - end) case_names, []) thy; - -val trfun_setup = - Sign.add_advanced_trfuns ([], - [("_case_syntax", DatatypeCase.case_tr true info_of_constr)], - [], []); +val In0_inject = @{thm In0_inject}; +val In1_inject = @{thm In1_inject}; +val Scons_inject = @{thm Scons_inject}; +val Leaf_inject = @{thm Leaf_inject}; +val In0_eq = @{thm In0_eq}; +val In1_eq = @{thm In1_eq}; +val In0_not_In1 = @{thm In0_not_In1}; +val In1_not_In0 = @{thm In1_not_In0}; +val Lim_inject = @{thm Lim_inject}; +val Inl_inject = @{thm Inl_inject}; +val Inr_inject = @{thm Inr_inject}; +val Suml_inject = @{thm Suml_inject}; +val Sumr_inject = @{thm Sumr_inject}; -(** document antiquotation **) +(** proof of characteristic theorems **) + +fun representation_proofs (config : config) (dt_info : info Symtab.table) + new_type_names descr sorts types_syntax constr_syntax case_names_induct thy = + let + val descr' = flat descr; + val big_name = space_implode "_" new_type_names; + val thy1 = Sign.add_path big_name thy; + val big_rec_name = big_name ^ "_rep_set"; + val rep_set_names' = + (if length descr' = 1 then [big_rec_name] else + (map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int) + (1 upto (length descr')))); + val rep_set_names = map (Sign.full_bname thy1) rep_set_names'; + + val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr); + val leafTs' = get_nonrec_types descr' sorts; + val branchTs = get_branching_types descr' sorts; + val branchT = if null branchTs then HOLogic.unitT + else Balanced_Tree.make (fn (T, U) => Type ("+", [T, U])) branchTs; + val arities = remove (op =) 0 (get_arities descr'); + val unneeded_vars = + subtract (op =) (List.foldr OldTerm.add_typ_tfree_names [] (leafTs' @ branchTs)) (hd tyvars); + val leafTs = leafTs' @ map (fn n => TFree (n, (the o AList.lookup (op =) sorts) n)) unneeded_vars; + val recTs = get_rec_types descr' sorts; + val (newTs, oldTs) = chop (length (hd descr)) recTs; + val sumT = if null leafTs then HOLogic.unitT + else Balanced_Tree.make (fn (T, U) => Type ("+", [T, U])) leafTs; + val Univ_elT = HOLogic.mk_setT (Type (node_name, [sumT, branchT])); + val UnivT = HOLogic.mk_setT Univ_elT; + val UnivT' = Univ_elT --> HOLogic.boolT; + val Collect = Const (@{const_name Collect}, UnivT' --> UnivT); + + val In0 = Const (In0_name, Univ_elT --> Univ_elT); + val In1 = Const (In1_name, Univ_elT --> Univ_elT); + val Leaf = Const (Leaf_name, sumT --> Univ_elT); + val Lim = Const (Lim_name, (branchT --> Univ_elT) --> Univ_elT); + + (* make injections needed for embedding types in leaves *) + + fun mk_inj T' x = + let + fun mk_inj' T n i = + if n = 1 then x else + let val n2 = n div 2; + val Type (_, [T1, T2]) = T + in + if i <= n2 then + Const (@{const_name "Sum_Type.Inl"}, T1 --> T) $ (mk_inj' T1 n2 i) + else + Const (@{const_name "Sum_Type.Inr"}, T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2)) + end + in mk_inj' sumT (length leafTs) (1 + find_index (fn T'' => T'' = T') leafTs) + end; + + (* make injections for constructors *) + + fun mk_univ_inj ts = Balanced_Tree.access + {left = fn t => In0 $ t, + right = fn t => In1 $ t, + init = + if ts = [] then Const (@{const_name undefined}, Univ_elT) + else foldr1 (HOLogic.mk_binop Scons_name) ts}; + + (* function spaces *) + + fun mk_fun_inj T' x = + let + fun mk_inj T n i = + if n = 1 then x else + let + val n2 = n div 2; + val Type (_, [T1, T2]) = T; + fun mkT U = (U --> Univ_elT) --> T --> Univ_elT + in + if i <= n2 then Const (Suml_name, mkT T1) $ mk_inj T1 n2 i + else Const (Sumr_name, mkT T2) $ mk_inj T2 (n - n2) (i - n2) + end + in mk_inj branchT (length branchTs) (1 + find_index (fn T'' => T'' = T') branchTs) + end; + + fun mk_lim t Ts = fold_rev (fn T => fn t => Lim $ mk_fun_inj T (Abs ("x", T, t))) Ts t; + + (************** generate introduction rules for representing set **********) + + val _ = message config "Constructing representing sets ..."; + + (* make introduction rule for a single constructor *) -val _ = ThyOutput.antiquotation "datatype" Args.tyname - (fn {source = src, context = ctxt, ...} => fn dtco => - let - val thy = ProofContext.theory_of ctxt; - val (vs, cos) = the_spec thy dtco; - val ty = Type (dtco, map TFree vs); - fun pretty_typ_bracket (ty as Type (_, _ :: _)) = - Pretty.enclose "(" ")" [Syntax.pretty_typ ctxt ty] - | pretty_typ_bracket ty = - Syntax.pretty_typ ctxt ty; - fun pretty_constr (co, tys) = - (Pretty.block o Pretty.breaks) - (Syntax.pretty_term ctxt (Const (co, tys ---> ty)) :: - map pretty_typ_bracket tys); - val pretty_datatype = - Pretty.block - (Pretty.command "datatype" :: Pretty.brk 1 :: - Syntax.pretty_typ ctxt ty :: - Pretty.str " =" :: Pretty.brk 1 :: - flat (separate [Pretty.brk 1, Pretty.str "| "] - (map (single o pretty_constr) cos))); - in ThyOutput.output (ThyOutput.maybe_pretty_source (K pretty_datatype) src [()]) end); + fun make_intr s n (i, (_, cargs)) = + let + fun mk_prem dt (j, prems, ts) = + (case strip_dtyp dt of + (dts, DtRec k) => + let + val Ts = map (typ_of_dtyp descr' sorts) dts; + val free_t = + app_bnds (mk_Free "x" (Ts ---> Univ_elT) j) (length Ts) + in (j + 1, list_all (map (pair "x") Ts, + HOLogic.mk_Trueprop + (Free (nth rep_set_names' k, UnivT') $ free_t)) :: prems, + mk_lim free_t Ts :: ts) + end + | _ => + let val T = typ_of_dtyp descr' sorts dt + in (j + 1, prems, (Leaf $ mk_inj T (mk_Free "x" T j))::ts) + end); + + val (_, prems, ts) = fold_rev mk_prem cargs (1, [], []); + val concl = HOLogic.mk_Trueprop + (Free (s, UnivT') $ mk_univ_inj ts n i) + in Logic.list_implies (prems, concl) + end; + + val intr_ts = maps (fn ((_, (_, _, constrs)), rep_set_name) => + map (make_intr rep_set_name (length constrs)) + ((1 upto (length constrs)) ~~ constrs)) (descr' ~~ rep_set_names'); + + val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy2) = + thy1 + |> Sign.map_naming Name_Space.conceal + |> Inductive.add_inductive_global + {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name, + coind = false, no_elim = true, no_ind = false, skip_mono = true, fork_mono = false} + (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') [] + (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] + ||> Sign.restore_naming thy1 + ||> Theory.checkpoint; + + (********************************* typedef ********************************) + + val (typedefs, thy3) = thy2 |> + Sign.parent_path |> + fold_map (fn ((((name, mx), tvs), c), name') => + Typedef.add_typedef false (SOME (Binding.name name')) (name, tvs, mx) + (Collect $ Const (c, UnivT')) NONE + (rtac exI 1 THEN rtac CollectI 1 THEN + QUIET_BREADTH_FIRST (has_fewer_prems 1) + (resolve_tac rep_intrs 1))) + (types_syntax ~~ tyvars ~~ + (take (length newTs) rep_set_names) ~~ new_type_names) ||> + Sign.add_path big_name; + + (*********************** definition of constructors ***********************) + val big_rep_name = (space_implode "_" new_type_names) ^ "_Rep_"; + val rep_names = map (curry op ^ "Rep_") new_type_names; + val rep_names' = map (fn i => big_rep_name ^ (string_of_int i)) + (1 upto (length (flat (tl descr)))); + val all_rep_names = map (Sign.intern_const thy3) rep_names @ + map (Sign.full_bname thy3) rep_names'; + + (* isomorphism declarations *) + + val iso_decls = map (fn (T, s) => (Binding.name s, T --> Univ_elT, NoSyn)) + (oldTs ~~ rep_names'); + + (* constructor definitions *) + + fun make_constr_def tname T n ((cname, cargs), (cname', mx)) (thy, defs, eqns, i) = + let + fun constr_arg dt (j, l_args, r_args) = + let val T = typ_of_dtyp descr' sorts dt; + val free_t = mk_Free "x" T j + in (case (strip_dtyp dt, strip_type T) of + ((_, DtRec m), (Us, U)) => (j + 1, free_t :: l_args, mk_lim + (Const (nth all_rep_names m, U --> Univ_elT) $ + app_bnds free_t (length Us)) Us :: r_args) + | _ => (j + 1, free_t::l_args, (Leaf $ mk_inj T free_t)::r_args)) + end; + + val (_, l_args, r_args) = fold_rev constr_arg cargs (1, [], []); + val constrT = (map (typ_of_dtyp descr' sorts) cargs) ---> T; + val abs_name = Sign.intern_const thy ("Abs_" ^ tname); + val rep_name = Sign.intern_const thy ("Rep_" ^ tname); + val lhs = list_comb (Const (cname, constrT), l_args); + val rhs = mk_univ_inj r_args n i; + val def = Logic.mk_equals (lhs, Const (abs_name, Univ_elT --> T) $ rhs); + val def_name = Long_Name.base_name cname ^ "_def"; + val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq + (Const (rep_name, T --> Univ_elT) $ lhs, rhs)); + val ([def_thm], thy') = + thy + |> Sign.add_consts_i [(cname', constrT, mx)] + |> (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)]; + + in (thy', defs @ [def_thm], eqns @ [eqn], i + 1) end; + + (* constructor definitions for datatype *) + + fun dt_constr_defs ((((_, (_, _, constrs)), tname), T), constr_syntax) + (thy, defs, eqns, rep_congs, dist_lemmas) = + let + val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong; + val rep_const = cterm_of thy + (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> Univ_elT)); + val cong' = + Drule.standard (cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong); + val dist = + Drule.standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma); + val (thy', defs', eqns', _) = fold ((make_constr_def tname T) (length constrs)) + (constrs ~~ constr_syntax) (Sign.add_path tname thy, defs, [], 1); + in + (Sign.parent_path thy', defs', eqns @ [eqns'], + rep_congs @ [cong'], dist_lemmas @ [dist]) + end; + + val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) = + fold dt_constr_defs + (hd descr ~~ new_type_names ~~ newTs ~~ constr_syntax) + (thy3 |> Sign.add_consts_i iso_decls |> Sign.parent_path, [], [], [], []); -(** abstract theory extensions relative to a datatype characterisation **) + (*********** isomorphisms for new types (introduced by typedef) ***********) + + val _ = message config "Proving isomorphism properties ..."; + + val newT_iso_axms = map (fn (_, td) => + (collect_simp (#Abs_inverse td), #Rep_inverse td, + collect_simp (#Rep td))) typedefs; + + val newT_iso_inj_thms = map (fn (_, td) => + (collect_simp (#Abs_inject td) RS iffD1, #Rep_inject td RS iffD1)) typedefs; + + (********* isomorphisms between existing types and "unfolded" types *******) -structure Datatype_Interpretation = Interpretation - (type T = config * string list val eq: T * T -> bool = eq_snd op =); -fun interpretation f = Datatype_Interpretation.interpretation (uncurry f); + (*---------------------------------------------------------------------*) + (* isomorphisms are defined using primrec-combinators: *) + (* generate appropriate functions for instantiating primrec-combinator *) + (* *) + (* e.g. dt_Rep_i = list_rec ... (%h t y. In1 (Scons (Leaf h) y)) *) + (* *) + (* also generate characteristic equations for isomorphisms *) + (* *) + (* e.g. dt_Rep_i (cons h t) = In1 (Scons (dt_Rep_j h) (dt_Rep_i t)) *) + (*---------------------------------------------------------------------*) + + fun make_iso_def k ks n (cname, cargs) (fs, eqns, i) = + let + val argTs = map (typ_of_dtyp descr' sorts) cargs; + val T = nth recTs k; + val rep_name = nth all_rep_names k; + val rep_const = Const (rep_name, T --> Univ_elT); + val constr = Const (cname, argTs ---> T); -fun make_dt_info alt_names descr sorts induct inducts rec_names rec_rewrites - (index, (((((((((((_, (tname, _, _))), inject), distinct), - exhaust), nchotomy), case_name), case_rewrites), case_cong), weak_case_cong), - (split, split_asm))) = - (tname, - {index = index, - alt_names = alt_names, - descr = descr, - sorts = sorts, - inject = inject, - distinct = distinct, - induct = induct, - inducts = inducts, - exhaust = exhaust, - nchotomy = nchotomy, - rec_names = rec_names, - rec_rewrites = rec_rewrites, - case_name = case_name, - case_rewrites = case_rewrites, - case_cong = case_cong, - weak_case_cong = weak_case_cong, - split = split, - split_asm = split_asm}); + fun process_arg ks' dt (i2, i2', ts, Ts) = + let + val T' = typ_of_dtyp descr' sorts dt; + val (Us, U) = strip_type T' + in (case strip_dtyp dt of + (_, DtRec j) => if j mem ks' then + (i2 + 1, i2' + 1, ts @ [mk_lim (app_bnds + (mk_Free "y" (Us ---> Univ_elT) i2') (length Us)) Us], + Ts @ [Us ---> Univ_elT]) + else + (i2 + 1, i2', ts @ [mk_lim + (Const (nth all_rep_names j, U --> Univ_elT) $ + app_bnds (mk_Free "x" T' i2) (length Us)) Us], Ts) + | _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (mk_Free "x" T' i2)], Ts)) + end; + + val (i2, i2', ts, Ts) = fold (process_arg ks) cargs (1, 1, [], []); + val xs = map (uncurry (mk_Free "x")) (argTs ~~ (1 upto (i2 - 1))); + val ys = map (uncurry (mk_Free "y")) (Ts ~~ (1 upto (i2' - 1))); + val f = list_abs_free (map dest_Free (xs @ ys), mk_univ_inj ts n i); + + val (_, _, ts', _) = fold (process_arg []) cargs (1, 1, [], []); + val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq + (rep_const $ list_comb (constr, xs), mk_univ_inj ts' n i)) + + in (fs @ [f], eqns @ [eqn], i + 1) end; + + (* define isomorphisms for all mutually recursive datatypes in list ds *) -fun derive_datatype_props config dt_names alt_names descr sorts - induct inject distinct thy1 = - let - val thy2 = thy1 |> Theory.checkpoint; - val flat_descr = flat descr; - val new_type_names = map Long_Name.base_name (the_default dt_names alt_names); - val _ = message config ("Deriving properties for datatype(s) " ^ commas_quote new_type_names); + fun make_iso_defs ds (thy, char_thms) = + let + val ks = map fst ds; + val (_, (tname, _, _)) = hd ds; + val {rec_rewrites, rec_names, ...} = the (Symtab.lookup dt_info tname); + + fun process_dt (k, (tname, _, constrs)) (fs, eqns, isos) = + let + val (fs', eqns', _) = + fold (make_iso_def k ks (length constrs)) constrs (fs, eqns, 1); + val iso = (nth recTs k, nth all_rep_names k) + in (fs', eqns', isos @ [iso]) end; + + val (fs, eqns, isos) = fold process_dt ds ([], [], []); + val fTs = map fastype_of fs; + val defs = map (fn (rec_name, (T, iso_name)) => (Binding.name (Long_Name.base_name iso_name ^ "_def"), + Logic.mk_equals (Const (iso_name, T --> Univ_elT), + list_comb (Const (rec_name, fTs @ [T] ---> Univ_elT), fs)))) (rec_names ~~ isos); + val (def_thms, thy') = + apsnd Theory.checkpoint ((PureThy.add_defs false o map Thm.no_attributes) defs thy); + + (* prove characteristic equations *) + + val rewrites = def_thms @ (map mk_meta_eq rec_rewrites); + val char_thms' = map (fn eqn => Skip_Proof.prove_global thy' [] [] eqn + (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns; + + in (thy', char_thms' @ char_thms) end; - val (exhaust, thy3) = DatatypeAbsProofs.prove_casedist_thms config new_type_names - descr sorts induct (mk_case_names_exhausts flat_descr dt_names) thy2; - val (nchotomys, thy4) = DatatypeAbsProofs.prove_nchotomys config new_type_names - descr sorts exhaust thy3; - val ((rec_names, rec_rewrites), thy5) = DatatypeAbsProofs.prove_primrec_thms - config new_type_names descr sorts (#inject o the o Symtab.lookup (get_all thy4)) - inject (distinct, all_distincts thy2 (get_rec_types flat_descr sorts)) - induct thy4; - val ((case_rewrites, case_names), thy6) = DatatypeAbsProofs.prove_case_thms - config new_type_names descr sorts rec_names rec_rewrites thy5; - val (case_congs, thy7) = DatatypeAbsProofs.prove_case_congs new_type_names - descr sorts nchotomys case_rewrites thy6; - val (weak_case_congs, thy8) = DatatypeAbsProofs.prove_weak_case_congs new_type_names - descr sorts thy7; - val (splits, thy9) = DatatypeAbsProofs.prove_split_thms - config new_type_names descr sorts inject distinct exhaust case_rewrites thy8; + val (thy5, iso_char_thms) = apfst Theory.checkpoint (fold_rev make_iso_defs + (tl descr) (Sign.add_path big_name thy4, [])); + + (* prove isomorphism properties *) + + fun mk_funs_inv thy thm = + let + val prop = Thm.prop_of thm; + val _ $ (_ $ ((S as Const (_, Type (_, [U, _]))) $ _ )) $ + (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.legacy_freeze prop; + val used = OldTerm.add_term_tfree_names (a, []); + + fun mk_thm i = + let + val Ts = map (TFree o rpair HOLogic.typeS) + (Name.variant_list used (replicate i "'t")); + val f = Free ("f", Ts ---> U) + in Skip_Proof.prove_global thy [] [] (Logic.mk_implies + (HOLogic.mk_Trueprop (HOLogic.list_all + (map (pair "x") Ts, S $ app_bnds f i)), + HOLogic.mk_Trueprop (HOLogic.mk_eq (list_abs (map (pair "x") Ts, + r $ (a $ app_bnds f i)), f)))) + (fn _ => EVERY [REPEAT_DETERM_N i (rtac ext 1), + REPEAT (etac allE 1), rtac thm 1, atac 1]) + end + in map (fn r => r RS subst) (thm :: map mk_thm arities) end; + + (* prove inj dt_Rep_i and dt_Rep_i x : dt_rep_set_i *) + + val fun_congs = map (fn T => make_elim (Drule.instantiate' + [SOME (ctyp_of thy5 T)] [] fun_cong)) branchTs; + + fun prove_iso_thms ds (inj_thms, elem_thms) = + let + val (_, (tname, _, _)) = hd ds; + val induct = (#induct o the o Symtab.lookup dt_info) tname; - val inducts = Project_Rule.projections (ProofContext.init thy2) induct; - val dt_infos = map_index - (make_dt_info alt_names flat_descr sorts induct inducts rec_names rec_rewrites) - (hd descr ~~ inject ~~ distinct ~~ exhaust ~~ nchotomys ~~ - case_names ~~ case_rewrites ~~ case_congs ~~ weak_case_congs ~~ splits); - val dt_names = map fst dt_infos; - val prfx = Binding.qualify true (space_implode "_" new_type_names); - val simps = flat (inject @ distinct @ case_rewrites) @ rec_rewrites; - val named_rules = flat (map_index (fn (index, tname) => - [((Binding.empty, [nth inducts index]), [Induct.induct_type tname]), - ((Binding.empty, [nth exhaust index]), [Induct.cases_type tname])]) dt_names); - val unnamed_rules = map (fn induct => - ((Binding.empty, [induct]), [Rule_Cases.inner_rule, Induct.induct_type ""])) - (Library.drop (length dt_names, inducts)); - in - thy9 - |> PureThy.add_thmss ([((prfx (Binding.name "simps"), simps), []), - ((prfx (Binding.name "inducts"), inducts), []), - ((prfx (Binding.name "splits"), maps (fn (x, y) => [x, y]) splits), []), - ((Binding.empty, flat case_rewrites @ flat distinct @ rec_rewrites), - [Simplifier.simp_add]), - ((Binding.empty, rec_rewrites), [Code.add_default_eqn_attribute]), - ((Binding.empty, flat inject), [iff_add]), - ((Binding.empty, map (fn th => th RS notE) (flat distinct)), - [Classical.safe_elim NONE]), - ((Binding.empty, weak_case_congs), [Simplifier.attrib (op addcongs)])] - @ named_rules @ unnamed_rules) - |> snd - |> add_case_tr' case_names - |> register dt_infos - |> Datatype_Interpretation.data (config, dt_names) - |> pair dt_names - end; + fun mk_ind_concl (i, _) = + let + val T = nth recTs i; + val Rep_t = Const (nth all_rep_names i, T --> Univ_elT); + val rep_set_name = nth rep_set_names i + in (HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $ + HOLogic.mk_eq (Rep_t $ mk_Free "x" T i, Rep_t $ Bound 0) $ + HOLogic.mk_eq (mk_Free "x" T i, Bound 0)), + Const (rep_set_name, UnivT') $ (Rep_t $ mk_Free "x" T i)) + end; + + val (ind_concl1, ind_concl2) = ListPair.unzip (map mk_ind_concl ds); + + val rewrites = map mk_meta_eq iso_char_thms; + val inj_thms' = map snd newT_iso_inj_thms @ + map (fn r => r RS @{thm injD}) inj_thms; + val inj_thm = Skip_Proof.prove_global thy5 [] [] + (HOLogic.mk_Trueprop (mk_conj ind_concl1)) (fn _ => EVERY + [(indtac induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, + REPEAT (EVERY + [rtac allI 1, rtac impI 1, + exh_tac (exh_thm_of dt_info) 1, + REPEAT (EVERY + [hyp_subst_tac 1, + rewrite_goals_tac rewrites, + REPEAT (dresolve_tac [In0_inject, In1_inject] 1), + (eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1) + ORELSE (EVERY + [REPEAT (eresolve_tac (Scons_inject :: + map make_elim [Leaf_inject, Inl_inject, Inr_inject]) 1), + REPEAT (cong_tac 1), rtac refl 1, + REPEAT (atac 1 ORELSE (EVERY + [REPEAT (rtac ext 1), + REPEAT (eresolve_tac (mp :: allE :: + map make_elim (Suml_inject :: Sumr_inject :: + Lim_inject :: inj_thms') @ fun_congs) 1), + atac 1]))])])])]); + + val inj_thms'' = map (fn r => r RS @{thm datatype_injI}) + (split_conj_thm inj_thm); + + val elem_thm = + Skip_Proof.prove_global thy5 [] [] (HOLogic.mk_Trueprop (mk_conj ind_concl2)) + (fn _ => + EVERY [(indtac induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, + rewrite_goals_tac rewrites, + REPEAT ((resolve_tac rep_intrs THEN_ALL_NEW + ((REPEAT o etac allE) THEN' ares_tac elem_thms)) 1)]); + + in (inj_thms'' @ inj_thms, elem_thms @ (split_conj_thm elem_thm)) + end; + val (iso_inj_thms_unfolded, iso_elem_thms) = + fold_rev prove_iso_thms (tl descr) ([], map #3 newT_iso_axms); + val iso_inj_thms = map snd newT_iso_inj_thms @ + map (fn r => r RS @{thm injD}) iso_inj_thms_unfolded; -(** declare existing type as datatype **) + (* prove dt_rep_set_i x --> x : range dt_Rep_i *) + + fun mk_iso_t (((set_name, iso_name), i), T) = + let val isoT = T --> Univ_elT + in HOLogic.imp $ + (Const (set_name, UnivT') $ mk_Free "x" Univ_elT i) $ + (if i < length newTs then HOLogic.true_const + else HOLogic.mk_mem (mk_Free "x" Univ_elT i, + Const (@{const_name image}, isoT --> HOLogic.mk_setT T --> UnivT) $ + Const (iso_name, isoT) $ Const (@{const_name UNIV}, HOLogic.mk_setT T))) + end; + + val iso_t = HOLogic.mk_Trueprop (mk_conj (map mk_iso_t + (rep_set_names ~~ all_rep_names ~~ (0 upto (length descr' - 1)) ~~ recTs))); + + (* all the theorems are proved by one single simultaneous induction *) + + val range_eqs = map (fn r => mk_meta_eq (r RS @{thm range_ex1_eq})) + iso_inj_thms_unfolded; + + val iso_thms = if length descr = 1 then [] else + drop (length newTs) (split_conj_thm + (Skip_Proof.prove_global thy5 [] [] iso_t (fn _ => EVERY + [(indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, + REPEAT (rtac TrueI 1), + rewrite_goals_tac (mk_meta_eq choice_eq :: + symmetric (mk_meta_eq @{thm expand_fun_eq}) :: range_eqs), + rewrite_goals_tac (map symmetric range_eqs), + REPEAT (EVERY + [REPEAT (eresolve_tac ([rangeE, ex1_implies_ex RS exE] @ + maps (mk_funs_inv thy5 o #1) newT_iso_axms) 1), + TRY (hyp_subst_tac 1), + rtac (sym RS range_eqI) 1, + resolve_tac iso_char_thms 1])]))); + + val Abs_inverse_thms' = + map #1 newT_iso_axms @ + map2 (fn r_inj => fn r => @{thm f_the_inv_into_f} OF [r_inj, r RS mp]) + iso_inj_thms_unfolded iso_thms; -fun prove_rep_datatype config dt_names alt_names descr sorts - raw_inject half_distinct raw_induct thy1 = - let - val raw_distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct; - val new_type_names = map Long_Name.base_name (the_default dt_names alt_names); - val (((inject, distinct), [induct]), thy2) = - thy1 - |> store_thmss "inject" new_type_names raw_inject - ||>> store_thmss "distinct" new_type_names raw_distinct - ||> Sign.add_path (space_implode "_" new_type_names) - ||>> PureThy.add_thms [((Binding.name "induct", raw_induct), [mk_case_names_induct descr])] - ||> Sign.restore_naming thy1; - in - thy2 - |> derive_datatype_props config dt_names alt_names [descr] sorts - induct inject distinct - end; + val Abs_inverse_thms = maps (mk_funs_inv thy5) Abs_inverse_thms'; + + (******************* freeness theorems for constructors *******************) + + val _ = message config "Proving freeness of constructors ..."; + + (* prove theorem Rep_i (Constr_j ...) = Inj_j ... *) + + fun prove_constr_rep_thm eqn = + let + val inj_thms = map fst newT_iso_inj_thms; + val rewrites = @{thm o_def} :: constr_defs @ (map (mk_meta_eq o #2) newT_iso_axms) + in Skip_Proof.prove_global thy5 [] [] eqn (fn _ => EVERY + [resolve_tac inj_thms 1, + rewrite_goals_tac rewrites, + rtac refl 3, + resolve_tac rep_intrs 2, + REPEAT (resolve_tac iso_elem_thms 1)]) + end; -fun gen_rep_datatype prep_term config after_qed alt_names raw_ts thy = - let - fun constr_of_term (Const (c, T)) = (c, T) - | constr_of_term t = - error ("Not a constant: " ^ Syntax.string_of_term_global thy t); - fun no_constr (c, T) = error ("Bad constructor: " - ^ Sign.extern_const thy c ^ "::" - ^ Syntax.string_of_typ_global thy T); - fun type_of_constr (cT as (_, T)) = + (*--------------------------------------------------------------*) + (* constr_rep_thms and rep_congs are used to prove distinctness *) + (* of constructors. *) + (*--------------------------------------------------------------*) + + val constr_rep_thms = map (map prove_constr_rep_thm) constr_rep_eqns; + + val dist_rewrites = map (fn (rep_thms, dist_lemma) => + dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0])) + (constr_rep_thms ~~ dist_lemmas); + + fun prove_distinct_thms dist_rewrites' (k, ts) = let - val frees = OldTerm.typ_tfrees T; - val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) T - handle TYPE _ => no_constr cT - val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else (); - val _ = if length frees <> length vs then no_constr cT else (); - in (tyco, (vs, cT)) end; + fun prove [] = [] + | prove (t :: ts) = + let + val dist_thm = Skip_Proof.prove_global thy5 [] [] t (fn _ => + EVERY [simp_tac (HOL_ss addsimps dist_rewrites') 1]) + in dist_thm :: Drule.standard (dist_thm RS not_sym) :: prove ts end; + in prove ts end; + + val distinct_thms = map2 (prove_distinct_thms) + dist_rewrites (Datatype_Prop.make_distincts descr sorts); + + (* prove injectivity of constructors *) - val raw_cs = AList.group (op =) (map (type_of_constr o constr_of_term o prep_term thy) raw_ts); - val _ = case map_filter (fn (tyco, _) => - if Symtab.defined (get_all thy) tyco then SOME tyco else NONE) raw_cs - of [] => () - | tycos => error ("Type(s) " ^ commas (map quote tycos) - ^ " already represented inductivly"); - val raw_vss = maps (map (map snd o fst) o snd) raw_cs; - val ms = case distinct (op =) (map length raw_vss) - of [n] => 0 upto n - 1 - | _ => error ("Different types in given constructors"); - fun inter_sort m = map (fn xs => nth xs m) raw_vss - |> Library.foldr1 (Sorts.inter_sort (Sign.classes_of thy)) - val sorts = map inter_sort ms; - val vs = Name.names Name.context Name.aT sorts; + fun prove_constr_inj_thm rep_thms t = + let val inj_thms = Scons_inject :: (map make_elim + (iso_inj_thms @ + [In0_inject, In1_inject, Leaf_inject, Inl_inject, Inr_inject, + Lim_inject, Suml_inject, Sumr_inject])) + in Skip_Proof.prove_global thy5 [] [] t (fn _ => EVERY + [rtac iffI 1, + REPEAT (etac conjE 2), hyp_subst_tac 2, rtac refl 2, + dresolve_tac rep_congs 1, dtac box_equals 1, + REPEAT (resolve_tac rep_thms 1), + REPEAT (eresolve_tac inj_thms 1), + REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (rtac ext 1), + REPEAT (eresolve_tac (make_elim fun_cong :: inj_thms) 1), + atac 1]))]) + end; + + val constr_inject = map (fn (ts, thms) => map (prove_constr_inj_thm thms) ts) + ((Datatype_Prop.make_injs descr sorts) ~~ constr_rep_thms); - fun norm_constr (raw_vs, (c, T)) = (c, map_atyps - (TFree o (the o AList.lookup (op =) (map fst raw_vs ~~ vs)) o fst o dest_TFree) T); + val ((constr_inject', distinct_thms'), thy6) = + thy5 + |> Sign.parent_path + |> store_thmss "inject" new_type_names constr_inject + ||>> store_thmss "distinct" new_type_names distinct_thms; + + (*************************** induction theorem ****************************) + + val _ = message config "Proving induction rule for datatypes ..."; - val cs = map (apsnd (map norm_constr)) raw_cs; - val dtyps_of_typ = map (dtyp_of_typ (map (rpair (map fst vs) o fst) cs)) - o fst o strip_type; - val dt_names = map fst cs; + val Rep_inverse_thms = (map (fn (_, iso, _) => iso RS subst) newT_iso_axms) @ + (map (fn r => r RS @{thm the_inv_f_f} RS subst) iso_inj_thms_unfolded); + val Rep_inverse_thms' = map (fn r => r RS @{thm the_inv_f_f}) iso_inj_thms_unfolded; + + fun mk_indrule_lemma ((i, _), T) (prems, concls) = + let + val Rep_t = Const (nth all_rep_names i, T --> Univ_elT) $ + mk_Free "x" T i; + + val Abs_t = if i < length newTs then + Const (Sign.intern_const thy6 + ("Abs_" ^ (nth new_type_names i)), Univ_elT --> T) + else Const (@{const_name the_inv_into}, + [HOLogic.mk_setT T, T --> Univ_elT, Univ_elT] ---> T) $ + HOLogic.mk_UNIV T $ Const (nth all_rep_names i, T --> Univ_elT) - fun mk_spec (i, (tyco, constr)) = (i, (tyco, - map (DtTFree o fst) vs, - (map o apsnd) dtyps_of_typ constr)) - val descr = map_index mk_spec cs; - val injs = DatatypeProp.make_injs [descr] vs; - val half_distincts = map snd (DatatypeProp.make_distincts [descr] vs); - val ind = DatatypeProp.make_ind [descr] vs; - val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts]; + in (prems @ [HOLogic.imp $ + (Const (nth rep_set_names i, UnivT') $ Rep_t) $ + (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))], + concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i]) + end; + + val (indrule_lemma_prems, indrule_lemma_concls) = + fold mk_indrule_lemma (descr' ~~ recTs) ([], []); + + val cert = cterm_of thy6; + + val indrule_lemma = Skip_Proof.prove_global thy6 [] [] + (Logic.mk_implies + (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems), + HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY + [REPEAT (etac conjE 1), + REPEAT (EVERY + [TRY (rtac conjI 1), resolve_tac Rep_inverse_thms 1, + etac mp 1, resolve_tac iso_elem_thms 1])]); - fun after_qed' raw_thms = - let - val [[[raw_induct]], raw_inject, half_distinct] = - unflat rules (map Drule.zero_var_indexes_list raw_thms); - (*FIXME somehow dubious*) - in - ProofContext.theory_result - (prove_rep_datatype config dt_names alt_names descr vs - raw_inject half_distinct raw_induct) - #-> after_qed - end; + val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma))); + val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else + map (Free o apfst fst o dest_Var) Ps; + val indrule_lemma' = cterm_instantiate (map cert Ps ~~ map cert frees) indrule_lemma; + + val dt_induct_prop = Datatype_Prop.make_ind descr sorts; + val dt_induct = Skip_Proof.prove_global thy6 [] + (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop) + (fn {prems, ...} => EVERY + [rtac indrule_lemma' 1, + (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, + EVERY (map (fn (prem, r) => (EVERY + [REPEAT (eresolve_tac Abs_inverse_thms 1), + simp_tac (HOL_basic_ss addsimps ((symmetric r)::Rep_inverse_thms')) 1, + DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)])) + (prems ~~ (constr_defs @ (map mk_meta_eq iso_char_thms))))]); + + val ([dt_induct'], thy7) = + thy6 + |> Sign.add_path big_name + |> PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] + ||> Sign.parent_path + ||> Theory.checkpoint; + in - thy - |> ProofContext.init - |> Proof.theorem_i NONE after_qed' ((map o map) (rpair []) (flat rules)) + ((constr_inject', distinct_thms', dt_induct'), thy7) end; -val rep_datatype = gen_rep_datatype Sign.cert_term; -val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global default_config (K I); - (** definitional introduction of datatypes **) @@ -472,7 +665,7 @@ [] => () | dups => error ("Duplicate datatypes: " ^ commas dups)); - fun prep_dt_spec ((tvs, tname, mx, constrs), tname') (dts', constr_syntax, sorts, i) = + fun prep_dt_spec (tvs, tname, mx, constrs) tname' (dts', constr_syntax, sorts, i) = let fun prep_constr (cname, cargs, mx') (constrs, constr_syntax', sorts') = let @@ -503,38 +696,27 @@ end; val (dts', constr_syntax, sorts', i) = - fold prep_dt_spec (dts ~~ new_type_names) ([], [], [], 0); - val sorts = sorts' @ (map (rpair (Sign.defaultS tmp_thy)) (subtract (op =) (map fst sorts') tyvars)); - val dt_info = get_all thy; + fold2 prep_dt_spec dts new_type_names ([], [], [], 0); + val sorts = sorts' @ map (rpair (Sign.defaultS tmp_thy)) (subtract (op =) (map fst sorts') tyvars); + val dt_info = Datatype_Data.get_all thy; val (descr, _) = unfold_datatypes tmp_thy dts' sorts dt_info dts' i; val _ = check_nonempty descr handle (exn as Datatype_Empty s) => if #strict config then error ("Nonemptiness check failed for datatype " ^ s) else raise exn; val _ = message config ("Constructing datatype(s) " ^ commas_quote new_type_names); - val ((inject, distinct, induct), thy') = thy |> - DatatypeRepProofs.representation_proofs config dt_info new_type_names descr sorts - types_syntax constr_syntax (mk_case_names_induct (flat descr)); + in - derive_datatype_props config dt_names (SOME new_type_names) descr sorts - induct inject distinct thy' + thy + |> representation_proofs config dt_info new_type_names descr sorts + types_syntax constr_syntax (Datatype_Data.mk_case_names_induct (flat descr)) + |-> (fn (inject, distinct, induct) => Datatype_Data.derive_datatype_props + config dt_names (SOME new_type_names) descr sorts + induct inject distinct) end; -val add_datatype = gen_add_datatype cert_typ; -val datatype_cmd = snd ooo gen_add_datatype read_typ default_config; - - - -(** package setup **) - -(* setup theory *) - -val setup = - trfun_setup #> - Datatype_Interpretation.init; - - -(* outer syntax *) +val add_datatype = gen_add_datatype Datatype_Data.cert_typ; +val datatype_cmd = snd ooo gen_add_datatype Datatype_Data.read_typ default_config; local @@ -560,12 +742,6 @@ OuterSyntax.command "datatype" "define inductive datatypes" K.thy_decl (parse_datatype_decls >> (fn (names, specs) => Toplevel.theory (datatype_cmd names specs))); -val _ = - OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_goal - (Scan.option (P.$$$ "(" |-- Scan.repeat1 P.name --| P.$$$ ")") -- Scan.repeat1 P.term - >> (fn (alt_names, ts) => Toplevel.print - o Toplevel.theory_to_proof (rep_datatype_cmd alt_names ts))); - end; end; diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Datatype/datatype_abs_proofs.ML --- a/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_abs_proofs.ML Fri Dec 04 15:20:24 2009 +0100 @@ -1,15 +1,9 @@ -(* Title: HOL/Tools/datatype_abs_proofs.ML +(* Title: HOL/Tools/Datatype/datatype_abs_proofs.ML Author: Stefan Berghofer, TU Muenchen -Proofs and defintions independent of concrete representation -of datatypes (i.e. requiring only abstract properties such as -injectivity / distinctness of constructors and induction) - - - case distinction (exhaustion) theorems - - characteristic equations for primrec combinators - - characteristic equations for case combinators - - equations for splitting "P (case ...)" expressions - - "nchotomy" and "case_cong" theorems for TFL +Datatype package: proofs and defintions independent of concrete +representation of datatypes (i.e. requiring only abstract +properties: injectivity / distinctness of constructors and induction). *) signature DATATYPE_ABS_PROOFS = @@ -38,10 +32,10 @@ thm list -> thm list list -> theory -> thm list * theory end; -structure DatatypeAbsProofs: DATATYPE_ABS_PROOFS = +structure Datatype_Abs_Proofs: DATATYPE_ABS_PROOFS = struct -open DatatypeAux; +open Datatype_Aux; (************************ case distinction theorems ***************************) @@ -51,7 +45,7 @@ val descr' = flat descr; val recTs = get_rec_types descr' sorts; - val newTs = Library.take (length (hd descr), recTs); + val newTs = take (length (hd descr)) recTs; val {maxidx, ...} = rep_thm induct; val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct))); @@ -62,7 +56,7 @@ Abs ("z", T', Const ("True", T''))) induct_Ps; val P = Abs ("z", T, HOLogic.imp $ HOLogic.mk_eq (Var (("a", maxidx+1), T), Bound 0) $ Var (("P", 0), HOLogic.boolT)) - val insts = Library.take (i, dummyPs) @ (P::(Library.drop (i + 1, dummyPs))); + val insts = take i dummyPs @ (P::(drop (i + 1) dummyPs)); val cert = cterm_of thy; val insts' = (map cert induct_Ps) ~~ (map cert insts); val induct' = refl RS ((nth @@ -78,7 +72,7 @@ end; val casedist_thms = map_index prove_casedist_thm - (newTs ~~ DatatypeProp.make_casedists descr sorts) + (newTs ~~ Datatype_Prop.make_casedists descr sorts) in thy |> store_thms_atts "exhaust" new_type_names (map single case_names_exhausts) casedist_thms @@ -98,7 +92,7 @@ val descr' = flat descr; val recTs = get_rec_types descr' sorts; val used = List.foldr OldTerm.add_typ_tfree_names [] recTs; - val newTs = Library.take (length (hd descr), recTs); + val newTs = take (length (hd descr)) recTs; val induct_Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of induct))); @@ -109,7 +103,7 @@ (1 upto (length descr')); val rec_set_names = map (Sign.full_bname thy0) rec_set_names'; - val (rec_result_Ts, reccomb_fn_Ts) = DatatypeProp.make_primrec_Ts descr sorts used; + val (rec_result_Ts, reccomb_fn_Ts) = Datatype_Prop.make_primrec_Ts descr sorts used; val rec_set_Ts = map (fn (T1, T2) => reccomb_fn_Ts @ [T1, T2] ---> HOLogic.boolT) (recTs ~~ rec_result_Ts); @@ -260,7 +254,7 @@ resolve_tac rec_unique_thms 1, resolve_tac rec_intrs 1, REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)])) - (DatatypeProp.make_primrecs new_type_names descr sorts thy2) + (Datatype_Prop.make_primrecs new_type_names descr sorts thy2) in thy2 @@ -283,7 +277,7 @@ val descr' = flat descr; val recTs = get_rec_types descr' sorts; val used = List.foldr OldTerm.add_typ_tfree_names [] recTs; - val newTs = Library.take (length (hd descr), recTs); + val newTs = take (length (hd descr)) recTs; val T' = TFree (Name.variant used "'t", HOLogic.typeS); fun mk_dummyT dt = binder_types (typ_of_dtyp descr' sorts dt) ---> T'; @@ -305,8 +299,8 @@ let val Ts = map (typ_of_dtyp descr' sorts) cargs; val Ts' = Ts @ map mk_dummyT (filter is_rec_type cargs); - val frees' = map (uncurry (mk_Free "x")) (Ts' ~~ (1 upto length Ts')); - val frees = Library.take (length cargs, frees'); + val frees' = map2 (mk_Free "x") Ts' (1 upto length Ts'); + val frees = take (length cargs) frees'; val free = mk_Free "f" (Ts ---> T') j in (free, list_abs_free (map dest_Free frees', @@ -314,14 +308,14 @@ end) (constrs ~~ (1 upto length constrs))); val caseT = (map (snd o dest_Free) fns1) @ [T] ---> T'; - val fns = flat (Library.take (i, case_dummy_fns)) @ - fns2 @ flat (Library.drop (i + 1, case_dummy_fns)); + val fns = flat (take i case_dummy_fns) @ + fns2 @ flat (drop (i + 1) case_dummy_fns); val reccomb = Const (recname, (map fastype_of fns) @ [T] ---> T'); val decl = ((Binding.name (Long_Name.base_name name), caseT), NoSyn); val def = (Binding.name (Long_Name.base_name name ^ "_def"), Logic.mk_equals (list_comb (Const (name, caseT), fns1), - list_comb (reccomb, (flat (Library.take (i, case_dummy_fns))) @ - fns2 @ (flat (Library.drop (i + 1, case_dummy_fns))) ))); + list_comb (reccomb, (flat (take i case_dummy_fns)) @ + fns2 @ (flat (drop (i + 1) case_dummy_fns))))); val ([def_thm], thy') = thy |> Sign.declare_const decl |> snd @@ -329,12 +323,12 @@ in (defs @ [def_thm], thy') end) (hd descr ~~ newTs ~~ case_names ~~ - Library.take (length newTs, reccomb_names)) ([], thy1) + take (length newTs) reccomb_names) ([], thy1) ||> Theory.checkpoint; val case_thms = map (map (fn t => Skip_Proof.prove_global thy2 [] [] t (fn _ => EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1]))) - (DatatypeProp.make_cases new_type_names descr sorts thy2) + (Datatype_Prop.make_cases new_type_names descr sorts thy2) in thy2 |> Context.theory_map ((fold o fold) Nitpick_Simps.add_thm case_thms) @@ -353,7 +347,7 @@ val descr' = flat descr; val recTs = get_rec_types descr' sorts; - val newTs = Library.take (length (hd descr), recTs); + val newTs = take (length (hd descr)) recTs; fun prove_split_thms ((((((t1, t2), inject), dist_rewrites'), exhaustion), case_thms'), T) = @@ -370,7 +364,7 @@ end; val split_thm_pairs = map prove_split_thms - ((DatatypeProp.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~ + ((Datatype_Prop.make_splits new_type_names descr sorts thy) ~~ constr_inject ~~ dist_rewrites ~~ casedist_thms ~~ case_thms ~~ newTs); val (split_thms, split_asm_thms) = ListPair.unzip split_thm_pairs @@ -388,7 +382,7 @@ Skip_Proof.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t) (fn {prems, ...} => EVERY [rtac ((hd prems) RS arg_cong) 1]) - val weak_case_congs = map prove_weak_case_cong (DatatypeProp.make_weak_case_congs + val weak_case_congs = map prove_weak_case_cong (Datatype_Prop.make_weak_case_congs new_type_names descr sorts thy) in thy |> store_thms "weak_case_cong" new_type_names weak_case_congs end; @@ -413,7 +407,7 @@ end; val nchotomys = - map prove_nchotomy (DatatypeProp.make_nchotomys descr sorts ~~ casedist_thms) + map prove_nchotomy (Datatype_Prop.make_nchotomys descr sorts ~~ casedist_thms) in thy |> store_thms "nchotomy" new_type_names nchotomys end; @@ -438,7 +432,7 @@ end) end; - val case_congs = map prove_case_cong (DatatypeProp.make_case_congs + val case_congs = map prove_case_cong (Datatype_Prop.make_case_congs new_type_names descr sorts thy ~~ nchotomys ~~ case_thms) in thy |> store_thms "case_cong" new_type_names case_congs end; diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Datatype/datatype_aux.ML --- a/src/HOL/Tools/Datatype/datatype_aux.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_aux.ML Fri Dec 04 15:20:24 2009 +0100 @@ -1,7 +1,7 @@ -(* Title: HOL/Tools/datatype_aux.ML +(* Title: HOL/Tools/Datatype/datatype_aux.ML Author: Stefan Berghofer, TU Muenchen -Auxiliary functions for defining datatypes. +Datatype package: auxiliary data structures and functions. *) signature DATATYPE_COMMON = @@ -79,9 +79,10 @@ val unfold_datatypes : theory -> descr -> (string * sort) list -> info Symtab.table -> descr -> int -> descr list * int + val find_shortest_path : descr -> int -> (string * int) option end; -structure DatatypeAux : DATATYPE_AUX = +structure Datatype_Aux : DATATYPE_AUX = struct (* datatype option flags *) @@ -365,4 +366,23 @@ in (descr' :: descrs, i') end; +(* find shortest path to constructor with no recursive arguments *) + +fun find_nonempty descr is i = + let + fun arg_nonempty (_, DtRec i) = if member (op =) is i + then NONE + else Option.map (Integer.add 1 o snd) (find_nonempty descr (i::is) i) + | arg_nonempty _ = SOME 0; + fun max_inf (SOME i) (SOME j) = SOME (Integer.max i j) + | max_inf _ _ = NONE; + fun max xs = fold max_inf xs (SOME 0); + val (_, _, constrs) = the (AList.lookup (op =) descr i); + val xs = sort (int_ord o pairself snd) + (map_filter (fn (s, dts) => Option.map (pair s) + (max (map (arg_nonempty o strip_dtyp) dts))) constrs) + in if null xs then NONE else SOME (hd xs) end; + +fun find_shortest_path descr i = find_nonempty descr [i] i; + end; diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Datatype/datatype_case.ML --- a/src/HOL/Tools/Datatype/datatype_case.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_case.ML Fri Dec 04 15:20:24 2009 +0100 @@ -1,30 +1,32 @@ -(* Title: HOL/Tools/datatype_case.ML +(* Title: HOL/Tools/Datatype/datatype_case.ML Author: Konrad Slind, Cambridge University Computer Laboratory Author: Stefan Berghofer, TU Muenchen -Nested case expressions on datatypes. +Datatype package: nested case expressions on datatypes. *) signature DATATYPE_CASE = sig - datatype config = Error | Warning | Quiet; - val make_case: (string * typ -> DatatypeAux.info option) -> + datatype config = Error | Warning | Quiet + type info = Datatype_Aux.info + val make_case: (string * typ -> info option) -> Proof.context -> config -> string list -> term -> (term * term) list -> term * (term * (int * bool)) list - val dest_case: (string -> DatatypeAux.info option) -> bool -> + val dest_case: (string -> info option) -> bool -> string list -> term -> (term * (term * term) list) option - val strip_case: (string -> DatatypeAux.info option) -> bool -> + val strip_case: (string -> info option) -> bool -> term -> (term * (term * term) list) option - val case_tr: bool -> (theory -> string * typ -> DatatypeAux.info option) - -> Proof.context -> term list -> term - val case_tr': (theory -> string -> DatatypeAux.info option) -> + val case_tr: bool -> (theory -> string * typ -> info option) -> + Proof.context -> term list -> term + val case_tr': (theory -> string -> info option) -> string -> Proof.context -> term list -> term end; -structure DatatypeCase : DATATYPE_CASE = +structure Datatype_Case : DATATYPE_CASE = struct datatype config = Error | Warning | Quiet; +type info = Datatype_Aux.info; exception CASE_ERROR of string * int; @@ -36,10 +38,10 @@ fun ty_info tab sT = case tab sT of - SOME ({descr, case_name, index, sorts, ...} : DatatypeAux.info) => + SOME ({descr, case_name, index, sorts, ...} : info) => let val (_, (tname, dts, constrs)) = nth descr index; - val mk_ty = DatatypeAux.typ_of_dtyp descr sorts; + val mk_ty = Datatype_Aux.typ_of_dtyp descr sorts; val T = Type (tname, map mk_ty dts) in SOME {case_name = case_name, @@ -84,7 +86,7 @@ val (_, Ty) = dest_Const c val Ts = binder_types Ty; val names = Name.variant_list used - (DatatypeProp.make_tnames (map Logic.unvarifyT Ts)); + (Datatype_Prop.make_tnames (map Logic.unvarifyT Ts)); val ty = body_type Ty; val ty_theta = ty_match ty colty handle Type.TYPE_MATCH => raise CASE_ERROR ("type mismatch", ~1) @@ -225,7 +227,6 @@ val nrows = maps (expand constructors used' pty) rows; val subproblems = partition ty_match ty_inst type_of used' constructors pty range_ty nrows; - val new_formals = map #new_formals subproblems val constructors' = map #constructor subproblems val news = map (fn {new_formals, group, ...} => {path = new_formals @ rstp, rows = group}) subproblems; diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Datatype/datatype_codegen.ML --- a/src/HOL/Tools/Datatype/datatype_codegen.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_codegen.ML Fri Dec 04 15:20:24 2009 +0100 @@ -1,4 +1,4 @@ -(* Title: HOL/Tools/datatype_codegen.ML +(* Title: HOL/Tools/Datatype/datatype_codegen.ML Author: Stefan Berghofer and Florian Haftmann, TU Muenchen Code generator facilities for inductive datatypes. @@ -6,45 +6,23 @@ signature DATATYPE_CODEGEN = sig - val find_shortest_path: Datatype.descr -> int -> (string * int) option val setup: theory -> theory end; -structure DatatypeCodegen : DATATYPE_CODEGEN = +structure Datatype_Codegen : DATATYPE_CODEGEN = struct -(** find shortest path to constructor with no recursive arguments **) - -fun find_nonempty (descr: Datatype.descr) is i = - let - val (_, _, constrs) = the (AList.lookup (op =) descr i); - fun arg_nonempty (_, DatatypeAux.DtRec i) = if member (op =) is i - then NONE - else Option.map (Integer.add 1 o snd) (find_nonempty descr (i::is) i) - | arg_nonempty _ = SOME 0; - fun max xs = Library.foldl - (fn (NONE, _) => NONE - | (SOME i, SOME j) => SOME (Int.max (i, j)) - | (_, NONE) => NONE) (SOME 0, xs); - val xs = sort (int_ord o pairself snd) - (map_filter (fn (s, dts) => Option.map (pair s) - (max (map (arg_nonempty o DatatypeAux.strip_dtyp) dts))) constrs) - in case xs of [] => NONE | x :: _ => SOME x end; - -fun find_shortest_path descr i = find_nonempty descr [i] i; - - (** SML code generator **) open Codegen; (* datatype definition *) -fun add_dt_defs thy defs dep module (descr: Datatype.descr) sorts gr = +fun add_dt_defs thy defs dep module descr sorts gr = let - val descr' = filter (can (map DatatypeAux.dest_DtTFree o #2 o snd)) descr; + val descr' = filter (can (map Datatype_Aux.dest_DtTFree o #2 o snd)) descr; val rtnames = map (#1 o snd) (filter (fn (_, (_, _, cs)) => - exists (exists DatatypeAux.is_rec_type o snd) cs) descr'); + exists (exists Datatype_Aux.is_rec_type o snd) cs) descr'); val (_, (tname, _, _)) :: _ = descr'; val node_id = tname ^ " (type)"; @@ -53,8 +31,8 @@ fun mk_dtdef prfx [] gr = ([], gr) | mk_dtdef prfx ((_, (tname, dts, cs))::xs) gr = let - val tvs = map DatatypeAux.dest_DtTFree dts; - val cs' = map (apsnd (map (DatatypeAux.typ_of_dtyp descr sorts))) cs; + val tvs = map Datatype_Aux.dest_DtTFree dts; + val cs' = map (apsnd (map (Datatype_Aux.typ_of_dtyp descr sorts))) cs; val ((_, type_id), gr') = mk_type_id module' tname gr; val (ps, gr'') = gr' |> fold_map (fn (cname, cargs) => @@ -84,8 +62,8 @@ fun mk_term_of_def gr prfx [] = [] | mk_term_of_def gr prfx ((_, (tname, dts, cs)) :: xs) = let - val cs' = map (apsnd (map (DatatypeAux.typ_of_dtyp descr sorts))) cs; - val dts' = map (DatatypeAux.typ_of_dtyp descr sorts) dts; + val cs' = map (apsnd (map (Datatype_Aux.typ_of_dtyp descr sorts))) cs; + val dts' = map (Datatype_Aux.typ_of_dtyp descr sorts) dts; val T = Type (tname, dts'); val rest = mk_term_of_def gr "and " xs; val (eqs, _) = fold_map (fn (cname, Ts) => fn prfx => @@ -107,12 +85,12 @@ fun mk_gen_of_def gr prfx [] = [] | mk_gen_of_def gr prfx ((i, (tname, dts, cs)) :: xs) = let - val tvs = map DatatypeAux.dest_DtTFree dts; - val Us = map (DatatypeAux.typ_of_dtyp descr sorts) dts; + val tvs = map Datatype_Aux.dest_DtTFree dts; + val Us = map (Datatype_Aux.typ_of_dtyp descr sorts) dts; val T = Type (tname, Us); val (cs1, cs2) = - List.partition (exists DatatypeAux.is_rec_type o snd) cs; - val SOME (cname, _) = find_shortest_path descr i; + List.partition (exists Datatype_Aux.is_rec_type o snd) cs; + val SOME (cname, _) = Datatype_Aux.find_shortest_path descr i; fun mk_delay p = Pretty.block [str "fn () =>", Pretty.brk 1, p]; @@ -122,14 +100,14 @@ fun mk_constr s b (cname, dts) = let val gs = map (fn dt => mk_app false (mk_gen gr module' false rtnames s - (DatatypeAux.typ_of_dtyp descr sorts dt)) - [str (if b andalso DatatypeAux.is_rec_type dt then "0" + (Datatype_Aux.typ_of_dtyp descr sorts dt)) + [str (if b andalso Datatype_Aux.is_rec_type dt then "0" else "j")]) dts; - val Ts = map (DatatypeAux.typ_of_dtyp descr sorts) dts; + val Ts = map (Datatype_Aux.typ_of_dtyp descr sorts) dts; val xs = map str - (DatatypeProp.indexify_names (replicate (length dts) "x")); + (Datatype_Prop.indexify_names (replicate (length dts) "x")); val ts = map str - (DatatypeProp.indexify_names (replicate (length dts) "t")); + (Datatype_Prop.indexify_names (replicate (length dts) "t")); val (_, id) = get_const_id gr cname in mk_let @@ -217,19 +195,19 @@ invoke_codegen thy defs dep module brack (eta_expand c ts (i+1)) gr else let - val ts1 = Library.take (i, ts); - val t :: ts2 = Library.drop (i, ts); + val ts1 = take i ts; + val t :: ts2 = drop i ts; val names = List.foldr OldTerm.add_term_names (map (fst o fst o dest_Var) (List.foldr OldTerm.add_term_vars [] ts1)) ts1; - val (Ts, dT) = split_last (Library.take (i+1, fst (strip_type T))); + val (Ts, dT) = split_last (take (i+1) (fst (strip_type T))); fun pcase [] [] [] gr = ([], gr) | pcase ((cname, cargs)::cs) (t::ts) (U::Us) gr = let val j = length cargs; val xs = Name.variant_list names (replicate j "x"); - val Us' = Library.take (j, fst (strip_type U)); - val frees = map Free (xs ~~ Us'); + val Us' = take j (fst (strip_type U)); + val frees = map2 (curry Free) xs Us'; val (cp, gr0) = invoke_codegen thy defs dep module false (list_comb (Const (cname, Us' ---> dT), frees)) gr; val t' = Envir.beta_norm (list_comb (t, frees)); @@ -274,12 +252,12 @@ fun datatype_codegen thy defs dep module brack t gr = (case strip_comb t of (c as Const (s, T), ts) => - (case Datatype.info_of_case thy s of + (case Datatype_Data.info_of_case thy s of SOME {index, descr, ...} => if is_some (get_assoc_code thy (s, T)) then NONE else SOME (pretty_case thy defs dep module brack (#3 (the (AList.lookup op = descr index))) c ts gr ) - | NONE => case (Datatype.info_of_constr thy (s, T), strip_type T) of + | NONE => case (Datatype_Data.info_of_constr thy (s, T), strip_type T) of (SOME {index, descr, ...}, (_, U as Type (tyname, _))) => if is_some (get_assoc_code thy (s, T)) then NONE else let @@ -294,7 +272,7 @@ | _ => NONE); fun datatype_tycodegen thy defs dep module brack (Type (s, Ts)) gr = - (case Datatype.get_info thy s of + (case Datatype_Data.get_info thy s of NONE => NONE | SOME {descr, sorts, ...} => if is_some (get_assoc_type thy s) then NONE else @@ -329,7 +307,7 @@ fun mk_case_cert thy tyco = let val raw_thms = - (#case_rewrites o Datatype.the_info thy) tyco; + (#case_rewrites o Datatype_Data.the_info thy) tyco; val thms as hd_thm :: _ = raw_thms |> Conjunction.intr_balanced |> Thm.unvarify @@ -362,9 +340,9 @@ fun mk_eq_eqns thy dtco = let - val (vs, cos) = Datatype.the_spec thy dtco; + val (vs, cos) = Datatype_Data.the_spec thy dtco; val { descr, index, inject = inject_thms, distinct = distinct_thms, ... } = - Datatype.the_info thy dtco; + Datatype_Data.the_info thy dtco; val ty = Type (dtco, map TFree vs); fun mk_eq (t1, t2) = Const (@{const_name eq_class.eq}, ty --> ty --> HOLogic.boolT) $ t1 $ t2; @@ -375,10 +353,10 @@ | _ => NONE) cos; fun prep_inject (trueprop $ (equiv $ (_ $ t1 $ t2) $ rhs)) = trueprop $ (equiv $ mk_eq (t1, t2) $ rhs); - val injects = map prep_inject (nth (DatatypeProp.make_injs [descr] vs) index); + val injects = map prep_inject (nth (Datatype_Prop.make_injs [descr] vs) index); fun prep_distinct (trueprop $ (not $ (_ $ t1 $ t2))) = [trueprop $ false_eq (t1, t2), trueprop $ false_eq (t2, t1)]; - val distincts = maps prep_distinct (snd (nth (DatatypeProp.make_distincts [descr] vs) index)); + val distincts = maps prep_distinct (snd (nth (Datatype_Prop.make_distincts [descr] vs) index)); val refl = HOLogic.mk_Trueprop (true_eq (Free ("x", ty), Free ("x", ty))); val simpset = Simplifier.context (ProofContext.init thy) (HOL_basic_ss addsimps (map Simpdata.mk_eq (@{thm eq} :: @{thm eq_True} :: inject_thms @ distinct_thms))); @@ -424,16 +402,16 @@ fun add_all_code config dtcos thy = let - val (vs :: _, coss) = (split_list o map (Datatype.the_spec thy)) dtcos; + val (vs :: _, coss) = (split_list o map (Datatype_Data.the_spec thy)) dtcos; val any_css = map2 (mk_constr_consts thy vs) dtcos coss; val css = if exists is_none any_css then [] else map_filter I any_css; - val case_rewrites = maps (#case_rewrites o Datatype.the_info thy) dtcos; + val case_rewrites = maps (#case_rewrites o Datatype_Data.the_info thy) dtcos; val certs = map (mk_case_cert thy) dtcos; in if null css then thy else thy - |> tap (fn _ => DatatypeAux.message config "Registering datatype for code generator ...") + |> tap (fn _ => Datatype_Aux.message config "Registering datatype for code generator ...") |> fold Code.add_datatype css |> fold_rev Code.add_default_eqn case_rewrites |> fold Code.add_case certs @@ -446,6 +424,6 @@ val setup = add_codegen "datatype" datatype_codegen #> add_tycodegen "datatype" datatype_tycodegen - #> Datatype.interpretation add_all_code + #> Datatype_Data.interpretation add_all_code end; diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Datatype/datatype_data.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Tools/Datatype/datatype_data.ML Fri Dec 04 15:20:24 2009 +0100 @@ -0,0 +1,471 @@ +(* Title: HOL/Tools/Datatype/datatype_data.ML + Author: Stefan Berghofer, TU Muenchen + +Datatype package: bookkeeping; interpretation of existing types as datatypes. +*) + +signature DATATYPE_DATA = +sig + include DATATYPE_COMMON + val derive_datatype_props : config -> string list -> string list option + -> descr list -> (string * sort) list -> thm -> thm list list -> thm list list + -> theory -> string list * theory + val rep_datatype : config -> (string list -> Proof.context -> Proof.context) + -> string list option -> term list -> theory -> Proof.state + val rep_datatype_cmd : string list option -> string list -> theory -> Proof.state + val get_info : theory -> string -> info option + val the_info : theory -> string -> info + val the_descr : theory -> string list + -> descr * (string * sort) list * string list + * string * (string list * string list) * (typ list * typ list) + val the_spec : theory -> string -> (string * sort) list * (string * typ list) list + val all_distincts : theory -> typ list -> thm list list + val get_constrs : theory -> string -> (string * typ) list option + val get_all : theory -> info Symtab.table + val info_of_constr : theory -> string * typ -> info option + val info_of_case : theory -> string -> info option + val interpretation : (config -> string list -> theory -> theory) -> theory -> theory + val make_case : Proof.context -> Datatype_Case.config -> string list -> term -> + (term * term) list -> term * (term * (int * bool)) list + val strip_case : Proof.context -> bool -> term -> (term * (term * term) list) option + val read_typ: theory -> string -> (string * sort) list -> typ * (string * sort) list + val cert_typ: theory -> typ -> (string * sort) list -> typ * (string * sort) list + val mk_case_names_induct: descr -> attribute + val setup: theory -> theory +end; + +structure Datatype_Data: DATATYPE_DATA = +struct + +open Datatype_Aux; + +(** theory data **) + +(* data management *) + +structure DatatypesData = Theory_Data +( + type T = + {types: info Symtab.table, + constrs: (string * info) list Symtab.table, + cases: info Symtab.table}; + + val empty = + {types = Symtab.empty, constrs = Symtab.empty, cases = Symtab.empty}; + val extend = I; + fun merge + ({types = types1, constrs = constrs1, cases = cases1}, + {types = types2, constrs = constrs2, cases = cases2}) : T = + {types = Symtab.merge (K true) (types1, types2), + constrs = Symtab.join (K (AList.merge (op =) (K true))) (constrs1, constrs2), + cases = Symtab.merge (K true) (cases1, cases2)}; +); + +val get_all = #types o DatatypesData.get; +val get_info = Symtab.lookup o get_all; + +fun the_info thy name = + (case get_info thy name of + SOME info => info + | NONE => error ("Unknown datatype " ^ quote name)); + +fun info_of_constr thy (c, T) = + let + val tab = Symtab.lookup_list ((#constrs o DatatypesData.get) thy) c; + val hint = case strip_type T of (_, Type (tyco, _)) => SOME tyco | _ => NONE; + val default = if null tab then NONE + else SOME (snd (Library.last_elem tab)) + (*conservative wrt. overloaded constructors*); + in case hint + of NONE => default + | SOME tyco => case AList.lookup (op =) tab tyco + of NONE => default (*permissive*) + | SOME info => SOME info + end; + +val info_of_case = Symtab.lookup o #cases o DatatypesData.get; + +fun register (dt_infos : (string * info) list) = + DatatypesData.map (fn {types, constrs, cases} => + {types = types |> fold Symtab.update dt_infos, + constrs = constrs |> fold (fn (constr, dtname_info) => + Symtab.map_default (constr, []) (cons dtname_info)) + (maps (fn (dtname, info as {descr, index, ...}) => + map (rpair (dtname, info) o fst) + (#3 (the (AList.lookup op = descr index)))) dt_infos), + cases = cases |> fold Symtab.update + (map (fn (_, info as {case_name, ...}) => (case_name, info)) dt_infos)}); + + +(* complex queries *) + +fun the_spec thy dtco = + let + val { descr, index, sorts = raw_sorts, ... } = the_info thy dtco; + val SOME (_, dtys, raw_cos) = AList.lookup (op =) descr index; + val sorts = map ((fn v => (v, (the o AList.lookup (op =) raw_sorts) v)) + o Datatype_Aux.dest_DtTFree) dtys; + val cos = map + (fn (co, tys) => (co, map (Datatype_Aux.typ_of_dtyp descr sorts) tys)) raw_cos; + in (sorts, cos) end; + +fun the_descr thy (raw_tycos as raw_tyco :: _) = + let + val info = the_info thy raw_tyco; + val descr = #descr info; + + val SOME (_, dtys, _) = AList.lookup (op =) descr (#index info); + val vs = map ((fn v => (v, (the o AList.lookup (op =) (#sorts info)) v)) + o dest_DtTFree) dtys; + + fun is_DtTFree (DtTFree _) = true + | is_DtTFree _ = false + val k = find_index (fn (_, (_, dTs, _)) => not (forall is_DtTFree dTs)) descr; + val protoTs as (dataTs, _) = chop k descr + |> (pairself o map) (fn (_, (tyco, dTs, _)) => (tyco, map (typ_of_dtyp descr vs) dTs)); + + val tycos = map fst dataTs; + val _ = if eq_set (op =) (tycos, raw_tycos) then () + else error ("Type constructors " ^ commas (map quote raw_tycos) + ^ " do not belong exhaustively to one mutual recursive datatype"); + + val (Ts, Us) = (pairself o map) Type protoTs; + + val names = map Long_Name.base_name (the_default tycos (#alt_names info)); + val (auxnames, _) = Name.make_context names + |> fold_map (yield_singleton Name.variants o name_of_typ) Us; + val prefix = space_implode "_" names; + + in (descr, vs, tycos, prefix, (names, auxnames), (Ts, Us)) end; + +fun all_distincts thy Ts = + let + fun add_tycos (Type (tyco, Ts)) = insert (op =) tyco #> fold add_tycos Ts + | add_tycos _ = I; + val tycos = fold add_tycos Ts []; + in map_filter (Option.map #distinct o get_info thy) tycos end; + +fun get_constrs thy dtco = + case try (the_spec thy) dtco + of SOME (sorts, cos) => + let + fun subst (v, sort) = TVar ((v, 0), sort); + fun subst_ty (TFree v) = subst v + | subst_ty ty = ty; + val dty = Type (dtco, map subst sorts); + fun mk_co (co, tys) = (co, map (Term.map_atyps subst_ty) tys ---> dty); + in SOME (map mk_co cos) end + | NONE => NONE; + + + +(** various auxiliary **) + +(* prepare datatype specifications *) + +fun read_typ thy str sorts = + let + val ctxt = ProofContext.init thy + |> fold (Variable.declare_typ o TFree) sorts; + val T = Syntax.read_typ ctxt str; + in (T, Term.add_tfreesT T sorts) end; + +fun cert_typ sign raw_T sorts = + let + val T = Type.no_tvars (Sign.certify_typ sign raw_T) + handle TYPE (msg, _, _) => error msg; + val sorts' = Term.add_tfreesT T sorts; + val _ = + case duplicates (op =) (map fst sorts') of + [] => () + | dups => error ("Inconsistent sort constraints for " ^ commas dups) + in (T, sorts') end; + + +(* case names *) + +local + +fun dt_recs (DtTFree _) = [] + | dt_recs (DtType (_, dts)) = maps dt_recs dts + | dt_recs (DtRec i) = [i]; + +fun dt_cases (descr: descr) (_, args, constrs) = + let + fun the_bname i = Long_Name.base_name (#1 (the (AList.lookup (op =) descr i))); + val bnames = map the_bname (distinct (op =) (maps dt_recs args)); + in map (fn (c, _) => space_implode "_" (Long_Name.base_name c :: bnames)) constrs end; + +fun induct_cases descr = + Datatype_Prop.indexify_names (maps (dt_cases descr) (map #2 descr)); + +fun exhaust_cases descr i = dt_cases descr (the (AList.lookup (op =) descr i)); + +in + +fun mk_case_names_induct descr = Rule_Cases.case_names (induct_cases descr); + +fun mk_case_names_exhausts descr new = + map (Rule_Cases.case_names o exhaust_cases descr o #1) + (filter (fn ((_, (name, _, _))) => member (op =) new name) descr); + +end; + + +(* translation rules for case *) + +fun make_case ctxt = Datatype_Case.make_case + (info_of_constr (ProofContext.theory_of ctxt)) ctxt; + +fun strip_case ctxt = Datatype_Case.strip_case + (info_of_case (ProofContext.theory_of ctxt)); + +fun add_case_tr' case_names thy = + Sign.add_advanced_trfuns ([], [], + map (fn case_name => + let val case_name' = Sign.const_syntax_name thy case_name + in (case_name', Datatype_Case.case_tr' info_of_case case_name') + end) case_names, []) thy; + +val trfun_setup = + Sign.add_advanced_trfuns ([], + [("_case_syntax", Datatype_Case.case_tr true info_of_constr)], + [], []); + + + +(** document antiquotation **) + +val _ = ThyOutput.antiquotation "datatype" Args.tyname + (fn {source = src, context = ctxt, ...} => fn dtco => + let + val thy = ProofContext.theory_of ctxt; + val (vs, cos) = the_spec thy dtco; + val ty = Type (dtco, map TFree vs); + fun pretty_typ_bracket (ty as Type (_, _ :: _)) = + Pretty.enclose "(" ")" [Syntax.pretty_typ ctxt ty] + | pretty_typ_bracket ty = + Syntax.pretty_typ ctxt ty; + fun pretty_constr (co, tys) = + (Pretty.block o Pretty.breaks) + (Syntax.pretty_term ctxt (Const (co, tys ---> ty)) :: + map pretty_typ_bracket tys); + val pretty_datatype = + Pretty.block + (Pretty.command "datatype" :: Pretty.brk 1 :: + Syntax.pretty_typ ctxt ty :: + Pretty.str " =" :: Pretty.brk 1 :: + flat (separate [Pretty.brk 1, Pretty.str "| "] + (map (single o pretty_constr) cos))); + in ThyOutput.output (ThyOutput.maybe_pretty_source (K pretty_datatype) src [()]) end); + + + +(** abstract theory extensions relative to a datatype characterisation **) + +structure Datatype_Interpretation = Interpretation + (type T = config * string list val eq: T * T -> bool = eq_snd op =); +fun interpretation f = Datatype_Interpretation.interpretation (uncurry f); + +fun make_dt_info alt_names descr sorts induct inducts rec_names rec_rewrites + (index, (((((((((((_, (tname, _, _))), inject), distinct), + exhaust), nchotomy), case_name), case_rewrites), case_cong), weak_case_cong), + (split, split_asm))) = + (tname, + {index = index, + alt_names = alt_names, + descr = descr, + sorts = sorts, + inject = inject, + distinct = distinct, + induct = induct, + inducts = inducts, + exhaust = exhaust, + nchotomy = nchotomy, + rec_names = rec_names, + rec_rewrites = rec_rewrites, + case_name = case_name, + case_rewrites = case_rewrites, + case_cong = case_cong, + weak_case_cong = weak_case_cong, + split = split, + split_asm = split_asm}); + +fun derive_datatype_props config dt_names alt_names descr sorts + induct inject distinct thy1 = + let + val thy2 = thy1 |> Theory.checkpoint; + val flat_descr = flat descr; + val new_type_names = map Long_Name.base_name (the_default dt_names alt_names); + val _ = message config ("Deriving properties for datatype(s) " ^ commas_quote new_type_names); + + val (exhaust, thy3) = Datatype_Abs_Proofs.prove_casedist_thms config new_type_names + descr sorts induct (mk_case_names_exhausts flat_descr dt_names) thy2; + val (nchotomys, thy4) = Datatype_Abs_Proofs.prove_nchotomys config new_type_names + descr sorts exhaust thy3; + val ((rec_names, rec_rewrites), thy5) = Datatype_Abs_Proofs.prove_primrec_thms + config new_type_names descr sorts (#inject o the o Symtab.lookup (get_all thy4)) + inject (distinct, all_distincts thy2 (get_rec_types flat_descr sorts)) + induct thy4; + val ((case_rewrites, case_names), thy6) = Datatype_Abs_Proofs.prove_case_thms + config new_type_names descr sorts rec_names rec_rewrites thy5; + val (case_congs, thy7) = Datatype_Abs_Proofs.prove_case_congs new_type_names + descr sorts nchotomys case_rewrites thy6; + val (weak_case_congs, thy8) = Datatype_Abs_Proofs.prove_weak_case_congs new_type_names + descr sorts thy7; + val (splits, thy9) = Datatype_Abs_Proofs.prove_split_thms + config new_type_names descr sorts inject distinct exhaust case_rewrites thy8; + + val inducts = Project_Rule.projections (ProofContext.init thy2) induct; + val dt_infos = map_index + (make_dt_info alt_names flat_descr sorts induct inducts rec_names rec_rewrites) + (hd descr ~~ inject ~~ distinct ~~ exhaust ~~ nchotomys ~~ + case_names ~~ case_rewrites ~~ case_congs ~~ weak_case_congs ~~ splits); + val dt_names = map fst dt_infos; + val prfx = Binding.qualify true (space_implode "_" new_type_names); + val simps = flat (inject @ distinct @ case_rewrites) @ rec_rewrites; + val named_rules = flat (map_index (fn (index, tname) => + [((Binding.empty, [nth inducts index]), [Induct.induct_type tname]), + ((Binding.empty, [nth exhaust index]), [Induct.cases_type tname])]) dt_names); + val unnamed_rules = map (fn induct => + ((Binding.empty, [induct]), [Rule_Cases.inner_rule, Induct.induct_type ""])) + (drop (length dt_names) inducts); + in + thy9 + |> PureThy.add_thmss ([((prfx (Binding.name "simps"), simps), []), + ((prfx (Binding.name "inducts"), inducts), []), + ((prfx (Binding.name "splits"), maps (fn (x, y) => [x, y]) splits), []), + ((Binding.empty, flat case_rewrites @ flat distinct @ rec_rewrites), + [Simplifier.simp_add]), + ((Binding.empty, rec_rewrites), [Code.add_default_eqn_attribute]), + ((Binding.empty, flat inject), [iff_add]), + ((Binding.empty, map (fn th => th RS notE) (flat distinct)), + [Classical.safe_elim NONE]), + ((Binding.empty, weak_case_congs), [Simplifier.attrib (op addcongs)])] + @ named_rules @ unnamed_rules) + |> snd + |> add_case_tr' case_names + |> register dt_infos + |> Datatype_Interpretation.data (config, dt_names) + |> pair dt_names + end; + + + +(** declare existing type as datatype **) + +fun prove_rep_datatype config dt_names alt_names descr sorts + raw_inject half_distinct raw_induct thy1 = + let + val raw_distinct = (map o maps) (fn thm => [thm, thm RS not_sym]) half_distinct; + val new_type_names = map Long_Name.base_name (the_default dt_names alt_names); + val (((inject, distinct), [induct]), thy2) = + thy1 + |> store_thmss "inject" new_type_names raw_inject + ||>> store_thmss "distinct" new_type_names raw_distinct + ||> Sign.add_path (space_implode "_" new_type_names) + ||>> PureThy.add_thms [((Binding.name "induct", raw_induct), [mk_case_names_induct descr])] + ||> Sign.restore_naming thy1; + in + thy2 + |> derive_datatype_props config dt_names alt_names [descr] sorts + induct inject distinct + end; + +fun gen_rep_datatype prep_term config after_qed alt_names raw_ts thy = + let + fun constr_of_term (Const (c, T)) = (c, T) + | constr_of_term t = + error ("Not a constant: " ^ Syntax.string_of_term_global thy t); + fun no_constr (c, T) = error ("Bad constructor: " + ^ Sign.extern_const thy c ^ "::" + ^ Syntax.string_of_typ_global thy T); + fun type_of_constr (cT as (_, T)) = + let + val frees = OldTerm.typ_tfrees T; + val (tyco, vs) = ((apsnd o map) (dest_TFree) o dest_Type o snd o strip_type) T + handle TYPE _ => no_constr cT + val _ = if has_duplicates (eq_fst (op =)) vs then no_constr cT else (); + val _ = if length frees <> length vs then no_constr cT else (); + in (tyco, (vs, cT)) end; + + val raw_cs = AList.group (op =) (map (type_of_constr o constr_of_term o prep_term thy) raw_ts); + val _ = case map_filter (fn (tyco, _) => + if Symtab.defined (get_all thy) tyco then SOME tyco else NONE) raw_cs + of [] => () + | tycos => error ("Type(s) " ^ commas (map quote tycos) + ^ " already represented inductivly"); + val raw_vss = maps (map (map snd o fst) o snd) raw_cs; + val ms = case distinct (op =) (map length raw_vss) + of [n] => 0 upto n - 1 + | _ => error ("Different types in given constructors"); + fun inter_sort m = map (fn xs => nth xs m) raw_vss + |> Library.foldr1 (Sorts.inter_sort (Sign.classes_of thy)) + val sorts = map inter_sort ms; + val vs = Name.names Name.context Name.aT sorts; + + fun norm_constr (raw_vs, (c, T)) = (c, map_atyps + (TFree o (the o AList.lookup (op =) (map fst raw_vs ~~ vs)) o fst o dest_TFree) T); + + val cs = map (apsnd (map norm_constr)) raw_cs; + val dtyps_of_typ = map (dtyp_of_typ (map (rpair (map fst vs) o fst) cs)) + o fst o strip_type; + val dt_names = map fst cs; + + fun mk_spec (i, (tyco, constr)) = (i, (tyco, + map (DtTFree o fst) vs, + (map o apsnd) dtyps_of_typ constr)) + val descr = map_index mk_spec cs; + val injs = Datatype_Prop.make_injs [descr] vs; + val half_distincts = map snd (Datatype_Prop.make_distincts [descr] vs); + val ind = Datatype_Prop.make_ind [descr] vs; + val rules = (map o map o map) Logic.close_form [[[ind]], injs, half_distincts]; + + fun after_qed' raw_thms = + let + val [[[raw_induct]], raw_inject, half_distinct] = + unflat rules (map Drule.zero_var_indexes_list raw_thms); + (*FIXME somehow dubious*) + in + ProofContext.theory_result + (prove_rep_datatype config dt_names alt_names descr vs + raw_inject half_distinct raw_induct) + #-> after_qed + end; + in + thy + |> ProofContext.init + |> Proof.theorem_i NONE after_qed' ((map o map) (rpair []) (flat rules)) + end; + +val rep_datatype = gen_rep_datatype Sign.cert_term; +val rep_datatype_cmd = gen_rep_datatype Syntax.read_term_global default_config (K I); + + + +(** package setup **) + +(* setup theory *) + +val setup = + trfun_setup #> + Datatype_Interpretation.init; + + +(* outer syntax *) + +local + +structure P = OuterParse and K = OuterKeyword + +in + +val _ = + OuterSyntax.command "rep_datatype" "represent existing types inductively" K.thy_goal + (Scan.option (P.$$$ "(" |-- Scan.repeat1 P.name --| P.$$$ ")") -- Scan.repeat1 P.term + >> (fn (alt_names, ts) => Toplevel.print + o Toplevel.theory_to_proof (rep_datatype_cmd alt_names ts))); + +end; + +end; diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Datatype/datatype_prop.ML --- a/src/HOL/Tools/Datatype/datatype_prop.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_prop.ML Fri Dec 04 15:20:24 2009 +0100 @@ -1,38 +1,39 @@ -(* Title: HOL/Tools/datatype_prop.ML +(* Title: HOL/Tools/Datatype/datatype_prop.ML Author: Stefan Berghofer, TU Muenchen -Characteristic properties of datatypes. +Datatype package: characteristic properties of datatypes. *) signature DATATYPE_PROP = sig + include DATATYPE_COMMON val indexify_names: string list -> string list val make_tnames: typ list -> string list - val make_injs : DatatypeAux.descr list -> (string * sort) list -> term list list - val make_distincts : DatatypeAux.descr list -> + val make_injs : descr list -> (string * sort) list -> term list list + val make_distincts : descr list -> (string * sort) list -> (int * term list) list (*no symmetric inequalities*) - val make_ind : DatatypeAux.descr list -> (string * sort) list -> term - val make_casedists : DatatypeAux.descr list -> (string * sort) list -> term list - val make_primrec_Ts : DatatypeAux.descr list -> (string * sort) list -> + val make_ind : descr list -> (string * sort) list -> term + val make_casedists : descr list -> (string * sort) list -> term list + val make_primrec_Ts : descr list -> (string * sort) list -> string list -> typ list * typ list - val make_primrecs : string list -> DatatypeAux.descr list -> + val make_primrecs : string list -> descr list -> (string * sort) list -> theory -> term list - val make_cases : string list -> DatatypeAux.descr list -> + val make_cases : string list -> descr list -> (string * sort) list -> theory -> term list list - val make_splits : string list -> DatatypeAux.descr list -> + val make_splits : string list -> descr list -> (string * sort) list -> theory -> (term * term) list - val make_weak_case_congs : string list -> DatatypeAux.descr list -> + val make_weak_case_congs : string list -> descr list -> (string * sort) list -> theory -> term list - val make_case_congs : string list -> DatatypeAux.descr list -> + val make_case_congs : string list -> descr list -> (string * sort) list -> theory -> term list - val make_nchotomys : DatatypeAux.descr list -> + val make_nchotomys : descr list -> (string * sort) list -> term list end; -structure DatatypeProp : DATATYPE_PROP = +structure Datatype_Prop : DATATYPE_PROP = struct -open DatatypeAux; +open Datatype_Aux; fun indexify_names names = let @@ -72,7 +73,7 @@ end; in map2 (fn d => fn T => fold_rev (make_inj T) (#3 (snd d)) []) - (hd descr) (Library.take (length (hd descr), get_rec_types descr' sorts)) + (hd descr) (take (length (hd descr)) (get_rec_types descr' sorts)) end; @@ -82,7 +83,7 @@ let val descr' = flat descr; val recTs = get_rec_types descr' sorts; - val newTs = Library.take (length (hd descr), recTs); + val newTs = take (length (hd descr)) recTs; fun prep_constr (cname, cargs) = (cname, map (typ_of_dtyp descr' sorts) cargs); @@ -168,14 +169,12 @@ HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT)))) end; - fun make_casedist ((_, (_, _, constrs)), T) = + fun make_casedist ((_, (_, _, constrs))) T = let val prems = map (make_casedist_prem T) constrs in Logic.list_implies (prems, HOLogic.mk_Trueprop (Free ("P", HOLogic.boolT))) end - in map make_casedist - ((hd descr) ~~ Library.take (length (hd descr), get_rec_types descr' sorts)) - end; + in map2 make_casedist (hd descr) (take (length (hd descr)) (get_rec_types descr' sorts)) end; (*************** characteristic equations for primrec combinator **************) @@ -257,7 +256,7 @@ val descr' = flat descr; val recTs = get_rec_types descr' sorts; val used = List.foldr OldTerm.add_typ_tfree_names [] recTs; - val newTs = Library.take (length (hd descr), recTs); + val newTs = take (length (hd descr)) recTs; val T' = TFree (Name.variant used "'t", HOLogic.typeS); val case_fn_Ts = map (fn (i, (_, _, constrs)) => @@ -280,7 +279,7 @@ let val descr' = flat descr; val recTs = get_rec_types descr' sorts; - val newTs = Library.take (length (hd descr), recTs); + val newTs = take (length (hd descr)) recTs; fun make_case T comb_t ((cname, cargs), f) = let @@ -304,7 +303,7 @@ val descr' = flat descr; val recTs = get_rec_types descr' sorts; val used' = List.foldr OldTerm.add_typ_tfree_names [] recTs; - val newTs = Library.take (length (hd descr), recTs); + val newTs = take (length (hd descr)) recTs; val T' = TFree (Name.variant used' "'t", HOLogic.typeS); val P = Free ("P", T' --> HOLogic.boolT); @@ -415,7 +414,7 @@ let val descr' = flat descr; val recTs = get_rec_types descr' sorts; - val newTs = Library.take (length (hd descr), recTs); + val newTs = take (length (hd descr)) recTs; fun mk_eqn T (cname, cargs) = let diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Datatype/datatype_realizer.ML --- a/src/HOL/Tools/Datatype/datatype_realizer.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/Datatype/datatype_realizer.ML Fri Dec 04 15:20:24 2009 +0100 @@ -1,8 +1,8 @@ -(* Title: HOL/Tools/datatype_realizer.ML +(* Title: HOL/Tools/Datatype/datatype_realizer.ML Author: Stefan Berghofer, TU Muenchen -Porgram extraction from proofs involving datatypes: -Realizers for induction and case analysis +Program extraction from proofs involving datatypes: +realizers for induction and case analysis. *) signature DATATYPE_REALIZER = @@ -11,10 +11,10 @@ val setup: theory -> theory end; -structure DatatypeRealizer : DATATYPE_REALIZER = +structure Datatype_Realizer : DATATYPE_REALIZER = struct -open DatatypeAux; +open Datatype_Aux; fun subsets i j = if i <= j then let val is = subsets (i+1) j @@ -28,16 +28,11 @@ fun prf_of thm = Reconstruct.reconstruct_proof (Thm.theory_of_thm thm) (Thm.prop_of thm) (Thm.proof_of thm); -fun prf_subst_vars inst = - Proofterm.map_proof_terms (subst_vars ([], inst)) I; - fun is_unit t = snd (strip_type (fastype_of t)) = HOLogic.unitT; fun tname_of (Type (s, _)) = s | tname_of _ = ""; -fun mk_realizes T = Const ("realizes", T --> HOLogic.boolT --> HOLogic.boolT); - fun make_ind sorts ({descr, rec_names, rec_rewrites, induct, ...} : info) is thy = let val recTs = get_rec_types descr sorts; @@ -60,7 +55,7 @@ (fn ((i, (_, _, constrs)), T) => fold_map (fn (cname, cargs) => fn j => let val Ts = map (typ_of_dtyp descr sorts) cargs; - val tnames = Name.variant_list pnames (DatatypeProp.make_tnames Ts); + val tnames = Name.variant_list pnames (Datatype_Prop.make_tnames Ts); val recs = filter (is_rec_type o fst o fst) (cargs ~~ tnames ~~ Ts); val frees = tnames ~~ Ts; @@ -97,7 +92,7 @@ if (j: int) = i then HOLogic.mk_fst t else mk_proj j is (HOLogic.mk_snd t); - val tnames = DatatypeProp.make_tnames recTs; + val tnames = Datatype_Prop.make_tnames recTs; val fTs = map fastype_of rec_fns; val ps = map (fn ((((i, _), T), U), s) => Abs ("x", T, make_pred i U T (list_comb (Const (s, fTs ---> T --> U), rec_fns) $ Bound 0) (Bound 0))) @@ -132,7 +127,7 @@ (Binding.qualified_name (space_implode "_" (ind_name :: vs @ ["correctness"])), thm) ||> Sign.restore_naming thy; - val ivs = rev (Term.add_vars (Logic.varify (DatatypeProp.make_ind [descr] sorts)) []); + val ivs = rev (Term.add_vars (Logic.varify (Datatype_Prop.make_ind [descr] sorts)) []); val rvs = rev (Thm.fold_terms Term.add_vars thm' []); val ivs1 = map Var (filter_out (fn (_, T) => tname_of (body_type T) mem ["set", "bool"]) ivs); @@ -169,7 +164,7 @@ fun make_casedist_prem T (cname, cargs) = let val Ts = map (typ_of_dtyp descr sorts) cargs; - val frees = Name.variant_list ["P", "y"] (DatatypeProp.make_tnames Ts) ~~ Ts; + val frees = Name.variant_list ["P", "y"] (Datatype_Prop.make_tnames Ts) ~~ Ts; val free_ts = map Free frees; val r = Free ("r" ^ Long_Name.base_name cname, Ts ---> rT) in (r, list_all_free (frees, Logic.mk_implies (HOLogic.mk_Trueprop diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Datatype/datatype_rep_proofs.ML --- a/src/HOL/Tools/Datatype/datatype_rep_proofs.ML Fri Dec 04 11:44:57 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,633 +0,0 @@ -(* Title: HOL/Tools/datatype_rep_proofs.ML - Author: Stefan Berghofer, TU Muenchen - -Definitional introduction of datatypes -Proof of characteristic theorems: - - - injectivity of constructors - - distinctness of constructors - - induction theorem -*) - -signature DATATYPE_REP_PROOFS = -sig - include DATATYPE_COMMON - val representation_proofs : config -> info Symtab.table -> - string list -> descr list -> (string * sort) list -> - (binding * mixfix) list -> (binding * mixfix) list list -> attribute - -> theory -> (thm list list * thm list list * thm) * theory -end; - -structure DatatypeRepProofs : DATATYPE_REP_PROOFS = -struct - -open DatatypeAux; - -val (_ $ (_ $ (_ $ (distinct_f $ _) $ _))) = hd (prems_of distinct_lemma); - -val collect_simp = rewrite_rule [mk_meta_eq mem_Collect_eq]; - - -(** theory context references **) - -fun exh_thm_of (dt_info : info Symtab.table) tname = - #exhaust (the (Symtab.lookup dt_info tname)); - -(******************************************************************************) - -fun representation_proofs (config : config) (dt_info : info Symtab.table) - new_type_names descr sorts types_syntax constr_syntax case_names_induct thy = - let - val Datatype_thy = ThyInfo.the_theory "Datatype" thy; - val node_name = "Datatype.node"; - val In0_name = "Datatype.In0"; - val In1_name = "Datatype.In1"; - val Scons_name = "Datatype.Scons"; - val Leaf_name = "Datatype.Leaf"; - val Numb_name = "Datatype.Numb"; - val Lim_name = "Datatype.Lim"; - val Suml_name = "Datatype.Suml"; - val Sumr_name = "Datatype.Sumr"; - - val [In0_inject, In1_inject, Scons_inject, Leaf_inject, - In0_eq, In1_eq, In0_not_In1, In1_not_In0, - Lim_inject, Suml_inject, Sumr_inject] = map (PureThy.get_thm Datatype_thy) - ["In0_inject", "In1_inject", "Scons_inject", "Leaf_inject", - "In0_eq", "In1_eq", "In0_not_In1", "In1_not_In0", - "Lim_inject", "Suml_inject", "Sumr_inject"]; - - val descr' = flat descr; - - val big_name = space_implode "_" new_type_names; - val thy1 = Sign.add_path big_name thy; - val big_rec_name = big_name ^ "_rep_set"; - val rep_set_names' = - (if length descr' = 1 then [big_rec_name] else - (map ((curry (op ^) (big_rec_name ^ "_")) o string_of_int) - (1 upto (length descr')))); - val rep_set_names = map (Sign.full_bname thy1) rep_set_names'; - - val tyvars = map (fn (_, (_, Ts, _)) => map dest_DtTFree Ts) (hd descr); - val leafTs' = get_nonrec_types descr' sorts; - val branchTs = get_branching_types descr' sorts; - val branchT = if null branchTs then HOLogic.unitT - else Balanced_Tree.make (fn (T, U) => Type ("+", [T, U])) branchTs; - val arities = remove (op =) 0 (get_arities descr'); - val unneeded_vars = - subtract (op =) (List.foldr OldTerm.add_typ_tfree_names [] (leafTs' @ branchTs)) (hd tyvars); - val leafTs = leafTs' @ map (fn n => TFree (n, (the o AList.lookup (op =) sorts) n)) unneeded_vars; - val recTs = get_rec_types descr' sorts; - val newTs = Library.take (length (hd descr), recTs); - val oldTs = Library.drop (length (hd descr), recTs); - val sumT = if null leafTs then HOLogic.unitT - else Balanced_Tree.make (fn (T, U) => Type ("+", [T, U])) leafTs; - val Univ_elT = HOLogic.mk_setT (Type (node_name, [sumT, branchT])); - val UnivT = HOLogic.mk_setT Univ_elT; - val UnivT' = Univ_elT --> HOLogic.boolT; - val Collect = Const ("Collect", UnivT' --> UnivT); - - val In0 = Const (In0_name, Univ_elT --> Univ_elT); - val In1 = Const (In1_name, Univ_elT --> Univ_elT); - val Leaf = Const (Leaf_name, sumT --> Univ_elT); - val Lim = Const (Lim_name, (branchT --> Univ_elT) --> Univ_elT); - - (* make injections needed for embedding types in leaves *) - - fun mk_inj T' x = - let - fun mk_inj' T n i = - if n = 1 then x else - let val n2 = n div 2; - val Type (_, [T1, T2]) = T - in - if i <= n2 then - Const ("Sum_Type.Inl", T1 --> T) $ (mk_inj' T1 n2 i) - else - Const ("Sum_Type.Inr", T2 --> T) $ (mk_inj' T2 (n - n2) (i - n2)) - end - in mk_inj' sumT (length leafTs) (1 + find_index (fn T'' => T'' = T') leafTs) - end; - - (* make injections for constructors *) - - fun mk_univ_inj ts = Balanced_Tree.access - {left = fn t => In0 $ t, - right = fn t => In1 $ t, - init = - if ts = [] then Const (@{const_name undefined}, Univ_elT) - else foldr1 (HOLogic.mk_binop Scons_name) ts}; - - (* function spaces *) - - fun mk_fun_inj T' x = - let - fun mk_inj T n i = - if n = 1 then x else - let - val n2 = n div 2; - val Type (_, [T1, T2]) = T; - fun mkT U = (U --> Univ_elT) --> T --> Univ_elT - in - if i <= n2 then Const (Suml_name, mkT T1) $ mk_inj T1 n2 i - else Const (Sumr_name, mkT T2) $ mk_inj T2 (n - n2) (i - n2) - end - in mk_inj branchT (length branchTs) (1 + find_index (fn T'' => T'' = T') branchTs) - end; - - fun mk_lim t Ts = fold_rev (fn T => fn t => Lim $ mk_fun_inj T (Abs ("x", T, t))) Ts t; - - (************** generate introduction rules for representing set **********) - - val _ = message config "Constructing representing sets ..."; - - (* make introduction rule for a single constructor *) - - fun make_intr s n (i, (_, cargs)) = - let - fun mk_prem dt (j, prems, ts) = - (case strip_dtyp dt of - (dts, DtRec k) => - let - val Ts = map (typ_of_dtyp descr' sorts) dts; - val free_t = - app_bnds (mk_Free "x" (Ts ---> Univ_elT) j) (length Ts) - in (j + 1, list_all (map (pair "x") Ts, - HOLogic.mk_Trueprop - (Free (nth rep_set_names' k, UnivT') $ free_t)) :: prems, - mk_lim free_t Ts :: ts) - end - | _ => - let val T = typ_of_dtyp descr' sorts dt - in (j + 1, prems, (Leaf $ mk_inj T (mk_Free "x" T j))::ts) - end); - - val (_, prems, ts) = fold_rev mk_prem cargs (1, [], []); - val concl = HOLogic.mk_Trueprop - (Free (s, UnivT') $ mk_univ_inj ts n i) - in Logic.list_implies (prems, concl) - end; - - val intr_ts = maps (fn ((_, (_, _, constrs)), rep_set_name) => - map (make_intr rep_set_name (length constrs)) - ((1 upto (length constrs)) ~~ constrs)) (descr' ~~ rep_set_names'); - - val ({raw_induct = rep_induct, intrs = rep_intrs, ...}, thy2) = - thy1 - |> Sign.map_naming Name_Space.conceal - |> Inductive.add_inductive_global - {quiet_mode = #quiet config, verbose = false, alt_name = Binding.name big_rec_name, - coind = false, no_elim = true, no_ind = false, skip_mono = true, fork_mono = false} - (map (fn s => ((Binding.name s, UnivT'), NoSyn)) rep_set_names') [] - (map (fn x => (Attrib.empty_binding, x)) intr_ts) [] - ||> Sign.restore_naming thy1 - ||> Theory.checkpoint; - - (********************************* typedef ********************************) - - val (typedefs, thy3) = thy2 |> - Sign.parent_path |> - fold_map (fn ((((name, mx), tvs), c), name') => - Typedef.add_typedef false (SOME (Binding.name name')) (name, tvs, mx) - (Collect $ Const (c, UnivT')) NONE - (rtac exI 1 THEN rtac CollectI 1 THEN - QUIET_BREADTH_FIRST (has_fewer_prems 1) - (resolve_tac rep_intrs 1))) - (types_syntax ~~ tyvars ~~ - (Library.take (length newTs, rep_set_names)) ~~ new_type_names) ||> - Sign.add_path big_name; - - (*********************** definition of constructors ***********************) - - val big_rep_name = (space_implode "_" new_type_names) ^ "_Rep_"; - val rep_names = map (curry op ^ "Rep_") new_type_names; - val rep_names' = map (fn i => big_rep_name ^ (string_of_int i)) - (1 upto (length (flat (tl descr)))); - val all_rep_names = map (Sign.intern_const thy3) rep_names @ - map (Sign.full_bname thy3) rep_names'; - - (* isomorphism declarations *) - - val iso_decls = map (fn (T, s) => (Binding.name s, T --> Univ_elT, NoSyn)) - (oldTs ~~ rep_names'); - - (* constructor definitions *) - - fun make_constr_def tname T n ((cname, cargs), (cname', mx)) (thy, defs, eqns, i) = - let - fun constr_arg dt (j, l_args, r_args) = - let val T = typ_of_dtyp descr' sorts dt; - val free_t = mk_Free "x" T j - in (case (strip_dtyp dt, strip_type T) of - ((_, DtRec m), (Us, U)) => (j + 1, free_t :: l_args, mk_lim - (Const (nth all_rep_names m, U --> Univ_elT) $ - app_bnds free_t (length Us)) Us :: r_args) - | _ => (j + 1, free_t::l_args, (Leaf $ mk_inj T free_t)::r_args)) - end; - - val (_, l_args, r_args) = fold_rev constr_arg cargs (1, [], []); - val constrT = (map (typ_of_dtyp descr' sorts) cargs) ---> T; - val abs_name = Sign.intern_const thy ("Abs_" ^ tname); - val rep_name = Sign.intern_const thy ("Rep_" ^ tname); - val lhs = list_comb (Const (cname, constrT), l_args); - val rhs = mk_univ_inj r_args n i; - val def = Logic.mk_equals (lhs, Const (abs_name, Univ_elT --> T) $ rhs); - val def_name = Long_Name.base_name cname ^ "_def"; - val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq - (Const (rep_name, T --> Univ_elT) $ lhs, rhs)); - val ([def_thm], thy') = - thy - |> Sign.add_consts_i [(cname', constrT, mx)] - |> (PureThy.add_defs false o map Thm.no_attributes) [(Binding.name def_name, def)]; - - in (thy', defs @ [def_thm], eqns @ [eqn], i + 1) end; - - (* constructor definitions for datatype *) - - fun dt_constr_defs ((((_, (_, _, constrs)), tname), T), constr_syntax) - (thy, defs, eqns, rep_congs, dist_lemmas) = - let - val _ $ (_ $ (cong_f $ _) $ _) = concl_of arg_cong; - val rep_const = cterm_of thy - (Const (Sign.intern_const thy ("Rep_" ^ tname), T --> Univ_elT)); - val cong' = - Drule.standard (cterm_instantiate [(cterm_of thy cong_f, rep_const)] arg_cong); - val dist = - Drule.standard (cterm_instantiate [(cterm_of thy distinct_f, rep_const)] distinct_lemma); - val (thy', defs', eqns', _) = fold ((make_constr_def tname T) (length constrs)) - (constrs ~~ constr_syntax) (Sign.add_path tname thy, defs, [], 1); - in - (Sign.parent_path thy', defs', eqns @ [eqns'], - rep_congs @ [cong'], dist_lemmas @ [dist]) - end; - - val (thy4, constr_defs, constr_rep_eqns, rep_congs, dist_lemmas) = - fold dt_constr_defs - (hd descr ~~ new_type_names ~~ newTs ~~ constr_syntax) - (thy3 |> Sign.add_consts_i iso_decls |> Sign.parent_path, [], [], [], []); - - - (*********** isomorphisms for new types (introduced by typedef) ***********) - - val _ = message config "Proving isomorphism properties ..."; - - val newT_iso_axms = map (fn (_, td) => - (collect_simp (#Abs_inverse td), #Rep_inverse td, - collect_simp (#Rep td))) typedefs; - - val newT_iso_inj_thms = map (fn (_, td) => - (collect_simp (#Abs_inject td) RS iffD1, #Rep_inject td RS iffD1)) typedefs; - - (********* isomorphisms between existing types and "unfolded" types *******) - - (*---------------------------------------------------------------------*) - (* isomorphisms are defined using primrec-combinators: *) - (* generate appropriate functions for instantiating primrec-combinator *) - (* *) - (* e.g. dt_Rep_i = list_rec ... (%h t y. In1 (Scons (Leaf h) y)) *) - (* *) - (* also generate characteristic equations for isomorphisms *) - (* *) - (* e.g. dt_Rep_i (cons h t) = In1 (Scons (dt_Rep_j h) (dt_Rep_i t)) *) - (*---------------------------------------------------------------------*) - - fun make_iso_def k ks n (cname, cargs) (fs, eqns, i) = - let - val argTs = map (typ_of_dtyp descr' sorts) cargs; - val T = nth recTs k; - val rep_name = nth all_rep_names k; - val rep_const = Const (rep_name, T --> Univ_elT); - val constr = Const (cname, argTs ---> T); - - fun process_arg ks' dt (i2, i2', ts, Ts) = - let - val T' = typ_of_dtyp descr' sorts dt; - val (Us, U) = strip_type T' - in (case strip_dtyp dt of - (_, DtRec j) => if j mem ks' then - (i2 + 1, i2' + 1, ts @ [mk_lim (app_bnds - (mk_Free "y" (Us ---> Univ_elT) i2') (length Us)) Us], - Ts @ [Us ---> Univ_elT]) - else - (i2 + 1, i2', ts @ [mk_lim - (Const (nth all_rep_names j, U --> Univ_elT) $ - app_bnds (mk_Free "x" T' i2) (length Us)) Us], Ts) - | _ => (i2 + 1, i2', ts @ [Leaf $ mk_inj T' (mk_Free "x" T' i2)], Ts)) - end; - - val (i2, i2', ts, Ts) = fold (process_arg ks) cargs (1, 1, [], []); - val xs = map (uncurry (mk_Free "x")) (argTs ~~ (1 upto (i2 - 1))); - val ys = map (uncurry (mk_Free "y")) (Ts ~~ (1 upto (i2' - 1))); - val f = list_abs_free (map dest_Free (xs @ ys), mk_univ_inj ts n i); - - val (_, _, ts', _) = fold (process_arg []) cargs (1, 1, [], []); - val eqn = HOLogic.mk_Trueprop (HOLogic.mk_eq - (rep_const $ list_comb (constr, xs), mk_univ_inj ts' n i)) - - in (fs @ [f], eqns @ [eqn], i + 1) end; - - (* define isomorphisms for all mutually recursive datatypes in list ds *) - - fun make_iso_defs ds (thy, char_thms) = - let - val ks = map fst ds; - val (_, (tname, _, _)) = hd ds; - val {rec_rewrites, rec_names, ...} = the (Symtab.lookup dt_info tname); - - fun process_dt (k, (tname, _, constrs)) (fs, eqns, isos) = - let - val (fs', eqns', _) = - fold (make_iso_def k ks (length constrs)) constrs (fs, eqns, 1); - val iso = (nth recTs k, nth all_rep_names k) - in (fs', eqns', isos @ [iso]) end; - - val (fs, eqns, isos) = fold process_dt ds ([], [], []); - val fTs = map fastype_of fs; - val defs = map (fn (rec_name, (T, iso_name)) => (Binding.name (Long_Name.base_name iso_name ^ "_def"), - Logic.mk_equals (Const (iso_name, T --> Univ_elT), - list_comb (Const (rec_name, fTs @ [T] ---> Univ_elT), fs)))) (rec_names ~~ isos); - val (def_thms, thy') = - apsnd Theory.checkpoint ((PureThy.add_defs false o map Thm.no_attributes) defs thy); - - (* prove characteristic equations *) - - val rewrites = def_thms @ (map mk_meta_eq rec_rewrites); - val char_thms' = map (fn eqn => Skip_Proof.prove_global thy' [] [] eqn - (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])) eqns; - - in (thy', char_thms' @ char_thms) end; - - val (thy5, iso_char_thms) = apfst Theory.checkpoint (fold_rev make_iso_defs - (tl descr) (Sign.add_path big_name thy4, [])); - - (* prove isomorphism properties *) - - fun mk_funs_inv thy thm = - let - val prop = Thm.prop_of thm; - val _ $ (_ $ ((S as Const (_, Type (_, [U, _]))) $ _ )) $ - (_ $ (_ $ (r $ (a $ _)) $ _)) = Type.legacy_freeze prop; - val used = OldTerm.add_term_tfree_names (a, []); - - fun mk_thm i = - let - val Ts = map (TFree o rpair HOLogic.typeS) - (Name.variant_list used (replicate i "'t")); - val f = Free ("f", Ts ---> U) - in Skip_Proof.prove_global thy [] [] (Logic.mk_implies - (HOLogic.mk_Trueprop (HOLogic.list_all - (map (pair "x") Ts, S $ app_bnds f i)), - HOLogic.mk_Trueprop (HOLogic.mk_eq (list_abs (map (pair "x") Ts, - r $ (a $ app_bnds f i)), f)))) - (fn _ => EVERY [REPEAT_DETERM_N i (rtac ext 1), - REPEAT (etac allE 1), rtac thm 1, atac 1]) - end - in map (fn r => r RS subst) (thm :: map mk_thm arities) end; - - (* prove inj dt_Rep_i and dt_Rep_i x : dt_rep_set_i *) - - val fun_congs = map (fn T => make_elim (Drule.instantiate' - [SOME (ctyp_of thy5 T)] [] fun_cong)) branchTs; - - fun prove_iso_thms ds (inj_thms, elem_thms) = - let - val (_, (tname, _, _)) = hd ds; - val induct = (#induct o the o Symtab.lookup dt_info) tname; - - fun mk_ind_concl (i, _) = - let - val T = nth recTs i; - val Rep_t = Const (nth all_rep_names i, T --> Univ_elT); - val rep_set_name = nth rep_set_names i - in (HOLogic.all_const T $ Abs ("y", T, HOLogic.imp $ - HOLogic.mk_eq (Rep_t $ mk_Free "x" T i, Rep_t $ Bound 0) $ - HOLogic.mk_eq (mk_Free "x" T i, Bound 0)), - Const (rep_set_name, UnivT') $ (Rep_t $ mk_Free "x" T i)) - end; - - val (ind_concl1, ind_concl2) = ListPair.unzip (map mk_ind_concl ds); - - val rewrites = map mk_meta_eq iso_char_thms; - val inj_thms' = map snd newT_iso_inj_thms @ - map (fn r => r RS @{thm injD}) inj_thms; - - val inj_thm = Skip_Proof.prove_global thy5 [] [] - (HOLogic.mk_Trueprop (mk_conj ind_concl1)) (fn _ => EVERY - [(indtac induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, - REPEAT (EVERY - [rtac allI 1, rtac impI 1, - exh_tac (exh_thm_of dt_info) 1, - REPEAT (EVERY - [hyp_subst_tac 1, - rewrite_goals_tac rewrites, - REPEAT (dresolve_tac [In0_inject, In1_inject] 1), - (eresolve_tac [In0_not_In1 RS notE, In1_not_In0 RS notE] 1) - ORELSE (EVERY - [REPEAT (eresolve_tac (Scons_inject :: - map make_elim [Leaf_inject, Inl_inject, Inr_inject]) 1), - REPEAT (cong_tac 1), rtac refl 1, - REPEAT (atac 1 ORELSE (EVERY - [REPEAT (rtac ext 1), - REPEAT (eresolve_tac (mp :: allE :: - map make_elim (Suml_inject :: Sumr_inject :: - Lim_inject :: inj_thms') @ fun_congs) 1), - atac 1]))])])])]); - - val inj_thms'' = map (fn r => r RS @{thm datatype_injI}) - (split_conj_thm inj_thm); - - val elem_thm = - Skip_Proof.prove_global thy5 [] [] (HOLogic.mk_Trueprop (mk_conj ind_concl2)) - (fn _ => - EVERY [(indtac induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, - rewrite_goals_tac rewrites, - REPEAT ((resolve_tac rep_intrs THEN_ALL_NEW - ((REPEAT o etac allE) THEN' ares_tac elem_thms)) 1)]); - - in (inj_thms'' @ inj_thms, elem_thms @ (split_conj_thm elem_thm)) - end; - - val (iso_inj_thms_unfolded, iso_elem_thms) = - fold_rev prove_iso_thms (tl descr) ([], map #3 newT_iso_axms); - val iso_inj_thms = map snd newT_iso_inj_thms @ - map (fn r => r RS @{thm injD}) iso_inj_thms_unfolded; - - (* prove dt_rep_set_i x --> x : range dt_Rep_i *) - - fun mk_iso_t (((set_name, iso_name), i), T) = - let val isoT = T --> Univ_elT - in HOLogic.imp $ - (Const (set_name, UnivT') $ mk_Free "x" Univ_elT i) $ - (if i < length newTs then HOLogic.true_const - else HOLogic.mk_mem (mk_Free "x" Univ_elT i, - Const (@{const_name image}, isoT --> HOLogic.mk_setT T --> UnivT) $ - Const (iso_name, isoT) $ Const (@{const_name UNIV}, HOLogic.mk_setT T))) - end; - - val iso_t = HOLogic.mk_Trueprop (mk_conj (map mk_iso_t - (rep_set_names ~~ all_rep_names ~~ (0 upto (length descr' - 1)) ~~ recTs))); - - (* all the theorems are proved by one single simultaneous induction *) - - val range_eqs = map (fn r => mk_meta_eq (r RS @{thm range_ex1_eq})) - iso_inj_thms_unfolded; - - val iso_thms = if length descr = 1 then [] else - Library.drop (length newTs, split_conj_thm - (Skip_Proof.prove_global thy5 [] [] iso_t (fn _ => EVERY - [(indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, - REPEAT (rtac TrueI 1), - rewrite_goals_tac (mk_meta_eq choice_eq :: - symmetric (mk_meta_eq @{thm expand_fun_eq}) :: range_eqs), - rewrite_goals_tac (map symmetric range_eqs), - REPEAT (EVERY - [REPEAT (eresolve_tac ([rangeE, ex1_implies_ex RS exE] @ - maps (mk_funs_inv thy5 o #1) newT_iso_axms) 1), - TRY (hyp_subst_tac 1), - rtac (sym RS range_eqI) 1, - resolve_tac iso_char_thms 1])]))); - - val Abs_inverse_thms' = - map #1 newT_iso_axms @ - map2 (fn r_inj => fn r => @{thm f_the_inv_into_f} OF [r_inj, r RS mp]) - iso_inj_thms_unfolded iso_thms; - - val Abs_inverse_thms = maps (mk_funs_inv thy5) Abs_inverse_thms'; - - (******************* freeness theorems for constructors *******************) - - val _ = message config "Proving freeness of constructors ..."; - - (* prove theorem Rep_i (Constr_j ...) = Inj_j ... *) - - fun prove_constr_rep_thm eqn = - let - val inj_thms = map fst newT_iso_inj_thms; - val rewrites = @{thm o_def} :: constr_defs @ (map (mk_meta_eq o #2) newT_iso_axms) - in Skip_Proof.prove_global thy5 [] [] eqn (fn _ => EVERY - [resolve_tac inj_thms 1, - rewrite_goals_tac rewrites, - rtac refl 3, - resolve_tac rep_intrs 2, - REPEAT (resolve_tac iso_elem_thms 1)]) - end; - - (*--------------------------------------------------------------*) - (* constr_rep_thms and rep_congs are used to prove distinctness *) - (* of constructors. *) - (*--------------------------------------------------------------*) - - val constr_rep_thms = map (map prove_constr_rep_thm) constr_rep_eqns; - - val dist_rewrites = map (fn (rep_thms, dist_lemma) => - dist_lemma::(rep_thms @ [In0_eq, In1_eq, In0_not_In1, In1_not_In0])) - (constr_rep_thms ~~ dist_lemmas); - - fun prove_distinct_thms dist_rewrites' (k, ts) = - let - fun prove [] = [] - | prove (t :: ts) = - let - val dist_thm = Skip_Proof.prove_global thy5 [] [] t (fn _ => - EVERY [simp_tac (HOL_ss addsimps dist_rewrites') 1]) - in dist_thm :: Drule.standard (dist_thm RS not_sym) :: prove ts end; - in prove ts end; - - val distinct_thms = map2 (prove_distinct_thms) - dist_rewrites (DatatypeProp.make_distincts descr sorts); - - (* prove injectivity of constructors *) - - fun prove_constr_inj_thm rep_thms t = - let val inj_thms = Scons_inject :: (map make_elim - (iso_inj_thms @ - [In0_inject, In1_inject, Leaf_inject, Inl_inject, Inr_inject, - Lim_inject, Suml_inject, Sumr_inject])) - in Skip_Proof.prove_global thy5 [] [] t (fn _ => EVERY - [rtac iffI 1, - REPEAT (etac conjE 2), hyp_subst_tac 2, rtac refl 2, - dresolve_tac rep_congs 1, dtac box_equals 1, - REPEAT (resolve_tac rep_thms 1), - REPEAT (eresolve_tac inj_thms 1), - REPEAT (ares_tac [conjI] 1 ORELSE (EVERY [REPEAT (rtac ext 1), - REPEAT (eresolve_tac (make_elim fun_cong :: inj_thms) 1), - atac 1]))]) - end; - - val constr_inject = map (fn (ts, thms) => map (prove_constr_inj_thm thms) ts) - ((DatatypeProp.make_injs descr sorts) ~~ constr_rep_thms); - - val ((constr_inject', distinct_thms'), thy6) = - thy5 - |> Sign.parent_path - |> store_thmss "inject" new_type_names constr_inject - ||>> store_thmss "distinct" new_type_names distinct_thms; - - (*************************** induction theorem ****************************) - - val _ = message config "Proving induction rule for datatypes ..."; - - val Rep_inverse_thms = (map (fn (_, iso, _) => iso RS subst) newT_iso_axms) @ - (map (fn r => r RS @{thm the_inv_f_f} RS subst) iso_inj_thms_unfolded); - val Rep_inverse_thms' = map (fn r => r RS @{thm the_inv_f_f}) iso_inj_thms_unfolded; - - fun mk_indrule_lemma ((i, _), T) (prems, concls) = - let - val Rep_t = Const (nth all_rep_names i, T --> Univ_elT) $ - mk_Free "x" T i; - - val Abs_t = if i < length newTs then - Const (Sign.intern_const thy6 - ("Abs_" ^ (nth new_type_names i)), Univ_elT --> T) - else Const (@{const_name the_inv_into}, - [HOLogic.mk_setT T, T --> Univ_elT, Univ_elT] ---> T) $ - HOLogic.mk_UNIV T $ Const (nth all_rep_names i, T --> Univ_elT) - - in (prems @ [HOLogic.imp $ - (Const (nth rep_set_names i, UnivT') $ Rep_t) $ - (mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ (Abs_t $ Rep_t))], - concls @ [mk_Free "P" (T --> HOLogic.boolT) (i + 1) $ mk_Free "x" T i]) - end; - - val (indrule_lemma_prems, indrule_lemma_concls) = - fold mk_indrule_lemma (descr' ~~ recTs) ([], []); - - val cert = cterm_of thy6; - - val indrule_lemma = Skip_Proof.prove_global thy6 [] [] - (Logic.mk_implies - (HOLogic.mk_Trueprop (mk_conj indrule_lemma_prems), - HOLogic.mk_Trueprop (mk_conj indrule_lemma_concls))) (fn _ => EVERY - [REPEAT (etac conjE 1), - REPEAT (EVERY - [TRY (rtac conjI 1), resolve_tac Rep_inverse_thms 1, - etac mp 1, resolve_tac iso_elem_thms 1])]); - - val Ps = map head_of (HOLogic.dest_conj (HOLogic.dest_Trueprop (concl_of indrule_lemma))); - val frees = if length Ps = 1 then [Free ("P", snd (dest_Var (hd Ps)))] else - map (Free o apfst fst o dest_Var) Ps; - val indrule_lemma' = cterm_instantiate (map cert Ps ~~ map cert frees) indrule_lemma; - - val dt_induct_prop = DatatypeProp.make_ind descr sorts; - val dt_induct = Skip_Proof.prove_global thy6 [] - (Logic.strip_imp_prems dt_induct_prop) (Logic.strip_imp_concl dt_induct_prop) - (fn {prems, ...} => EVERY - [rtac indrule_lemma' 1, - (indtac rep_induct [] THEN_ALL_NEW ObjectLogic.atomize_prems_tac) 1, - EVERY (map (fn (prem, r) => (EVERY - [REPEAT (eresolve_tac Abs_inverse_thms 1), - simp_tac (HOL_basic_ss addsimps ((symmetric r)::Rep_inverse_thms')) 1, - DEPTH_SOLVE_1 (ares_tac [prem] 1 ORELSE etac allE 1)])) - (prems ~~ (constr_defs @ (map mk_meta_eq iso_char_thms))))]); - - val ([dt_induct'], thy7) = - thy6 - |> Sign.add_path big_name - |> PureThy.add_thms [((Binding.name "induct", dt_induct), [case_names_induct])] - ||> Sign.parent_path - ||> Theory.checkpoint; - - in - ((constr_inject', distinct_thms', dt_induct'), thy7) - end; - -end; diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Function/fun.ML --- a/src/HOL/Tools/Function/fun.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/Function/fun.ML Fri Dec 04 15:20:24 2009 +0100 @@ -121,9 +121,9 @@ (Function_Split.split_all_equations ctxt compleqs) fun restore_spec thms = - bnds ~~ Library.take (length bnds, Library.unflat spliteqs thms) + bnds ~~ take (length bnds) (unflat spliteqs thms) - val spliteqs' = flat (Library.take (length bnds, spliteqs)) + val spliteqs' = flat (take (length bnds) spliteqs) val fnames = map (fst o fst) fixes val indices = map (fn eq => find_index (curry op = (fname_of eq)) fnames) spliteqs' diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Function/size.ML --- a/src/HOL/Tools/Function/size.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/Function/size.ML Fri Dec 04 15:20:24 2009 +0100 @@ -13,7 +13,7 @@ structure Size: SIZE = struct -open DatatypeAux; +open Datatype_Aux; structure SizeData = Theory_Data ( @@ -177,7 +177,7 @@ fun gen_mk_size_eq p size_of size_ofp size_const T (cname, cargs) = let val Ts = map (typ_of_dtyp descr sorts) cargs; - val tnames = Name.variant_list f_names (DatatypeProp.make_tnames Ts); + val tnames = Name.variant_list f_names (Datatype_Prop.make_tnames Ts); val ts = map_filter (fn (sT as (s, T), dt) => Option.map (fn sz => sz $ Free sT) (if p dt then size_of_type size_of extra_size size_ofp T diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Function/sum_tree.ML --- a/src/HOL/Tools/Function/sum_tree.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/Function/sum_tree.ML Fri Dec 04 15:20:24 2009 +0100 @@ -8,8 +8,8 @@ struct (* Theory dependencies *) -val proj_in_rules = [@{thm "Datatype.Projl_Inl"}, @{thm "Datatype.Projr_Inr"}] -val sumcase_split_ss = HOL_basic_ss addsimps (@{thm "Product_Type.split"} :: @{thms "sum.cases"}) +val proj_in_rules = [@{thm Projl_Inl}, @{thm Projr_Inr}] +val sumcase_split_ss = HOL_basic_ss addsimps (@{thm Product_Type.split} :: @{thms sum.cases}) (* top-down access in balanced tree *) fun access_top_down {left, right, init} len i = @@ -31,8 +31,8 @@ fun mk_proj ST n i = access_top_down { init = (ST, I : term -> term), - left = (fn (T as Type ("+", [LT, RT]), proj) => (LT, App (Const (@{const_name Datatype.Projl}, T --> LT)) o proj)), - right =(fn (T as Type ("+", [LT, RT]), proj) => (RT, App (Const (@{const_name Datatype.Projr}, T --> RT)) o proj))} n i + left = (fn (T as Type ("+", [LT, RT]), proj) => (LT, App (Const (@{const_name Sum_Type.Projl}, T --> LT)) o proj)), + right =(fn (T as Type ("+", [LT, RT]), proj) => (RT, App (Const (@{const_name Sum_Type.Projr}, T --> RT)) o proj))} n i |> snd fun mk_sumcases T fs = diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Groebner_Basis/groebner.ML --- a/src/HOL/Tools/Groebner_Basis/groebner.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/Groebner_Basis/groebner.ML Fri Dec 04 15:20:24 2009 +0100 @@ -899,7 +899,7 @@ val (qs,p) = isolate_monomials vars eq val rs = ideal (qs @ ps) p (fn (s,t) => TermOrd.term_ord (term_of s, term_of t)) - in (eq,Library.take (length qs, rs) ~~ vars) + in (eq, take (length qs) rs ~~ vars) end; fun subst_in_poly i p = Thm.rhs_of (ring_normalize_conv (vsubst i p)); in diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Nitpick/nitpick.ML --- a/src/HOL/Tools/Nitpick/nitpick.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/Nitpick/nitpick.ML Fri Dec 04 15:20:24 2009 +0100 @@ -655,7 +655,7 @@ update_checked_problems problems (unsat_js @ map fst lib_ps); if null con_ps then let - val num_genuine = Library.take (max_potential, lib_ps) + val num_genuine = take max_potential lib_ps |> map (print_and_check false) |> filter (equal (SOME true)) |> length val max_genuine = max_genuine - num_genuine @@ -689,7 +689,7 @@ end else let - val _ = Library.take (max_genuine, con_ps) + val _ = take max_genuine con_ps |> List.app (ignore o print_and_check true) val max_genuine = max_genuine - length con_ps in diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Nitpick/nitpick_hol.ML --- a/src/HOL/Tools/Nitpick/nitpick_hol.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/Nitpick/nitpick_hol.ML Fri Dec 04 15:20:24 2009 +0100 @@ -516,7 +516,7 @@ | NONE => NONE (* FIXME: use antiquotation for "code_numeral" below or detect "rep_datatype", - e.g., by adding a field to "DatatypeAux.info". *) + e.g., by adding a field to "Datatype_Aux.info". *) (* string -> bool *) fun is_basic_datatype s = s mem [@{type_name "*"}, @{type_name bool}, @{type_name unit}, diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Nitpick/nitpick_scope.ML --- a/src/HOL/Tools/Nitpick/nitpick_scope.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/Nitpick/nitpick_scope.ML Fri Dec 04 15:20:24 2009 +0100 @@ -496,7 +496,7 @@ bisim_depths mono_Ts nonmono_Ts val ranks = map rank_of_block blocks val all = all_combinations_ordered_smartly (map (rpair 0) ranks) - val head = Library.take (max_scopes, all) + val head = take max_scopes all val descs = map_filter (scope_descriptor_from_combination ext_ctxt blocks) head in diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML --- a/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/Predicate_Compile/predicate_compile_core.ML Fri Dec 04 15:20:24 2009 +0100 @@ -1282,7 +1282,7 @@ val v' = Free (name', T); in lambda v (fst (Datatype.make_case - (ProofContext.init thy) DatatypeCase.Quiet [] v + (ProofContext.init thy) Datatype_Case.Quiet [] v [(HOLogic.mk_tuple out_ts, if null eqs'' then success_t else Const (@{const_name HOL.If}, HOLogic.boolT --> U --> U --> U) $ diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/TFL/tfl.ML --- a/src/HOL/Tools/TFL/tfl.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/TFL/tfl.ML Fri Dec 04 15:20:24 2009 +0100 @@ -193,8 +193,7 @@ fun mk_pat (c,l) = let val L = length (binder_types (type_of c)) fun build (prfx,tag,plist) = - let val args = Library.take (L,plist) - and plist' = Library.drop(L,plist) + let val (args, plist') = chop L plist in (prfx,tag,list_comb(c,args)::plist') end in map build l end; diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/inductive.ML --- a/src/HOL/Tools/inductive.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/inductive.ML Fri Dec 04 15:20:24 2009 +0100 @@ -213,7 +213,7 @@ fun make_bool_args' xs = make_bool_args (K HOLogic.false_const) (K HOLogic.true_const) xs; -fun arg_types_of k c = Library.drop (k, binder_types (fastype_of c)); +fun arg_types_of k c = drop k (binder_types (fastype_of c)); fun find_arg T x [] = sys_error "find_arg" | find_arg T x ((p as (_, (SOME _, _))) :: ps) = diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/inductive_codegen.ML --- a/src/HOL/Tools/inductive_codegen.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/inductive_codegen.ML Fri Dec 04 15:20:24 2009 +0100 @@ -104,7 +104,7 @@ let val cnstrs = flat (maps (map (fn (_, (_, _, cs)) => map (apsnd length) cs) o #descr o snd) - (Symtab.dest (Datatype.get_all thy))); + (Symtab.dest (Datatype_Data.get_all thy))); fun check t = (case strip_comb t of (Var _, []) => true | (Const (s, _), ts) => (case AList.lookup (op =) cnstrs s of diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/inductive_realizer.ML --- a/src/HOL/Tools/inductive_realizer.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/inductive_realizer.ML Fri Dec 04 15:20:24 2009 +0100 @@ -67,7 +67,7 @@ val Tvs = map TVar iTs; val (Const (s, _), ts) = strip_comb (HOLogic.dest_Trueprop (Logic.strip_imp_concl (prop_of (hd intrs)))); - val params = map dest_Var (Library.take (nparms, ts)); + val params = map dest_Var (take nparms ts); val tname = Binding.name (space_implode "_" (Long_Name.base_name s ^ "T" :: vs)); fun constr_of_intr intr = (Binding.name (Long_Name.base_name (name_of_thm intr)), map (Logic.unvarifyT o snd) (subtract (op =) params (rev (Term.add_vars (prop_of intr) []))) @ @@ -243,7 +243,7 @@ thy |> f (map snd dts) |-> (fn dtinfo => pair (map fst dts, SOME dtinfo)) - handle DatatypeAux.Datatype_Empty name' => + handle Datatype_Aux.Datatype_Empty name' => let val name = Long_Name.base_name name'; val dname = Name.variant used "Dummy"; @@ -398,7 +398,7 @@ val (thm', thy') = PureThy.store_thm (Binding.qualified_name (space_implode "_" (Long_Name.qualify qualifier "induct" :: vs' @ Ps @ ["correctness"])), thm) thy; val thms = map (fn th => zero_var_indexes (rotate_prems ~1 (th RS mp))) - (DatatypeAux.split_conj_thm thm'); + (Datatype_Aux.split_conj_thm thm'); val ([thms'], thy'') = PureThy.add_thmss [((Binding.qualified_name (space_implode "_" (Long_Name.qualify qualifier "inducts" :: vs' @ Ps @ diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/old_primrec.ML --- a/src/HOL/Tools/old_primrec.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/old_primrec.ML Fri Dec 04 15:20:24 2009 +0100 @@ -21,7 +21,7 @@ structure OldPrimrec : OLD_PRIMREC = struct -open DatatypeAux; +open Datatype_Aux; exception RecError of string; @@ -124,8 +124,8 @@ let val fnameT' as (fname', _) = dest_Const f; val (_, rpos, _) = the (AList.lookup (op =) rec_eqns fnameT'); - val ls = Library.take (rpos, ts); - val rest = Library.drop (rpos, ts); + val ls = take rpos ts; + val rest = drop rpos ts; val (x', rs) = (hd rest, tl rest) handle Empty => raise RecError ("not enough arguments\ \ in recursive application\nof function " ^ quote fname' ^ " on rhs"); @@ -246,7 +246,7 @@ fun gen_primrec_i note def alt_name eqns_atts thy = let val (eqns, atts) = split_list eqns_atts; - val dt_info = Datatype.get_all thy; + val dt_info = Datatype_Data.get_all thy; val rec_eqns = fold_rev (process_eqn thy o snd) eqns [] ; val tnames = distinct (op =) (map (#1 o snd) rec_eqns); val dts = find_dts dt_info tnames tnames; diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/primrec.ML --- a/src/HOL/Tools/primrec.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/primrec.ML Fri Dec 04 15:20:24 2009 +0100 @@ -23,7 +23,7 @@ structure Primrec : PRIMREC = struct -open DatatypeAux; +open Datatype_Aux; exception PrimrecError of string * term option; @@ -220,7 +220,7 @@ val eqns = fold_rev (process_eqn (fn v => Variable.is_fixed lthy v orelse exists (fn ((w, _), _) => v = Binding.name_of w) fixes)) eqs []; val tnames = distinct (op =) (map (#1 o snd) eqns); - val dts = find_dts (Datatype.get_all (ProofContext.theory_of lthy)) tnames tnames; + val dts = find_dts (Datatype_Data.get_all (ProofContext.theory_of lthy)) tnames tnames; val main_fns = map (fn (tname, {index, ...}) => (index, (fst o the o find_first (fn (_, x) => #1 x = tname)) eqns)) dts; val {descr, rec_names, rec_rewrites, ...} = diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/quickcheck_generators.ML --- a/src/HOL/Tools/quickcheck_generators.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/quickcheck_generators.ML Fri Dec 04 15:20:24 2009 +0100 @@ -246,10 +246,10 @@ mk_const @{const_name random_fun_lift} [fTs ---> T, fT] $ mk_random_fun_lift fTs t; val t = mk_random_fun_lift fTs (nth random_auxs k $ size_pred $ size'); - val size = Option.map snd (DatatypeCodegen.find_shortest_path descr k) + val size = Option.map snd (Datatype_Aux.find_shortest_path descr k) |> the_default 0; in (SOME size, (t, fTs ---> T)) end; - val tss = DatatypeAux.interpret_construction descr vs + val tss = Datatype_Aux.interpret_construction descr vs { atyp = mk_random_call, dtyp = mk_random_aux_call }; fun mk_consexpr simpleT (c, xs) = let @@ -287,9 +287,9 @@ fun mk_random_datatype config descr vs tycos prfx (names, auxnames) (Ts, Us) thy = let - val _ = DatatypeAux.message config "Creating quickcheck generators ..."; + val _ = Datatype_Aux.message config "Creating quickcheck generators ..."; val mk_prop_eq = HOLogic.mk_Trueprop o HOLogic.mk_eq; - fun mk_size_arg k = case DatatypeCodegen.find_shortest_path descr k + fun mk_size_arg k = case Datatype_Aux.find_shortest_path descr k of SOME (_, l) => if l = 0 then size else @{term "max :: code_numeral \ code_numeral \ code_numeral"} $ HOLogic.mk_number @{typ code_numeral} l $ size @@ -328,10 +328,10 @@ val typerep_vs = (map o apsnd) (curry (Sorts.inter_sort algebra) @{sort typerep}) raw_vs; val random_insts = (map (rpair @{sort random}) o flat o maps snd o maps snd) - (DatatypeAux.interpret_construction descr typerep_vs + (Datatype_Aux.interpret_construction descr typerep_vs { atyp = single, dtyp = (K o K o K) [] }); val term_of_insts = (map (rpair @{sort term_of}) o flat o maps snd o maps snd) - (DatatypeAux.interpret_construction descr typerep_vs + (Datatype_Aux.interpret_construction descr typerep_vs { atyp = K [], dtyp = K o K }); val has_inst = exists (fn tyco => can (Sorts.mg_domain algebra tyco) @{sort random}) tycos; diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/record.ML --- a/src/HOL/Tools/record.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/record.ML Fri Dec 04 15:20:24 2009 +0100 @@ -321,7 +321,7 @@ fun rec_id i T = let val rTs = dest_recTs T; - val rTs' = if i < 0 then rTs else Library.take (i, rTs); + val rTs' = if i < 0 then rTs else take i rTs; in implode (map #1 rTs') end; @@ -1582,7 +1582,7 @@ (Logic.strip_assums_concl (prop_of rule'))); (*ca indicates if rule is a case analysis or induction rule*) val (x, ca) = - (case rev (Library.drop (length params, ys)) of + (case rev (drop (length params) ys) of [] => (head_of (fst (HOLogic.dest_eq (HOLogic.dest_Trueprop (hd (rev (Logic.strip_assums_hyp (hd (prems_of rule')))))))), true) | [x] => (head_of x, false)); @@ -1635,7 +1635,7 @@ else if len > 16 then let fun group16 [] = [] - | group16 xs = Library.take (16, xs) :: group16 (Library.drop (16, xs)); + | group16 xs = take 16 xs :: group16 (drop 16 xs); val vars' = group16 vars; val (composites, (thy', i')) = fold_map mk_even_istuple vars' (thy, i); in @@ -1772,7 +1772,7 @@ fun chunks [] [] = [] | chunks [] xs = [xs] - | chunks (l :: ls) xs = Library.take (l, xs) :: chunks ls (Library.drop (l, xs)); + | chunks (l :: ls) xs = take l xs :: chunks ls (drop l xs); fun chop_last [] = error "chop_last: list should not be empty" | chop_last [x] = ([], x) @@ -1881,12 +1881,12 @@ val extension_names = map (unsuffix ext_typeN o fst o #extension) parents @ [extN]; val extension_id = implode extension_names; - fun rec_schemeT n = mk_recordT (map #extension (Library.drop (n, parents))) extT; + fun rec_schemeT n = mk_recordT (map #extension (drop n parents)) extT; val rec_schemeT0 = rec_schemeT 0; fun recT n = let val (c, Ts) = extension in - mk_recordT (map #extension (Library.drop (n, parents))) + mk_recordT (map #extension (drop n parents)) (Type (c, subst_last HOLogic.unitT Ts)) end; val recT0 = recT 0; @@ -1896,7 +1896,7 @@ val (args', more) = chop_last args; fun mk_ext' ((name, T), args) more = mk_ext (name, T) (args @ [more]); fun build Ts = - fold_rev mk_ext' (Library.drop (n, (extension_names ~~ Ts) ~~ chunks parent_chunks args')) + fold_rev mk_ext' (drop n ((extension_names ~~ Ts) ~~ chunks parent_chunks args')) more; in if more = HOLogic.unit @@ -1989,7 +1989,7 @@ val (accessor_thms, updator_thms, upd_acc_cong_assists) = timeit_msg "record getting tree access/updates:" get_access_update_thms; - fun lastN xs = Library.drop (parent_fields_len, xs); + fun lastN xs = drop parent_fields_len xs; (*selectors*) fun mk_sel_spec ((c, T), thm) = diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/refute.ML --- a/src/HOL/Tools/refute.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/refute.ML Fri Dec 04 15:20:24 2009 +0100 @@ -406,12 +406,12 @@ (* TRANSLATION HOL -> PROPOSITIONAL LOGIC, BOOLEAN ASSIGNMENT -> MODEL *) (* ------------------------------------------------------------------------- *) - fun typ_of_dtyp descr typ_assoc (DatatypeAux.DtTFree a) = + fun typ_of_dtyp descr typ_assoc (Datatype_Aux.DtTFree a) = (* replace a 'DtTFree' variable by the associated type *) - the (AList.lookup (op =) typ_assoc (DatatypeAux.DtTFree a)) - | typ_of_dtyp descr typ_assoc (DatatypeAux.DtType (s, ds)) = + the (AList.lookup (op =) typ_assoc (Datatype_Aux.DtTFree a)) + | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtType (s, ds)) = Type (s, map (typ_of_dtyp descr typ_assoc) ds) - | typ_of_dtyp descr typ_assoc (DatatypeAux.DtRec i) = + | typ_of_dtyp descr typ_assoc (Datatype_Aux.DtRec i) = let val (s, ds, _) = the (AList.lookup (op =) descr i) in @@ -946,7 +946,7 @@ (* sanity check: every element in 'dtyps' must be a *) (* 'DtTFree' *) val _ = if Library.exists (fn d => - case d of DatatypeAux.DtTFree _ => false | _ => true) typs then + case d of Datatype_Aux.DtTFree _ => false | _ => true) typs then raise REFUTE ("ground_types", "datatype argument (for type " ^ Syntax.string_of_typ_global thy T ^ ") is not a variable") else () @@ -958,11 +958,11 @@ val dT = typ_of_dtyp descr typ_assoc d in case d of - DatatypeAux.DtTFree _ => + Datatype_Aux.DtTFree _ => collect_types dT acc - | DatatypeAux.DtType (_, ds) => + | Datatype_Aux.DtType (_, ds) => collect_types dT (fold_rev collect_dtyp ds acc) - | DatatypeAux.DtRec i => + | Datatype_Aux.DtRec i => if dT mem acc then acc (* prevent infinite recursion *) else @@ -971,7 +971,7 @@ (* if the current type is a recursive IDT (i.e. a depth *) (* is required), add it to 'acc' *) val acc_dT = if Library.exists (fn (_, ds) => - Library.exists DatatypeAux.is_rec_type ds) dconstrs then + Library.exists Datatype_Aux.is_rec_type ds) dconstrs then insert (op =) dT acc else acc (* collect argument types *) @@ -985,7 +985,7 @@ in (* argument types 'Ts' could be added here, but they are also *) (* added by 'collect_dtyp' automatically *) - collect_dtyp (DatatypeAux.DtRec index) acc + collect_dtyp (Datatype_Aux.DtRec index) acc end | NONE => (* not an inductive datatype, e.g. defined via "typedef" or *) @@ -1157,7 +1157,7 @@ in (* recursive datatype? *) Library.exists (fn (_, ds) => - Library.exists DatatypeAux.is_rec_type ds) constrs + Library.exists Datatype_Aux.is_rec_type ds) constrs end | NONE => false) | _ => false) types @@ -1952,7 +1952,7 @@ val typ_assoc = dtyps ~~ Ts (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) val _ = if Library.exists (fn d => - case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps + case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps then raise REFUTE ("IDT_interpreter", "datatype argument (for type " @@ -2076,7 +2076,7 @@ val typ_assoc = dtyps ~~ Ts (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) val _ = if Library.exists (fn d => - case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps + case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps then raise REFUTE ("IDT_constructor_interpreter", "datatype argument (for type " @@ -2135,7 +2135,7 @@ val typ_assoc = dtyps ~~ Ts' (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) val _ = if Library.exists (fn d => - case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps + case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps then raise REFUTE ("IDT_constructor_interpreter", "datatype argument (for type " @@ -2350,7 +2350,7 @@ (* sanity check: every element in 'dtyps' must be a *) (* 'DtTFree' *) val _ = if Library.exists (fn d => - case d of DatatypeAux.DtTFree _ => false + case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps then raise REFUTE ("IDT_recursion_interpreter", @@ -2372,10 +2372,10 @@ (case AList.lookup op= acc d of NONE => (case d of - DatatypeAux.DtTFree _ => + Datatype_Aux.DtTFree _ => (* add the association, proceed *) rec_typ_assoc ((d, T)::acc) xs - | DatatypeAux.DtType (s, ds) => + | Datatype_Aux.DtType (s, ds) => let val (s', Ts) = dest_Type T in @@ -2385,7 +2385,7 @@ raise REFUTE ("IDT_recursion_interpreter", "DtType/Type mismatch") end - | DatatypeAux.DtRec i => + | Datatype_Aux.DtRec i => let val (_, ds, _) = the (AList.lookup (op =) descr i) val (_, Ts) = dest_Type T @@ -2401,7 +2401,7 @@ raise REFUTE ("IDT_recursion_interpreter", "different type associations for the same dtyp")) val typ_assoc = filter - (fn (DatatypeAux.DtTFree _, _) => true | (_, _) => false) + (fn (Datatype_Aux.DtTFree _, _) => true | (_, _) => false) (rec_typ_assoc [] (#2 (the (AList.lookup (op =) descr idt_index)) ~~ (snd o dest_Type) IDT)) (* sanity check: typ_assoc must associate types to the *) @@ -2417,7 +2417,7 @@ val mc_intrs = map (fn (idx, (_, _, cs)) => let val c_return_typ = typ_of_dtyp descr typ_assoc - (DatatypeAux.DtRec idx) + (Datatype_Aux.DtRec idx) in (idx, map (fn (cname, cargs) => (#1 o interpret thy (typs, []) {maxvars=0, @@ -2471,7 +2471,7 @@ val _ = map (fn (idx, intrs) => let val T = typ_of_dtyp descr typ_assoc - (DatatypeAux.DtRec idx) + (Datatype_Aux.DtRec idx) in if length intrs <> size_of_type thy (typs, []) T then raise REFUTE ("IDT_recursion_interpreter", @@ -2552,7 +2552,7 @@ val (_, _, constrs) = the (AList.lookup (op =) descr idx) val (_, dtyps) = List.nth (constrs, c) val rec_dtyps_args = filter - (DatatypeAux.is_rec_type o fst) (dtyps ~~ args) + (Datatype_Aux.is_rec_type o fst) (dtyps ~~ args) (* map those indices to interpretations *) val rec_dtyps_intrs = map (fn (dtyp, arg) => let @@ -2565,18 +2565,18 @@ (* takes the dtyp and interpretation of an element, *) (* and computes the interpretation for the *) (* corresponding recursive argument *) - fun rec_intr (DatatypeAux.DtRec i) (Leaf xs) = + fun rec_intr (Datatype_Aux.DtRec i) (Leaf xs) = (* recursive argument is "rec_i params elem" *) compute_array_entry i (find_index (fn x => x = True) xs) - | rec_intr (DatatypeAux.DtRec _) (Node _) = + | rec_intr (Datatype_Aux.DtRec _) (Node _) = raise REFUTE ("IDT_recursion_interpreter", "interpretation for IDT is a node") - | rec_intr (DatatypeAux.DtType ("fun", [dt1, dt2])) + | rec_intr (Datatype_Aux.DtType ("fun", [dt1, dt2])) (Node xs) = (* recursive argument is something like *) (* "\x::dt1. rec_? params (elem x)" *) Node (map (rec_intr dt2) xs) - | rec_intr (DatatypeAux.DtType ("fun", [_, _])) + | rec_intr (Datatype_Aux.DtType ("fun", [_, _])) (Leaf _) = raise REFUTE ("IDT_recursion_interpreter", "interpretation for function dtyp is a leaf") @@ -2889,7 +2889,7 @@ (* go back up the stack until we find a level where we can go *) (* to the next sibling node *) let - val offsets' = Library.dropwhile + val offsets' = dropwhile (fn off' => off' mod size_elem = size_elem - 1) offsets in case offsets' of @@ -3169,7 +3169,7 @@ val typ_assoc = dtyps ~~ Ts (* sanity check: every element in 'dtyps' must be a 'DtTFree' *) val _ = if Library.exists (fn d => - case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps + case d of Datatype_Aux.DtTFree _ => false | _ => true) dtyps then raise REFUTE ("IDT_printer", "datatype argument (for type " ^ Syntax.string_of_typ_global thy (Type (s, Ts)) ^ ") is not a variable") diff -r 651028e34b5d -r 01dcd9b926bf src/HOL/Tools/split_rule.ML --- a/src/HOL/Tools/split_rule.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOL/Tools/split_rule.ML Fri Dec 04 15:20:24 2009 +0100 @@ -74,8 +74,8 @@ let val cterm = Thm.cterm_of (Thm.theory_of_thm rl) val (Us', U') = strip_type T; - val Us = Library.take (length ts, Us'); - val U = Library.drop (length ts, Us') ---> U'; + val Us = take (length ts) Us'; + val U = drop (length ts) Us' ---> U'; val T' = maps HOLogic.flatten_tupleT Us ---> U; fun mk_tuple (v as Var ((a, _), T)) (xs, insts) = let diff -r 651028e34b5d -r 01dcd9b926bf src/HOLCF/Tools/Domain/domain_axioms.ML --- a/src/HOLCF/Tools/Domain/domain_axioms.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOLCF/Tools/Domain/domain_axioms.ML Fri Dec 04 15:20:24 2009 +0100 @@ -38,10 +38,10 @@ (@{type_name "u"}, @{const_name "u_map"})]; fun copy_of_dtyp tab r dt = - if DatatypeAux.is_rec_type dt then copy tab r dt else ID -and copy tab r (DatatypeAux.DtRec i) = r i - | copy tab r (DatatypeAux.DtTFree a) = ID - | copy tab r (DatatypeAux.DtType (c, ds)) = + if Datatype_Aux.is_rec_type dt then copy tab r dt else ID +and copy tab r (Datatype_Aux.DtRec i) = r i + | copy tab r (Datatype_Aux.DtTFree a) = ID + | copy tab r (Datatype_Aux.DtType (c, ds)) = case Symtab.lookup tab c of SOME f => list_ccomb (%%:f, map (copy_of_dtyp tab r) ds) | NONE => (warning ("copy_of_dtyp: unknown type constructor " ^ c); ID); diff -r 651028e34b5d -r 01dcd9b926bf src/HOLCF/Tools/Domain/domain_extender.ML --- a/src/HOLCF/Tools/Domain/domain_extender.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOLCF/Tools/Domain/domain_extender.ML Fri Dec 04 15:20:24 2009 +0100 @@ -153,7 +153,7 @@ val thy' = thy'' |> Domain_Syntax.add_syntax false comp_dnam eqs'; val dts = map (Type o fst) eqs'; val new_dts = map (fn ((s,Ts),_) => (s, map (fst o dest_TFree) Ts)) eqs'; - fun strip ss = Library.drop (find_index (fn s => s = "'") ss + 1, ss); + fun strip ss = drop (find_index (fn s => s = "'") ss + 1) ss; fun typid (Type (id,_)) = let val c = hd (Symbol.explode (Long_Name.base_name id)) in if Symbol.is_letter c then c else "t" end @@ -162,7 +162,7 @@ fun one_con (con,args,mx) = ((Syntax.const_name mx (Binding.name_of con)), ListPair.map (fn ((lazy,sel,tp),vn) => - mk_arg ((lazy, DatatypeAux.dtyp_of_typ new_dts tp), + mk_arg ((lazy, Datatype_Aux.dtyp_of_typ new_dts tp), Option.map Binding.name_of sel,vn)) (args,(mk_var_names(map (typid o third) args))) ) : cons; @@ -228,7 +228,7 @@ val thy' = thy'' |> Domain_Syntax.add_syntax true comp_dnam eqs'; val dts = map (Type o fst) eqs'; val new_dts = map (fn ((s,Ts),_) => (s, map (fst o dest_TFree) Ts)) eqs'; - fun strip ss = Library.drop (find_index (fn s => s = "'") ss + 1, ss); + fun strip ss = drop (find_index (fn s => s = "'") ss + 1) ss; fun typid (Type (id,_)) = let val c = hd (Symbol.explode (Long_Name.base_name id)) in if Symbol.is_letter c then c else "t" end @@ -237,7 +237,7 @@ fun one_con (con,args,mx) = ((Syntax.const_name mx (Binding.name_of con)), ListPair.map (fn ((lazy,sel,tp),vn) => - mk_arg ((lazy, DatatypeAux.dtyp_of_typ new_dts tp), + mk_arg ((lazy, Datatype_Aux.dtyp_of_typ new_dts tp), Option.map Binding.name_of sel,vn)) (args,(mk_var_names(map (typid o third) args))) ) : cons; diff -r 651028e34b5d -r 01dcd9b926bf src/HOLCF/Tools/Domain/domain_isomorphism.ML --- a/src/HOLCF/Tools/Domain/domain_isomorphism.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOLCF/Tools/Domain/domain_isomorphism.ML Fri Dec 04 15:20:24 2009 +0100 @@ -581,12 +581,12 @@ mk_fst t :: mk_copy_args xs (mk_snd t); in mk_copy_args doms copy_arg end; fun copy_of_dtyp (T, dt) = - if DatatypeAux.is_rec_type dt + if Datatype_Aux.is_rec_type dt then copy_of_dtyp' (T, dt) else ID_const T - and copy_of_dtyp' (T, DatatypeAux.DtRec i) = nth copy_args i - | copy_of_dtyp' (T, DatatypeAux.DtTFree a) = ID_const T - | copy_of_dtyp' (T as Type (_, Ts), DatatypeAux.DtType (c, ds)) = + and copy_of_dtyp' (T, Datatype_Aux.DtRec i) = nth copy_args i + | copy_of_dtyp' (T, Datatype_Aux.DtTFree a) = ID_const T + | copy_of_dtyp' (T as Type (_, Ts), Datatype_Aux.DtType (c, ds)) = case Symtab.lookup map_tab' c of SOME f => Library.foldl mk_capply @@ -599,7 +599,7 @@ val copy_bind = Binding.suffix_name "_copy" tbind; val (copy_const, thy) = thy |> Sign.declare_const ((copy_bind, copy_type), NoSyn); - val dtyp = DatatypeAux.dtyp_of_typ new_dts rhsT; + val dtyp = Datatype_Aux.dtyp_of_typ new_dts rhsT; val body = copy_of_dtyp (rhsT, dtyp); val comp = mk_cfcomp (abs_const, mk_cfcomp (body, rep_const)); val rhs = big_lambda copy_arg comp; diff -r 651028e34b5d -r 01dcd9b926bf src/HOLCF/Tools/Domain/domain_library.ML --- a/src/HOLCF/Tools/Domain/domain_library.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOLCF/Tools/Domain/domain_library.ML Fri Dec 04 15:20:24 2009 +0100 @@ -230,7 +230,7 @@ val mk_arg = I; fun rec_of ((_,dtyp),_,_) = - case dtyp of DatatypeAux.DtRec i => i | _ => ~1; + case dtyp of Datatype_Aux.DtRec i => i | _ => ~1; (* FIXME: what about indirect recursion? *) fun is_lazy arg = fst (first arg); @@ -246,12 +246,12 @@ (* ----- combinators for making dtyps ----- *) -fun mk_uD T = DatatypeAux.DtType(@{type_name "u"}, [T]); -fun mk_sprodD (T, U) = DatatypeAux.DtType(@{type_name "**"}, [T, U]); -fun mk_ssumD (T, U) = DatatypeAux.DtType(@{type_name "++"}, [T, U]); -fun mk_liftD T = DatatypeAux.DtType(@{type_name "lift"}, [T]); -val unitD = DatatypeAux.DtType(@{type_name "unit"}, []); -val boolD = DatatypeAux.DtType(@{type_name "bool"}, []); +fun mk_uD T = Datatype_Aux.DtType(@{type_name "u"}, [T]); +fun mk_sprodD (T, U) = Datatype_Aux.DtType(@{type_name "**"}, [T, U]); +fun mk_ssumD (T, U) = Datatype_Aux.DtType(@{type_name "++"}, [T, U]); +fun mk_liftD T = Datatype_Aux.DtType(@{type_name "lift"}, [T]); +val unitD = Datatype_Aux.DtType(@{type_name "unit"}, []); +val boolD = Datatype_Aux.DtType(@{type_name "bool"}, []); val oneD = mk_liftD unitD; val trD = mk_liftD boolD; fun big_sprodD ds = case ds of [] => oneD | _ => foldr1 mk_sprodD ds; diff -r 651028e34b5d -r 01dcd9b926bf src/HOLCF/Tools/Domain/domain_theorems.ML --- a/src/HOLCF/Tools/Domain/domain_theorems.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/HOLCF/Tools/Domain/domain_theorems.ML Fri Dec 04 15:20:24 2009 +0100 @@ -600,7 +600,7 @@ let val lhs = dc_copy`%"f"`(con_app con args); fun one_rhs arg = - if DatatypeAux.is_rec_type (dtyp_of arg) + if Datatype_Aux.is_rec_type (dtyp_of arg) then Domain_Axioms.copy_of_dtyp map_tab (proj (%:"f") eqs) (dtyp_of arg) ` (%# arg) else (%# arg); @@ -730,7 +730,7 @@ let fun mk_take n = dc_take (List.nth (dnames, n)) $ %:"n"; fun one_rhs arg = - if DatatypeAux.is_rec_type (dtyp_of arg) + if Datatype_Aux.is_rec_type (dtyp_of arg) then Domain_Axioms.copy_of_dtyp map_tab mk_take (dtyp_of arg) ` (%# arg) else (%# arg); diff -r 651028e34b5d -r 01dcd9b926bf src/Provers/splitter.ML --- a/src/Provers/splitter.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Provers/splitter.ML Fri Dec 04 15:20:24 2009 +0100 @@ -122,8 +122,8 @@ fun down [] t i = Bound 0 | down (p::ps) t i = let val (h,ts) = strip_comb t - val v1 = ListPair.map var (Library.take(p,ts), i upto (i+p-1)) - val u::us = Library.drop(p,ts) + val v1 = ListPair.map var (take p ts, i upto (i+p-1)) + val u::us = drop p ts val v2 = ListPair.map var (us, (i+p) upto (i+length(ts)-2)) in list_comb(h,v1@[down ps u (i+length ts)]@v2) end; in Abs("", T, down (rev pos) t maxi) end; @@ -191,7 +191,7 @@ fun mk_split_pack (thm, T: typ, T', n, ts, apsns, pos, TB, t) = if n > length ts then [] else let val lev = length apsns - val lbnos = fold add_lbnos (Library.take (n, ts)) [] + val lbnos = fold add_lbnos (take n ts) [] val flbnos = filter (fn i => i < lev) lbnos val tt = incr_boundvars (~lev) t in if null flbnos then @@ -259,7 +259,7 @@ Const(c, cT) => let fun find [] = [] | find ((gcT, pat, thm, T, n)::tups) = - let val t2 = list_comb (h, Library.take (n, ts)) + let val t2 = list_comb (h, take n ts) in if Sign.typ_instance sg (cT, gcT) andalso fomatch sg (Ts,pat,t2) then mk_split_pack(thm,T,T',n,ts,apsns,pos,type_of1(Ts,t2),t2) diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/General/path.ML --- a/src/Pure/General/path.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/General/path.ML Fri Dec 04 15:20:24 2009 +0100 @@ -139,7 +139,7 @@ val split_ext = split_path (fn (prfx, s) => apfst (append (Path prfx)) (case take_suffix (fn c => c <> ".") (explode s) of ([], _) => (Path [Basic s], "") - | (cs, e) => (Path [Basic (implode (Library.take (length cs - 1, cs)))], implode e))); + | (cs, e) => (Path [Basic (implode (take (length cs - 1) cs))], implode e))); (* expand variables *) diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/General/scan.ML --- a/src/Pure/General/scan.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/General/scan.ML Fri Dec 04 15:20:24 2009 +0100 @@ -221,7 +221,7 @@ fun trace scan xs = let val (y, xs') = scan xs - in ((y, Library.take (length xs - length xs', xs)), xs') end; + in ((y, take (length xs - length xs') xs), xs') end; (* stopper *) diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/General/symbol.ML --- a/src/Pure/General/symbol.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/General/symbol.ML Fri Dec 04 15:20:24 2009 +0100 @@ -183,7 +183,7 @@ val raw1 = raw0 o implode; val raw2 = enclose "\\<^raw" ">" o string_of_int o ord; - fun encode cs = enc (Library.take_prefix raw_chr cs) + fun encode cs = enc (take_prefix raw_chr cs) and enc ([], []) = [] | enc (cs, []) = [raw1 cs] | enc ([], d :: ds) = raw2 d :: encode ds @@ -198,11 +198,11 @@ fun beginning n cs = let - val drop_blanks = #1 o Library.take_suffix is_ascii_blank; + val drop_blanks = #1 o take_suffix is_ascii_blank; val all_cs = drop_blanks cs; val dots = if length all_cs > n then " ..." else ""; in - (drop_blanks (Library.take (n, all_cs)) + (drop_blanks (take n all_cs) |> map (fn c => if is_ascii_blank c then space else c) |> implode) ^ dots end; @@ -491,8 +491,8 @@ fun strip_blanks s = sym_explode s - |> Library.take_prefix is_blank |> #2 - |> Library.take_suffix is_blank |> #1 + |> take_prefix is_blank |> #2 + |> take_suffix is_blank |> #1 |> implode; @@ -516,7 +516,7 @@ then chr (ord s + 1) :: ss else "a" :: s :: ss; - val (ss, qs) = apfst rev (Library.take_suffix is_quasi (sym_explode str)); + val (ss, qs) = apfst rev (take_suffix is_quasi (sym_explode str)); val ss' = if symbolic_end ss then "'" :: ss else bump ss; in implode (rev ss' @ qs) end; diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/Isar/class_target.ML --- a/src/Pure/Isar/class_target.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/Isar/class_target.ML Fri Dec 04 15:20:24 2009 +0100 @@ -278,7 +278,7 @@ fun improve (c, ty) = (case AList.lookup (op =) primary_constraints c of SOME ty' => (case try (Type.raw_match (ty', ty)) Vartab.empty of SOME tyenv => (case Vartab.lookup tyenv (Name.aT, 0) - of SOME (_, ty' as TVar (tvar as (vi, sort))) => + of SOME (_, ty' as TVar (vi, sort)) => if TypeInfer.is_param vi andalso Sorts.sort_le algebra (base_sort, sort) then SOME (ty', TFree (Name.aT, base_sort)) @@ -434,9 +434,10 @@ fun synchronize_inst_syntax ctxt = let - val Instantiation { arities = (_, _, sort), params = params } = Instantiation.get ctxt; + val Instantiation { params, ... } = Instantiation.get ctxt; - val lookup_inst_param = AxClass.lookup_inst_param (Sign.consts_of (ProofContext.theory_of ctxt)) params; + val lookup_inst_param = AxClass.lookup_inst_param + (Sign.consts_of (ProofContext.theory_of ctxt)) params; fun subst (c, ty) = case lookup_inst_param (c, ty) of SOME (v_ty as (_, ty)) => SOME (ty, Free v_ty) | NONE => NONE; @@ -502,7 +503,7 @@ val consts = Sign.consts_of thy; val improve_constraints = AList.lookup (op =) (map (fn (_, (class, (c, _))) => (c, [[class]])) class_params); - fun resort_check ts ctxt = case resort_terms pp algebra consts improve_constraints ts + fun resort_check ts ctxt = case resort_terms (Syntax.pp ctxt) algebra consts improve_constraints ts of NONE => NONE | SOME ts' => SOME (ts', ctxt); val lookup_inst_param = AxClass.lookup_inst_param consts params; @@ -563,8 +564,7 @@ fun conclude_instantiation lthy = let - val { arities, params } = the_instantiation lthy; - val (tycos, vs, sort) = arities; + val (tycos, vs, sort) = (#arities o the_instantiation) lthy; val thy = ProofContext.theory_of lthy; val _ = map (fn tyco => if Sign.of_sort thy (Type (tyco, map TFree vs), sort) @@ -574,8 +574,7 @@ fun pretty_instantiation lthy = let - val { arities, params } = the_instantiation lthy; - val (tycos, vs, sort) = arities; + val { arities = (tycos, vs, sort), params } = the_instantiation lthy; val thy = ProofContext.theory_of lthy; fun pr_arity tyco = Syntax.pretty_arity lthy (tyco, map snd vs, sort); fun pr_param ((c, _), (v, ty)) = diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/Isar/code.ML --- a/src/Pure/Isar/code.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/Isar/code.ML Fri Dec 04 15:20:24 2009 +0100 @@ -138,7 +138,7 @@ val args = args_of thm; val incr_idx = Logic.incr_indexes ([], Thm.maxidx_of thm + 1); fun matches_args args' = length args <= length args' andalso - Pattern.matchess thy (args, (map incr_idx o curry Library.take (length args)) args'); + Pattern.matchess thy (args, (map incr_idx o take (length args)) args'); fun drop (thm', proper') = if (proper orelse not proper') andalso matches_args (args_of thm') then (warning ("Code generator: dropping redundant code equation\n" ^ @@ -256,7 +256,7 @@ (Symtab.empty, (Symtab.empty, Symtab.empty)))), Unsynchronized.ref empty_data); fun copy (spec, data) = (spec, Unsynchronized.ref (! data)); val extend = copy; - fun merge _ ((spec1, data1), (spec2, data2)) = + fun merge _ ((spec1, _), (spec2, _)) = (merge_spec (spec1, spec2), Unsynchronized.ref empty_data); ); diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/Isar/obtain.ML --- a/src/Pure/Isar/obtain.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/Isar/obtain.ML Fri Dec 04 15:20:24 2009 +0100 @@ -232,12 +232,12 @@ err ("Failed to unify variable " ^ string_of_term (Free (x, Envir.norm_type tyenv T)) ^ " against parameter " ^ string_of_term (Syntax.mark_boundT (y, Envir.norm_type tyenv U)) ^ " in") rule; - val (tyenv, _) = fold unify (map #1 vars ~~ Library.take (m, params)) + val (tyenv, _) = fold unify (map #1 vars ~~ take m params) (Vartab.empty, Int.max (maxidx, Thm.maxidx_of rule)); val norm_type = Envir.norm_type tyenv; val xs = map (apsnd norm_type o fst) vars; - val ys = map (apsnd norm_type) (Library.drop (m, params)); + val ys = map (apsnd norm_type) (drop m params); val ys' = map Name.internal (Name.variant_list (map fst xs) (map fst ys)) ~~ map #2 ys; val terms = map (Drule.mk_term o cert o Free) (xs @ ys'); diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/Isar/proof_context.ML --- a/src/Pure/Isar/proof_context.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/Isar/proof_context.ML Fri Dec 04 15:20:24 2009 +0100 @@ -1340,7 +1340,7 @@ val suppressed = len - ! prems_limit; val prt_prems = if null prems then [] else [Pretty.big_list "prems:" ((if suppressed <= 0 then [] else [Pretty.str "..."]) @ - map (Display.pretty_thm ctxt) (Library.drop (suppressed, prems)))]; + map (Display.pretty_thm ctxt) (drop suppressed prems))]; in prt_structs @ prt_fixes @ prt_prems end; diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/Isar/rule_cases.ML --- a/src/Pure/Isar/rule_cases.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/Isar/rule_cases.ML Fri Dec 04 15:20:24 2009 +0100 @@ -144,7 +144,7 @@ (Option.map (curry Logic.nth_prem i) rule_struct, Logic.nth_prem (i, prop))) () of NONE => (name, NONE) | SOME p => (name, extract_case is_open thy p name concls)) :: cs, i - 1); - in fold_rev add_case (Library.drop (n - nprems, cases)) ([], n) |> #1 end; + in fold_rev add_case (drop (n - nprems) cases) ([], n) |> #1 end; in @@ -205,8 +205,8 @@ th |> unfold_prems n defs |> unfold_prems_concls defs - |> Drule.multi_resolve (Library.take (m, facts)) - |> Seq.map (pair (xx, (n - m, Library.drop (m, facts)))) + |> Drule.multi_resolve (take m facts) + |> Seq.map (pair (xx, (n - m, drop m facts))) end; end; @@ -345,7 +345,7 @@ fun prep_rule n th = let val th' = Thm.permute_prems 0 n th; - val prems = Library.take (Thm.nprems_of th' - n, Drule.cprems_of th'); + val prems = take (Thm.nprems_of th' - n) (Drule.cprems_of th'); val th'' = Drule.implies_elim_list th' (map Thm.assume prems); in (prems, (n, th'')) end; diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/Proof/extraction.ML --- a/src/Pure/Proof/extraction.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/Proof/extraction.ML Fri Dec 04 15:20:24 2009 +0100 @@ -461,7 +461,7 @@ end else (vs', tye) - in fold_rev add_args (Library.take (n, vars) ~~ Library.take (n, ts)) ([], []) end; + in fold_rev add_args (take n vars ~~ take n ts) ([], []) end; fun find (vs: string list) = Option.map snd o find_first (curry (eq_set (op =)) vs o fst); fun find' (s: string) = map_filter (fn (s', x) => if s = s' then SOME x else NONE); diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/Syntax/ast.ML --- a/src/Pure/Syntax/ast.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/Syntax/ast.ML Fri Dec 04 15:20:24 2009 +0100 @@ -158,7 +158,7 @@ (case (ast, pat) of (Appl asts, Appl pats) => let val a = length asts and p = length pats in - if a > p then (Appl (Library.take (p, asts)), Library.drop (p, asts)) + if a > p then (Appl (take p asts), drop p asts) else (ast, []) end | _ => (ast, [])); diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/Syntax/syntax.ML --- a/src/Pure/Syntax/syntax.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/Syntax/syntax.ML Fri Dec 04 15:20:24 2009 +0100 @@ -499,7 +499,7 @@ (("Ambiguous input" ^ Position.str_of pos ^ "\nproduces " ^ string_of_int len ^ " parse trees" ^ (if len <= limit then "" else " (" ^ string_of_int limit ^ " displayed)") ^ ":") :: - map show_pt (Library.take (limit, pts))))); + map show_pt (take limit pts)))); SynTrans.pts_to_asts ctxt (lookup_tr parse_ast_trtab) pts end; @@ -545,7 +545,7 @@ else cat_error (ambig_msg ()) (cat_lines (("Ambiguous input, " ^ string_of_int len ^ " terms are type correct" ^ (if len <= limit then "" else " (" ^ string_of_int limit ^ " displayed)") ^ ":") :: - map (Pretty.string_of_term pp) (Library.take (limit, results)))) + map (Pretty.string_of_term pp) (take limit results))) end; fun standard_parse_term pp check get_sort map_const map_free map_type map_sort diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/Thy/present.ML --- a/src/Pure/Thy/present.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/Thy/present.ML Fri Dec 04 15:20:24 2009 +0100 @@ -285,7 +285,7 @@ (browser_info := empty_browser_info; session_info := NONE) else let - val parent_name = name_of_session (Library.take (length path - 1, path)); + val parent_name = name_of_session (take (length path - 1) path); val session_name = name_of_session path; val sess_prefix = Path.make path; val html_prefix = Path.append (Path.expand output_path) sess_prefix; diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/Thy/thy_output.ML --- a/src/Pure/Thy/thy_output.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/Thy/thy_output.ML Fri Dec 04 15:20:24 2009 +0100 @@ -265,7 +265,7 @@ if nesting = 0 andalso not (Toplevel.is_proof state) then tag_stack else if nesting >= 0 then (tag', replicate nesting tag @ tags) else - (case Library.drop (~ nesting - 1, tags) of + (case drop (~ nesting - 1) tags of tgs :: tgss => (tgs, tgss) | [] => err_bad_nesting (Position.str_of cmd_pos)); diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/Tools/find_theorems.ML --- a/src/Pure/Tools/find_theorems.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/Tools/find_theorems.ML Fri Dec 04 15:20:24 2009 +0100 @@ -409,7 +409,7 @@ val len = length matches; val lim = the_default (! limit) opt_limit; - in (SOME len, Library.drop (len - lim, matches)) end; + in (SOME len, drop (len - lim) matches) end; val find = if rem_dups orelse is_none opt_limit diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/assumption.ML --- a/src/Pure/assumption.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/assumption.ML Fri Dec 04 15:20:24 2009 +0100 @@ -88,12 +88,12 @@ (* local assumptions *) fun local_assumptions_of inner outer = - Library.drop (length (all_assumptions_of outer), all_assumptions_of inner); + drop (length (all_assumptions_of outer)) (all_assumptions_of inner); val local_assms_of = maps #2 oo local_assumptions_of; fun local_prems_of inner outer = - Library.drop (length (all_prems_of outer), all_prems_of inner); + drop (length (all_prems_of outer)) (all_prems_of inner); (* add assumptions *) diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/axclass.ML --- a/src/Pure/axclass.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/axclass.ML Fri Dec 04 15:20:24 2009 +0100 @@ -257,12 +257,11 @@ | NONE => NONE; fun unoverload_const thy (c_ty as (c, _)) = - case class_of_param thy c - of SOME class (* FIXME unused? *) => - (case get_inst_tyco (Sign.consts_of thy) c_ty - of SOME tyco => try (param_of_inst thy) (c, tyco) |> the_default c - | NONE => c) - | NONE => c; + if is_some (class_of_param thy c) + then case get_inst_tyco (Sign.consts_of thy) c_ty + of SOME tyco => try (param_of_inst thy) (c, tyco) |> the_default c + | NONE => c + else c; (** instances **) diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/codegen.ML --- a/src/Pure/codegen.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/codegen.ML Fri Dec 04 15:20:24 2009 +0100 @@ -579,11 +579,11 @@ fun eta_expand t ts i = let val k = length ts; - val Ts = Library.drop (k, binder_types (fastype_of t)); + val Ts = drop k (binder_types (fastype_of t)); val j = i - k in List.foldr (fn (T, t) => Abs ("x", T, t)) - (list_comb (t, ts @ map Bound (j-1 downto 0))) (Library.take (j, Ts)) + (list_comb (t, ts @ map Bound (j-1 downto 0))) (take j Ts) end; fun mk_app _ p [] = p diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/goal_display.ML --- a/src/Pure/goal_display.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/goal_display.ML Fri Dec 04 15:20:24 2009 +0100 @@ -99,7 +99,7 @@ (if main then [prt_term B] else []) @ (if ngoals = 0 then [Pretty.str "No subgoals!"] else if ngoals > maxgoals then - pretty_subgoals (Library.take (maxgoals, As)) @ + pretty_subgoals (take maxgoals As) @ (if total then [Pretty.str ("A total of " ^ string_of_int ngoals ^ " subgoals...")] else []) else pretty_subgoals As) @ diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/library.ML --- a/src/Pure/library.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/library.ML Fri Dec 04 15:20:24 2009 +0100 @@ -81,6 +81,8 @@ val filter: ('a -> bool) -> 'a list -> 'a list val filter_out: ('a -> bool) -> 'a list -> 'a list val map_filter: ('a -> 'b option) -> 'a list -> 'b list + val take: int -> 'a list -> 'a list + val drop: int -> 'a list -> 'a list val chop: int -> 'a list -> 'a list * 'a list val dropwhile: ('a -> bool) -> 'a list -> 'a list val nth: 'a list -> int -> 'a @@ -224,8 +226,6 @@ val foldl: ('a * 'b -> 'a) -> 'a * 'b list -> 'a val foldr: ('a * 'b -> 'b) -> 'a list * 'b -> 'b val foldl_map: ('a * 'b -> 'a * 'c) -> 'a * 'b list -> 'a * 'c list - val take: int * 'a list -> 'a list - val drop: int * 'a list -> 'a list val last_elem: 'a list -> 'a end; @@ -423,20 +423,20 @@ fun filter_out f = filter (not o f); val map_filter = List.mapPartial; +fun take (0: int) xs = [] + | take _ [] = [] + | take n (x :: xs) = + if n > 0 then x :: take (n - 1) xs else []; + +fun drop (0: int) xs = xs + | drop _ [] = [] + | drop n (x :: xs) = + if n > 0 then drop (n - 1) xs else x :: xs; + fun chop (0: int) xs = ([], xs) | chop _ [] = ([], []) | chop n (x :: xs) = chop (n - 1) xs |>> cons x; -(*take the first n elements from a list*) -fun take (n: int, []) = [] - | take (n, x :: xs) = - if n > 0 then x :: take (n - 1, xs) else []; - -(*drop the first n elements from a list*) -fun drop (n: int, []) = [] - | drop (n, x :: xs) = - if n > 0 then drop (n - 1, xs) else x :: xs; - fun dropwhile P [] = [] | dropwhile P (ys as x::xs) = if P x then dropwhile P xs else ys; diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/meta_simplifier.ML --- a/src/Pure/meta_simplifier.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/meta_simplifier.ML Fri Dec 04 15:20:24 2009 +0100 @@ -1163,9 +1163,9 @@ val (rrs', asm') = rules_of_prem ss prem' in mut_impc prems concl rrss asms (prem' :: prems') (rrs' :: rrss') (asm' :: asms') (SOME (fold_rev (disch true) - (Library.take (i, prems)) + (take i prems) (Drule.imp_cong_rule eqn (reflexive (Drule.list_implies - (Library.drop (i, prems), concl))))) :: eqns) + (drop i prems, concl))))) :: eqns) ss (length prems') ~1 end diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/proofterm.ML --- a/src/Pure/proofterm.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/proofterm.ML Fri Dec 04 15:20:24 2009 +0100 @@ -1003,7 +1003,7 @@ | _ => error "shrink: proof not in normal form"); val vs = vars_of prop; val (ts', ts'') = chop (length vs) ts; - val insts = Library.take (length ts', map (fst o dest_Var) vs) ~~ ts'; + val insts = take (length ts') (map (fst o dest_Var) vs) ~~ ts'; val nvs = Library.foldl (fn (ixns', (ixn, ixns)) => insert (op =) ixn (case AList.lookup (op =) insts ixn of SOME (SOME t) => if is_proj t then union (op =) ixns ixns' else ixns' diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/tactic.ML --- a/src/Pure/tactic.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/tactic.ML Fri Dec 04 15:20:24 2009 +0100 @@ -173,7 +173,7 @@ perm_tac 0 (1 - i); fun distinct_subgoal_tac i st = - (case Library.drop (i - 1, Thm.prems_of st) of + (case drop (i - 1) (Thm.prems_of st) of [] => no_tac st | A :: Bs => st |> EVERY (fold (fn (B, k) => diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/thm.ML --- a/src/Pure/thm.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/thm.ML Fri Dec 04 15:20:24 2009 +0100 @@ -1457,7 +1457,7 @@ val short = length iparams - length cs; val newnames = if short < 0 then error "More names than abstractions!" - else Name.variant_list cs (Library.take (short, iparams)) @ cs; + else Name.variant_list cs (take short iparams) @ cs; val freenames = Term.fold_aterms (fn Free (x, _) => insert (op =) x | _ => I) Bi []; val newBi = Logic.list_rename_params (newnames, Bi); in diff -r 651028e34b5d -r 01dcd9b926bf src/Pure/variable.ML --- a/src/Pure/variable.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Pure/variable.ML Fri Dec 04 15:20:24 2009 +0100 @@ -345,7 +345,7 @@ val fixes_inner = fixes_of inner; val fixes_outer = fixes_of outer; - val gen_fixes = map #2 (Library.take (length fixes_inner - length fixes_outer, fixes_inner)); + val gen_fixes = map #2 (take (length fixes_inner - length fixes_outer) fixes_inner); val still_fixed = not o member (op =) gen_fixes; val type_occs_inner = type_occs_of inner; diff -r 651028e34b5d -r 01dcd9b926bf src/Tools/Code/code_haskell.ML --- a/src/Tools/Code/code_haskell.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Tools/Code/code_haskell.ML Fri Dec 04 15:20:24 2009 +0100 @@ -42,7 +42,7 @@ (map (str o lookup_var tyvars) vnames) @ str "." @@ Pretty.brk 1; fun pr_tycoexpr tyvars fxy (tyco, tys) = brackify fxy (str tyco :: map (pr_typ tyvars BR) tys) - and pr_typ tyvars fxy (tycoexpr as tyco `%% tys) = (case syntax_tyco tyco + and pr_typ tyvars fxy (tyco `%% tys) = (case syntax_tyco tyco of NONE => pr_tycoexpr tyvars fxy (deresolve tyco, tys) | SOME (i, pr) => pr (pr_typ tyvars) fxy tys) | pr_typ tyvars fxy (ITyVar v) = (str o lookup_var tyvars) v; @@ -78,7 +78,7 @@ and pr_app' tyvars thm vars ((c, (_, tys)), ts) = case contr_classparam_typs c of [] => (str o deresolve) c :: map (pr_term tyvars thm vars BR) ts | fingerprint => let - val ts_fingerprint = ts ~~ curry Library.take (length ts) fingerprint; + val ts_fingerprint = ts ~~ take (length ts) fingerprint; val needs_annotation = forall (fn (_, NONE) => true | (t, SOME _) => (not o Code_Thingol.locally_monomorphic) t) ts_fingerprint; fun pr_term_anno (t, NONE) _ = pr_term tyvars thm vars BR t @@ -86,7 +86,7 @@ brackets [pr_term tyvars thm vars NOBR t, str "::", pr_typ tyvars NOBR ty]; in if needs_annotation then - (str o deresolve) c :: map2 pr_term_anno ts_fingerprint (curry Library.take (length ts) tys) + (str o deresolve) c :: map2 pr_term_anno ts_fingerprint (take (length ts) tys) else (str o deresolve) c :: map (pr_term tyvars thm vars BR) ts end and pr_app tyvars = gen_pr_app (pr_app' tyvars) (pr_term tyvars) syntax_const diff -r 651028e34b5d -r 01dcd9b926bf src/Tools/Code/code_printer.ML --- a/src/Tools/Code/code_printer.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Tools/Code/code_printer.ML Fri Dec 04 15:20:24 2009 +0100 @@ -231,7 +231,7 @@ of NONE => brackify fxy (pr_app thm vars app) | SOME (k, pr) => let - fun pr' fxy ts = pr (pr_term thm) thm vars fxy (ts ~~ curry Library.take k tys); + fun pr' fxy ts = pr (pr_term thm) thm vars fxy (ts ~~ take k tys); in if k = length ts then pr' fxy ts else if k < length ts diff -r 651028e34b5d -r 01dcd9b926bf src/Tools/Code/code_target.ML --- a/src/Tools/Code/code_target.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Tools/Code/code_target.ML Fri Dec 04 15:20:24 2009 +0100 @@ -46,6 +46,7 @@ val add_syntax_tyco: string -> string -> tyco_syntax option -> theory -> theory val add_syntax_const: string -> string -> proto_const_syntax option -> theory -> theory val add_reserved: string -> string -> theory -> theory + val add_include: string -> string * (string * string list) option -> theory -> theory end; structure Code_Target : CODE_TARGET = @@ -165,10 +166,9 @@ val abort_allowed = snd o CodeTargetData.get; -fun assert_target thy target = - case Symtab.lookup (fst (CodeTargetData.get thy)) target - of SOME data => target - | NONE => error ("Unknown code target language: " ^ quote target); +fun assert_target thy target = if Symtab.defined (fst (CodeTargetData.get thy)) target + then target + else error ("Unknown code target language: " ^ quote target); fun put_target (target, seri) thy = let @@ -238,7 +238,7 @@ fun read_tyco thy = cert_tyco thy o Sign.intern_type thy; -fun gen_add_syntax_class prep_class prep_const target raw_class raw_syn thy = +fun gen_add_syntax_class prep_class target raw_class raw_syn thy = let val class = prep_class thy raw_class; in case raw_syn @@ -318,7 +318,7 @@ | add (name, NONE) incls = Symtab.delete name incls; in map_includes target (add args) thy end; -val add_include = gen_add_include Code.check_const; +val add_include = gen_add_include (K I); val add_include_cmd = gen_add_include Code.read_const; fun add_module_alias target (thyname, modlname) = @@ -355,14 +355,15 @@ in -val add_syntax_class = gen_add_syntax_class cert_class (K I); +val add_syntax_class = gen_add_syntax_class cert_class; val add_syntax_inst = gen_add_syntax_inst cert_class cert_tyco; val add_syntax_tyco = gen_add_syntax_tyco cert_tyco; val add_syntax_const = gen_add_syntax_const (K I); val allow_abort = gen_allow_abort (K I); val add_reserved = add_reserved; +val add_include = add_include; -val add_syntax_class_cmd = gen_add_syntax_class read_class Code.read_const; +val add_syntax_class_cmd = gen_add_syntax_class read_class; val add_syntax_inst_cmd = gen_add_syntax_inst read_class read_tyco; val add_syntax_tyco_cmd = gen_add_syntax_tyco read_tyco; val add_syntax_const_cmd = gen_add_syntax_const Code.read_const; @@ -606,7 +607,7 @@ ((filter OuterLex.is_proper o OuterSyntax.scan Position.none) cmd) of SOME f => (writeln "Now generating code..."; f (theory thyname)) | NONE => error ("Bad directive " ^ quote cmd))) - handle TOPLEVEL_ERROR => OS.Process.exit OS.Process.failure; + handle Runtime.TOPLEVEL_ERROR => OS.Process.exit OS.Process.failure; end; (*local*) diff -r 651028e34b5d -r 01dcd9b926bf src/Tools/Code/code_thingol.ML --- a/src/Tools/Code/code_thingol.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Tools/Code/code_thingol.ML Fri Dec 04 15:20:24 2009 +0100 @@ -226,7 +226,7 @@ val l = k - j; val ctxt = (fold o fold_varnames) Name.declare ts Name.context; val vs_tys = (map o apfst) SOME - (Name.names ctxt "a" ((curry Library.take l o curry Library.drop j) tys)); + (Name.names ctxt "a" ((take l o drop j) tys)); in vs_tys `|==> IConst c `$$ ts @ map (IVar o fst) vs_tys end; fun contains_dictvar t = @@ -773,7 +773,7 @@ val clauses = if null case_pats then mk_clause (fn ([t], body) => (t, body)) [ty] (the_single ts_clause) else maps (fn ((constr as IConst (_, (_, tys)), n), t) => - mk_clause (fn (ts, body) => (constr `$$ ts, body)) (curry Library.take n tys) t) + mk_clause (fn (ts, body) => (constr `$$ ts, body)) (take n tys) t) (constrs ~~ ts_clause); in ((t, ty), clauses) end; in @@ -788,7 +788,7 @@ if length ts < num_args then let val k = length ts; - val tys = (curry Library.take (num_args - k) o curry Library.drop k o fst o strip_type) ty; + val tys = (take (num_args - k) o drop k o fst o strip_type) ty; val ctxt = (fold o fold_aterms) Term.declare_term_frees ts Name.context; val vs = Name.names ctxt "a" tys; in @@ -797,8 +797,8 @@ #>> (fn (tys, t) => map2 (fn (v, _) => pair (SOME v)) vs tys `|==> t) end else if length ts > num_args then - translate_case thy algbr eqngr thm case_scheme ((c, ty), Library.take (num_args, ts)) - ##>> fold_map (translate_term thy algbr eqngr thm) (Library.drop (num_args, ts)) + translate_case thy algbr eqngr thm case_scheme ((c, ty), take num_args ts) + ##>> fold_map (translate_term thy algbr eqngr thm) (drop num_args ts) #>> (fn (t, ts) => t `$$ ts) else translate_case thy algbr eqngr thm case_scheme ((c, ty), ts) @@ -930,10 +930,7 @@ ((snd o #constants o Consts.dest o #consts o Sign.rep_sg) thy') []; fun belongs_here c = not (exists (fn thy'' => Sign.declared_const thy'' c) (Theory.parents_of thy')) - in case some_thyname - of NONE => cs - | SOME thyname => filter belongs_here cs - end; + in if is_some some_thyname then cs else filter belongs_here cs end; fun read_const_expr "*" = ([], consts_of NONE) | read_const_expr s = if String.isSuffix ".*" s then ([], consts_of (SOME (unsuffix ".*" s))) diff -r 651028e34b5d -r 01dcd9b926bf src/Tools/induct.ML --- a/src/Tools/induct.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Tools/induct.ML Fri Dec 04 15:20:24 2009 +0100 @@ -280,11 +280,11 @@ fun align_left msg xs ys = let val m = length xs and n = length ys - in if m < n then error msg else (Library.take (n, xs) ~~ ys) end; + in if m < n then error msg else (take n xs ~~ ys) end; fun align_right msg xs ys = let val m = length xs and n = length ys - in if m < n then error msg else (Library.drop (m - n, xs) ~~ ys) end; + in if m < n then error msg else (drop (m - n) xs ~~ ys) end; (* prep_inst *) diff -r 651028e34b5d -r 01dcd9b926bf src/Tools/induct_tacs.ML --- a/src/Tools/induct_tacs.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/Tools/induct_tacs.ML Fri Dec 04 15:20:24 2009 +0100 @@ -59,7 +59,7 @@ fun prep_inst (concl, xs) = let val vs = Induct.vars_of concl - in map_filter prep_var (Library.drop (length vs - length xs, vs) ~~ xs) end; + in map_filter prep_var (drop (length vs - length xs) vs ~~ xs) end; in diff -r 651028e34b5d -r 01dcd9b926bf src/ZF/Tools/cartprod.ML --- a/src/ZF/Tools/cartprod.ML Fri Dec 04 11:44:57 2009 +0100 +++ b/src/ZF/Tools/cartprod.ML Fri Dec 04 15:20:24 2009 +0100 @@ -96,8 +96,8 @@ (*Makes a nested tuple from a list, following the product type structure*) fun mk_tuple pair (Type("*", [T1,T2])) tms = - pair $ (mk_tuple pair T1 tms) - $ (mk_tuple pair T2 (Library.drop (length (factors T1), tms))) + pair $ mk_tuple pair T1 tms + $ mk_tuple pair T2 (drop (length (factors T1)) tms) | mk_tuple pair T (t::_) = t; (*Attempts to remove occurrences of split, and pair-valued parameters*)