# HG changeset patch # User huffman # Date 1278428915 25200 # Node ID bf6c1216db43bc086af7f23eddd312c753297a9f # Parent 6432bf0d719117a6e18983d4984c1b0ce38a1ad8# Parent 8244558af8a55148b3c3ce8f0067795024110fa8 merged diff -r 6432bf0d7191 -r bf6c1216db43 doc-src/Classes/Thy/Classes.thy --- a/doc-src/Classes/Thy/Classes.thy Mon Jul 05 09:14:51 2010 -0700 +++ b/doc-src/Classes/Thy/Classes.thy Tue Jul 06 08:08:35 2010 -0700 @@ -194,7 +194,7 @@ using our simple algebra: *} -instantiation %quote * :: (semigroup, semigroup) semigroup +instantiation %quote prod :: (semigroup, semigroup) semigroup begin definition %quote @@ -260,7 +260,7 @@ end %quote -instantiation %quote * :: (monoidl, monoidl) monoidl +instantiation %quote prod :: (monoidl, monoidl) monoidl begin definition %quote @@ -297,7 +297,7 @@ end %quote -instantiation %quote * :: (monoid, monoid) monoid +instantiation %quote prod :: (monoid, monoid) monoid begin instance %quote proof diff -r 6432bf0d7191 -r bf6c1216db43 doc-src/Classes/Thy/document/Classes.tex --- a/doc-src/Classes/Thy/document/Classes.tex Mon Jul 05 09:14:51 2010 -0700 +++ b/doc-src/Classes/Thy/document/Classes.tex Tue Jul 06 08:08:35 2010 -0700 @@ -286,7 +286,7 @@ % \isatagquote \isacommand{instantiation}\isamarkupfalse% -\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}semigroup{\isacharcomma}\ semigroup{\isacharparenright}\ semigroup\isanewline +\ prod\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}semigroup{\isacharcomma}\ semigroup{\isacharparenright}\ semigroup\isanewline \isakeyword{begin}\isanewline \isanewline \isacommand{definition}\isamarkupfalse% @@ -405,7 +405,7 @@ \isanewline \isanewline \isacommand{instantiation}\isamarkupfalse% -\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}monoidl{\isacharcomma}\ monoidl{\isacharparenright}\ monoidl\isanewline +\ prod\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}monoidl{\isacharcomma}\ monoidl{\isacharparenright}\ monoidl\isanewline \isakeyword{begin}\isanewline \isanewline \isacommand{definition}\isamarkupfalse% @@ -479,7 +479,7 @@ \isanewline \isanewline \isacommand{instantiation}\isamarkupfalse% -\ {\isacharasterisk}\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}monoid{\isacharcomma}\ monoid{\isacharparenright}\ monoid\isanewline +\ prod\ {\isacharcolon}{\isacharcolon}\ {\isacharparenleft}monoid{\isacharcomma}\ monoid{\isacharparenright}\ monoid\isanewline \isakeyword{begin}\isanewline \isanewline \isacommand{instance}\isamarkupfalse% diff -r 6432bf0d7191 -r bf6c1216db43 src/HOL/Codegenerator_Test/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Codegenerator_Test/ROOT.ML Tue Jul 06 08:08:35 2010 -0700 @@ -0,0 +1,1 @@ +use_thys ["Generate", "Generate_Pretty"]; diff -r 6432bf0d7191 -r bf6c1216db43 src/HOL/Imperative_HOL/Array.thy --- a/src/HOL/Imperative_HOL/Array.thy Mon Jul 05 09:14:51 2010 -0700 +++ b/src/HOL/Imperative_HOL/Array.thy Tue Jul 06 08:08:35 2010 -0700 @@ -8,47 +8,149 @@ imports Heap_Monad begin +subsection {* Primitive layer *} + +definition + array_present :: "'a\heap array \ heap \ bool" where + "array_present a h \ addr_of_array a < lim h" + +definition + get_array :: "'a\heap array \ heap \ 'a list" where + "get_array a h = map from_nat (arrays h (TYPEREP('a)) (addr_of_array a))" + +definition + set_array :: "'a\heap array \ 'a list \ heap \ heap" where + "set_array a x = + arrays_update (\h. h(TYPEREP('a) := ((h(TYPEREP('a))) (addr_of_array a:=map to_nat x))))" + +definition array :: "'a list \ heap \ 'a\heap array \ heap" where + "array xs h = (let + l = lim h; + r = Array l; + h'' = set_array r xs (h\lim := l + 1\) + in (r, h''))" + +definition length :: "'a\heap array \ heap \ nat" where + "length a h = List.length (get_array a h)" + +definition change :: "'a\heap array \ nat \ 'a \ heap \ heap" where + "change a i x h = set_array a ((get_array a h)[i:=x]) h" + +text {* Properties of imperative arrays *} + +text {* FIXME: Does there exist a "canonical" array axiomatisation in +the literature? *} + +definition noteq_arrs :: "('a\heap) array \ ('b\heap) array \ bool" (infix "=!!=" 70) where + "r =!!= s \ TYPEREP('a) \ TYPEREP('b) \ addr_of_array r \ addr_of_array s" + +lemma noteq_arrs_sym: "a =!!= b \ b =!!= a" + and unequal_arrs [simp]: "a \ a' \ a =!!= a'" + unfolding noteq_arrs_def by auto + +lemma noteq_arrs_irrefl: "r =!!= r \ False" + unfolding noteq_arrs_def by auto + +lemma present_new_arr: "array_present a h \ a =!!= fst (array xs h)" + by (simp add: array_present_def noteq_arrs_def array_def Let_def) + +lemma array_get_set_eq [simp]: "get_array r (set_array r x h) = x" + by (simp add: get_array_def set_array_def o_def) + +lemma array_get_set_neq [simp]: "r =!!= s \ get_array r (set_array s x h) = get_array r h" + by (simp add: noteq_arrs_def get_array_def set_array_def) + +lemma set_array_same [simp]: + "set_array r x (set_array r y h) = set_array r x h" + by (simp add: set_array_def) + +lemma array_set_set_swap: + "r =!!= r' \ set_array r x (set_array r' x' h) = set_array r' x' (set_array r x h)" + by (simp add: Let_def expand_fun_eq noteq_arrs_def set_array_def) + +lemma get_array_change_eq [simp]: + "get_array a (change a i v h) = (get_array a h) [i := v]" + by (simp add: change_def) + +lemma nth_change_array_neq_array [simp]: + "a =!!= b \ get_array a (change b j v h) ! i = get_array a h ! i" + by (simp add: change_def noteq_arrs_def) + +lemma get_arry_array_change_elem_neqIndex [simp]: + "i \ j \ get_array a (change a j v h) ! i = get_array a h ! i" + by simp + +lemma length_change [simp]: + "length a (change b i v h) = length a h" + by (simp add: change_def length_def set_array_def get_array_def) + +lemma change_swap_neqArray: + "a =!!= a' \ + change a i v (change a' i' v' h) + = change a' i' v' (change a i v h)" +apply (unfold change_def) +apply simp +apply (subst array_set_set_swap, assumption) +apply (subst array_get_set_neq) +apply (erule noteq_arrs_sym) +apply (simp) +done + +lemma change_swap_neqIndex: + "\ i \ i' \ \ change a i v (change a i' v' h) = change a i' v' (change a i v h)" + by (auto simp add: change_def array_set_set_swap list_update_swap) + +lemma get_array_init_array_list: + "get_array (fst (array ls h)) (snd (array ls' h)) = ls'" + by (simp add: Let_def split_def array_def) + +lemma set_array: + "set_array (fst (array ls h)) + new_ls (snd (array ls h)) + = snd (array new_ls h)" + by (simp add: Let_def split_def array_def) + +lemma array_present_change [simp]: + "array_present a (change b i v h) = array_present a h" + by (simp add: change_def array_present_def set_array_def get_array_def) + + + subsection {* Primitives *} definition new :: "nat \ 'a\heap \ 'a array Heap" where - [code del]: "new n x = Heap_Monad.heap (Heap.array n x)" + [code del]: "new n x = Heap_Monad.heap (Array.array (replicate n x))" definition of_list :: "'a\heap list \ 'a array Heap" where - [code del]: "of_list xs = Heap_Monad.heap (Heap.array_of_list xs)" + [code del]: "of_list xs = Heap_Monad.heap (Array.array xs)" definition - length :: "'a\heap array \ nat Heap" where - [code del]: "length arr = Heap_Monad.heap (\h. (Heap.length arr h, h))" + len :: "'a\heap array \ nat Heap" where + [code del]: "len arr = Heap_Monad.heap (\h. (Array.length arr h, h))" definition nth :: "'a\heap array \ nat \ 'a Heap" where - [code del]: "nth a i = (do len \ length a; + [code del]: "nth a i = (do len \ len a; (if i < len then Heap_Monad.heap (\h. (get_array a h ! i, h)) - else raise (''array lookup: index out of range'')) + else raise ''array lookup: index out of range'') done)" definition upd :: "nat \ 'a \ 'a\heap array \ 'a\heap array Heap" where - [code del]: "upd i x a = (do len \ length a; + [code del]: "upd i x a = (do len \ len a; (if i < len - then Heap_Monad.heap (\h. (a, Heap.upd a i x h)) - else raise (''array update: index out of range'')) + then Heap_Monad.heap (\h. (a, change a i x h)) + else raise ''array update: index out of range'') done)" lemma upd_return: "upd i x a \ return a = upd i x a" -proof (rule Heap_eqI) - fix h - obtain len h' where "Heap_Monad.execute (Array.length a) h = (len, h')" - by (cases "Heap_Monad.execute (Array.length a) h") - then show "Heap_Monad.execute (upd i x a \ return a) h = Heap_Monad.execute (upd i x a) h" - by (auto simp add: upd_def bindM_def split: sum.split) -qed + by (rule Heap_eqI) (simp add: upd_def bindM_def split: option.split) subsection {* Derivates *} @@ -79,7 +181,7 @@ freeze :: "'a\heap array \ 'a list Heap" where "freeze a = (do - n \ length a; + n \ len a; mapM (nth a) [0..heap \ 'a) \ 'a array \ 'a array Heap" where "map f a = (do - n \ length a; + n \ len a; mapM (\n. map_entry n f a) [0.._. x)" - by (induct n) (simp_all add: make_def new_def Heap_Monad.heap_def - monad_simp array_of_list_replicate [symmetric] - map_replicate_trivial replicate_append_same - of_list_def) + by (rule Heap_eqI) (simp add: make_def new_def map_replicate_trivial of_list_def) lemma array_of_list_make [code]: "of_list xs = make (List.length xs) (\n. xs ! n)" - unfolding make_def map_nth .. + by (rule Heap_eqI) (simp add: make_def map_nth) subsection {* Code generator setup *} @@ -134,14 +232,12 @@ "Array.make n f = Array.make' (Code_Numeral.of_nat n) (f o Code_Numeral.nat_of)" by (simp add: make'_def o_def) -definition length' where - [code del]: "length' = Array.length \== liftM Code_Numeral.of_nat" -hide_const (open) length' +definition len' where + [code del]: "len' a = Array.len a \= (\n. return (Code_Numeral.of_nat n))" +hide_const (open) len' lemma [code]: - "Array.length = Array.length' \== liftM Code_Numeral.nat_of" - by (simp add: length'_def monad_simp', - simp add: liftM_def comp_def monad_simp, - simp add: monad_simp') + "Array.len a = Array.len' a \= (\i. return (Code_Numeral.nat_of i))" + by (simp add: len'_def) definition nth' where [code del]: "nth' a = Array.nth a o Code_Numeral.nat_of" @@ -155,7 +251,7 @@ hide_const (open) upd' lemma [code]: "Array.upd i x a = Array.upd' a (Code_Numeral.of_nat i) x \ return a" - by (simp add: upd'_def monad_simp upd_return) + by (simp add: upd'_def upd_return) subsubsection {* SML *} @@ -165,7 +261,7 @@ code_const Array.new' (SML "(fn/ ()/ =>/ Array.array/ ((_),/ (_)))") code_const Array.of_list' (SML "(fn/ ()/ =>/ Array.fromList/ _)") code_const Array.make' (SML "(fn/ ()/ =>/ Array.tabulate/ ((_),/ (_)))") -code_const Array.length' (SML "(fn/ ()/ =>/ Array.length/ _)") +code_const Array.len' (SML "(fn/ ()/ =>/ Array.length/ _)") code_const Array.nth' (SML "(fn/ ()/ =>/ Array.sub/ ((_),/ (_)))") code_const Array.upd' (SML "(fn/ ()/ =>/ Array.update/ ((_),/ (_),/ (_)))") @@ -178,7 +274,7 @@ code_const Array (OCaml "failwith/ \"bare Array\"") code_const Array.new' (OCaml "(fun/ ()/ ->/ Array.make/ (Big'_int.int'_of'_big'_int/ _)/ _)") code_const Array.of_list' (OCaml "(fun/ ()/ ->/ Array.of'_list/ _)") -code_const Array.length' (OCaml "(fun/ ()/ ->/ Big'_int.big'_int'_of'_int/ (Array.length/ _))") +code_const Array.len' (OCaml "(fun/ ()/ ->/ Big'_int.big'_int'_of'_int/ (Array.length/ _))") code_const Array.nth' (OCaml "(fun/ ()/ ->/ Array.get/ _/ (Big'_int.int'_of'_big'_int/ _))") code_const Array.upd' (OCaml "(fun/ ()/ ->/ Array.set/ _/ (Big'_int.int'_of'_big'_int/ _)/ _)") @@ -191,8 +287,10 @@ code_const Array (Haskell "error/ \"bare Array\"") code_const Array.new' (Haskell "Heap.newArray/ (0,/ _)") code_const Array.of_list' (Haskell "Heap.newListArray/ (0,/ _)") -code_const Array.length' (Haskell "Heap.lengthArray") +code_const Array.len' (Haskell "Heap.lengthArray") code_const Array.nth' (Haskell "Heap.readArray") code_const Array.upd' (Haskell "Heap.writeArray") +hide_const (open) new map -- {* avoid clashed with some popular names *} + end diff -r 6432bf0d7191 -r bf6c1216db43 src/HOL/Imperative_HOL/Heap.thy --- a/src/HOL/Imperative_HOL/Heap.thy Mon Jul 05 09:14:51 2010 -0700 +++ b/src/HOL/Imperative_HOL/Heap.thy Tue Jul 06 08:08:35 2010 -0700 @@ -1,4 +1,4 @@ -(* Title: HOL/Library/Heap.thy +(* Title: HOL/Imperative_HOL/Heap.thy Author: John Matthews, Galois Connections; Alexander Krauss, TU Muenchen *) @@ -14,8 +14,6 @@ class heap = typerep + countable -text {* Instances for common HOL types *} - instance nat :: heap .. instance prod :: (heap, heap) heap .. @@ -34,47 +32,26 @@ instance String.literal :: heap .. -text {* Reflected types themselves are heap-representable *} - -instantiation typerep :: countable -begin - -fun to_nat_typerep :: "typerep \ nat" where - "to_nat_typerep (Typerep.Typerep c ts) = to_nat (to_nat c, to_nat (map to_nat_typerep ts))" - -instance -proof (rule countable_classI) - fix t t' :: typerep and ts - have "(\t'. to_nat_typerep t = to_nat_typerep t' \ t = t') - \ (\ts'. map to_nat_typerep ts = map to_nat_typerep ts' \ ts = ts')" - proof (induct rule: typerep.induct) - case (Typerep c ts) show ?case - proof (rule allI, rule impI) - fix t' - assume hyp: "to_nat_typerep (Typerep.Typerep c ts) = to_nat_typerep t'" - then obtain c' ts' where t': "t' = (Typerep.Typerep c' ts')" - by (cases t') auto - with Typerep hyp have "c = c'" and "ts = ts'" by simp_all - with t' show "Typerep.Typerep c ts = t'" by simp - qed - next - case Nil_typerep then show ?case by simp - next - case (Cons_typerep t ts) then show ?case by auto - qed - then have "to_nat_typerep t = to_nat_typerep t' \ t = t'" by auto - moreover assume "to_nat_typerep t = to_nat_typerep t'" - ultimately show "t = t'" by simp -qed - -end - instance typerep :: heap .. subsection {* A polymorphic heap with dynamic arrays and references *} +text {* + References and arrays are developed in parallel, + but keeping them separate makes some later proofs simpler. +*} + types addr = nat -- "untyped heap references" +types heap_rep = nat -- "representable values" + +record heap = + arrays :: "typerep \ addr \ heap_rep list" + refs :: "typerep \ addr \ heap_rep" + lim :: addr + +definition empty :: heap where + "empty = \arrays = (\_ _. []), refs = (\_ _. 0), lim = 0\" datatype 'a array = Array addr datatype 'a ref = Ref addr -- "note the phantom type 'a " @@ -99,6 +76,8 @@ instance ref :: (type) countable by (rule countable_classI [of addr_of_ref]) simp +text {* Syntactic convenience *} + setup {* Sign.add_const_constraint (@{const_name Array}, SOME @{typ "nat \ 'a\heap array"}) #> Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \ 'a\heap ref"}) @@ -106,335 +85,6 @@ #> Sign.add_const_constraint (@{const_name addr_of_ref}, SOME @{typ "'a\heap ref \ nat"}) *} -types heap_rep = nat -- "representable values" - -record heap = - arrays :: "typerep \ addr \ heap_rep list" - refs :: "typerep \ addr \ heap_rep" - lim :: addr - -definition empty :: heap where - "empty = \arrays = (\_. undefined), refs = (\_. undefined), lim = 0\" -- "why undefined?" - - -subsection {* Imperative references and arrays *} - -text {* - References and arrays are developed in parallel, - but keeping them separate makes some later proofs simpler. -*} - -subsubsection {* Primitive operations *} - -definition - new_ref :: "heap \ ('a\heap) ref \ heap" where - "new_ref h = (let l = lim h in (Ref l, h\lim := l + 1\))" - -definition - new_array :: "heap \ ('a\heap) array \ heap" where - "new_array h = (let l = lim h in (Array l, h\lim := l + 1\))" - -definition - ref_present :: "'a\heap ref \ heap \ bool" where - "ref_present r h \ addr_of_ref r < lim h" - -definition - array_present :: "'a\heap array \ heap \ bool" where - "array_present a h \ addr_of_array a < lim h" - -definition - get_ref :: "'a\heap ref \ heap \ 'a" where - "get_ref r h = from_nat (refs h (TYPEREP('a)) (addr_of_ref r))" - -definition - get_array :: "'a\heap array \ heap \ 'a list" where - "get_array a h = map from_nat (arrays h (TYPEREP('a)) (addr_of_array a))" - -definition - set_ref :: "'a\heap ref \ 'a \ heap \ heap" where - "set_ref r x = - refs_update (\h. h(TYPEREP('a) := ((h (TYPEREP('a))) (addr_of_ref r:=to_nat x))))" - -definition - set_array :: "'a\heap array \ 'a list \ heap \ heap" where - "set_array a x = - arrays_update (\h. h(TYPEREP('a) := ((h(TYPEREP('a))) (addr_of_array a:=map to_nat x))))" - -subsubsection {* Interface operations *} - -definition - ref :: "'a \ heap \ 'a\heap ref \ heap" where - "ref x h = (let (r, h') = new_ref h; - h'' = set_ref r x h' - in (r, h''))" - -definition - array :: "nat \ 'a \ heap \ 'a\heap array \ heap" where - "array n x h = (let (r, h') = new_array h; - h'' = set_array r (replicate n x) h' - in (r, h''))" - -definition - array_of_list :: "'a list \ heap \ 'a\heap array \ heap" where - "array_of_list xs h = (let (r, h') = new_array h; - h'' = set_array r xs h' - in (r, h''))" - -definition - upd :: "'a\heap array \ nat \ 'a \ heap \ heap" where - "upd a i x h = set_array a ((get_array a h)[i:=x]) h" - -definition - length :: "'a\heap array \ heap \ nat" where - "length a h = size (get_array a h)" - -definition - array_ran :: "('a\heap) option array \ heap \ 'a set" where - "array_ran a h = {e. Some e \ set (get_array a h)}" - -- {*FIXME*} - - -subsubsection {* Reference equality *} - -text {* - The following relations are useful for comparing arrays and references. -*} - -definition - noteq_refs :: "('a\heap) ref \ ('b\heap) ref \ bool" (infix "=!=" 70) -where - "r =!= s \ TYPEREP('a) \ TYPEREP('b) \ addr_of_ref r \ addr_of_ref s" - -definition - noteq_arrs :: "('a\heap) array \ ('b\heap) array \ bool" (infix "=!!=" 70) -where - "r =!!= s \ TYPEREP('a) \ TYPEREP('b) \ addr_of_array r \ addr_of_array s" - -lemma noteq_refs_sym: "r =!= s \ s =!= r" - and noteq_arrs_sym: "a =!!= b \ b =!!= a" - and unequal_refs [simp]: "r \ r' \ r =!= r'" -- "same types!" - and unequal_arrs [simp]: "a \ a' \ a =!!= a'" -unfolding noteq_refs_def noteq_arrs_def by auto - -lemma noteq_refs_irrefl: "r =!= r \ False" - unfolding noteq_refs_def by auto - -lemma present_new_ref: "ref_present r h \ r =!= fst (ref v h)" - by (simp add: ref_present_def new_ref_def ref_def Let_def noteq_refs_def) - -lemma present_new_arr: "array_present a h \ a =!!= fst (array v x h)" - by (simp add: array_present_def noteq_arrs_def new_array_def array_def Let_def) - - -subsubsection {* Properties of heap containers *} - -text {* Properties of imperative arrays *} - -text {* FIXME: Does there exist a "canonical" array axiomatisation in -the literature? *} - -lemma array_get_set_eq [simp]: "get_array r (set_array r x h) = x" - by (simp add: get_array_def set_array_def o_def) - -lemma array_get_set_neq [simp]: "r =!!= s \ get_array r (set_array s x h) = get_array r h" - by (simp add: noteq_arrs_def get_array_def set_array_def) - -lemma set_array_same [simp]: - "set_array r x (set_array r y h) = set_array r x h" - by (simp add: set_array_def) - -lemma array_set_set_swap: - "r =!!= r' \ set_array r x (set_array r' x' h) = set_array r' x' (set_array r x h)" - by (simp add: Let_def expand_fun_eq noteq_arrs_def set_array_def) - -lemma array_ref_set_set_swap: - "set_array r x (set_ref r' x' h) = set_ref r' x' (set_array r x h)" - by (simp add: Let_def expand_fun_eq set_array_def set_ref_def) - -lemma get_array_upd_eq [simp]: - "get_array a (upd a i v h) = (get_array a h) [i := v]" - by (simp add: upd_def) - -lemma nth_upd_array_neq_array [simp]: - "a =!!= b \ get_array a (upd b j v h) ! i = get_array a h ! i" - by (simp add: upd_def noteq_arrs_def) - -lemma get_arry_array_upd_elem_neqIndex [simp]: - "i \ j \ get_array a (upd a j v h) ! i = get_array a h ! i" - by simp - -lemma length_upd_eq [simp]: - "length a (upd a i v h) = length a h" - by (simp add: length_def upd_def) - -lemma length_upd_neq [simp]: - "length a (upd b i v h) = length a h" - by (simp add: upd_def length_def set_array_def get_array_def) - -lemma upd_swap_neqArray: - "a =!!= a' \ - upd a i v (upd a' i' v' h) - = upd a' i' v' (upd a i v h)" -apply (unfold upd_def) -apply simp -apply (subst array_set_set_swap, assumption) -apply (subst array_get_set_neq) -apply (erule noteq_arrs_sym) -apply (simp) -done - -lemma upd_swap_neqIndex: - "\ i \ i' \ \ upd a i v (upd a i' v' h) = upd a i' v' (upd a i v h)" -by (auto simp add: upd_def array_set_set_swap list_update_swap) - -lemma get_array_init_array_list: - "get_array (fst (array_of_list ls h)) (snd (array_of_list ls' h)) = ls'" - by (simp add: Let_def split_def array_of_list_def) - -lemma set_array: - "set_array (fst (array_of_list ls h)) - new_ls (snd (array_of_list ls h)) - = snd (array_of_list new_ls h)" - by (simp add: Let_def split_def array_of_list_def) - -lemma array_present_upd [simp]: - "array_present a (upd b i v h) = array_present a h" - by (simp add: upd_def array_present_def set_array_def get_array_def) - -lemma array_of_list_replicate: - "array_of_list (replicate n x) = array n x" - by (simp add: expand_fun_eq array_of_list_def array_def) - - -text {* Properties of imperative references *} - -lemma next_ref_fresh [simp]: - assumes "(r, h') = new_ref h" - shows "\ ref_present r h" - using assms by (cases h) (auto simp add: new_ref_def ref_present_def Let_def) - -lemma next_ref_present [simp]: - assumes "(r, h') = new_ref h" - shows "ref_present r h'" - using assms by (cases h) (auto simp add: new_ref_def ref_present_def Let_def) - -lemma ref_get_set_eq [simp]: "get_ref r (set_ref r x h) = x" - by (simp add: get_ref_def set_ref_def) - -lemma ref_get_set_neq [simp]: "r =!= s \ get_ref r (set_ref s x h) = get_ref r h" - by (simp add: noteq_refs_def get_ref_def set_ref_def) - -(* FIXME: We need some infrastructure to infer that locally generated - new refs (by new_ref(_no_init), new_array(')) are distinct - from all existing refs. -*) - -lemma ref_set_get: "set_ref r (get_ref r h) h = h" -apply (simp add: set_ref_def get_ref_def) -oops - -lemma set_ref_same[simp]: - "set_ref r x (set_ref r y h) = set_ref r x h" - by (simp add: set_ref_def) - -lemma ref_set_set_swap: - "r =!= r' \ set_ref r x (set_ref r' x' h) = set_ref r' x' (set_ref r x h)" - by (simp add: Let_def expand_fun_eq noteq_refs_def set_ref_def) - -lemma ref_new_set: "fst (ref v (set_ref r v' h)) = fst (ref v h)" - by (simp add: ref_def new_ref_def set_ref_def Let_def) - -lemma ref_get_new [simp]: - "get_ref (fst (ref v h)) (snd (ref v' h)) = v'" - by (simp add: ref_def Let_def split_def) - -lemma ref_set_new [simp]: - "set_ref (fst (ref v h)) new_v (snd (ref v h)) = snd (ref new_v h)" - by (simp add: ref_def Let_def split_def) - -lemma ref_get_new_neq: "r =!= (fst (ref v h)) \ - get_ref r (snd (ref v h)) = get_ref r h" - by (simp add: get_ref_def set_ref_def ref_def Let_def new_ref_def noteq_refs_def) - -lemma lim_set_ref [simp]: - "lim (set_ref r v h) = lim h" - by (simp add: set_ref_def) - -lemma ref_present_new_ref [simp]: - "ref_present r h \ ref_present r (snd (ref v h))" - by (simp add: new_ref_def ref_present_def ref_def Let_def) - -lemma ref_present_set_ref [simp]: - "ref_present r (set_ref r' v h) = ref_present r h" - by (simp add: set_ref_def ref_present_def) - -lemma noteq_refsI: "\ ref_present r h; \ref_present r' h \ \ r =!= r'" - unfolding noteq_refs_def ref_present_def - by auto - -lemma array_ranI: "\ Some b = get_array a h ! i; i < Heap.length a h \ \ b \ array_ran a h" -unfolding array_ran_def Heap.length_def by simp - -lemma array_ran_upd_array_Some: - assumes "cl \ array_ran a (Heap.upd a i (Some b) h)" - shows "cl \ array_ran a h \ cl = b" -proof - - have "set (get_array a h[i := Some b]) \ insert (Some b) (set (get_array a h))" by (rule set_update_subset_insert) - with assms show ?thesis - unfolding array_ran_def Heap.upd_def by fastsimp -qed - -lemma array_ran_upd_array_None: - assumes "cl \ array_ran a (Heap.upd a i None h)" - shows "cl \ array_ran a h" -proof - - have "set (get_array a h[i := None]) \ insert None (set (get_array a h))" by (rule set_update_subset_insert) - with assms show ?thesis - unfolding array_ran_def Heap.upd_def by auto -qed - - -text {* Non-interaction between imperative array and imperative references *} - -lemma get_array_set_ref [simp]: "get_array a (set_ref r v h) = get_array a h" - by (simp add: get_array_def set_ref_def) - -lemma nth_set_ref [simp]: "get_array a (set_ref r v h) ! i = get_array a h ! i" - by simp - -lemma get_ref_upd [simp]: "get_ref r (upd a i v h) = get_ref r h" - by (simp add: get_ref_def set_array_def upd_def) - -lemma new_ref_upd: "fst (ref v (upd a i v' h)) = fst (ref v h)" - by (simp add: set_array_def get_array_def Let_def ref_new_set upd_def ref_def new_ref_def) - -text {*not actually true ???*} -lemma upd_set_ref_swap: "upd a i v (set_ref r v' h) = set_ref r v' (upd a i v h)" -apply (case_tac a) -apply (simp add: Let_def upd_def) -apply auto -oops - -lemma length_new_ref[simp]: - "length a (snd (ref v h)) = length a h" - by (simp add: get_array_def set_ref_def length_def new_ref_def ref_def Let_def) - -lemma get_array_new_ref [simp]: - "get_array a (snd (ref v h)) = get_array a h" - by (simp add: new_ref_def ref_def set_ref_def get_array_def Let_def) - -lemma ref_present_upd [simp]: - "ref_present r (upd a i v h) = ref_present r h" - by (simp add: upd_def ref_present_def set_array_def get_array_def) - -lemma array_present_set_ref [simp]: - "array_present a (set_ref r v h) = array_present a h" - by (simp add: array_present_def set_ref_def) - -lemma array_present_new_ref [simp]: - "array_present a h \ array_present a (snd (ref v h))" - by (simp add: array_present_def new_ref_def ref_def Let_def) - -hide_const (open) empty array array_of_list upd length ref +hide_const (open) empty end diff -r 6432bf0d7191 -r bf6c1216db43 src/HOL/Imperative_HOL/Heap_Monad.thy --- a/src/HOL/Imperative_HOL/Heap_Monad.thy Mon Jul 05 09:14:51 2010 -0700 +++ b/src/HOL/Imperative_HOL/Heap_Monad.thy Tue Jul 06 08:08:35 2010 -0700 @@ -12,16 +12,12 @@ subsubsection {* Monad combinators *} -datatype exception = Exn - text {* Monadic heap actions either produce values and transform the heap, or fail *} -datatype 'a Heap = Heap "heap \ ('a + exception) \ heap" +datatype 'a Heap = Heap "heap \ ('a \ heap) option" -primrec - execute :: "'a Heap \ heap \ ('a + exception) \ heap" where - "execute (Heap f) = f" -lemmas [code del] = execute.simps +primrec execute :: "'a Heap \ heap \ ('a \ heap) option" where + [code del]: "execute (Heap f) = f" lemma Heap_execute [simp]: "Heap (execute f) = f" by (cases f) simp_all @@ -34,58 +30,67 @@ "(\h. (\x. execute (f x) h) = (\y. execute (g y) h)) \ f = g" by (auto simp: expand_fun_eq intro: Heap_eqI) -lemma Heap_strip: "(\f. PROP P f) \ (\g. PROP P (Heap g))" -proof - fix g :: "heap \ ('a + exception) \ heap" - assume "\f. PROP P f" - then show "PROP P (Heap g)" . -next - fix f :: "'a Heap" - assume assm: "\g. PROP P (Heap g)" - then have "PROP P (Heap (execute f))" . - then show "PROP P f" by simp -qed - -definition - heap :: "(heap \ 'a \ heap) \ 'a Heap" where - [code del]: "heap f = Heap (\h. apfst Inl (f h))" +definition heap :: "(heap \ 'a \ heap) \ 'a Heap" where + [code del]: "heap f = Heap (Some \ f)" lemma execute_heap [simp]: - "execute (heap f) h = apfst Inl (f h)" + "execute (heap f) = Some \ f" by (simp add: heap_def) -definition - bindM :: "'a Heap \ ('a \ 'b Heap) \ 'b Heap" (infixl ">>=" 54) where - [code del]: "f >>= g = Heap (\h. case execute f h of - (Inl x, h') \ execute (g x) h' - | r \ r)" - -notation - bindM (infixl "\=" 54) +lemma heap_cases [case_names succeed fail]: + fixes f and h + assumes succeed: "\x h'. execute f h = Some (x, h') \ P" + assumes fail: "execute f h = None \ P" + shows P + using assms by (cases "execute f h") auto -abbreviation - chainM :: "'a Heap \ 'b Heap \ 'b Heap" (infixl ">>" 54) where - "f >> g \ f >>= (\_. g)" - -notation - chainM (infixl "\" 54) - -definition - return :: "'a \ 'a Heap" where +definition return :: "'a \ 'a Heap" where [code del]: "return x = heap (Pair x)" lemma execute_return [simp]: - "execute (return x) h = apfst Inl (x, h)" + "execute (return x) = Some \ Pair x" by (simp add: return_def) -definition - raise :: "string \ 'a Heap" where -- {* the string is just decoration *} - [code del]: "raise s = Heap (Pair (Inr Exn))" +definition raise :: "string \ 'a Heap" where -- {* the string is just decoration *} + [code del]: "raise s = Heap (\_. None)" lemma execute_raise [simp]: - "execute (raise s) h = (Inr Exn, h)" + "execute (raise s) = (\_. None)" by (simp add: raise_def) +definition bindM :: "'a Heap \ ('a \ 'b Heap) \ 'b Heap" (infixl ">>=" 54) where + [code del]: "f >>= g = Heap (\h. case execute f h of + Some (x, h') \ execute (g x) h' + | None \ None)" + +notation bindM (infixl "\=" 54) + +lemma execute_bind [simp]: + "execute f h = Some (x, h') \ execute (f \= g) h = execute (g x) h'" + "execute f h = None \ execute (f \= g) h = None" + by (simp_all add: bindM_def) + +lemma execute_bind_heap [simp]: + "execute (heap f \= g) h = execute (g (fst (f h))) (snd (f h))" + by (simp add: bindM_def split_def) + +lemma return_bind [simp]: "return x \= f = f x" + by (rule Heap_eqI) simp + +lemma bind_return [simp]: "f \= return = f" + by (rule Heap_eqI) (simp add: bindM_def split: option.splits) + +lemma bind_bind [simp]: "(f \= g) \= k = f \= (\x. g x \= k)" + by (rule Heap_eqI) (simp add: bindM_def split: option.splits) + +lemma raise_bind [simp]: "raise e \= f = raise e" + by (rule Heap_eqI) simp + +abbreviation chainM :: "'a Heap \ 'b Heap \ 'b Heap" (infixl ">>" 54) where + "f >> g \ f >>= (\_. g)" + +notation chainM (infixl "\" 54) + subsubsection {* do-syntax *} @@ -170,88 +175,10 @@ subsection {* Monad properties *} -subsubsection {* Monad laws *} - -lemma return_bind: "return x \= f = f x" - by (simp add: bindM_def return_def) - -lemma bind_return: "f \= return = f" -proof (rule Heap_eqI) - fix h - show "execute (f \= return) h = execute f h" - by (auto simp add: bindM_def return_def split: sum.splits prod.splits) -qed - -lemma bind_bind: "(f \= g) \= h = f \= (\x. g x \= h)" - by (rule Heap_eqI) (auto simp add: bindM_def split: split: sum.splits prod.splits) - -lemma bind_bind': "f \= (\x. g x \= h x) = f \= (\x. g x \= (\y. return (x, y))) \= (\(x, y). h x y)" - by (rule Heap_eqI) (auto simp add: bindM_def split: split: sum.splits prod.splits) - -lemma raise_bind: "raise e \= f = raise e" - by (simp add: raise_def bindM_def) - - -lemmas monad_simp = return_bind bind_return bind_bind raise_bind - - subsection {* Generic combinators *} -definition - liftM :: "('a \ 'b) \ 'a \ 'b Heap" -where - "liftM f = return o f" - -definition - compM :: "('a \ 'b Heap) \ ('b \ 'c Heap) \ 'a \ 'c Heap" (infixl ">>==" 54) -where - "(f >>== g) = (\x. f x \= g)" - -notation - compM (infixl "\==" 54) - -lemma liftM_collapse: "liftM f x = return (f x)" - by (simp add: liftM_def) - -lemma liftM_compM: "liftM f \== g = g o f" - by (auto intro: Heap_eqI' simp add: expand_fun_eq liftM_def compM_def bindM_def) - -lemma compM_return: "f \== return = f" - by (simp add: compM_def monad_simp) - -lemma compM_compM: "(f \== g) \== h = f \== (g \== h)" - by (simp add: compM_def monad_simp) - -lemma liftM_bind: - "(\x. liftM f x \= liftM g) = liftM (\x. g (f x))" - by (rule Heap_eqI') (simp add: monad_simp liftM_def bindM_def) - -lemma liftM_comp: - "liftM f o g = liftM (f o g)" - by (rule Heap_eqI') (simp add: liftM_def) - -lemmas monad_simp' = monad_simp liftM_compM compM_return - compM_compM liftM_bind liftM_comp - -primrec - mapM :: "('a \ 'b Heap) \ 'a list \ 'b list Heap" -where - "mapM f [] = return []" - | "mapM f (x#xs) = do y \ f x; - ys \ mapM f xs; - return (y # ys) - done" - -primrec - foldM :: "('a \ 'b \ 'b Heap) \ 'a list \ 'b \ 'b Heap" -where - "foldM f [] s = return s" - | "foldM f (x#xs) s = f x s \= foldM f xs" - -definition - assert :: "('a \ bool) \ 'a \ 'a Heap" -where - "assert P x = (if P x then return x else raise (''assert''))" +definition assert :: "('a \ bool) \ 'a \ 'a Heap" where + "assert P x = (if P x then return x else raise ''assert'')" lemma assert_cong [fundef_cong]: assumes "P = P'" @@ -259,28 +186,42 @@ shows "(assert P x >>= f) = (assert P' x >>= f')" using assms by (auto simp add: assert_def return_bind raise_bind) +definition liftM :: "('a \ 'b) \ 'a \ 'b Heap" where + "liftM f = return o f" + +lemma liftM_collapse [simp]: + "liftM f x = return (f x)" + by (simp add: liftM_def) + +lemma bind_liftM: + "(f \= liftM g) = (f \= (\x. return (g x)))" + by (simp add: liftM_def comp_def) + +primrec mapM :: "('a \ 'b Heap) \ 'a list \ 'b list Heap" where + "mapM f [] = return []" +| "mapM f (x#xs) = do + y \ f x; + ys \ mapM f xs; + return (y # ys) + done" + + subsubsection {* A monadic combinator for simple recursive functions *} text {* Using a locale to fix arguments f and g of MREC *} locale mrec = -fixes - f :: "'a => ('b + 'a) Heap" - and g :: "'a => 'a => 'b => 'b Heap" + fixes f :: "'a \ ('b + 'a) Heap" + and g :: "'a \ 'a \ 'b \ 'b Heap" begin -function (default "\(x,h). (Inr Exn, undefined)") - mrec -where - "mrec x h = - (case Heap_Monad.execute (f x) h of - (Inl (Inl r), h') \ (Inl r, h') - | (Inl (Inr s), h') \ - (case mrec s h' of - (Inl z, h'') \ Heap_Monad.execute (g x s z) h'' - | (Inr e, h'') \ (Inr e, h'')) - | (Inr e, h') \ (Inr e, h') - )" +function (default "\(x, h). None") mrec :: "'a \ heap \ ('b \ heap) option" where + "mrec x h = (case execute (f x) h of + Some (Inl r, h') \ Some (r, h') + | Some (Inr s, h') \ (case mrec s h' of + Some (z, h'') \ execute (g x s z) h'' + | None \ None) + | None \ None)" by auto lemma graph_implies_dom: @@ -290,38 +231,37 @@ apply (erule mrec_rel.cases) by simp -lemma mrec_default: "\ mrec_dom (x, h) \ mrec x h = (Inr Exn, undefined)" +lemma mrec_default: "\ mrec_dom (x, h) \ mrec x h = None" unfolding mrec_def by (rule fundef_default_value[OF mrec_sumC_def graph_implies_dom, of _ _ "(x, h)", simplified]) lemma mrec_di_reverse: assumes "\ mrec_dom (x, h)" shows " - (case Heap_Monad.execute (f x) h of - (Inl (Inl r), h') \ False - | (Inl (Inr s), h') \ \ mrec_dom (s, h') - | (Inr e, h') \ False + (case execute (f x) h of + Some (Inl r, h') \ False + | Some (Inr s, h') \ \ mrec_dom (s, h') + | None \ False )" -using assms -by (auto split:prod.splits sum.splits) - (erule notE, rule accpI, elim mrec_rel.cases, simp)+ - +using assms apply (auto split: option.split sum.split) +apply (rule ccontr) +apply (erule notE, rule accpI, elim mrec_rel.cases, auto)+ +done lemma mrec_rule: "mrec x h = - (case Heap_Monad.execute (f x) h of - (Inl (Inl r), h') \ (Inl r, h') - | (Inl (Inr s), h') \ + (case execute (f x) h of + Some (Inl r, h') \ Some (r, h') + | Some (Inr s, h') \ (case mrec s h' of - (Inl z, h'') \ Heap_Monad.execute (g x s z) h'' - | (Inr e, h'') \ (Inr e, h'')) - | (Inr e, h') \ (Inr e, h') + Some (z, h'') \ execute (g x s z) h'' + | None \ None) + | None \ None )" apply (cases "mrec_dom (x,h)", simp) apply (frule mrec_default) apply (frule mrec_di_reverse, simp) -by (auto split: sum.split prod.split simp: mrec_default) - +by (auto split: sum.split option.split simp: mrec_default) definition "MREC x = Heap (mrec x)" @@ -340,32 +280,31 @@ apply simp apply (rule ext) apply (unfold mrec_rule[of x]) - by (auto split:prod.splits sum.splits) - + by (auto split: option.splits prod.splits sum.splits) lemma MREC_pinduct: - assumes "Heap_Monad.execute (MREC x) h = (Inl r, h')" - assumes non_rec_case: "\ x h h' r. Heap_Monad.execute (f x) h = (Inl (Inl r), h') \ P x h h' r" - assumes rec_case: "\ x h h1 h2 h' s z r. Heap_Monad.execute (f x) h = (Inl (Inr s), h1) \ Heap_Monad.execute (MREC s) h1 = (Inl z, h2) \ P s h1 h2 z - \ Heap_Monad.execute (g x s z) h2 = (Inl r, h') \ P x h h' r" + assumes "execute (MREC x) h = Some (r, h')" + assumes non_rec_case: "\ x h h' r. execute (f x) h = Some (Inl r, h') \ P x h h' r" + assumes rec_case: "\ x h h1 h2 h' s z r. execute (f x) h = Some (Inr s, h1) \ execute (MREC s) h1 = Some (z, h2) \ P s h1 h2 z + \ execute (g x s z) h2 = Some (r, h') \ P x h h' r" shows "P x h h' r" proof - - from assms(1) have mrec: "mrec x h = (Inl r, h')" + from assms(1) have mrec: "mrec x h = Some (r, h')" unfolding MREC_def execute.simps . from mrec have dom: "mrec_dom (x, h)" apply - apply (rule ccontr) apply (drule mrec_default) by auto - from mrec have h'_r: "h' = (snd (mrec x h))" "r = (Sum_Type.Projl (fst (mrec x h)))" + from mrec have h'_r: "h' = snd (the (mrec x h))" "r = fst (the (mrec x h))" by auto - from mrec have "P x h (snd (mrec x h)) (Sum_Type.Projl (fst (mrec x h)))" + from mrec have "P x h (snd (the (mrec x h))) (fst (the (mrec x h)))" proof (induct arbitrary: r h' rule: mrec.pinduct[OF dom]) case (1 x h) - obtain rr h' where "mrec x h = (rr, h')" by fastsimp - obtain fret h1 where exec_f: "Heap_Monad.execute (f x) h = (fret, h1)" by fastsimp + obtain rr h' where "the (mrec x h) = (rr, h')" by fastsimp show ?case - proof (cases fret) - case (Inl a) + proof (cases "execute (f x) h") + case (Some result) + then obtain a h1 where exec_f: "execute (f x) h = Some (a, h1)" by fastsimp note Inl' = this show ?thesis proof (cases a) @@ -375,23 +314,28 @@ next case (Inr b) note Inr' = this - obtain ret_mrec h2 where mrec_rec: "mrec b h1 = (ret_mrec, h2)" by fastsimp - from this Inl 1(1) exec_f mrec show ?thesis - proof (cases "ret_mrec") - case (Inl aaa) - from this mrec exec_f Inl' Inr' 1(1) mrec_rec 1(2) [OF exec_f [symmetric] Inl' Inr', of "aaa" "h2"] 1(3) - show ?thesis - apply auto - apply (rule rec_case) - unfolding MREC_def by auto + show ?thesis + proof (cases "mrec b h1") + case (Some result) + then obtain aaa h2 where mrec_rec: "mrec b h1 = Some (aaa, h2)" by fastsimp + moreover from this have "P b h1 (snd (the (mrec b h1))) (fst (the (mrec b h1)))" + apply (intro 1(2)) + apply (auto simp add: Inr Inl') + done + moreover note mrec mrec_rec exec_f Inl' Inr' 1(1) 1(3) + ultimately show ?thesis + apply auto + apply (rule rec_case) + apply auto + unfolding MREC_def by auto next - case (Inr b) - from this Inl 1(1) exec_f mrec Inr' mrec_rec 1(3) show ?thesis by auto + case None + from this 1(1) exec_f mrec Inr' 1(3) show ?thesis by auto qed qed next - case (Inr b) - from this 1(1) mrec exec_f 1(3) show ?thesis by simp + case None + from this 1(1) mrec 1(3) show ?thesis by simp qed qed from this h'_r show ?thesis by simp @@ -405,45 +349,32 @@ lemmas MREC_rule = mrec.MREC_rule lemmas MREC_pinduct = mrec.MREC_pinduct -hide_const (open) heap execute - subsection {* Code generator setup *} subsubsection {* Logical intermediate layer *} -definition - Fail :: "String.literal \ exception" -where - [code del]: "Fail s = Exn" +primrec raise' :: "String.literal \ 'a Heap" where + [code del, code_post]: "raise' (STR s) = raise s" -definition - raise_exc :: "exception \ 'a Heap" -where - [code del]: "raise_exc e = raise []" +lemma raise_raise' [code_inline]: + "raise s = raise' (STR s)" + by simp -lemma raise_raise_exc [code, code_unfold]: - "raise s = raise_exc (Fail (STR s))" - unfolding Fail_def raise_exc_def raise_def .. - -hide_const (open) Fail raise_exc +code_datatype raise' -- {* avoid @{const "Heap"} formally *} subsubsection {* SML and OCaml *} code_type Heap (SML "unit/ ->/ _") -code_const Heap (SML "raise/ (Fail/ \"bare Heap\")") code_const "op \=" (SML "!(fn/ f'_/ =>/ fn/ ()/ =>/ f'_/ (_/ ())/ ())") code_const return (SML "!(fn/ ()/ =>/ _)") -code_const "Heap_Monad.Fail" (SML "Fail") -code_const "Heap_Monad.raise_exc" (SML "!(fn/ ()/ =>/ raise/ _)") +code_const Heap_Monad.raise' (SML "!(raise/ Fail/ _)") code_type Heap (OCaml "_") -code_const Heap (OCaml "failwith/ \"bare Heap\"") code_const "op \=" (OCaml "!(fun/ f'_/ ()/ ->/ f'_/ (_/ ())/ ())") code_const return (OCaml "!(fun/ ()/ ->/ _)") -code_const "Heap_Monad.Fail" (OCaml "Failure") -code_const "Heap_Monad.raise_exc" (OCaml "!(fun/ ()/ ->/ raise/ _)") +code_const Heap_Monad.raise' (OCaml "failwith/ _") setup {* @@ -514,8 +445,6 @@ *} -code_reserved OCaml Failure raise - subsubsection {* Haskell *} @@ -556,10 +485,10 @@ text {* Monad *} code_type Heap (Haskell "Heap.ST/ Heap.RealWorld/ _") -code_const Heap (Haskell "error/ \"bare Heap\"") code_monad "op \=" Haskell code_const return (Haskell "return") -code_const "Heap_Monad.Fail" (Haskell "_") -code_const "Heap_Monad.raise_exc" (Haskell "error") +code_const Heap_Monad.raise' (Haskell "error/ _") + +hide_const (open) Heap heap execute raise' end diff -r 6432bf0d7191 -r bf6c1216db43 src/HOL/Imperative_HOL/Ref.thy --- a/src/HOL/Imperative_HOL/Ref.thy Mon Jul 05 09:14:51 2010 -0700 +++ b/src/HOL/Imperative_HOL/Ref.thy Tue Jul 06 08:08:35 2010 -0700 @@ -1,12 +1,11 @@ (* Title: HOL/Library/Ref.thy - ID: $Id$ Author: John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen *) header {* Monadic references *} theory Ref -imports Heap_Monad +imports Array begin text {* @@ -15,45 +14,177 @@ and http://www.smlnj.org/doc/Conversion/top-level-comparison.html *} +subsection {* Primitive layer *} + +definition present :: "heap \ 'a\heap ref \ bool" where + "present h r \ addr_of_ref r < lim h" + +definition get :: "heap \ 'a\heap ref \ 'a" where + "get h = from_nat \ refs h TYPEREP('a) \ addr_of_ref" + +definition set :: "'a\heap ref \ 'a \ heap \ heap" where + "set r x = refs_update + (\h. h(TYPEREP('a) := ((h (TYPEREP('a))) (addr_of_ref r := to_nat x))))" + +definition alloc :: "'a \ heap \ 'a\heap ref \ heap" where + "alloc x h = (let + l = lim h; + r = Ref l + in (r, set r x (h\lim := l + 1\)))" + +definition noteq :: "'a\heap ref \ 'b\heap ref \ bool" (infix "=!=" 70) where + "r =!= s \ TYPEREP('a) \ TYPEREP('b) \ addr_of_ref r \ addr_of_ref s" + +lemma noteq_sym: "r =!= s \ s =!= r" + and unequal [simp]: "r \ r' \ r =!= r'" -- "same types!" + by (auto simp add: noteq_def) + +lemma noteq_irrefl: "r =!= r \ False" + by (auto simp add: noteq_def) + +lemma present_alloc_neq: "present h r \ r =!= fst (alloc v h)" + by (simp add: present_def alloc_def noteq_def Let_def) + +lemma next_fresh [simp]: + assumes "(r, h') = alloc x h" + shows "\ present h r" + using assms by (cases h) (auto simp add: alloc_def present_def Let_def) + +lemma next_present [simp]: + assumes "(r, h') = alloc x h" + shows "present h' r" + using assms by (cases h) (auto simp add: alloc_def set_def present_def Let_def) + +lemma get_set_eq [simp]: + "get (set r x h) r = x" + by (simp add: get_def set_def) + +lemma get_set_neq [simp]: + "r =!= s \ get (set s x h) r = get h r" + by (simp add: noteq_def get_def set_def) + +lemma set_same [simp]: + "set r x (set r y h) = set r x h" + by (simp add: set_def) + +lemma set_set_swap: + "r =!= r' \ set r x (set r' x' h) = set r' x' (set r x h)" + by (simp add: noteq_def set_def expand_fun_eq) + +lemma alloc_set: + "fst (alloc x (set r x' h)) = fst (alloc x h)" + by (simp add: alloc_def set_def Let_def) + +lemma get_alloc [simp]: + "get (snd (alloc x h)) (fst (alloc x' h)) = x" + by (simp add: alloc_def Let_def) + +lemma set_alloc [simp]: + "set (fst (alloc v h)) v' (snd (alloc v h)) = snd (alloc v' h)" + by (simp add: alloc_def Let_def) + +lemma get_alloc_neq: "r =!= fst (alloc v h) \ + get (snd (alloc v h)) r = get h r" + by (simp add: get_def set_def alloc_def Let_def noteq_def) + +lemma lim_set [simp]: + "lim (set r v h) = lim h" + by (simp add: set_def) + +lemma present_alloc [simp]: + "present h r \ present (snd (alloc v h)) r" + by (simp add: present_def alloc_def Let_def) + +lemma present_set [simp]: + "present (set r v h) = present h" + by (simp add: present_def expand_fun_eq) + +lemma noteq_I: + "present h r \ \ present h r' \ r =!= r'" + by (auto simp add: noteq_def present_def) + + subsection {* Primitives *} -definition - new :: "'a\heap \ 'a ref Heap" where - [code del]: "new v = Heap_Monad.heap (Heap.ref v)" +definition ref :: "'a\heap \ 'a ref Heap" where + [code del]: "ref v = Heap_Monad.heap (alloc v)" -definition - lookup :: "'a\heap ref \ 'a Heap" ("!_" 61) where - [code del]: "lookup r = Heap_Monad.heap (\h. (get_ref r h, h))" +definition lookup :: "'a\heap ref \ 'a Heap" ("!_" 61) where + [code del]: "lookup r = Heap_Monad.heap (\h. (get h r, h))" -definition - update :: "'a ref \ ('a\heap) \ unit Heap" ("_ := _" 62) where - [code del]: "update r e = Heap_Monad.heap (\h. ((), set_ref r e h))" +definition update :: "'a ref \ 'a\heap \ unit Heap" ("_ := _" 62) where + [code del]: "update r v = Heap_Monad.heap (\h. ((), set r v h))" subsection {* Derivates *} -definition - change :: "('a\heap \ 'a) \ 'a ref \ 'a Heap" -where - "change f r = (do x \ ! r; - let y = f x; - r := y; - return y - done)" - -hide_const (open) new lookup update change +definition change :: "('a\heap \ 'a) \ 'a ref \ 'a Heap" where + "change f r = (do + x \ ! r; + let y = f x; + r := y; + return y + done)" subsection {* Properties *} lemma lookup_chain: "(!r \ f) = f" - by (cases f) - (auto simp add: Let_def bindM_def lookup_def expand_fun_eq) + by (rule Heap_eqI) (simp add: lookup_def) lemma update_change [code]: - "r := e = Ref.change (\_. e) r \ return ()" - by (auto simp add: monad_simp change_def lookup_chain) + "r := e = change (\_. e) r \ return ()" + by (rule Heap_eqI) (simp add: change_def lookup_chain) + + +text {* Non-interaction between imperative array and imperative references *} + +lemma get_array_set [simp]: + "get_array a (set r v h) = get_array a h" + by (simp add: get_array_def set_def) + +lemma nth_set [simp]: + "get_array a (set r v h) ! i = get_array a h ! i" + by simp + +lemma get_change [simp]: + "get (Array.change a i v h) r = get h r" + by (simp add: get_def Array.change_def set_array_def) + +lemma alloc_change: + "fst (alloc v (Array.change a i v' h)) = fst (alloc v h)" + by (simp add: Array.change_def get_array_def set_array_def alloc_def Let_def) + +lemma change_set_swap: + "Array.change a i v (set r v' h) = set r v' (Array.change a i v h)" + by (simp add: Array.change_def get_array_def set_array_def set_def) + +lemma length_alloc [simp]: + "Array.length a (snd (alloc v h)) = Array.length a h" + by (simp add: Array.length_def get_array_def alloc_def set_def Let_def) + +lemma get_array_alloc [simp]: + "get_array a (snd (alloc v h)) = get_array a h" + by (simp add: get_array_def alloc_def set_def Let_def) + +lemma present_change [simp]: + "present (Array.change a i v h) = present h" + by (simp add: Array.change_def set_array_def expand_fun_eq present_def) + +lemma array_present_set [simp]: + "array_present a (set r v h) = array_present a h" + by (simp add: array_present_def set_def) + +lemma array_present_alloc [simp]: + "array_present a h \ array_present a (snd (alloc v h))" + by (simp add: array_present_def alloc_def Let_def) + +lemma set_array_set_swap: + "set_array a xs (set r x' h) = set r x' (set_array a xs h)" + by (simp add: set_array_def set_def) + +hide_const (open) present get set alloc lookup update change subsection {* Code generator setup *} @@ -62,7 +193,7 @@ code_type ref (SML "_/ Unsynchronized.ref") code_const Ref (SML "raise/ (Fail/ \"bare Ref\")") -code_const Ref.new (SML "(fn/ ()/ =>/ Unsynchronized.ref/ _)") +code_const ref (SML "(fn/ ()/ =>/ Unsynchronized.ref/ _)") code_const Ref.lookup (SML "(fn/ ()/ =>/ !/ _)") code_const Ref.update (SML "(fn/ ()/ =>/ _/ :=/ _)") @@ -73,7 +204,7 @@ code_type ref (OCaml "_/ ref") code_const Ref (OCaml "failwith/ \"bare Ref\")") -code_const Ref.new (OCaml "(fn/ ()/ =>/ ref/ _)") +code_const ref (OCaml "(fn/ ()/ =>/ ref/ _)") code_const Ref.lookup (OCaml "(fn/ ()/ =>/ !/ _)") code_const Ref.update (OCaml "(fn/ ()/ =>/ _/ :=/ _)") @@ -84,7 +215,7 @@ code_type ref (Haskell "Heap.STRef/ Heap.RealWorld/ _") code_const Ref (Haskell "error/ \"bare Ref\"") -code_const Ref.new (Haskell "Heap.newSTRef") +code_const ref (Haskell "Heap.newSTRef") code_const Ref.lookup (Haskell "Heap.readSTRef") code_const Ref.update (Haskell "Heap.writeSTRef") diff -r 6432bf0d7191 -r bf6c1216db43 src/HOL/Imperative_HOL/Relational.thy --- a/src/HOL/Imperative_HOL/Relational.thy Mon Jul 05 09:14:51 2010 -0700 +++ b/src/HOL/Imperative_HOL/Relational.thy Tue Jul 06 08:08:35 2010 -0700 @@ -9,10 +9,10 @@ definition crel :: "'a Heap \ heap \ heap \ 'a \ bool" where - crel_def': "crel c h h' r \ Heap_Monad.execute c h = (Inl r, h')" + crel_def': "crel c h h' r \ Heap_Monad.execute c h = Some (r, h')" lemma crel_def: -- FIXME - "crel c h h' r \ (Inl r, h') = Heap_Monad.execute c h" + "crel c h h' r \ Some (r, h') = Heap_Monad.execute c h" unfolding crel_def' by auto lemma crel_deterministic: "\ crel f h h' a; crel f h h'' b \ \ (a = b) \ (h' = h'')" @@ -28,8 +28,7 @@ lemma crelE[consumes 1]: assumes "crel (f >>= g) h h'' r'" obtains h' r where "crel f h h' r" "crel (g r) h' h'' r'" - using assms - by (auto simp add: crel_def bindM_def Let_def prod_case_beta split_def Pair_fst_snd_eq split add: sum.split_asm) + using assms by (auto simp add: crel_def bindM_def split: option.split_asm) lemma crelE'[consumes 1]: assumes "crel (f >> g) h h'' r'" @@ -86,35 +85,35 @@ lemma crel_heap: assumes "crel (Heap_Monad.heap f) h h' r" obtains "h' = snd (f h)" "r = fst (f h)" - using assms - unfolding heap_def crel_def apfst_def split_def prod_fun_def by simp_all + using assms by (cases "f h") (simp add: crel_def) + subsection {* Elimination rules for array commands *} lemma crel_length: - assumes "crel (length a) h h' r" - obtains "h = h'" "r = Heap.length a h'" + assumes "crel (len a) h h' r" + obtains "h = h'" "r = Array.length a h'" using assms - unfolding length_def + unfolding Array.len_def by (elim crel_heap) simp (* Strong version of the lemma for this operation is missing *) lemma crel_new_weak: assumes "crel (Array.new n v) h h' r" obtains "get_array r h' = List.replicate n v" - using assms unfolding Array.new_def - by (elim crel_heap) (auto simp:Heap.array_def Let_def split_def) + using assms unfolding Array.new_def + by (elim crel_heap) (auto simp: array_def Let_def split_def) lemma crel_nth[consumes 1]: assumes "crel (nth a i) h h' r" - obtains "r = (get_array a h) ! i" "h = h'" "i < Heap.length a h" + obtains "r = (get_array a h) ! i" "h = h'" "i < Array.length a h" using assms unfolding nth_def by (auto elim!: crelE crel_if crel_raise crel_length crel_heap) lemma crel_upd[consumes 1]: assumes "crel (upd i v a) h h' r" - obtains "r = a" "h' = Heap.upd a i v h" + obtains "r = a" "h' = Array.change a i v h" using assms unfolding upd_def by (elim crelE crel_if crel_return crel_raise @@ -130,14 +129,14 @@ lemma crel_map_entry: assumes "crel (Array.map_entry i f a) h h' r" - obtains "r = a" "h' = Heap.upd a i (f (get_array a h ! i)) h" + obtains "r = a" "h' = Array.change a i (f (get_array a h ! i)) h" using assms unfolding Array.map_entry_def by (elim crelE crel_upd crel_nth) auto lemma crel_swap: assumes "crel (Array.swap i x a) h h' r" - obtains "r = get_array a h ! i" "h' = Heap.upd a i x h" + obtains "r = get_array a h ! i" "h' = Array.change a i x h" using assms unfolding Array.swap_def by (elim crelE crel_upd crel_nth crel_return) auto @@ -161,25 +160,25 @@ lemma crel_mapM_nth: assumes - "crel (mapM (Array.nth a) [Heap.length a h - n.. Heap.length a h" - shows "h = h' \ xs = drop (Heap.length a h - n) (get_array a h)" + "crel (mapM (Array.nth a) [Array.length a h - n.. Array.length a h" + shows "h = h' \ xs = drop (Array.length a h - n) (get_array a h)" using assms proof (induct n arbitrary: xs h h') case 0 thus ?case - by (auto elim!: crel_return simp add: Heap.length_def) + by (auto elim!: crel_return simp add: Array.length_def) next case (Suc n) - from Suc(3) have "[Heap.length a h - Suc n..n. map_entry n f a) [Heap.length a h - n..n. map_entry n f a) [Array.length a h - n..n. map_entry n f a) [Heap.length a h - n..n. map_entry n f a) [Array.length a h - n..n. map_entry n f a) [Heap.length a ?h1 - n..n. map_entry n f a) [Array.length a ?h1 - n..n. map_entry n f a) [Heap.length a h - n.. Heap.length a h" - assumes "i \ Heap.length a h - n" - assumes "i < Heap.length a h" + assumes "crel (mapM (\n. map_entry n f a) [Array.length a h - n.. Array.length a h" + assumes "i \ Array.length a h - n" + assumes "i < Array.length a h" shows "get_array a h' ! i = f (get_array a h ! i)" using assms proof (induct n arbitrary: h h' r) @@ -231,54 +230,54 @@ by (auto elim!: crel_return) next case (Suc n) - let ?h1 = "Heap.upd a (Heap.length a h - Suc n) (f (get_array a h ! (Heap.length a h - Suc n))) h" - from Suc(3) have "[Heap.length a h - Suc n..n. map_entry n f a) [Heap.length a h - n..n. map_entry n f a) [Array.length a h - n.. i \ Heap.length a ?h1 - n" by arith - from crel_mapM have crel_mapM': "crel (mapM (\n. map_entry n f a) [Heap.length a ?h1 - n.. i \ Array.length a ?h1 - n" by arith + from crel_mapM have crel_mapM': "crel (mapM (\n. map_entry n f a) [Array.length a ?h1 - n..n. map_entry n f a) [Heap.length a h - n.. Heap.length a h" - shows "Heap.length a h' = Heap.length a h" + assumes "crel (mapM (\n. map_entry n f a) [Array.length a h - n.. Array.length a h" + shows "Array.length a h' = Array.length a h" using assms proof (induct n arbitrary: h h' r) case 0 thus ?case by (auto elim!: crel_return) next case (Suc n) - let ?h1 = "Heap.upd a (Heap.length a h - Suc n) (f (get_array a h ! (Heap.length a h - Suc n))) h" - from Suc(3) have "[Heap.length a h - Suc n..n. map_entry n f a) [Heap.length a h - n..n. map_entry n f a) [Array.length a h - n..n. map_entry n f a) [Heap.length a ?h1 - n..n. map_entry n f a) [Array.length a ?h1 - n..n. map_entry n f a) [0..n. map_entry n f a) [0..n. map_entry n f a) [Heap.length a h - Heap.length a h..n. map_entry n f a) [Array.length a h - Array.length a h.. lim h' = Suc (lim h) *) -lemma crel_Ref_new: - assumes "crel (Ref.new v) h h' x" - obtains "get_ref x h' = v" - and "\ ref_present x h" - and "ref_present x h'" - and "\y. ref_present y h \ get_ref y h = get_ref y h'" +lemma crel_ref: + assumes "crel (ref v) h h' x" + obtains "Ref.get h' x = v" + and "\ Ref.present h x" + and "Ref.present h' x" + and "\y. Ref.present h y \ Ref.get h y = Ref.get h' y" (* and "lim h' = Suc (lim h)" *) - and "\y. ref_present y h \ ref_present y h'" + and "\y. Ref.present h y \ Ref.present h' y" using assms - unfolding Ref.new_def + unfolding Ref.ref_def apply (elim crel_heap) - unfolding Heap.ref_def + unfolding Ref.alloc_def apply (simp add: Let_def) - unfolding Heap.new_ref_def - apply (simp add: Let_def) - unfolding ref_present_def + unfolding Ref.present_def apply auto - unfolding get_ref_def set_ref_def + unfolding Ref.get_def Ref.set_def apply auto done lemma crel_lookup: - assumes "crel (!ref) h h' r" - obtains "h = h'" "r = get_ref ref h" + assumes "crel (!r') h h' r" + obtains "h = h'" "r = Ref.get h r'" using assms unfolding Ref.lookup_def by (auto elim: crel_heap) lemma crel_update: - assumes "crel (ref := v) h h' r" - obtains "h' = set_ref ref v h" "r = ()" + assumes "crel (r' := v) h h' r" + obtains "h' = Ref.set r' v h" "r = ()" using assms unfolding Ref.update_def by (auto elim: crel_heap) lemma crel_change: - assumes "crel (Ref.change f ref) h h' r" - obtains "h' = set_ref ref (f (get_ref ref h)) h" "r = f (get_ref ref h)" + assumes "crel (Ref.change f r') h h' r" + obtains "h' = Ref.set r' (f (Ref.get h r')) h" "r = f (Ref.get h r')" using assms unfolding Ref.change_def Let_def by (auto elim!: crelE crel_lookup crel_update crel_return) @@ -369,11 +366,9 @@ apply (cases f) apply simp apply (simp add: expand_fun_eq split_def) +apply (auto split: option.split) +apply (erule_tac x="x" in meta_allE) apply auto -apply (case_tac "fst (fun x)") -apply (simp_all add: Pair_fst_snd_eq) -apply (erule_tac x="x" in meta_allE) -apply fastsimp done section {* Introduction rules *} @@ -436,22 +431,22 @@ subsection {* Introduction rules for array commands *} lemma crel_lengthI: - shows "crel (length a) h h (Heap.length a h)" - unfolding length_def + shows "crel (Array.len a) h h (Array.length a h)" + unfolding len_def by (rule crel_heapI') auto (* thm crel_newI for Array.new is missing *) lemma crel_nthI: - assumes "i < Heap.length a h" + assumes "i < Array.length a h" shows "crel (nth a i) h h ((get_array a h) ! i)" using assms unfolding nth_def by (auto intro!: crelI crel_ifI crel_raiseI crel_lengthI crel_heapI') lemma crel_updI: - assumes "i < Heap.length a h" - shows "crel (upd i v a) h (Heap.upd a i v h) a" + assumes "i < Array.length a h" + shows "crel (upd i v a) h (Array.change a i v h) a" using assms unfolding upd_def by (auto intro!: crelI crel_ifI crel_returnI crel_raiseI @@ -472,15 +467,15 @@ subsubsection {* Introduction rules for reference commands *} lemma crel_lookupI: - shows "crel (!ref) h h (get_ref ref h)" + shows "crel (!r) h h (Ref.get h r)" unfolding lookup_def by (auto intro!: crel_heapI') lemma crel_updateI: - shows "crel (ref := v) h (set_ref ref v h) ()" + shows "crel (r := v) h (Ref.set r v h) ()" unfolding update_def by (auto intro!: crel_heapI') lemma crel_changeI: - shows "crel (Ref.change f ref) h (set_ref ref (f (get_ref ref h)) h) (f (get_ref ref h))" + shows "crel (Ref.change f r) h (Ref.set r (f (Ref.get h r)) h) (f (Ref.get h r))" unfolding change_def Let_def by (auto intro!: crelI crel_returnI crel_lookupI crel_updateI) subsubsection {* Introduction rules for the assert command *} @@ -502,10 +497,10 @@ shows "P x h h' r" proof (rule MREC_pinduct[OF assms(1)[unfolded crel_def, symmetric]]) fix x h h1 h2 h' s z r - assume "Heap_Monad.execute (f x) h = (Inl (Inr s), h1)" - "Heap_Monad.execute (MREC f g s) h1 = (Inl z, h2)" + assume "Heap_Monad.execute (f x) h = Some (Inr s, h1)" + "Heap_Monad.execute (MREC f g s) h1 = Some (z, h2)" "P s h1 h2 z" - "Heap_Monad.execute (g x s z) h2 = (Inl r, h')" + "Heap_Monad.execute (g x s z) h2 = Some (r, h')" from assms(3)[unfolded crel_def, OF this(1)[symmetric] this(2)[symmetric] this(3) this(4)[symmetric]] show "P x h h' r" . next @@ -519,15 +514,15 @@ definition noError :: "'a Heap \ heap \ bool" where - "noError c h \ (\r h'. (Inl r, h') = Heap_Monad.execute c h)" + "noError c h \ (\r h'. Some (r, h') = Heap_Monad.execute c h)" lemma noError_def': -- FIXME - "noError c h \ (\r h'. Heap_Monad.execute c h = (Inl r, h'))" + "noError c h \ (\r h'. Heap_Monad.execute c h = Some (r, h'))" unfolding noError_def apply auto proof - fix r h' - assume "(Inl r, h') = Heap_Monad.execute c h" - then have "Heap_Monad.execute c h = (Inl r, h')" .. - then show "\r h'. Heap_Monad.execute c h = (Inl r, h')" by blast + assume "Some (r, h') = Heap_Monad.execute c h" + then have "Heap_Monad.execute c h = Some (r, h')" .. + then show "\r h'. Heap_Monad.execute c h = Some (r, h')" by blast qed subsection {* Introduction rules for basic monadic commands *} @@ -592,8 +587,8 @@ subsection {* Introduction rules for array commands *} lemma noError_length: - shows "noError (Array.length a) h" - unfolding length_def + shows "noError (Array.len a) h" + unfolding len_def by (intro noError_heap) lemma noError_new: @@ -601,14 +596,14 @@ unfolding Array.new_def by (intro noError_heap) lemma noError_upd: - assumes "i < Heap.length a h" + assumes "i < Array.length a h" shows "noError (Array.upd i v a) h" using assms unfolding upd_def by (auto intro!: noErrorI noError_if noError_return noError_length noError_heap) (auto elim: crel_length) lemma noError_nth: -assumes "i < Heap.length a h" +assumes "i < Array.length a h" shows "noError (Array.nth a i) h" using assms unfolding nth_def @@ -619,14 +614,14 @@ unfolding of_list_def by (rule noError_heap) lemma noError_map_entry: - assumes "i < Heap.length a h" + assumes "i < Array.length a h" shows "noError (map_entry i f a) h" using assms unfolding map_entry_def by (auto elim: crel_nth intro!: noErrorI noError_nth noError_upd) lemma noError_swap: - assumes "i < Heap.length a h" + assumes "i < Array.length a h" shows "noError (swap i x a) h" using assms unfolding swap_def @@ -640,7 +635,7 @@ (*TODO: move to HeapMonad *) lemma mapM_append: "mapM f (xs @ ys) = mapM f xs \= (\xs. mapM f ys \= (\ys. return (xs @ ys)))" - by (induct xs) (simp_all add: monad_simp) + by (induct xs) simp_all lemma noError_freeze: shows "noError (freeze a) h" @@ -649,42 +644,42 @@ noError_nth crel_nthI elim: crel_length) lemma noError_mapM_map_entry: - assumes "n \ Heap.length a h" - shows "noError (mapM (\n. map_entry n f a) [Heap.length a h - n.. Array.length a h" + shows "noError (mapM (\n. map_entry n f a) [Array.length a h - n.. nat \ nat \ nat \ nat Heap" where @@ -101,7 +101,7 @@ lemma part_length_remains: assumes "crel (part1 a l r p) h h' rs" - shows "Heap.length a h = Heap.length a h'" + shows "Array.length a h = Array.length a h'" using assms proof (induct a l r p arbitrary: h h' rs rule:part1.induct) case (1 a l r p h h' rs) @@ -207,7 +207,7 @@ by (elim crelE crel_nth crel_if crel_return) auto from swp False have "get_array a h1 ! r \ p" unfolding swap_def - by (auto simp add: Heap.length_def elim!: crelE crel_nth crel_upd crel_return) + by (auto simp add: Array.length_def elim!: crelE crel_nth crel_upd crel_return) with part_outer_remains [OF rec2] lr have a_r: "get_array a h' ! r \ p" by fastsimp have "\i. (i \ r = (i = r \ i \ r - 1))" by arith @@ -243,7 +243,7 @@ lemma partition_length_remains: assumes "crel (partition a l r) h h' rs" - shows "Heap.length a h = Heap.length a h'" + shows "Array.length a h = Array.length a h'" proof - from assms part_length_remains show ?thesis unfolding partition.simps swap_def @@ -287,14 +287,14 @@ else middle)" unfolding partition.simps by (elim crelE crel_return crel_nth crel_if crel_upd) simp - from swap have h'_def: "h' = Heap.upd a r (get_array a h1 ! rs) - (Heap.upd a rs (get_array a h1 ! r) h1)" + from swap have h'_def: "h' = Array.change a r (get_array a h1 ! rs) + (Array.change a rs (get_array a h1 ! r) h1)" unfolding swap_def by (elim crelE crel_return crel_nth crel_upd) simp - from swap have in_bounds: "r < Heap.length a h1 \ rs < Heap.length a h1" + from swap have in_bounds: "r < Array.length a h1 \ rs < Array.length a h1" unfolding swap_def by (elim crelE crel_return crel_nth crel_upd) simp - from swap have swap_length_remains: "Heap.length a h1 = Heap.length a h'" + from swap have swap_length_remains: "Array.length a h1 = Array.length a h'" unfolding swap_def by (elim crelE crel_return crel_nth crel_upd) auto from `l < r` have "l \ r - 1" by simp note middle_in_bounds = part_returns_index_in_bounds[OF part this] @@ -304,7 +304,7 @@ with swap have right_remains: "get_array a h ! r = get_array a h' ! rs" unfolding swap_def - by (auto simp add: Heap.length_def elim!: crelE crel_return crel_nth crel_upd) (cases "r = rs", auto) + by (auto simp add: Array.length_def elim!: crelE crel_return crel_nth crel_upd) (cases "r = rs", auto) from part_partitions [OF part] show ?thesis proof (cases "get_array a h1 ! middle \ ?pivot") @@ -314,12 +314,12 @@ fix i assume i_is_left: "l \ i \ i < rs" with swap_length_remains in_bounds middle_in_bounds rs_equals `l < r` - have i_props: "i < Heap.length a h'" "i \ r" "i \ rs" by auto + have i_props: "i < Array.length a h'" "i \ r" "i \ rs" by auto from i_is_left rs_equals have "l \ i \ i < middle \ i = middle" by arith with part_partitions[OF part] right_remains True have "get_array a h1 ! i \ get_array a h' ! rs" by fastsimp with i_props h'_def in_bounds have "get_array a h' ! i \ get_array a h' ! rs" - unfolding Heap.upd_def Heap.length_def by simp + unfolding Array.change_def Array.length_def by simp } moreover { @@ -331,7 +331,7 @@ proof assume i_is: "rs < i \ i \ r - 1" with swap_length_remains in_bounds middle_in_bounds rs_equals - have i_props: "i < Heap.length a h'" "i \ r" "i \ rs" by auto + have i_props: "i < Array.length a h'" "i \ r" "i \ rs" by auto from part_partitions[OF part] rs_equals right_remains i_is have "get_array a h' ! rs \ get_array a h1 ! i" by fastsimp @@ -345,7 +345,7 @@ by fastsimp with i_is True rs_equals right_remains h'_def show ?thesis using in_bounds - unfolding Heap.upd_def Heap.length_def + unfolding Array.change_def Array.length_def by auto qed } @@ -357,11 +357,11 @@ fix i assume i_is_left: "l \ i \ i < rs" with swap_length_remains in_bounds middle_in_bounds rs_equals - have i_props: "i < Heap.length a h'" "i \ r" "i \ rs" by auto + have i_props: "i < Array.length a h'" "i \ r" "i \ rs" by auto from part_partitions[OF part] rs_equals right_remains i_is_left have "get_array a h1 ! i \ get_array a h' ! rs" by fastsimp with i_props h'_def have "get_array a h' ! i \ get_array a h' ! rs" - unfolding Heap.upd_def by simp + unfolding Array.change_def by simp } moreover { @@ -372,7 +372,7 @@ proof assume i_is: "rs < i \ i \ r - 1" with swap_length_remains in_bounds middle_in_bounds rs_equals - have i_props: "i < Heap.length a h'" "i \ r" "i \ rs" by auto + have i_props: "i < Array.length a h'" "i \ r" "i \ rs" by auto from part_partitions[OF part] rs_equals right_remains i_is have "get_array a h' ! rs \ get_array a h1 ! i" by fastsimp @@ -381,7 +381,7 @@ assume i_is: "i = r" from i_is False rs_equals right_remains h'_def show ?thesis using in_bounds - unfolding Heap.upd_def Heap.length_def + unfolding Array.change_def Array.length_def by auto qed } @@ -425,7 +425,7 @@ lemma length_remains: assumes "crel (quicksort a l r) h h' rs" - shows "Heap.length a h = Heap.length a h'" + shows "Array.length a h = Array.length a h'" using assms proof (induct a l r arbitrary: h h' rs rule: quicksort.induct) case (1 a l r h h' rs) @@ -482,7 +482,7 @@ lemma quicksort_sorts: assumes "crel (quicksort a l r) h h' rs" - assumes l_r_length: "l < Heap.length a h" "r < Heap.length a h" + assumes l_r_length: "l < Array.length a h" "r < Array.length a h" shows "sorted (subarray l (r + 1) a h')" using assms proof (induct a l r arbitrary: h h' rs rule: quicksort.induct) @@ -524,7 +524,7 @@ from quicksort_outer_remains [OF qs1] quicksort_permutes [OF qs1] True length_remains 1(5) pivot multiset_of_sublist [of l p "get_array a h1" "get_array a h2"] have multiset_partconds1: "multiset_of (subarray l p a h2) = multiset_of (subarray l p a h1)" - unfolding Heap.length_def subarray_def by (cases p, auto) + unfolding Array.length_def subarray_def by (cases p, auto) with left_subarray_remains part_conds1 pivot_unchanged have part_conds2': "\j. j \ set (subarray l p a h') \ j \ get_array a h' ! p" by (simp, subst set_of_multiset_of[symmetric], simp) @@ -535,7 +535,7 @@ from quicksort_outer_remains [OF qs2] quicksort_permutes [OF qs2] True length_remains 1(5) pivot multiset_of_sublist [of "p + 1" "r + 1" "get_array a h2" "get_array a h'"] have multiset_partconds2: "multiset_of (subarray (p + 1) (r + 1) a h') = multiset_of (subarray (p + 1) (r + 1) a h2)" - unfolding Heap.length_def subarray_def by auto + unfolding Array.length_def subarray_def by auto with right_subarray_remains part_conds2 pivot_unchanged have part_conds1': "\j. j \ set (subarray (p + 1) (r + 1) a h') \ get_array a h' ! p \ j" by (simp, subst set_of_multiset_of[symmetric], simp) @@ -556,18 +556,18 @@ lemma quicksort_is_sort: - assumes crel: "crel (quicksort a 0 (Heap.length a h - 1)) h h' rs" + assumes crel: "crel (quicksort a 0 (Array.length a h - 1)) h h' rs" shows "get_array a h' = sort (get_array a h)" proof (cases "get_array a h = []") case True with quicksort_is_skip[OF crel] show ?thesis - unfolding Heap.length_def by simp + unfolding Array.length_def by simp next case False from quicksort_sorts [OF crel] False have "sorted (sublist' 0 (List.length (get_array a h)) (get_array a h'))" - unfolding Heap.length_def subarray_def by auto + unfolding Array.length_def subarray_def by auto with length_remains[OF crel] have "sorted (get_array a h')" - unfolding Heap.length_def by simp + unfolding Array.length_def by simp with quicksort_permutes [OF crel] properties_for_sort show ?thesis by fastsimp qed @@ -576,7 +576,7 @@ We will now show that exceptions do not occur. *} lemma noError_part1: - assumes "l < Heap.length a h" "r < Heap.length a h" + assumes "l < Array.length a h" "r < Array.length a h" shows "noError (part1 a l r p) h" using assms proof (induct a l r p arbitrary: h rule: part1.induct) @@ -587,7 +587,7 @@ qed lemma noError_partition: - assumes "l < r" "l < Heap.length a h" "r < Heap.length a h" + assumes "l < r" "l < Array.length a h" "r < Array.length a h" shows "noError (partition a l r) h" using assms unfolding partition.simps swap_def @@ -603,7 +603,7 @@ done lemma noError_quicksort: - assumes "l < Heap.length a h" "r < Heap.length a h" + assumes "l < Array.length a h" "r < Array.length a h" shows "noError (quicksort a l r) h" using assms proof (induct a l r arbitrary: h rule: quicksort.induct) @@ -628,7 +628,7 @@ subsection {* Example *} definition "qsort a = do - k \ length a; + k \ len a; quicksort a 0 (k - 1); return a done" diff -r 6432bf0d7191 -r bf6c1216db43 src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy --- a/src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy Mon Jul 05 09:14:51 2010 -0700 +++ b/src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy Tue Jul 06 08:08:35 2010 -0700 @@ -37,7 +37,7 @@ else get_array a h ! k)" using assms unfolding swap.simps by (elim crel_elim_all) - (auto simp: Heap.length_def) + (auto simp: length_def) lemma rev_pointwise: assumes "crel (rev a i j) h h' r" shows "get_array a h' ! k = (if k < i then get_array a h ! k @@ -69,7 +69,7 @@ lemma rev_length: assumes "crel (rev a i j) h h' r" - shows "Heap.length a h = Heap.length a h'" + shows "Array.length a h = Array.length a h'" using assms proof (induct a i j arbitrary: h h' rule: rev.induct) case (1 a i j h h'') @@ -93,7 +93,7 @@ qed lemma rev2_rev': assumes "crel (rev a i j) h h' u" - assumes "j < Heap.length a h" + assumes "j < Array.length a h" shows "subarray i (j + 1) a h' = List.rev (subarray i (j + 1) a h)" proof - { @@ -103,15 +103,15 @@ by auto } with assms(2) rev_length[OF assms(1)] show ?thesis - unfolding subarray_def Heap.length_def + unfolding subarray_def Array.length_def by (auto simp add: length_sublist' rev_nth min_def nth_sublist' intro!: nth_equalityI) qed lemma rev2_rev: - assumes "crel (rev a 0 (Heap.length a h - 1)) h h' u" + assumes "crel (rev a 0 (Array.length a h - 1)) h h' u" shows "get_array a h' = List.rev (get_array a h)" using rev2_rev'[OF assms] rev_length[OF assms] assms - by (cases "Heap.length a h = 0", auto simp add: Heap.length_def + by (cases "Array.length a h = 0", auto simp add: Array.length_def subarray_def sublist'_all rev.simps[where j=0] elim!: crel_elim_all) (drule sym[of "List.length (get_array a h)"], simp) diff -r 6432bf0d7191 -r bf6c1216db43 src/HOL/Imperative_HOL/ex/Linked_Lists.thy --- a/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Mon Jul 05 09:14:51 2010 -0700 +++ b/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Tue Jul 06 08:08:35 2010 -0700 @@ -13,7 +13,7 @@ setup {* Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \ 'a\type ref"}) *} datatype 'a node = Empty | Node 'a "('a node) ref" -fun +primrec node_encode :: "'a\countable node \ nat" where "node_encode Empty = 0" @@ -28,11 +28,11 @@ instance node :: (heap) heap .. -fun make_llist :: "'a\heap list \ 'a node Heap" +primrec make_llist :: "'a\heap list \ 'a node Heap" where [simp del]: "make_llist [] = return Empty" | "make_llist (x#xs) = do tl \ make_llist xs; - next \ Ref.new tl; + next \ ref tl; return (Node x next) done" @@ -56,31 +56,31 @@ return (x#xs) done" unfolding traverse_def -by (auto simp: traverse_def monad_simp MREC_rule) +by (auto simp: traverse_def MREC_rule) section {* Proving correctness with relational abstraction *} subsection {* Definition of list_of, list_of', refs_of and refs_of' *} -fun list_of :: "heap \ ('a::heap) node \ 'a list \ bool" +primrec list_of :: "heap \ ('a::heap) node \ 'a list \ bool" where "list_of h r [] = (r = Empty)" -| "list_of h r (a#as) = (case r of Empty \ False | Node b bs \ (a = b \ list_of h (get_ref bs h) as))" +| "list_of h r (a#as) = (case r of Empty \ False | Node b bs \ (a = b \ list_of h (Ref.get h bs) as))" definition list_of' :: "heap \ ('a::heap) node ref \ 'a list \ bool" where - "list_of' h r xs = list_of h (get_ref r h) xs" + "list_of' h r xs = list_of h (Ref.get h r) xs" -fun refs_of :: "heap \ ('a::heap) node \ 'a node ref list \ bool" +primrec refs_of :: "heap \ ('a::heap) node \ 'a node ref list \ bool" where "refs_of h r [] = (r = Empty)" -| "refs_of h r (x#xs) = (case r of Empty \ False | Node b bs \ (x = bs) \ refs_of h (get_ref bs h) xs)" +| "refs_of h r (x#xs) = (case r of Empty \ False | Node b bs \ (x = bs) \ refs_of h (Ref.get h bs) xs)" -fun refs_of' :: "heap \ ('a::heap) node ref \ 'a node ref list \ bool" +primrec refs_of' :: "heap \ ('a::heap) node ref \ 'a node ref list \ bool" where "refs_of' h r [] = False" -| "refs_of' h r (x#xs) = ((x = r) \ refs_of h (get_ref x h) xs)" +| "refs_of' h r (x#xs) = ((x = r) \ refs_of h (Ref.get h x) xs)" subsection {* Properties of these definitions *} @@ -88,35 +88,35 @@ lemma list_of_Empty[simp]: "list_of h Empty xs = (xs = [])" by (cases xs, auto) -lemma list_of_Node[simp]: "list_of h (Node x ps) xs = (\xs'. (xs = x # xs') \ list_of h (get_ref ps h) xs')" +lemma list_of_Node[simp]: "list_of h (Node x ps) xs = (\xs'. (xs = x # xs') \ list_of h (Ref.get h ps) xs')" by (cases xs, auto) -lemma list_of'_Empty[simp]: "get_ref q h = Empty \ list_of' h q xs = (xs = [])" +lemma list_of'_Empty[simp]: "Ref.get h q = Empty \ list_of' h q xs = (xs = [])" unfolding list_of'_def by simp -lemma list_of'_Node[simp]: "get_ref q h = Node x ps \ list_of' h q xs = (\xs'. (xs = x # xs') \ list_of' h ps xs')" +lemma list_of'_Node[simp]: "Ref.get h q = Node x ps \ list_of' h q xs = (\xs'. (xs = x # xs') \ list_of' h ps xs')" unfolding list_of'_def by simp -lemma list_of'_Nil: "list_of' h q [] \ get_ref q h = Empty" +lemma list_of'_Nil: "list_of' h q [] \ Ref.get h q = Empty" unfolding list_of'_def by simp lemma list_of'_Cons: assumes "list_of' h q (x#xs)" -obtains n where "get_ref q h = Node x n" and "list_of' h n xs" +obtains n where "Ref.get h q = Node x n" and "list_of' h n xs" using assms unfolding list_of'_def by (auto split: node.split_asm) lemma refs_of_Empty[simp] : "refs_of h Empty xs = (xs = [])" by (cases xs, auto) -lemma refs_of_Node[simp]: "refs_of h (Node x ps) xs = (\prs. xs = ps # prs \ refs_of h (get_ref ps h) prs)" +lemma refs_of_Node[simp]: "refs_of h (Node x ps) xs = (\prs. xs = ps # prs \ refs_of h (Ref.get h ps) prs)" by (cases xs, auto) -lemma refs_of'_def': "refs_of' h p ps = (\prs. (ps = (p # prs)) \ refs_of h (get_ref p h) prs)" +lemma refs_of'_def': "refs_of' h p ps = (\prs. (ps = (p # prs)) \ refs_of h (Ref.get h p) prs)" by (cases ps, auto) lemma refs_of'_Node: assumes "refs_of' h p xs" - assumes "get_ref p h = Node x pn" + assumes "Ref.get h p = Node x pn" obtains pnrs where "xs = p # pnrs" and "refs_of' h pn pnrs" using assms @@ -166,7 +166,7 @@ assumes "list_of' h r xs" shows "\rs. refs_of' h r rs" proof - - from assms obtain rs' where "refs_of h (get_ref r h) rs'" + from assms obtain rs' where "refs_of h (Ref.get h r) rs'" unfolding list_of'_def by (rule list_of_refs_of) thus ?thesis unfolding refs_of'_def' by auto qed @@ -238,7 +238,7 @@ done lemma refs_of_next: -assumes "refs_of h (get_ref p h) rs" +assumes "refs_of h (Ref.get h p) rs" shows "p \ set rs" proof (rule ccontr) assume a: "\ (p \ set rs)" @@ -264,7 +264,7 @@ subsection {* Interaction of these predicates with our heap transitions *} -lemma list_of_set_ref: "refs_of h q rs \ p \ set rs \ list_of (set_ref p v h) q as = list_of h q as" +lemma list_of_set_ref: "refs_of h q rs \ p \ set rs \ list_of (Ref.set p v h) q as = list_of h q as" using assms proof (induct as arbitrary: q rs) case Nil thus ?case by simp @@ -275,15 +275,15 @@ case Empty thus ?thesis by auto next case (Node a ref) - from Cons(2) Node obtain rs' where 1: "refs_of h (get_ref ref h) rs'" and rs_rs': "rs = ref # rs'" by auto + from Cons(2) Node obtain rs' where 1: "refs_of h (Ref.get h ref) rs'" and rs_rs': "rs = ref # rs'" by auto from Cons(3) rs_rs' have "ref \ p" by fastsimp - hence ref_eq: "get_ref ref (set_ref p v h) = (get_ref ref h)" by (auto simp add: ref_get_set_neq) + hence ref_eq: "Ref.get (Ref.set p v h) ref = (Ref.get h ref)" by (auto simp add: Ref.get_set_neq) from rs_rs' Cons(3) have 2: "p \ set rs'" by simp from Cons.hyps[OF 1 2] Node ref_eq show ?thesis by simp qed qed -lemma refs_of_set_ref: "refs_of h q rs \ p \ set rs \ refs_of (set_ref p v h) q as = refs_of h q as" +lemma refs_of_set_ref: "refs_of h q rs \ p \ set rs \ refs_of (Ref.set p v h) q as = refs_of h q as" proof (induct as arbitrary: q rs) case Nil thus ?case by simp next @@ -293,15 +293,15 @@ case Empty thus ?thesis by auto next case (Node a ref) - from Cons(2) Node obtain rs' where 1: "refs_of h (get_ref ref h) rs'" and rs_rs': "rs = ref # rs'" by auto + from Cons(2) Node obtain rs' where 1: "refs_of h (Ref.get h ref) rs'" and rs_rs': "rs = ref # rs'" by auto from Cons(3) rs_rs' have "ref \ p" by fastsimp - hence ref_eq: "get_ref ref (set_ref p v h) = (get_ref ref h)" by (auto simp add: ref_get_set_neq) + hence ref_eq: "Ref.get (Ref.set p v h) ref = (Ref.get h ref)" by (auto simp add: Ref.get_set_neq) from rs_rs' Cons(3) have 2: "p \ set rs'" by simp from Cons.hyps[OF 1 2] Node ref_eq show ?thesis by auto qed qed -lemma refs_of_set_ref2: "refs_of (set_ref p v h) q rs \ p \ set rs \ refs_of (set_ref p v h) q rs = refs_of h q rs" +lemma refs_of_set_ref2: "refs_of (Ref.set p v h) q rs \ p \ set rs \ refs_of (Ref.set p v h) q rs = refs_of h q rs" proof (induct rs arbitrary: q) case Nil thus ?case by simp next @@ -311,9 +311,9 @@ case Empty thus ?thesis by auto next case (Node a ref) - from Cons(2) Node have 1:"refs_of (set_ref p v h) (get_ref ref (set_ref p v h)) xs" and x_ref: "x = ref" by auto + from Cons(2) Node have 1:"refs_of (Ref.set p v h) (Ref.get (Ref.set p v h) ref) xs" and x_ref: "x = ref" by auto from Cons(3) this have "ref \ p" by fastsimp - hence ref_eq: "get_ref ref (set_ref p v h) = (get_ref ref h)" by (auto simp add: ref_get_set_neq) + hence ref_eq: "Ref.get (Ref.set p v h) ref = (Ref.get h ref)" by (auto simp add: Ref.get_set_neq) from Cons(3) have 2: "p \ set xs" by simp with Cons.hyps 1 2 Node ref_eq show ?thesis by simp @@ -323,7 +323,7 @@ lemma list_of'_set_ref: assumes "refs_of' h q rs" assumes "p \ set rs" - shows "list_of' (set_ref p v h) q as = list_of' h q as" + shows "list_of' (Ref.set p v h) q as = list_of' h q as" proof - from assms have "q \ p" by (auto simp only: dest!: refs_of'E) with assms show ?thesis @@ -333,18 +333,18 @@ lemma list_of'_set_next_ref_Node[simp]: assumes "list_of' h r xs" - assumes "get_ref p h = Node x r'" + assumes "Ref.get h p = Node x r'" assumes "refs_of' h r rs" assumes "p \ set rs" - shows "list_of' (set_ref p (Node x r) h) p (x#xs) = list_of' h r xs" + shows "list_of' (Ref.set p (Node x r) h) p (x#xs) = list_of' h r xs" using assms unfolding list_of'_def refs_of'_def' -by (auto simp add: list_of_set_ref noteq_refs_sym) +by (auto simp add: list_of_set_ref Ref.noteq_sym) lemma refs_of'_set_ref: assumes "refs_of' h q rs" assumes "p \ set rs" - shows "refs_of' (set_ref p v h) q as = refs_of' h q as" + shows "refs_of' (Ref.set p v h) q as = refs_of' h q as" using assms proof - from assms have "q \ p" by (auto simp only: dest!: refs_of'E) @@ -354,9 +354,9 @@ qed lemma refs_of'_set_ref2: - assumes "refs_of' (set_ref p v h) q rs" + assumes "refs_of' (Ref.set p v h) q rs" assumes "p \ set rs" - shows "refs_of' (set_ref p v h) q as = refs_of' h q as" + shows "refs_of' (Ref.set p v h) q as = refs_of' h q as" using assms proof - from assms have "q \ p" by (auto simp only: dest!: refs_of'E) @@ -364,7 +364,7 @@ unfolding refs_of'_def' apply auto apply (subgoal_tac "prs = prsa") - apply (insert refs_of_set_ref2[of p v h "get_ref q h"]) + apply (insert refs_of_set_ref2[of p v h "Ref.get h q"]) apply (erule_tac x="prs" in meta_allE) apply auto apply (auto dest: refs_of_is_fun) @@ -372,15 +372,15 @@ qed lemma refs_of'_set_next_ref: -assumes "get_ref p h1 = Node x pn" -assumes "refs_of' (set_ref p (Node x r1) h1) p rs" +assumes "Ref.get h1 p = Node x pn" +assumes "refs_of' (Ref.set p (Node x r1) h1) p rs" obtains r1s where "rs = (p#r1s)" and "refs_of' h1 r1 r1s" using assms proof - from assms refs_of'_distinct[OF assms(2)] have "\ r1s. rs = (p # r1s) \ refs_of' h1 r1 r1s" apply - unfolding refs_of'_def'[of _ p] - apply (auto, frule refs_of_set_ref2) by (auto dest: noteq_refs_sym) + apply (auto, frule refs_of_set_ref2) by (auto dest: Ref.noteq_sym) with prems show thesis by auto qed @@ -388,7 +388,7 @@ lemma refs_of_invariant: assumes "refs_of h (r::('a::heap) node) xs" - assumes "\refs. refs_of h r refs \ (\ref \ set refs. ref_present ref h \ ref_present ref h' \ get_ref ref h = get_ref ref h')" + assumes "\refs. refs_of h r refs \ (\ref \ set refs. Ref.present h ref \ Ref.present h' ref \ Ref.get h ref = Ref.get h' ref)" shows "refs_of h' r xs" using assms proof (induct xs arbitrary: r) @@ -396,28 +396,28 @@ next case (Cons x xs') from Cons(2) obtain v where Node: "r = Node v x" by (cases r, auto) - from Cons(2) Node have refs_of_next: "refs_of h (get_ref x h) xs'" by simp - from Cons(2-3) Node have ref_eq: "get_ref x h = get_ref x h'" by auto - from ref_eq refs_of_next have 1: "refs_of h (get_ref x h') xs'" by simp - from Cons(2) Cons(3) have "\ref \ set xs'. ref_present ref h \ ref_present ref h' \ get_ref ref h = get_ref ref h'" + from Cons(2) Node have refs_of_next: "refs_of h (Ref.get h x) xs'" by simp + from Cons(2-3) Node have ref_eq: "Ref.get h x = Ref.get h' x" by auto + from ref_eq refs_of_next have 1: "refs_of h (Ref.get h' x) xs'" by simp + from Cons(2) Cons(3) have "\ref \ set xs'. Ref.present h ref \ Ref.present h' ref \ Ref.get h ref = Ref.get h' ref" by fastsimp - with Cons(3) 1 have 2: "\refs. refs_of h (get_ref x h') refs \ (\ref \ set refs. ref_present ref h \ ref_present ref h' \ get_ref ref h = get_ref ref h')" + with Cons(3) 1 have 2: "\refs. refs_of h (Ref.get h' x) refs \ (\ref \ set refs. Ref.present h ref \ Ref.present h' ref \ Ref.get h ref = Ref.get h' ref)" by (fastsimp dest: refs_of_is_fun) - from Cons.hyps[OF 1 2] have "refs_of h' (get_ref x h') xs'" . + from Cons.hyps[OF 1 2] have "refs_of h' (Ref.get h' x) xs'" . with Node show ?case by simp qed lemma refs_of'_invariant: assumes "refs_of' h r xs" - assumes "\refs. refs_of' h r refs \ (\ref \ set refs. ref_present ref h \ ref_present ref h' \ get_ref ref h = get_ref ref h')" + assumes "\refs. refs_of' h r refs \ (\ref \ set refs. Ref.present h ref \ Ref.present h' ref \ Ref.get h ref = Ref.get h' ref)" shows "refs_of' h' r xs" using assms proof - - from assms obtain prs where refs:"refs_of h (get_ref r h) prs" and xs_def: "xs = r # prs" + from assms obtain prs where refs:"refs_of h (Ref.get h r) prs" and xs_def: "xs = r # prs" unfolding refs_of'_def' by auto - from xs_def assms have x_eq: "get_ref r h = get_ref r h'" by fastsimp - from refs assms xs_def have 2: "\refs. refs_of h (get_ref r h) refs \ - (\ref\set refs. ref_present ref h \ ref_present ref h' \ get_ref ref h = get_ref ref h')" + from xs_def assms have x_eq: "Ref.get h r = Ref.get h' r" by fastsimp + from refs assms xs_def have 2: "\refs. refs_of h (Ref.get h r) refs \ + (\ref\set refs. Ref.present h ref \ Ref.present h' ref \ Ref.get h ref = Ref.get h' ref)" by (fastsimp dest: refs_of_is_fun) from refs_of_invariant [OF refs 2] xs_def x_eq show ?thesis unfolding refs_of'_def' by auto @@ -425,7 +425,7 @@ lemma list_of_invariant: assumes "list_of h (r::('a::heap) node) xs" - assumes "\refs. refs_of h r refs \ (\ref \ set refs. ref_present ref h \ ref_present ref h' \ get_ref ref h = get_ref ref h')" + assumes "\refs. refs_of h r refs \ (\ref \ set refs. Ref.present h ref \ Ref.present h' ref \ Ref.get h ref = Ref.get h' ref)" shows "list_of h' r xs" using assms proof (induct xs arbitrary: r) @@ -437,16 +437,16 @@ by (cases r, auto) from Cons(2) obtain rs where rs_def: "refs_of h r rs" by (rule list_of_refs_of) from Node rs_def obtain rss where refs_of: "refs_of h r (ref#rss)" and rss_def: "rs = ref#rss" by auto - from Cons(3) Node refs_of have ref_eq: "get_ref ref h = get_ref ref h'" + from Cons(3) Node refs_of have ref_eq: "Ref.get h ref = Ref.get h' ref" by auto - from Cons(2) ref_eq Node have 1: "list_of h (get_ref ref h') xs'" by simp - from refs_of Node ref_eq have refs_of_ref: "refs_of h (get_ref ref h') rss" by simp - from Cons(3) rs_def have rs_heap_eq: "\ref\set rs. ref_present ref h \ ref_present ref h' \ get_ref ref h = get_ref ref h'" by simp - from refs_of_ref rs_heap_eq rss_def have 2: "\refs. refs_of h (get_ref ref h') refs \ - (\ref\set refs. ref_present ref h \ ref_present ref h' \ get_ref ref h = get_ref ref h')" + from Cons(2) ref_eq Node have 1: "list_of h (Ref.get h' ref) xs'" by simp + from refs_of Node ref_eq have refs_of_ref: "refs_of h (Ref.get h' ref) rss" by simp + from Cons(3) rs_def have rs_heap_eq: "\ref\set rs. Ref.present h ref \ Ref.present h' ref \ Ref.get h ref = Ref.get h' ref" by simp + from refs_of_ref rs_heap_eq rss_def have 2: "\refs. refs_of h (Ref.get h' ref) refs \ + (\ref\set refs. Ref.present h ref \ Ref.present h' ref \ Ref.get h ref = Ref.get h' ref)" by (auto dest: refs_of_is_fun) from Cons(1)[OF 1 2] - have "list_of h' (get_ref ref h') xs'" . + have "list_of h' (Ref.get h' ref) xs'" . with Node show ?case unfolding list_of'_def by simp @@ -454,29 +454,29 @@ lemma make_llist: assumes "crel (make_llist xs) h h' r" -shows "list_of h' r xs \ (\rs. refs_of h' r rs \ (\ref \ (set rs). ref_present ref h'))" +shows "list_of h' r xs \ (\rs. refs_of h' r rs \ (\ref \ (set rs). Ref.present h' ref))" using assms proof (induct xs arbitrary: h h' r) case Nil thus ?case by (auto elim: crel_return simp add: make_llist.simps) next case (Cons x xs') from Cons.prems obtain h1 r1 r' where make_llist: "crel (make_llist xs') h h1 r1" - and crel_refnew:"crel (Ref.new r1) h1 h' r'" and Node: "r = Node x r'" + and crel_refnew:"crel (ref r1) h1 h' r'" and Node: "r = Node x r'" unfolding make_llist.simps by (auto elim!: crelE crel_return) from Cons.hyps[OF make_llist] have list_of_h1: "list_of h1 r1 xs'" .. from Cons.hyps[OF make_llist] obtain rs' where rs'_def: "refs_of h1 r1 rs'" by (auto intro: list_of_refs_of) - from Cons.hyps[OF make_llist] rs'_def have refs_present: "\ref\set rs'. ref_present ref h1" by simp + from Cons.hyps[OF make_llist] rs'_def have refs_present: "\ref\set rs'. Ref.present h1 ref" by simp from crel_refnew rs'_def refs_present have refs_unchanged: "\refs. refs_of h1 r1 refs \ - (\ref\set refs. ref_present ref h1 \ ref_present ref h' \ get_ref ref h1 = get_ref ref h')" - by (auto elim!: crel_Ref_new dest: refs_of_is_fun) + (\ref\set refs. Ref.present h1 ref \ Ref.present h' ref \ Ref.get h1 ref = Ref.get h' ref)" + by (auto elim!: crel_ref dest: refs_of_is_fun) with list_of_invariant[OF list_of_h1 refs_unchanged] Node crel_refnew have fstgoal: "list_of h' r (x # xs')" unfolding list_of.simps - by (auto elim!: crel_Ref_new) - from refs_unchanged rs'_def have refs_still_present: "\ref\set rs'. ref_present ref h'" by auto + by (auto elim!: crel_ref) + from refs_unchanged rs'_def have refs_still_present: "\ref\set rs'. Ref.present h' ref" by auto from refs_of_invariant[OF rs'_def refs_unchanged] refs_unchanged Node crel_refnew refs_still_present - have sndgoal: "\rs. refs_of h' r rs \ (\ref\set rs. ref_present ref h')" - by (fastsimp elim!: crel_Ref_new dest: refs_of_is_fun) + have sndgoal: "\rs. refs_of h' r rs \ (\ref\set rs. Ref.present h' ref)" + by (fastsimp elim!: crel_ref dest: refs_of_is_fun) from fstgoal sndgoal show ?case .. qed @@ -531,12 +531,12 @@ done" unfolding rev'_def MREC_rule[of _ _ "(q, p)"] unfolding rev'_def[symmetric] thm arg_cong2 - by (auto simp add: monad_simp expand_fun_eq intro: arg_cong2[where f = "op \="] split: node.split) + by (auto simp add: expand_fun_eq intro: arg_cong2[where f = "op \="] split: node.split) -fun rev :: "('a:: heap) node \ 'a node Heap" +primrec rev :: "('a:: heap) node \ 'a node Heap" where "rev Empty = return Empty" -| "rev (Node x n) = (do q \ Ref.new Empty; p \ Ref.new (Node x n); v \ rev' (q, p); !v done)" +| "rev (Node x n) = (do q \ ref Empty; p \ ref (Node x n); v \ rev' (q, p); !v done)" subsection {* Correctness Proof *} @@ -556,17 +556,17 @@ case (Cons x xs) (*"LinkedList.list_of h' (get_ref v h') (List.rev xs @ x # qsa)"*) from Cons(4) obtain ref where - p_is_Node: "get_ref p h = Node x ref" + p_is_Node: "Ref.get h p = Node x ref" (*and "ref_present ref h"*) and list_of'_ref: "list_of' h ref xs" - unfolding list_of'_def by (cases "get_ref p h", auto) - from p_is_Node Cons(2) have crel_rev': "crel (rev' (p, ref)) (set_ref p (Node x q) h) h' v" + unfolding list_of'_def by (cases "Ref.get h p", auto) + from p_is_Node Cons(2) have crel_rev': "crel (rev' (p, ref)) (Ref.set p (Node x q) h) h' v" by (auto simp add: rev'_simps [of q p] elim!: crelE crel_lookup crel_update) from Cons(3) obtain qrs where qrs_def: "refs_of' h q qrs" by (elim list_of'_refs_of') from Cons(4) obtain prs where prs_def: "refs_of' h p prs" by (elim list_of'_refs_of') from qrs_def prs_def Cons(5) have distinct_pointers: "set qrs \ set prs = {}" by fastsimp from qrs_def prs_def distinct_pointers refs_of'E have p_notin_qrs: "p \ set qrs" by fastsimp - from Cons(3) qrs_def this have 1: "list_of' (set_ref p (Node x q) h) p (x#qs)" + from Cons(3) qrs_def this have 1: "list_of' (Ref.set p (Node x q) h) p (x#qs)" unfolding list_of'_def apply (simp) unfolding list_of'_def[symmetric] @@ -575,16 +575,16 @@ unfolding refs_of'_def' by auto from prs_refs prs_def have p_not_in_refs: "p \ set refs" by (fastsimp dest!: refs_of'_distinct) - with refs_def p_is_Node list_of'_ref have 2: "list_of' (set_ref p (Node x q) h) ref xs" + with refs_def p_is_Node list_of'_ref have 2: "list_of' (Ref.set p (Node x q) h) ref xs" by (auto simp add: list_of'_set_ref) - from p_notin_qrs qrs_def have refs_of1: "refs_of' (set_ref p (Node x q) h) p (p#qrs)" + from p_notin_qrs qrs_def have refs_of1: "refs_of' (Ref.set p (Node x q) h) p (p#qrs)" unfolding refs_of'_def' apply (simp) unfolding refs_of'_def'[symmetric] by (simp add: refs_of'_set_ref) - from p_not_in_refs p_is_Node refs_def have refs_of2: "refs_of' (set_ref p (Node x q) h) ref refs" + from p_not_in_refs p_is_Node refs_def have refs_of2: "refs_of' (Ref.set p (Node x q) h) ref refs" by (simp add: refs_of'_set_ref) - from p_not_in_refs refs_of1 refs_of2 distinct_pointers prs_refs have 3: "\qrs prs. refs_of' (set_ref p (Node x q) h) p qrs \ refs_of' (set_ref p (Node x q) h) ref prs \ set prs \ set qrs = {}" + from p_not_in_refs refs_of1 refs_of2 distinct_pointers prs_refs have 3: "\qrs prs. refs_of' (Ref.set p (Node x q) h) p qrs \ refs_of' (Ref.set p (Node x q) h) ref prs \ set prs \ set qrs = {}" apply - apply (rule allI)+ apply (rule impI) apply (erule conjE) apply (drule refs_of'_is_fun) back back apply assumption apply (drule refs_of'_is_fun) back back apply assumption @@ -595,7 +595,7 @@ lemma rev_correctness: assumes list_of_h: "list_of h r xs" - assumes validHeap: "\refs. refs_of h r refs \ (\r \ set refs. ref_present r h)" + assumes validHeap: "\refs. refs_of h r refs \ (\r \ set refs. Ref.present h r)" assumes crel_rev: "crel (rev r) h h' r'" shows "list_of h' r' (List.rev xs)" using assms @@ -606,39 +606,39 @@ next case (Node x ps) with crel_rev obtain p q h1 h2 h3 v where - init: "crel (Ref.new Empty) h h1 q" - "crel (Ref.new (Node x ps)) h1 h2 p" + init: "crel (ref Empty) h h1 q" + "crel (ref (Node x ps)) h1 h2 p" and crel_rev':"crel (rev' (q, p)) h2 h3 v" and lookup: "crel (!v) h3 h' r'" using rev.simps by (auto elim!: crelE) from init have a1:"list_of' h2 q []" unfolding list_of'_def - by (auto elim!: crel_Ref_new) + by (auto elim!: crel_ref) from list_of_h obtain refs where refs_def: "refs_of h r refs" by (rule list_of_refs_of) - from validHeap init refs_def have heap_eq: "\refs. refs_of h r refs \ (\ref\set refs. ref_present ref h \ ref_present ref h2 \ get_ref ref h = get_ref ref h2)" - by (fastsimp elim!: crel_Ref_new dest: refs_of_is_fun) + from validHeap init refs_def have heap_eq: "\refs. refs_of h r refs \ (\ref\set refs. Ref.present h ref \ Ref.present h2 ref \ Ref.get h ref = Ref.get h2 ref)" + by (fastsimp elim!: crel_ref dest: refs_of_is_fun) from list_of_invariant[OF list_of_h heap_eq] have "list_of h2 r xs" . from init this Node have a2: "list_of' h2 p xs" apply - unfolding list_of'_def - apply (auto elim!: crel_Ref_new) + apply (auto elim!: crel_ref) done from init have refs_of_q: "refs_of' h2 q [q]" - by (auto elim!: crel_Ref_new) + by (auto elim!: crel_ref) from refs_def Node have refs_of'_ps: "refs_of' h ps refs" by (auto simp add: refs_of'_def'[symmetric]) - from validHeap refs_def have all_ref_present: "\r\set refs. ref_present r h" by simp - from init refs_of'_ps Node this have heap_eq: "\refs. refs_of' h ps refs \ (\ref\set refs. ref_present ref h \ ref_present ref h2 \ get_ref ref h = get_ref ref h2)" - by (fastsimp elim!: crel_Ref_new dest: refs_of'_is_fun) + from validHeap refs_def have all_ref_present: "\r\set refs. Ref.present h r" by simp + from init refs_of'_ps Node this have heap_eq: "\refs. refs_of' h ps refs \ (\ref\set refs. Ref.present h ref \ Ref.present h2 ref \ Ref.get h ref = Ref.get h2 ref)" + by (fastsimp elim!: crel_ref dest: refs_of'_is_fun) from refs_of'_invariant[OF refs_of'_ps this] have "refs_of' h2 ps refs" . with init have refs_of_p: "refs_of' h2 p (p#refs)" - by (auto elim!: crel_Ref_new simp add: refs_of'_def') + by (auto elim!: crel_ref simp add: refs_of'_def') with init all_ref_present have q_is_new: "q \ set (p#refs)" - by (auto elim!: crel_Ref_new intro!: noteq_refsI) + by (auto elim!: crel_ref intro!: Ref.noteq_I) from refs_of_p refs_of_q q_is_new have a3: "\qrs prs. refs_of' h2 q qrs \ refs_of' h2 p prs \ set prs \ set qrs = {}" by (fastsimp simp only: set.simps dest: refs_of'_is_fun) - from rev'_invariant [OF crel_rev' a1 a2 a3] have "list_of h3 (get_ref v h3) (List.rev xs)" + from rev'_invariant [OF crel_rev' a1 a2 a3] have "list_of h3 (Ref.get h3 v) (List.rev xs)" unfolding list_of'_def by auto with lookup show ?thesis by (auto elim: crel_lookup) @@ -734,32 +734,32 @@ lemma merge_induct2: assumes "list_of' h (p::'a::{heap, ord} node ref) xs" assumes "list_of' h q ys" - assumes "\ ys p q. \ list_of' h p []; list_of' h q ys; get_ref p h = Empty \ \ P p q [] ys" - assumes "\ x xs' p q pn. \ list_of' h p (x#xs'); list_of' h q []; get_ref p h = Node x pn; get_ref q h = Empty \ \ P p q (x#xs') []" + assumes "\ ys p q. \ list_of' h p []; list_of' h q ys; Ref.get h p = Empty \ \ P p q [] ys" + assumes "\ x xs' p q pn. \ list_of' h p (x#xs'); list_of' h q []; Ref.get h p = Node x pn; Ref.get h q = Empty \ \ P p q (x#xs') []" assumes "\ x xs' y ys' p q pn qn. - \ list_of' h p (x#xs'); list_of' h q (y#ys'); get_ref p h = Node x pn; get_ref q h = Node y qn; + \ list_of' h p (x#xs'); list_of' h q (y#ys'); Ref.get h p = Node x pn; Ref.get h q = Node y qn; x \ y; P pn q xs' (y#ys') \ \ P p q (x#xs') (y#ys')" assumes "\ x xs' y ys' p q pn qn. - \ list_of' h p (x#xs'); list_of' h q (y#ys'); get_ref p h = Node x pn; get_ref q h = Node y qn; + \ list_of' h p (x#xs'); list_of' h q (y#ys'); Ref.get h p = Node x pn; Ref.get h q = Node y qn; \ x \ y; P p qn (x#xs') ys'\ \ P p q (x#xs') (y#ys')" shows "P p q xs ys" using assms(1-2) proof (induct xs ys arbitrary: p q rule: Lmerge.induct) case (2 ys) - from 2(1) have "get_ref p h = Empty" unfolding list_of'_def by simp + from 2(1) have "Ref.get h p = Empty" unfolding list_of'_def by simp with 2(1-2) assms(3) show ?case by blast next case (3 x xs') - from 3(1) obtain pn where Node: "get_ref p h = Node x pn" by (rule list_of'_Cons) - from 3(2) have "get_ref q h = Empty" unfolding list_of'_def by simp + from 3(1) obtain pn where Node: "Ref.get h p = Node x pn" by (rule list_of'_Cons) + from 3(2) have "Ref.get h q = Empty" unfolding list_of'_def by simp with Node 3(1-2) assms(4) show ?case by blast next case (1 x xs' y ys') - from 1(3) obtain pn where pNode:"get_ref p h = Node x pn" + from 1(3) obtain pn where pNode:"Ref.get h p = Node x pn" and list_of'_pn: "list_of' h pn xs'" by (rule list_of'_Cons) - from 1(4) obtain qn where qNode:"get_ref q h = Node y qn" + from 1(4) obtain qn where qNode:"Ref.get h q = Node y qn" and list_of'_qn: "list_of' h qn ys'" by (rule list_of'_Cons) show ?case proof (cases "x \ y") @@ -780,15 +780,15 @@ assumes "list_of' h p xs" assumes "list_of' h q ys" assumes "crel (merge p q) h h' r" -assumes "\ ys p q. \ list_of' h p []; list_of' h q ys; get_ref p h = Empty \ \ P p q h h q [] ys" -assumes "\ x xs' p q pn. \ list_of' h p (x#xs'); list_of' h q []; get_ref p h = Node x pn; get_ref q h = Empty \ \ P p q h h p (x#xs') []" +assumes "\ ys p q. \ list_of' h p []; list_of' h q ys; Ref.get h p = Empty \ \ P p q h h q [] ys" +assumes "\ x xs' p q pn. \ list_of' h p (x#xs'); list_of' h q []; Ref.get h p = Node x pn; Ref.get h q = Empty \ \ P p q h h p (x#xs') []" assumes "\ x xs' y ys' p q pn qn h1 r1 h'. - \ list_of' h p (x#xs'); list_of' h q (y#ys');get_ref p h = Node x pn; get_ref q h = Node y qn; - x \ y; crel (merge pn q) h h1 r1 ; P pn q h h1 r1 xs' (y#ys'); h' = set_ref p (Node x r1) h1 \ + \ list_of' h p (x#xs'); list_of' h q (y#ys');Ref.get h p = Node x pn; Ref.get h q = Node y qn; + x \ y; crel (merge pn q) h h1 r1 ; P pn q h h1 r1 xs' (y#ys'); h' = Ref.set p (Node x r1) h1 \ \ P p q h h' p (x#xs') (y#ys')" assumes "\ x xs' y ys' p q pn qn h1 r1 h'. - \ list_of' h p (x#xs'); list_of' h q (y#ys'); get_ref p h = Node x pn; get_ref q h = Node y qn; - \ x \ y; crel (merge p qn) h h1 r1; P p qn h h1 r1 (x#xs') ys'; h' = set_ref q (Node y r1) h1 \ + \ list_of' h p (x#xs'); list_of' h q (y#ys'); Ref.get h p = Node x pn; Ref.get h q = Node y qn; + \ x \ y; crel (merge p qn) h h1 r1; P p qn h h1 r1 (x#xs') ys'; h' = Ref.set q (Node y r1) h1 \ \ P p q h h' q (x#xs') (y#ys')" shows "P p q h h' r xs ys" using assms(3) @@ -808,7 +808,7 @@ case (3 x xs' y ys' p q pn qn) from 3(3-5) 3(7) obtain h1 r1 where 1: "crel (merge pn q) h h1 r1" - and 2: "h' = set_ref p (Node x r1) h1 \ r = p" + and 2: "h' = Ref.set p (Node x r1) h1 \ r = p" unfolding merge_simps[of p q] by (auto elim!: crel_lookup crelE crel_return crel_if crel_update) from 3(6)[OF 1] assms(6) [OF 3(1-5)] 1 2 show ?case by simp @@ -816,7 +816,7 @@ case (4 x xs' y ys' p q pn qn) from 4(3-5) 4(7) obtain h1 r1 where 1: "crel (merge p qn) h h1 r1" - and 2: "h' = set_ref q (Node y r1) h1 \ r = q" + and 2: "h' = Ref.set q (Node y r1) h1 \ r = q" unfolding merge_simps[of p q] by (auto elim!: crel_lookup crelE crel_return crel_if crel_update) from 4(6)[OF 1] assms(7) [OF 4(1-5)] 1 2 show ?case by simp @@ -834,7 +834,7 @@ assumes "crel (merge p q) h h' r'" assumes "set xs \ set ys = {}" assumes "r \ set xs \ set ys" - shows "get_ref r h = get_ref r h'" + shows "Ref.get h r = Ref.get h' r" proof - from assms(1) obtain ps where ps_def: "list_of' h p ps" by (rule refs_of'_list_of') from assms(2) obtain qs where qs_def: "list_of' h q qs" by (rule refs_of'_list_of') @@ -853,7 +853,8 @@ from pnrs_def 3(12) have "r \ p" by auto with 3(11) 3(12) pnrs_def refs_of'_distinct[OF 3(9)] have p_in: "p \ set pnrs \ set ys" by auto from 3(11) pnrs_def have no_inter: "set pnrs \ set ys = {}" by auto - from 3(7)[OF refs_of'_pn 3(10) this p_in] 3(3) have p_is_Node: "get_ref p h1 = Node x pn" by simp + from 3(7)[OF refs_of'_pn 3(10) this p_in] 3(3) have p_is_Node: "Ref.get h1 p = Node x pn" + by simp from 3(7)[OF refs_of'_pn 3(10) no_inter r_in] 3(8) `r \ p` show ?case by simp next @@ -866,7 +867,7 @@ from qnrs_def 4(12) have "r \ q" by auto with 4(11) 4(12) qnrs_def refs_of'_distinct[OF 4(10)] have q_in: "q \ set xs \ set qnrs" by auto from 4(11) qnrs_def have no_inter: "set xs \ set qnrs = {}" by auto - from 4(7)[OF 4(9) refs_of'_qn this q_in] 4(4) have q_is_Node: "get_ref q h1 = Node y qn" by simp + from 4(7)[OF 4(9) refs_of'_qn this q_in] 4(4) have q_is_Node: "Ref.get h1 q = Node y qn" by simp from 4(7)[OF 4(9) refs_of'_qn no_inter r_in] 4(8) `r \ q` show ?case by simp qed @@ -899,7 +900,7 @@ by (rule refs_of'_Node) from 3(10) 3(9) 3(11) pnrs_def refs_of'_distinct[OF 3(9)] have p_in: "p \ set pnrs \ set ys" by auto from 3(11) pnrs_def have no_inter: "set pnrs \ set ys = {}" by auto - from merge_unchanged[OF refs_of'_pn 3(10) 3(6) no_inter p_in] have p_stays: "get_ref p h1 = get_ref p h" .. + from merge_unchanged[OF refs_of'_pn 3(10) 3(6) no_inter p_in] have p_stays: "Ref.get h1 p = Ref.get h p" .. from 3 p_stays obtain r1s where rs_def: "rs = p#r1s" and refs_of'_r1:"refs_of' h1 r1 r1s" by (auto elim: refs_of'_set_next_ref) @@ -912,7 +913,7 @@ by (rule refs_of'_Node) from 4(10) 4(9) 4(11) qnrs_def refs_of'_distinct[OF 4(10)] have q_in: "q \ set xs \ set qnrs" by auto from 4(11) qnrs_def have no_inter: "set xs \ set qnrs = {}" by auto - from merge_unchanged[OF 4(9) refs_of'_qn 4(6) no_inter q_in] have q_stays: "get_ref q h1 = get_ref q h" .. + from merge_unchanged[OF 4(9) refs_of'_qn 4(6) no_inter q_in] have q_stays: "Ref.get h1 q = Ref.get h q" .. from 4 q_stays obtain r1s where rs_def: "rs = q#r1s" and refs_of'_r1:"refs_of' h1 r1 r1s" by (auto elim: refs_of'_set_next_ref) @@ -945,7 +946,7 @@ from prs_def qrs_def 3(9) pnrs_def have no_inter: "set pnrs \ set qrs = {}" by fastsimp from no_inter refs_of'_pn qrs_def have no_inter2: "\qrs prs. refs_of' h q qrs \ refs_of' h pn prs \ set prs \ set qrs = {}" by (fastsimp dest: refs_of'_is_fun) - from merge_unchanged[OF refs_of'_pn qrs_def 3(6) no_inter p_in] have p_stays: "get_ref p h1 = get_ref p h" .. + from merge_unchanged[OF refs_of'_pn qrs_def 3(6) no_inter p_in] have p_stays: "Ref.get h1 p = Ref.get h p" .. from 3(7)[OF no_inter2] obtain rs where rs_def: "refs_of' h1 r1 rs" by (rule list_of'_refs_of') from refs_of'_merge[OF refs_of'_pn qrs_def 3(6) no_inter this] p_in have p_rs: "p \ set rs" by auto with 3(7)[OF no_inter2] 3(1-5) 3(8) p_rs rs_def p_stays @@ -962,7 +963,7 @@ from prs_def qrs_def 4(9) qnrs_def have no_inter: "set prs \ set qnrs = {}" by fastsimp from no_inter refs_of'_qn prs_def have no_inter2: "\qrs prs. refs_of' h qn qrs \ refs_of' h p prs \ set prs \ set qrs = {}" by (fastsimp dest: refs_of'_is_fun) - from merge_unchanged[OF prs_def refs_of'_qn 4(6) no_inter q_in] have q_stays: "get_ref q h1 = get_ref q h" .. + from merge_unchanged[OF prs_def refs_of'_qn 4(6) no_inter q_in] have q_stays: "Ref.get h1 q = Ref.get h q" .. from 4(7)[OF no_inter2] obtain rs where rs_def: "refs_of' h1 r1 rs" by (rule list_of'_refs_of') from refs_of'_merge[OF prs_def refs_of'_qn 4(6) no_inter this] q_in have q_rs: "q \ set rs" by auto with 4(7)[OF no_inter2] 4(1-5) 4(8) q_rs rs_def q_stays @@ -984,8 +985,8 @@ (do ll_xs \ make_llist (filter (%n. n mod 2 = 0) [2..8]); ll_ys \ make_llist (filter (%n. n mod 2 = 1) [5..11]); - r \ Ref.new ll_xs; - q \ Ref.new ll_ys; + r \ ref ll_xs; + q \ ref ll_ys; p \ merge r q; ll_zs \ !p; zs \ traverse ll_zs; @@ -998,4 +999,4 @@ ML {* @{code test_2} () *} ML {* @{code test_3} () *} -end \ No newline at end of file +end diff -r 6432bf0d7191 -r bf6c1216db43 src/HOL/Imperative_HOL/ex/SatChecker.thy --- a/src/HOL/Imperative_HOL/ex/SatChecker.thy Mon Jul 05 09:14:51 2010 -0700 +++ b/src/HOL/Imperative_HOL/ex/SatChecker.thy Tue Jul 06 08:08:35 2010 -0700 @@ -118,6 +118,32 @@ text {* Specific definition for derived clauses in the Array *} +definition + array_ran :: "('a\heap) option array \ heap \ 'a set" where + "array_ran a h = {e. Some e \ set (get_array a h)}" + -- {*FIXME*} + +lemma array_ranI: "\ Some b = get_array a h ! i; i < Array.length a h \ \ b \ array_ran a h" +unfolding array_ran_def Array.length_def by simp + +lemma array_ran_upd_array_Some: + assumes "cl \ array_ran a (Array.change a i (Some b) h)" + shows "cl \ array_ran a h \ cl = b" +proof - + have "set (get_array a h[i := Some b]) \ insert (Some b) (set (get_array a h))" by (rule set_update_subset_insert) + with assms show ?thesis + unfolding array_ran_def Array.change_def by fastsimp +qed + +lemma array_ran_upd_array_None: + assumes "cl \ array_ran a (Array.change a i None h)" + shows "cl \ array_ran a h" +proof - + have "set (get_array a h[i := None]) \ insert None (set (get_array a h))" by (rule set_update_subset_insert) + with assms show ?thesis + unfolding array_ran_def Array.change_def by auto +qed + definition correctArray :: "Clause list \ Clause option array \ heap \ bool" where "correctArray rootcls a h = @@ -126,7 +152,7 @@ lemma correctArray_update: assumes "correctArray rcs a h" assumes "correctClause rcs c" "sorted c" "distinct c" - shows "correctArray rcs a (Heap.upd a i (Some c) h)" + shows "correctArray rcs a (Array.change a i (Some c) h)" using assms unfolding correctArray_def by (auto dest:array_ran_upd_array_Some) @@ -145,7 +171,7 @@ subsection{* Function definitions *} -fun res_mem :: "Lit \ Clause \ Clause Heap" +primrec res_mem :: "Lit \ Clause \ Clause Heap" where "res_mem l [] = raise ''MiniSatChecked.res_thm: Cannot find literal''" | "res_mem l (x#xs) = (if (x = l) then return xs else (do v \ res_mem l xs; return (x # v) done))" @@ -393,7 +419,7 @@ done)" -fun res_thm2 :: "Clause option array \ (Lit * ClauseId) \ Clause \ Clause Heap" +primrec res_thm2 :: "Clause option array \ (Lit * ClauseId) \ Clause \ Clause Heap" where "res_thm2 a (l, j) cli = ( if l = 0 then raise(''Illegal literal'') @@ -402,6 +428,12 @@ res_thm' l cli clj done))" +primrec + foldM :: "('a \ 'b \ 'b Heap) \ 'a list \ 'b \ 'b Heap" +where + "foldM f [] s = return s" + | "foldM f (x#xs) s = f x s \= foldM f xs" + fun doProofStep2 :: "Clause option array \ ProofStep \ Clause list \ Clause list Heap" where "doProofStep2 a (Conflict saveTo (i, rs)) rcs = @@ -439,7 +471,7 @@ fix clj let ?rs = "merge (remove l cli) (remove (compl l) clj)" let ?rs' = "merge (remove (compl l) cli) (remove l clj)" - assume "h = h'" "Some clj = get_array a h' ! j" "j < Heap.length a h'" + assume "h = h'" "Some clj = get_array a h' ! j" "j < Array.length a h'" with correct_a have clj: "correctClause r clj" "sorted clj" "distinct clj" unfolding correctArray_def by (auto intro: array_ranI) with clj l_not_zero correct_cli @@ -453,7 +485,7 @@ } { fix v clj - assume "Some clj = get_array a h ! j" "j < Heap.length a h" + assume "Some clj = get_array a h ! j" "j < Array.length a h" with correct_a have clj: "correctClause r clj \ sorted clj \ distinct clj" unfolding correctArray_def by (auto intro: array_ranI) assume "crel (res_thm' l cli clj) h h' rs" @@ -600,7 +632,7 @@ subsection {* Checker functions *} -fun lres_thm :: "Clause option list \ (Lit * ClauseId) \ Clause \ Clause Heap" +primrec lres_thm :: "Clause option list \ (Lit * ClauseId) \ Clause \ Clause Heap" where "lres_thm xs (l, j) cli = (if (j < List.length xs) then (case (xs ! j) of None \ raise (''MiniSatChecked.res_thm: No resolvant clause in thms array for Conflict step.'') @@ -634,7 +666,7 @@ section {* Functional version with RedBlackTrees *} -fun tres_thm :: "(ClauseId, Clause) RBT_Impl.rbt \ Lit \ ClauseId \ Clause \ Clause Heap" +primrec tres_thm :: "(ClauseId, Clause) RBT_Impl.rbt \ Lit \ ClauseId \ Clause \ Clause Heap" where "tres_thm t (l, j) cli = (case (RBT_Impl.lookup t j) of diff -r 6432bf0d7191 -r bf6c1216db43 src/HOL/Imperative_HOL/ex/Sorted_List.thy --- a/src/HOL/Imperative_HOL/ex/Sorted_List.thy Mon Jul 05 09:14:51 2010 -0700 +++ b/src/HOL/Imperative_HOL/ex/Sorted_List.thy Tue Jul 06 08:08:35 2010 -0700 @@ -33,7 +33,7 @@ text {* The remove function removes an element from a sorted list *} -fun remove :: "('a :: linorder) \ 'a list \ 'a list" +primrec remove :: "('a :: linorder) \ 'a list \ 'a list" where "remove a [] = []" | "remove a (x#xs) = (if a > x then (x # remove a xs) else (if a = x then xs else x#xs))" @@ -86,16 +86,13 @@ apply (auto simp add: sorted_Cons) done -subsection {* Efficient member function for sorted lists: smem *} +subsection {* Efficient member function for sorted lists *} -fun smember :: "('a::linorder) \ 'a list \ bool" (infixl "smem" 55) -where - "x smem [] = False" -| "x smem (y#ys) = (if x = y then True else (if (x > y) then x smem ys else False))" +primrec smember :: "'a list \ 'a::linorder \ bool" where + "smember [] x \ False" +| "smember (y#ys) x \ x = y \ (x > y \ smember ys x)" -lemma "sorted xs \ x smem xs = (x \ set xs)" -apply (induct xs) -apply (auto simp add: sorted_Cons) -done +lemma "sorted xs \ smember xs x \ (x \ set xs)" + by (induct xs) (auto simp add: sorted_Cons) end \ No newline at end of file diff -r 6432bf0d7191 -r bf6c1216db43 src/HOL/Imperative_HOL/ex/Subarray.thy --- a/src/HOL/Imperative_HOL/ex/Subarray.thy Mon Jul 05 09:14:51 2010 -0700 +++ b/src/HOL/Imperative_HOL/ex/Subarray.thy Tue Jul 06 08:08:35 2010 -0700 @@ -8,65 +8,64 @@ imports Array Sublist begin -definition subarray :: "nat \ nat \ ('a::heap) array \ heap \ 'a list" -where +definition subarray :: "nat \ nat \ ('a::heap) array \ heap \ 'a list" where "subarray n m a h \ sublist' n m (get_array a h)" -lemma subarray_upd: "i \ m \ subarray n m a (Heap.upd a i v h) = subarray n m a h" -apply (simp add: subarray_def Heap.upd_def) +lemma subarray_upd: "i \ m \ subarray n m a (Array.change a i v h) = subarray n m a h" +apply (simp add: subarray_def Array.change_def) apply (simp add: sublist'_update1) done -lemma subarray_upd2: " i < n \ subarray n m a (Heap.upd a i v h) = subarray n m a h" -apply (simp add: subarray_def Heap.upd_def) +lemma subarray_upd2: " i < n \ subarray n m a (Array.change a i v h) = subarray n m a h" +apply (simp add: subarray_def Array.change_def) apply (subst sublist'_update2) apply fastsimp apply simp done -lemma subarray_upd3: "\ n \ i; i < m\ \ subarray n m a (Heap.upd a i v h) = subarray n m a h[i - n := v]" -unfolding subarray_def Heap.upd_def +lemma subarray_upd3: "\ n \ i; i < m\ \ subarray n m a (Array.change a i v h) = subarray n m a h[i - n := v]" +unfolding subarray_def Array.change_def by (simp add: sublist'_update3) lemma subarray_Nil: "n \ m \ subarray n m a h = []" by (simp add: subarray_def sublist'_Nil') -lemma subarray_single: "\ n < Heap.length a h \ \ subarray n (Suc n) a h = [get_array a h ! n]" -by (simp add: subarray_def Heap.length_def sublist'_single) +lemma subarray_single: "\ n < Array.length a h \ \ subarray n (Suc n) a h = [get_array a h ! n]" +by (simp add: subarray_def length_def sublist'_single) -lemma length_subarray: "m \ Heap.length a h \ List.length (subarray n m a h) = m - n" -by (simp add: subarray_def Heap.length_def length_sublist') +lemma length_subarray: "m \ Array.length a h \ List.length (subarray n m a h) = m - n" +by (simp add: subarray_def length_def length_sublist') -lemma length_subarray_0: "m \ Heap.length a h \ List.length (subarray 0 m a h) = m" +lemma length_subarray_0: "m \ Array.length a h \ List.length (subarray 0 m a h) = m" by (simp add: length_subarray) -lemma subarray_nth_array_Cons: "\ i < Heap.length a h; i < j \ \ (get_array a h ! i) # subarray (Suc i) j a h = subarray i j a h" -unfolding Heap.length_def subarray_def +lemma subarray_nth_array_Cons: "\ i < Array.length a h; i < j \ \ (get_array a h ! i) # subarray (Suc i) j a h = subarray i j a h" +unfolding Array.length_def subarray_def by (simp add: sublist'_front) -lemma subarray_nth_array_back: "\ i < j; j \ Heap.length a h\ \ subarray i j a h = subarray i (j - 1) a h @ [get_array a h ! (j - 1)]" -unfolding Heap.length_def subarray_def +lemma subarray_nth_array_back: "\ i < j; j \ Array.length a h\ \ subarray i j a h = subarray i (j - 1) a h @ [get_array a h ! (j - 1)]" +unfolding Array.length_def subarray_def by (simp add: sublist'_back) lemma subarray_append: "\ i < j; j < k \ \ subarray i j a h @ subarray j k a h = subarray i k a h" unfolding subarray_def by (simp add: sublist'_append) -lemma subarray_all: "subarray 0 (Heap.length a h) a h = get_array a h" -unfolding Heap.length_def subarray_def +lemma subarray_all: "subarray 0 (Array.length a h) a h = get_array a h" +unfolding Array.length_def subarray_def by (simp add: sublist'_all) -lemma nth_subarray: "\ k < j - i; j \ Heap.length a h \ \ subarray i j a h ! k = get_array a h ! (i + k)" -unfolding Heap.length_def subarray_def +lemma nth_subarray: "\ k < j - i; j \ Array.length a h \ \ subarray i j a h ! k = get_array a h ! (i + k)" +unfolding Array.length_def subarray_def by (simp add: nth_sublist') -lemma subarray_eq_samelength_iff: "Heap.length a h = Heap.length a h' \ (subarray i j a h = subarray i j a h') = (\i'. i \ i' \ i' < j \ get_array a h ! i' = get_array a h' ! i')" -unfolding Heap.length_def subarray_def by (rule sublist'_eq_samelength_iff) +lemma subarray_eq_samelength_iff: "Array.length a h = Array.length a h' \ (subarray i j a h = subarray i j a h') = (\i'. i \ i' \ i' < j \ get_array a h ! i' = get_array a h' ! i')" +unfolding Array.length_def subarray_def by (rule sublist'_eq_samelength_iff) -lemma all_in_set_subarray_conv: "(\j. j \ set (subarray l r a h) \ P j) = (\k. l \ k \ k < r \ k < Heap.length a h \ P (get_array a h ! k))" -unfolding subarray_def Heap.length_def by (rule all_in_set_sublist'_conv) +lemma all_in_set_subarray_conv: "(\j. j \ set (subarray l r a h) \ P j) = (\k. l \ k \ k < r \ k < Array.length a h \ P (get_array a h ! k))" +unfolding subarray_def Array.length_def by (rule all_in_set_sublist'_conv) -lemma ball_in_set_subarray_conv: "(\j \ set (subarray l r a h). P j) = (\k. l \ k \ k < r \ k < Heap.length a h \ P (get_array a h ! k))" -unfolding subarray_def Heap.length_def by (rule ball_in_set_sublist'_conv) +lemma ball_in_set_subarray_conv: "(\j \ set (subarray l r a h). P j) = (\k. l \ k \ k < r \ k < Array.length a h \ P (get_array a h ! k))" +unfolding subarray_def Array.length_def by (rule ball_in_set_sublist'_conv) end \ No newline at end of file diff -r 6432bf0d7191 -r bf6c1216db43 src/HOL/Library/Countable.thy --- a/src/HOL/Library/Countable.thy Mon Jul 05 09:14:51 2010 -0700 +++ b/src/HOL/Library/Countable.thy Tue Jul 06 08:08:35 2010 -0700 @@ -110,26 +110,20 @@ "to_nat_typerep (Typerep.Typerep c ts) = to_nat (to_nat c, to_nat (map to_nat_typerep ts))" instance proof (rule countable_classI) - fix t t' :: typerep and ts - have "(\t'. to_nat_typerep t = to_nat_typerep t' \ t = t') - \ (\ts'. map to_nat_typerep ts = map to_nat_typerep ts' \ ts = ts')" - proof (induct rule: typerep.induct) - case (Typerep c ts) show ?case - proof (rule allI, rule impI) - fix t' - assume hyp: "to_nat_typerep (Typerep.Typerep c ts) = to_nat_typerep t'" - then obtain c' ts' where t': "t' = (Typerep.Typerep c' ts')" - by (cases t') auto - with Typerep hyp have "c = c'" and "ts = ts'" by simp_all - with t' show "Typerep.Typerep c ts = t'" by simp - qed + fix t t' :: typerep and ts ts' :: "typerep list" + assume "to_nat_typerep t = to_nat_typerep t'" + moreover have "to_nat_typerep t = to_nat_typerep t' \ t = t'" + and "map to_nat_typerep ts = map to_nat_typerep ts' \ ts = ts'" + proof (induct t and ts arbitrary: t' and ts' rule: typerep.inducts) + case (Typerep c ts t') + then obtain c' ts' where t': "t' = Typerep.Typerep c' ts'" by (cases t') auto + with Typerep have "c = c'" and "ts = ts'" by simp_all + with t' show "Typerep.Typerep c ts = t'" by simp next case Nil_typerep then show ?case by simp next case (Cons_typerep t ts) then show ?case by auto qed - then have "to_nat_typerep t = to_nat_typerep t' \ t = t'" by auto - moreover assume "to_nat_typerep t = to_nat_typerep t'" ultimately show "t = t'" by simp qed diff -r 6432bf0d7191 -r bf6c1216db43 src/Provers/clasimp.ML --- a/src/Provers/clasimp.ML Mon Jul 05 09:14:51 2010 -0700 +++ b/src/Provers/clasimp.ML Tue Jul 06 08:08:35 2010 -0700 @@ -203,7 +203,7 @@ (CHANGED o nodup_depth_tac cs' n); (* slower but more general *) in EVERY [ALLGOALS (Simplifier.asm_full_simp_tac ss), TRY (Classical.safe_tac cs), - REPEAT (FIRSTGOAL maintac), + REPEAT_DETERM (FIRSTGOAL maintac), TRY (Classical.safe_tac (cs addSss ss)), prune_params_tac] end; diff -r 6432bf0d7191 -r bf6c1216db43 src/Pure/Concurrent/future.ML --- a/src/Pure/Concurrent/future.ML Mon Jul 05 09:14:51 2010 -0700 +++ b/src/Pure/Concurrent/future.ML Tue Jul 06 08:08:35 2010 -0700 @@ -59,6 +59,7 @@ val cancel_group: group -> unit val cancel: 'a future -> unit val shutdown: unit -> unit + val report: (unit -> 'a) -> 'a end; structure Future: FUTURE = @@ -284,7 +285,7 @@ if forall (Thread.isActive o #1) (! workers) then () else let - val (alive, dead) = List.partition (Thread.isActive o #1) (! workers); + val (alive, dead) = List.partition (Thread.isActive o #1) (! workers); val _ = workers := alive; in Multithreading.tracing 0 (fn () => @@ -523,6 +524,16 @@ else (); +(* report markup *) + +fun report e = + let + val _ = Output.status (Markup.markup Markup.forked ""); + val x = e (); (*sic -- report "joined" only for success*) + val _ = Output.status (Markup.markup Markup.joined ""); + in x end; + + (*final declarations of this structure!*) val map = map_future; diff -r 6432bf0d7191 -r bf6c1216db43 src/Pure/Isar/outer_syntax.ML --- a/src/Pure/Isar/outer_syntax.ML Mon Jul 05 09:14:51 2010 -0700 +++ b/src/Pure/Isar/outer_syntax.ML Tue Jul 06 08:08:35 2010 -0700 @@ -240,7 +240,10 @@ val _ = List.app Thy_Syntax.report_token toks; in (case Source.exhaust (toplevel_source false NONE (K commands) (Source.of_list toks)) of - [tr] => (tr, true) + [tr] => + if Keyword.is_control (Toplevel.name_of tr) then + (Toplevel.malformed range_pos "Illegal control command", true) + else (tr, true) | [] => (Toplevel.ignored range_pos, false) | _ => (Toplevel.malformed range_pos not_singleton, true)) handle ERROR msg => (Toplevel.malformed range_pos msg, true) diff -r 6432bf0d7191 -r bf6c1216db43 src/Pure/Isar/toplevel.ML --- a/src/Pure/Isar/toplevel.ML Mon Jul 05 09:14:51 2010 -0700 +++ b/src/Pure/Isar/toplevel.ML Tue Jul 06 08:08:35 2010 -0700 @@ -86,9 +86,9 @@ val error_msg: transition -> exn * string -> unit val add_hook: (transition -> state -> state -> unit) -> unit val transition: bool -> transition -> state -> (state * (exn * string) option) option + val run_command: string -> transition -> state -> state option val commit_exit: Position.T -> transition val command: transition -> state -> state - val run_command: string -> transition -> state -> state option val excursion: (transition * transition list) list -> (transition * state) list lazy end; @@ -561,6 +561,14 @@ fun status tr m = setmp_thread_position tr (fn () => Output.status (Markup.markup m "")) (); +fun async_state (tr as Transition {print, ...}) st = + if print then + ignore + (Future.fork_group (Task_Queue.new_group (Future.worker_group ())) + (fn () => + setmp_thread_position tr (fn () => Future.report (fn () => print_state false st)) ())) + else (); + fun error_msg tr exn_info = setmp_thread_position tr (fn () => Output.error_msg (ML_Compiler.exn_message (EXCURSION_FAIL exn_info))) (); @@ -614,6 +622,22 @@ end; +(* managed execution *) + +fun run_command thy_name (tr as Transition {print, ...}) st = + (case + (case init_of tr of + SOME name => Exn.capture (fn () => Thy_Load.check_name thy_name name) () + | NONE => Exn.Result ()) of + Exn.Result () => + (case transition false tr st of + SOME (st', NONE) => (status tr Markup.finished; async_state tr st'; SOME st') + | SOME (_, SOME (exn as Exn.Interrupt, _)) => reraise exn + | SOME (_, SOME exn_info) => (error_msg tr exn_info; status tr Markup.failed; NONE) + | NONE => (error_msg tr (TERMINATE, at_command tr); status tr Markup.failed; NONE)) + | Exn.Exn exn => (error_msg tr (exn, at_command tr); status tr Markup.failed; NONE)); + + (* commit final exit *) fun commit_exit pos = @@ -635,19 +659,6 @@ let val st' = command tr st in (st', st') end; -fun run_command thy_name tr st = - (case - (case init_of tr of - SOME name => Exn.capture (fn () => Thy_Load.check_name thy_name name) () - | NONE => Exn.Result ()) of - Exn.Result () => - (case transition true tr st of - SOME (st', NONE) => (status tr Markup.finished; SOME st') - | SOME (_, SOME (exn as Exn.Interrupt, _)) => reraise exn - | SOME (_, SOME exn_info) => (error_msg tr exn_info; status tr Markup.failed; NONE) - | NONE => (error_msg tr (TERMINATE, at_command tr); status tr Markup.failed; NONE)) - | Exn.Exn exn => (error_msg tr (exn, at_command tr); status tr Markup.failed; NONE)); - (* excursion of units, consisting of commands with proof *) diff -r 6432bf0d7191 -r bf6c1216db43 src/Pure/PIDE/document.scala --- a/src/Pure/PIDE/document.scala Mon Jul 05 09:14:51 2010 -0700 +++ b/src/Pure/PIDE/document.scala Tue Jul 06 08:08:35 2010 -0700 @@ -14,13 +14,13 @@ { /* command start positions */ - def command_starts(commands: Linear_Set[Command]): Iterator[(Command, Int)] = + def command_starts(commands: Iterator[Command], offset: Int = 0): Iterator[(Command, Int)] = { - var offset = 0 - for (cmd <- commands.iterator) yield { - val start = offset - offset += cmd.length - (cmd, start) + var i = offset + for (command <- commands) yield { + val start = i + i += command.length + (command, start) } } @@ -60,9 +60,10 @@ { eds match { case e :: es => - command_starts(commands).find { // FIXME relative search! + command_starts(commands.iterator).find { case (cmd, cmd_start) => - e.can_edit(cmd.source, cmd_start) || e.is_insert && e.start == cmd_start + cmd.length + e.can_edit(cmd.source, cmd_start) || + e.is_insert && e.start == cmd_start + cmd.length } match { case Some((cmd, cmd_start)) if e.can_edit(cmd.source, cmd_start) => val (rest, text) = e.edit(cmd.source, cmd_start) @@ -144,7 +145,7 @@ { /* command ranges */ - def command_starts: Iterator[(Command, Int)] = Document.command_starts(commands) + def command_starts: Iterator[(Command, Int)] = Document.command_starts(commands.iterator) def command_start(cmd: Command): Option[Int] = command_starts.find(_._1 == cmd).map(_._2) diff -r 6432bf0d7191 -r bf6c1216db43 src/Pure/Syntax/parser.ML --- a/src/Pure/Syntax/parser.ML Mon Jul 05 09:14:51 2010 -0700 +++ b/src/Pure/Syntax/parser.ML Tue Jul 06 08:08:35 2010 -0700 @@ -8,9 +8,9 @@ sig type gram val empty_gram: gram - val extend_gram: gram -> Syn_Ext.xprod list -> gram + val extend_gram: Syn_Ext.xprod list -> gram -> gram val make_gram: Syn_Ext.xprod list -> gram - val merge_grams: gram -> gram -> gram + val merge_gram: gram * gram -> gram val pretty_gram: gram -> Pretty.T list datatype parsetree = Node of string * parsetree list | @@ -23,19 +23,16 @@ structure Parser: PARSER = struct -open Lexicon Syn_Ext; - - (** datatype gram **) type nt_tag = int; (*production for the NTs are stored in an array so we can identify NTs by their index*) -datatype symb = Terminal of token +datatype symb = Terminal of Lexicon.token | Nonterminal of nt_tag * int; (*(tag, precedence)*) -type nt_gram = ((nt_tag list * token list) * - (token option * (symb list * string * int) list) list); +type nt_gram = ((nt_tag list * Lexicon.token list) * + (Lexicon.token option * (symb list * string * int) list) list); (*(([dependent_nts], [start_tokens]), [(start_token, [(rhs, name, prio)])])*) (*depent_nts is a list of all NTs whose lookahead @@ -53,8 +50,8 @@ lambda productions are stored as normal productions and also as an entry in "lambdas"*) -val UnknownStart = eof; (*productions for which no starting token is - known yet are associated with this token*) +val UnknownStart = Lexicon.eof; (*productions for which no starting token is + known yet are associated with this token*) (* get all NTs that are connected with a list of NTs (used for expanding chain list)*) @@ -125,7 +122,7 @@ (*get all known starting tokens for a nonterminal*) fun starts_for_nt nt = snd (fst (Array.sub (prods, nt))); - val token_union = uncurry (union matching_tokens); + val token_union = uncurry (union Lexicon.matching_tokens); (*update prods, lookaheads, and lambdas according to new lambda NTs*) val (added_starts', lambdas') = @@ -158,7 +155,7 @@ val nt_dependencies' = union (op =) nts nt_dependencies; (*associate production with new starting tokens*) - fun copy ([]: token option list) nt_prods = nt_prods + fun copy ([]: Lexicon.token option list) nt_prods = nt_prods | copy (tk :: tks) nt_prods = let val old_prods = these (AList.lookup (op =) nt_prods tk); @@ -259,7 +256,7 @@ let val ((old_nts, old_tks), nt_prods) = Array.sub (prods, nt); - val new_tks = subtract matching_tokens old_tks start_tks; + val new_tks = subtract Lexicon.matching_tokens old_tks start_tks; (*store new production*) fun store [] prods is_new = @@ -278,7 +275,7 @@ else (new_prod :: tk_prods, true); val prods' = prods - |> is_new' ? AList.update (op =) (tk: token option, tk_prods'); + |> is_new' ? AList.update (op =) (tk: Lexicon.token option, tk_prods'); in store tks prods' (is_new orelse is_new') end; val (nt_prods', prod_count', changed) = @@ -329,10 +326,10 @@ andalso member (op =) new_tks (the tk); (*associate production with new starting tokens*) - fun copy ([]: token list) nt_prods = nt_prods + fun copy ([]: Lexicon.token list) nt_prods = nt_prods | copy (tk :: tks) nt_prods = let - val tk_prods = (these o AList.lookup (op =) nt_prods) (SOME tk); + val tk_prods = these (AList.lookup (op =) nt_prods (SOME tk)); val tk_prods' = if not lambda then p :: tk_prods @@ -359,7 +356,7 @@ val (lookahead as (old_nts, old_tks), nt_prods) = Array.sub (prods, nt); - val tk_prods = (these o AList.lookup (op =) nt_prods) key; + val tk_prods = these (AList.lookup (op =) nt_prods key); (*associate productions with new lookahead tokens*) val (tk_prods', nt_prods') = @@ -370,7 +367,7 @@ |> (key = SOME UnknownStart) ? AList.update (op =) (key, tk_prods') val added_tks = - subtract matching_tokens old_tks new_tks; + subtract Lexicon.matching_tokens old_tks new_tks; in if null added_tks then (Array.update (prods, nt, (lookahead, nt_prods')); process_nts nts added) @@ -381,7 +378,7 @@ end; val ((dependent, _), _) = Array.sub (prods, changed_nt); - in add_starts (starts @ (process_nts dependent [])) end; + in add_starts (starts @ process_nts dependent []) end; in add_starts added_starts' end; in add_prods prods chains' lambdas' prod_count ps end; @@ -394,8 +391,8 @@ val nt_name = the o Inttab.lookup (Inttab.make (map swap (Symtab.dest tags))); - fun pretty_symb (Terminal (Token (Literal, s, _))) = Pretty.quote (Pretty.str s) - | pretty_symb (Terminal tok) = Pretty.str (str_of_token tok) + fun pretty_symb (Terminal (Lexicon.Token (Lexicon.Literal, s, _))) = Pretty.quote (Pretty.str s) + | pretty_symb (Terminal tok) = Pretty.str (Lexicon.str_of_token tok) | pretty_symb (Nonterminal (tag, p)) = Pretty.str (nt_name tag ^ "[" ^ signed_string_of_int p ^ "]"); @@ -422,7 +419,6 @@ (** Operations on gramars **) -(*The mother of all grammars*) val empty_gram = Gram {nt_count = 0, prod_count = 0, tags = Symtab.empty, chains = [], lambdas = [], prods = Array.array (0, (([], []), []))}; @@ -439,75 +435,75 @@ (*Add productions to a grammar*) -fun extend_gram gram [] = gram - | extend_gram (Gram {nt_count, prod_count, tags, chains, lambdas, prods}) - xprods = - let - (*Get tag for existing nonterminal or create a new one*) - fun get_tag nt_count tags nt = - case Symtab.lookup tags nt of - SOME tag => (nt_count, tags, tag) - | NONE => (nt_count+1, Symtab.update_new (nt, nt_count) tags, - nt_count); +fun extend_gram [] gram = gram + | extend_gram xprods (Gram {nt_count, prod_count, tags, chains, lambdas, prods}) = + let + (*Get tag for existing nonterminal or create a new one*) + fun get_tag nt_count tags nt = + case Symtab.lookup tags nt of + SOME tag => (nt_count, tags, tag) + | NONE => (nt_count+1, Symtab.update_new (nt, nt_count) tags, + nt_count); - (*Convert symbols to the form used by the parser; - delimiters and predefined terms are stored as terminals, - nonterminals are converted to integer tags*) - fun symb_of [] nt_count tags result = (nt_count, tags, rev result) - | symb_of ((Delim s) :: ss) nt_count tags result = - symb_of ss nt_count tags (Terminal (Token (Literal, s, Position.no_range)) :: result) - | symb_of ((Argument (s, p)) :: ss) nt_count tags result = - let - val (nt_count', tags', new_symb) = - case predef_term s of - NONE => - let val (nt_count', tags', s_tag) = get_tag nt_count tags s; - in (nt_count', tags', Nonterminal (s_tag, p)) end - | SOME tk => (nt_count, tags, Terminal tk); - in symb_of ss nt_count' tags' (new_symb :: result) end - | symb_of (_ :: ss) nt_count tags result = - symb_of ss nt_count tags result; + (*Convert symbols to the form used by the parser; + delimiters and predefined terms are stored as terminals, + nonterminals are converted to integer tags*) + fun symb_of [] nt_count tags result = (nt_count, tags, rev result) + | symb_of ((Syn_Ext.Delim s) :: ss) nt_count tags result = + symb_of ss nt_count tags + (Terminal (Lexicon.Token (Lexicon.Literal, s, Position.no_range)) :: result) + | symb_of ((Syn_Ext.Argument (s, p)) :: ss) nt_count tags result = + let + val (nt_count', tags', new_symb) = + case Lexicon.predef_term s of + NONE => + let val (nt_count', tags', s_tag) = get_tag nt_count tags s; + in (nt_count', tags', Nonterminal (s_tag, p)) end + | SOME tk => (nt_count, tags, Terminal tk); + in symb_of ss nt_count' tags' (new_symb :: result) end + | symb_of (_ :: ss) nt_count tags result = + symb_of ss nt_count tags result; - (*Convert list of productions by invoking symb_of for each of them*) - fun prod_of [] nt_count prod_count tags result = - (nt_count, prod_count, tags, result) - | prod_of ((XProd (lhs, xsymbs, const, pri)) :: ps) - nt_count prod_count tags result = - let val (nt_count', tags', lhs_tag) = get_tag nt_count tags lhs; + (*Convert list of productions by invoking symb_of for each of them*) + fun prod_of [] nt_count prod_count tags result = + (nt_count, prod_count, tags, result) + | prod_of ((Syn_Ext.XProd (lhs, xsymbs, const, pri)) :: ps) + nt_count prod_count tags result = + let val (nt_count', tags', lhs_tag) = get_tag nt_count tags lhs; - val (nt_count'', tags'', prods) = - symb_of xsymbs nt_count' tags' []; - in prod_of ps nt_count'' (prod_count+1) tags'' - ((lhs_tag, (prods, const, pri)) :: result) - end; + val (nt_count'', tags'', prods) = + symb_of xsymbs nt_count' tags' []; + in prod_of ps nt_count'' (prod_count+1) tags'' + ((lhs_tag, (prods, const, pri)) :: result) + end; - val (nt_count', prod_count', tags', xprods') = - prod_of xprods nt_count prod_count tags []; + val (nt_count', prod_count', tags', xprods') = + prod_of xprods nt_count prod_count tags []; - (*Copy array containing productions of old grammar; - this has to be done to preserve the old grammar while being able - to change the array's content*) - val prods' = - let fun get_prod i = if i < nt_count then Array.sub (prods, i) - else (([], []), []); - in Array.tabulate (nt_count', get_prod) end; + (*Copy array containing productions of old grammar; + this has to be done to preserve the old grammar while being able + to change the array's content*) + val prods' = + let fun get_prod i = if i < nt_count then Array.sub (prods, i) + else (([], []), []); + in Array.tabulate (nt_count', get_prod) end; - val fromto_chains = inverse_chains chains []; + val fromto_chains = inverse_chains chains []; - (*Add new productions to old ones*) - val (fromto_chains', lambdas', _) = - add_prods prods' fromto_chains lambdas NONE xprods'; + (*Add new productions to old ones*) + val (fromto_chains', lambdas', _) = + add_prods prods' fromto_chains lambdas NONE xprods'; - val chains' = inverse_chains fromto_chains' []; - in Gram {nt_count = nt_count', prod_count = prod_count', tags = tags', - chains = chains', lambdas = lambdas', prods = prods'} - end; + val chains' = inverse_chains fromto_chains' []; + in Gram {nt_count = nt_count', prod_count = prod_count', tags = tags', + chains = chains', lambdas = lambdas', prods = prods'} + end; -val make_gram = extend_gram empty_gram; +fun make_gram xprods = extend_gram xprods empty_gram; (*Merge two grammars*) -fun merge_grams gram_a gram_b = +fun merge_gram (gram_a, gram_b) = let (*find out which grammar is bigger*) val (Gram {nt_count = nt_count1, prod_count = prod_count1, tags = tags1, @@ -604,7 +600,7 @@ datatype parsetree = Node of string * parsetree list | - Tip of token; + Tip of Lexicon.token; type state = nt_tag * int * (*identification and production precedence*) @@ -675,7 +671,7 @@ fun movedot_term (A, j, ts, Terminal a :: sa, id, i) c = - if valued_token c then + if Lexicon.valued_token c then (A, j, ts @ [Tip c], sa, id, i) else (A, j, ts, sa, id, i); @@ -695,105 +691,105 @@ (*get all productions of a NT and NTs chained to it which can be started by specified token*) fun prods_for prods chains include_none tk nts = -let - fun token_assoc (list, key) = - let fun assoc [] result = result - | assoc ((keyi, pi) :: pairs) result = - if is_some keyi andalso matching_tokens (the keyi, key) - orelse include_none andalso is_none keyi then - assoc pairs (pi @ result) - else assoc pairs result; - in assoc list [] end; + let + fun token_assoc (list, key) = + let fun assoc [] result = result + | assoc ((keyi, pi) :: pairs) result = + if is_some keyi andalso Lexicon.matching_tokens (the keyi, key) + orelse include_none andalso is_none keyi then + assoc pairs (pi @ result) + else assoc pairs result; + in assoc list [] end; - fun get_prods [] result = result - | get_prods (nt :: nts) result = - let val nt_prods = snd (Array.sub (prods, nt)); - in get_prods nts ((token_assoc (nt_prods, tk)) @ result) end; -in get_prods (connected_with chains nts nts) [] end; + fun get_prods [] result = result + | get_prods (nt :: nts) result = + let val nt_prods = snd (Array.sub (prods, nt)); + in get_prods nts ((token_assoc (nt_prods, tk)) @ result) end; + in get_prods (connected_with chains nts nts) [] end; fun PROCESSS warned prods chains Estate i c states = -let -fun all_prods_for nt = prods_for prods chains true c [nt]; + let + fun all_prods_for nt = prods_for prods chains true c [nt]; -fun processS used [] (Si, Sii) = (Si, Sii) - | processS used (S :: States) (Si, Sii) = - (case S of - (_, _, _, Nonterminal (nt, minPrec) :: _, _, _) => - let (*predictor operation*) - val (used', new_states) = - (case AList.lookup (op =) used nt of - SOME (usedPrec, l) => (*nonterminal has been processed*) - if usedPrec <= minPrec then - (*wanted precedence has been processed*) - (used, movedot_lambda S l) - else (*wanted precedence hasn't been parsed yet*) - let - val tk_prods = all_prods_for nt; + fun processS used [] (Si, Sii) = (Si, Sii) + | processS used (S :: States) (Si, Sii) = + (case S of + (_, _, _, Nonterminal (nt, minPrec) :: _, _, _) => + let (*predictor operation*) + val (used', new_states) = + (case AList.lookup (op =) used nt of + SOME (usedPrec, l) => (*nonterminal has been processed*) + if usedPrec <= minPrec then + (*wanted precedence has been processed*) + (used, movedot_lambda S l) + else (*wanted precedence hasn't been parsed yet*) + let + val tk_prods = all_prods_for nt; - val States' = mkStates i minPrec nt - (getRHS' minPrec usedPrec tk_prods); - in (update_prec (nt, minPrec) used, - movedot_lambda S l @ States') - end + val States' = mkStates i minPrec nt + (getRHS' minPrec usedPrec tk_prods); + in (update_prec (nt, minPrec) used, + movedot_lambda S l @ States') + end - | NONE => (*nonterminal is parsed for the first time*) - let val tk_prods = all_prods_for nt; - val States' = mkStates i minPrec nt - (getRHS minPrec tk_prods); - in ((nt, (minPrec, [])) :: used, States') end); + | NONE => (*nonterminal is parsed for the first time*) + let val tk_prods = all_prods_for nt; + val States' = mkStates i minPrec nt + (getRHS minPrec tk_prods); + in ((nt, (minPrec, [])) :: used, States') end); - val dummy = - if not (!warned) andalso - length (new_states @ States) > (!branching_level) then - (warning "Currently parsed expression could be extremely ambiguous."; - warned := true) - else (); - in - processS used' (new_states @ States) (S :: Si, Sii) - end - | (_, _, _, Terminal a :: _, _, _) => (*scanner operation*) - processS used States - (S :: Si, - if matching_tokens (a, c) then movedot_term S c :: Sii else Sii) - | (A, prec, ts, [], id, j) => (*completer operation*) - let val tt = if id = "" then ts else [Node (id, ts)] in - if j = i then (*lambda production?*) - let - val (used', O) = update_trees used (A, (tt, prec)); + val dummy = + if not (!warned) andalso + length (new_states @ States) > (!branching_level) then + (warning "Currently parsed expression could be extremely ambiguous."; + warned := true) + else (); in - case O of - NONE => - let val Slist = getS A prec Si; - val States' = map (movedot_nonterm tt) Slist; - in processS used' (States' @ States) (S :: Si, Sii) end - | SOME n => - if n >= prec then processS used' States (S :: Si, Sii) - else - let val Slist = getS' A prec n Si; - val States' = map (movedot_nonterm tt) Slist; - in processS used' (States' @ States) (S :: Si, Sii) end + processS used' (new_states @ States) (S :: Si, Sii) end - else - let val Slist = getStates Estate i j A prec - in processS used (map (movedot_nonterm tt) Slist @ States) - (S :: Si, Sii) - end - end) -in processS [] states ([], []) end; + | (_, _, _, Terminal a :: _, _, _) => (*scanner operation*) + processS used States + (S :: Si, + if Lexicon.matching_tokens (a, c) then movedot_term S c :: Sii else Sii) + | (A, prec, ts, [], id, j) => (*completer operation*) + let val tt = if id = "" then ts else [Node (id, ts)] in + if j = i then (*lambda production?*) + let + val (used', O) = update_trees used (A, (tt, prec)); + in + case O of + NONE => + let val Slist = getS A prec Si; + val States' = map (movedot_nonterm tt) Slist; + in processS used' (States' @ States) (S :: Si, Sii) end + | SOME n => + if n >= prec then processS used' States (S :: Si, Sii) + else + let val Slist = getS' A prec n Si; + val States' = map (movedot_nonterm tt) Slist; + in processS used' (States' @ States) (S :: Si, Sii) end + end + else + let val Slist = getStates Estate i j A prec + in processS used (map (movedot_nonterm tt) Slist @ States) + (S :: Si, Sii) + end + end) + in processS [] states ([], []) end; fun produce warned prods tags chains stateset i indata prev_token = (case Array.sub (stateset, i) of [] => let - val toks = if is_eof prev_token then indata else prev_token :: indata; - val pos = Position.str_of (pos_of_token prev_token); + val toks = if Lexicon.is_eof prev_token then indata else prev_token :: indata; + val pos = Position.str_of (Lexicon.pos_of_token prev_token); in if null toks then error ("Inner syntax error: unexpected end of input" ^ pos) else error (Pretty.string_of (Pretty.block (Pretty.str ("Inner syntax error" ^ pos ^ " at \"") :: - Pretty.breaks (map (Pretty.str o str_of_token) (#1 (split_last toks))) @ + Pretty.breaks (map (Pretty.str o Lexicon.str_of_token) (#1 (split_last toks))) @ [Pretty.str "\""]))) end | s => @@ -807,21 +803,20 @@ end)); -fun get_trees l = map_filter (fn (_, _, [pt], _, _, _) => SOME pt | _ => NONE) - l; +fun get_trees l = map_filter (fn (_, _, [pt], _, _, _) => SOME pt | _ => NONE) l; fun earley prods tags chains startsymbol indata = let - val start_tag = case Symtab.lookup tags startsymbol of - SOME tag => tag - | NONE => error ("parse: Unknown startsymbol " ^ - quote startsymbol); - val S0 = [(~1, 0, [], [Nonterminal (start_tag, 0), Terminal eof], "", 0)]; + val start_tag = + (case Symtab.lookup tags startsymbol of + SOME tag => tag + | NONE => error ("Inner syntax: unknown startsymbol " ^ quote startsymbol)); + val S0 = [(~1, 0, [], [Nonterminal (start_tag, 0), Terminal Lexicon.eof], "", 0)]; val s = length indata + 1; val Estate = Array.array (s, []); in Array.update (Estate, 0, S0); - get_trees (produce (Unsynchronized.ref false) prods tags chains Estate 0 indata eof) + get_trees (produce (Unsynchronized.ref false) prods tags chains Estate 0 indata Lexicon.eof) end; @@ -830,10 +825,10 @@ val end_pos = (case try List.last toks of NONE => Position.none - | SOME (Token (_, _, (_, end_pos))) => end_pos); + | SOME (Lexicon.Token (_, _, (_, end_pos))) => end_pos); val r = - (case earley prods tags chains start (toks @ [mk_eof end_pos]) of - [] => sys_error "parse: no parse trees" + (case earley prods tags chains start (toks @ [Lexicon.mk_eof end_pos]) of + [] => raise Fail "no parse trees" | pts => pts); in r end; @@ -842,7 +837,8 @@ let fun freeze a = map_range (curry Array.sub a) (Array.length a); val prods = maps snd (maps snd (freeze (#prods gram))); - fun guess (SOME ([Nonterminal (_, k), Terminal (Token (Literal, s, _)), Nonterminal (_, l)], _, j)) = + fun guess (SOME ([Nonterminal (_, k), + Terminal (Lexicon.Token (Lexicon.Literal, s, _)), Nonterminal (_, l)], _, j)) = if k = j andalso l = j + 1 then SOME (s, true, false, j) else if k = j + 1 then if l = j then SOME (s, false, true, j) else if l = j + 1 then SOME (s, false, false, j) diff -r 6432bf0d7191 -r bf6c1216db43 src/Pure/Syntax/syntax.ML --- a/src/Pure/Syntax/syntax.ML Mon Jul 05 09:14:51 2010 -0700 +++ b/src/Pure/Syntax/syntax.ML Tue Jul 06 08:08:35 2010 -0700 @@ -297,7 +297,7 @@ Syntax ({input = new_xprods @ input, lexicon = fold Scan.extend_lexicon (Syn_Ext.delims_of new_xprods) lexicon, - gram = Parser.extend_gram gram new_xprods, + gram = Parser.extend_gram new_xprods gram, consts = Library.merge (op =) (consts1, filter_out Lexicon.is_marked consts2), prmodes = insert (op =) mode (Library.merge (op =) (prmodes1, prmodes2)), parse_ast_trtab = @@ -362,7 +362,7 @@ Syntax ({input = Library.merge (op =) (input1, input2), lexicon = Scan.merge_lexicons (lexicon1, lexicon2), - gram = Parser.merge_grams gram1 gram2, + gram = Parser.merge_gram (gram1, gram2), consts = sort_distinct string_ord (consts1 @ consts2), prmodes = Library.merge (op =) (prmodes1, prmodes2), parse_ast_trtab = diff -r 6432bf0d7191 -r bf6c1216db43 src/Pure/System/isabelle_process.ML --- a/src/Pure/System/isabelle_process.ML Mon Jul 05 09:14:51 2010 -0700 +++ b/src/Pure/System/isabelle_process.ML Tue Jul 06 08:08:35 2010 -0700 @@ -91,6 +91,7 @@ (Unsynchronized.change print_mode (fold (update op =) [isabelle_processN, Keyword.keyword_status_reportN, Pretty.symbolicN]); setup_channels out |> init_message; + quick_and_dirty := true; (* FIXME !? *) Keyword.report (); Output.status (Markup.markup Markup.ready ""); Isar.toplevel_loop {init = true, welcome = false, sync = true, secure = true}); diff -r 6432bf0d7191 -r bf6c1216db43 src/Pure/System/isabelle_process.scala --- a/src/Pure/System/isabelle_process.scala Mon Jul 05 09:14:51 2010 -0700 +++ b/src/Pure/System/isabelle_process.scala Tue Jul 06 08:08:35 2010 -0700 @@ -19,89 +19,55 @@ { /* results */ - object Kind extends Enumeration { - //{{{ values and codes - // internal system notification - val SYSTEM = Value("SYSTEM") - // Posix channels/events - val STDIN = Value("STDIN") - val STDOUT = Value("STDOUT") - val SIGNAL = Value("SIGNAL") - val EXIT = Value("EXIT") - // Isabelle messages - val INIT = Value("INIT") - val STATUS = Value("STATUS") - val WRITELN = Value("WRITELN") - val TRACING = Value("TRACING") - val WARNING = Value("WARNING") - val ERROR = Value("ERROR") - val DEBUG = Value("DEBUG") - // messages codes - val code = Map( - ('A' : Int) -> Kind.INIT, - ('B' : Int) -> Kind.STATUS, - ('C' : Int) -> Kind.WRITELN, - ('D' : Int) -> Kind.TRACING, - ('E' : Int) -> Kind.WARNING, - ('F' : Int) -> Kind.ERROR, - ('G' : Int) -> Kind.DEBUG, - ('0' : Int) -> Kind.SYSTEM, - ('1' : Int) -> Kind.STDIN, - ('2' : Int) -> Kind.STDOUT, - ('3' : Int) -> Kind.SIGNAL, - ('4' : Int) -> Kind.EXIT) + object Kind { // message markup val markup = Map( - Kind.INIT -> Markup.INIT, - Kind.STATUS -> Markup.STATUS, - Kind.WRITELN -> Markup.WRITELN, - Kind.TRACING -> Markup.TRACING, - Kind.WARNING -> Markup.WARNING, - Kind.ERROR -> Markup.ERROR, - Kind.DEBUG -> Markup.DEBUG, - Kind.SYSTEM -> Markup.SYSTEM, - Kind.STDIN -> Markup.STDIN, - Kind.STDOUT -> Markup.STDOUT, - Kind.SIGNAL -> Markup.SIGNAL, - Kind.EXIT -> Markup.EXIT) - //}}} - def is_raw(kind: Value) = - kind == STDOUT - def is_control(kind: Value) = - kind == SYSTEM || - kind == SIGNAL || - kind == EXIT - def is_system(kind: Value) = - kind == SYSTEM || - kind == STDIN || - kind == SIGNAL || - kind == EXIT || - kind == STATUS + ('A' : Int) -> Markup.INIT, + ('B' : Int) -> Markup.STATUS, + ('C' : Int) -> Markup.WRITELN, + ('D' : Int) -> Markup.TRACING, + ('E' : Int) -> Markup.WARNING, + ('F' : Int) -> Markup.ERROR, + ('G' : Int) -> Markup.DEBUG) + def is_raw(kind: String) = + kind == Markup.STDOUT + def is_control(kind: String) = + kind == Markup.SYSTEM || + kind == Markup.SIGNAL || + kind == Markup.EXIT + def is_system(kind: String) = + kind == Markup.SYSTEM || + kind == Markup.STDIN || + kind == Markup.SIGNAL || + kind == Markup.EXIT || + kind == Markup.STATUS } - class Result(val kind: Kind.Value, val props: List[(String, String)], val body: List[XML.Tree]) + class Result(val message: XML.Elem) { - def message = XML.Elem(Kind.markup(kind), props, body) + def kind = message.name + def body = message.body + + def is_raw = Kind.is_raw(kind) + def is_control = Kind.is_control(kind) + def is_system = Kind.is_system(kind) + def is_status = kind == Markup.STATUS + def is_ready = is_status && body == List(XML.Elem(Markup.READY, Nil, Nil)) override def toString: String = { val res = - if (kind == Kind.STATUS) body.map(_.toString).mkString - else Pretty.string_of(body) + if (is_status) message.body.map(_.toString).mkString + else Pretty.string_of(message.body) + val props = message.attributes if (props.isEmpty) kind.toString + " [[" + res + "]]" else kind.toString + " " + (for ((x, y) <- props) yield x + "=" + y).mkString("{", ",", "}") + " [[" + res + "]]" } - def is_raw = Kind.is_raw(kind) - def is_control = Kind.is_control(kind) - def is_system = Kind.is_system(kind) - def is_ready = kind == Kind.STATUS && body == List(XML.Elem(Markup.READY, Nil, Nil)) - - def cache(c: XML.Cache): Result = - new Result(kind, c.cache_props(props), c.cache_trees(body)) + def cache(c: XML.Cache): Result = new Result(c.cache_tree(message).asInstanceOf[XML.Elem]) } } @@ -127,15 +93,15 @@ /* results */ - private def put_result(kind: Kind.Value, props: List[(String, String)], body: List[XML.Tree]) + private def put_result(kind: String, props: List[(String, String)], body: List[XML.Tree]) { - if (kind == Kind.INIT) { + if (kind == Markup.INIT) { for ((Markup.PID, p) <- props) pid = p } - receiver ! new Result(kind, props, body) + receiver ! new Result(XML.Elem(kind, props, body)) } - private def put_result(kind: Kind.Value, text: String) + private def put_result(kind: String, text: String) { put_result(kind, Nil, List(XML.Text(system.symbols.decode(text)))) } @@ -145,13 +111,13 @@ def interrupt() = synchronized { if (proc == null) error("Cannot interrupt Isabelle: no process") - if (pid == null) put_result(Kind.SYSTEM, "Cannot interrupt: unknown pid") + if (pid == null) put_result(Markup.SYSTEM, "Cannot interrupt: unknown pid") else { try { if (system.execute(true, "kill", "-INT", pid).waitFor == 0) - put_result(Kind.SIGNAL, "INT") + put_result(Markup.SIGNAL, "INT") else - put_result(Kind.SYSTEM, "Cannot interrupt: kill command failed") + put_result(Markup.SYSTEM, "Cannot interrupt: kill command failed") } catch { case e: IOException => error("Cannot interrupt Isabelle: " + e.getMessage) } } @@ -162,7 +128,7 @@ else { try_close() Thread.sleep(500) // FIXME property!? - put_result(Kind.SIGNAL, "KILL") + put_result(Markup.SIGNAL, "KILL") proc.destroy proc = null pid = null @@ -191,9 +157,12 @@ output_sync("Isabelle.command " + Isabelle_Syntax.encode_properties(props) + " " + Isabelle_Syntax.encode_string(text)) - def ML(text: String) = + def ML_val(text: String) = output_sync("ML_val " + Isabelle_Syntax.encode_string(text)) + def ML_command(text: String) = + output_sync("ML_command " + Isabelle_Syntax.encode_string(text)) + def close() = synchronized { // FIXME watchdog/timeout output_raw("\u0000") closing = true @@ -222,17 +191,17 @@ finished = true } else { - put_result(Kind.STDIN, s) + put_result(Markup.STDIN, s) writer.write(s) writer.flush } //}}} } catch { - case e: IOException => put_result(Kind.SYSTEM, "Stdin thread: " + e.getMessage) + case e: IOException => put_result(Markup.SYSTEM, "Stdin thread: " + e.getMessage) } } - put_result(Kind.SYSTEM, "Stdin thread terminated") + put_result(Markup.SYSTEM, "Stdin thread terminated") } } @@ -256,7 +225,7 @@ else done = true } if (result.length > 0) { - put_result(Kind.STDOUT, result.toString) + put_result(Markup.STDOUT, result.toString) result.length = 0 } else { @@ -267,10 +236,10 @@ //}}} } catch { - case e: IOException => put_result(Kind.SYSTEM, "Stdout thread: " + e.getMessage) + case e: IOException => put_result(Markup.SYSTEM, "Stdout thread: " + e.getMessage) } } - put_result(Kind.SYSTEM, "Stdout thread terminated") + put_result(Markup.SYSTEM, "Stdout thread terminated") } } @@ -332,8 +301,8 @@ val body = read_chunk() header match { case List(XML.Elem(name, props, Nil)) - if name.size == 1 && Kind.code.isDefinedAt(name(0)) => - put_result(Kind.code(name(0)), props, body) + if name.size == 1 && Kind.markup.isDefinedAt(name(0)) => + put_result(Kind.markup(name(0)), props, body) case _ => throw new Protocol_Error("bad header: " + header.toString) } } @@ -341,15 +310,15 @@ } catch { case e: IOException => - put_result(Kind.SYSTEM, "Cannot read message:\n" + e.getMessage) + put_result(Markup.SYSTEM, "Cannot read message:\n" + e.getMessage) case e: Protocol_Error => - put_result(Kind.SYSTEM, "Malformed message:\n" + e.getMessage) + put_result(Markup.SYSTEM, "Malformed message:\n" + e.getMessage) } } while (c != -1) stream.close try_close() - put_result(Kind.SYSTEM, "Message thread terminated") + put_result(Markup.SYSTEM, "Message thread terminated") } } @@ -392,8 +361,8 @@ override def run() = { val rc = proc.waitFor() Thread.sleep(300) // FIXME property!? - put_result(Kind.SYSTEM, "Exit thread terminated") - put_result(Kind.EXIT, rc.toString) + put_result(Markup.SYSTEM, "Exit thread terminated") + put_result(Markup.EXIT, rc.toString) rm_fifo() } }.start diff -r 6432bf0d7191 -r bf6c1216db43 src/Pure/System/session.scala --- a/src/Pure/System/session.scala Mon Jul 05 09:14:51 2010 -0700 +++ b/src/Pure/System/session.scala Tue Jul 06 08:08:35 2010 -0700 @@ -110,14 +110,14 @@ { raw_results.event(result) - val target_id: Option[Session.Entity_ID] = Position.get_id(result.props) + val target_id: Option[Session.Entity_ID] = Position.get_id(result.message.attributes) val target: Option[Session.Entity] = target_id match { case None => None case Some(id) => lookup_entity(id) } if (target.isDefined) target.get.consume(result.message, indicate_command_change) - else if (result.kind == Isabelle_Process.Kind.STATUS) { + else if (result.is_status) { // global status message result.body match { @@ -145,7 +145,7 @@ case _ => if (!result.is_ready) bad_result(result) } } - else if (result.kind == Isabelle_Process.Kind.EXIT) + else if (result.kind == Markup.EXIT) prover = null else if (result.is_raw) raw_output.event(result) @@ -176,7 +176,7 @@ { receiveWithin(timeout) { case result: Isabelle_Process.Result - if result.kind == Isabelle_Process.Kind.INIT => + if result.kind == Markup.INIT => while (receive { case result: Isabelle_Process.Result => handle_result(result); !result.is_ready @@ -184,7 +184,7 @@ None case result: Isabelle_Process.Result - if result.kind == Isabelle_Process.Kind.EXIT => + if result.kind == Markup.EXIT => Some(startup_error()) case TIMEOUT => // FIXME clarify diff -r 6432bf0d7191 -r bf6c1216db43 src/Pure/goal.ML --- a/src/Pure/goal.ML Mon Jul 05 09:14:51 2010 -0700 +++ b/src/Pure/goal.ML Tue Jul 06 08:08:35 2010 -0700 @@ -104,12 +104,7 @@ (* parallel proofs *) -fun fork e = Future.fork_pri ~1 (fn () => - let - val _ = Output.status (Markup.markup Markup.forked ""); - val x = e (); (*sic*) - val _ = Output.status (Markup.markup Markup.joined ""); - in x end); +fun fork e = Future.fork_pri ~1 (fn () => Future.report e); val parallel_proofs = Unsynchronized.ref 1; diff -r 6432bf0d7191 -r bf6c1216db43 src/Pure/library.scala --- a/src/Pure/library.scala Mon Jul 05 09:14:51 2010 -0700 +++ b/src/Pure/library.scala Tue Jul 06 08:08:35 2010 -0700 @@ -129,11 +129,12 @@ def timeit[A](message: String)(e: => A) = { - val start = System.currentTimeMillis() + val start = System.nanoTime() val result = Exn.capture(e) - val stop = System.currentTimeMillis() + val stop = System.nanoTime() System.err.println( - (if (message.isEmpty) "" else message + ": ") + (stop - start) + "ms elapsed time") + (if (message == null || message.isEmpty) "" else message + ": ") + + ((stop - start).toDouble / 1000000) + "ms elapsed time") Exn.release(result) } } diff -r 6432bf0d7191 -r bf6c1216db43 src/Pure/unify.ML --- a/src/Pure/unify.ML Mon Jul 05 09:14:51 2010 -0700 +++ b/src/Pure/unify.ML Tue Jul 06 08:08:35 2010 -0700 @@ -451,20 +451,23 @@ end; +(*If an argument contains a banned Bound, then it should be deleted. + But if the only path is flexible, this is difficult; the code gives up! + In %x y.?a(x) =?= %x y.?b(?c(y)) should we instantiate ?b or ?c *) +exception CHANGE_FAIL; (*flexible occurrence of banned variable, or other reason to quit*) + + (*Flex argument: a term, its type, and the index that refers to it.*) type flarg = {t: term, T: typ, j: int}; (*Form the arguments into records for deletion/sorting.*) fun flexargs ([], [], []) = [] : flarg list | flexargs (j :: js, t :: ts, T :: Ts) = {j = j, t = t, T = T} :: flexargs (js, ts, Ts) - | flexargs _ = raise Fail "flexargs"; - - -(*If an argument contains a banned Bound, then it should be deleted. - But if the only path is flexible, this is difficult; the code gives up! - In %x y.?a(x) =?= %x y.?b(?c(y)) should we instantiate ?b or ?c *) -exception CHANGE_FAIL; (*flexible occurrence of banned variable*) - + | flexargs _ = raise CHANGE_FAIL; +(*We give up if we see a variable of function type not applied to a full list of + arguments (remember, this code assumes that terms are fully eta-expanded). This situation + can occur if a type variable is instantiated with a function type. +*) (*Check whether the 'banned' bound var indices occur rigidly in t*) fun rigid_bound (lev, banned) t = @@ -516,7 +519,7 @@ val (Var (v, T), ts) = strip_comb t; val (Ts, U) = strip_type env T and js = length ts - 1 downto 0; - val args = sort (make_ord arg_less) (List.foldr (change_arg banned) [] (flexargs (js, ts, Ts))) + val args = sort (make_ord arg_less) (List.foldr (change_arg banned) [] (flexargs (js, ts, Ts))) val ts' = map #t args; in if decreasing (length Ts) args then (env, (list_comb (Var (v, T), ts'))) diff -r 6432bf0d7191 -r bf6c1216db43 src/Tools/Code/code_thingol.ML --- a/src/Tools/Code/code_thingol.ML Mon Jul 05 09:14:51 2010 -0700 +++ b/src/Tools/Code/code_thingol.ML Tue Jul 06 08:08:35 2010 -0700 @@ -267,10 +267,7 @@ | purify_base "op &" = "and" | purify_base "op |" = "or" | purify_base "op -->" = "implies" - | purify_base "op :" = "member" | purify_base "op =" = "eq" - | purify_base "*" = "product" - | purify_base "+" = "sum" | purify_base s = Name.desymbolize false s; fun namify thy get_basename get_thyname name = let diff -r 6432bf0d7191 -r bf6c1216db43 src/Tools/jEdit/src/jedit/document_model.scala --- a/src/Tools/jEdit/src/jedit/document_model.scala Mon Jul 05 09:14:51 2010 -0700 +++ b/src/Tools/jEdit/src/jedit/document_model.scala Tue Jul 06 08:08:35 2010 -0700 @@ -95,14 +95,6 @@ def to_current(doc: Document, offset: Int): Int = (offset /: changes_from(doc)) ((i, change) => change after i) - def lines_of_command(doc: Document, cmd: Command): (Int, Int) = - { - val start = doc.command_start(cmd).get // FIXME total? - val stop = start + cmd.length - (buffer.getLineOfOffset(to_current(doc, start)), - buffer.getLineOfOffset(to_current(doc, stop))) - } - /* text edits */ diff -r 6432bf0d7191 -r bf6c1216db43 src/Tools/jEdit/src/jedit/document_view.scala --- a/src/Tools/jEdit/src/jedit/document_view.scala Mon Jul 05 09:14:51 2010 -0700 +++ b/src/Tools/jEdit/src/jedit/document_view.scala Tue Jul 06 08:08:35 2010 -0700 @@ -104,13 +104,10 @@ react { case Command_Set(changed) => Swing_Thread.now { + // FIXME cover doc states as well!!? val document = model.recent_document() - // FIXME cover doc states as well!!? - for (command <- changed if document.commands.contains(command)) { - update_syntax(document, command) - invalidate_line(document, command) - overview.repaint() - } + if (changed.exists(document.commands.contains)) + full_repaint(document, changed) } case bad => System.err.println("command_change_actor: ignoring bad message " + bad) @@ -118,34 +115,82 @@ } } + def full_repaint(document: Document, changed: Set[Command]) + { + Swing_Thread.assert() + + val buffer = model.buffer + var visible_change = false + + for ((command, start) <- document.command_starts) { + if (changed(command)) { + val stop = start + command.length + val line1 = buffer.getLineOfOffset(model.to_current(document, start)) + val line2 = buffer.getLineOfOffset(model.to_current(document, stop)) + if (line2 >= text_area.getFirstLine && + line1 <= text_area.getFirstLine + text_area.getVisibleLines) + visible_change = true + text_area.invalidateLineRange(line1, line2) + } + } + if (visible_change) model.buffer.propertiesChanged() + + overview.repaint() // FIXME paint here!? + } + /* text_area_extension */ private val text_area_extension = new TextAreaExtension { - override def paintValidLine(gfx: Graphics2D, - screen_line: Int, physical_line: Int, start: Int, end: Int, y: Int) + override def paintScreenLineRange(gfx: Graphics2D, + first_line: Int, last_line: Int, physical_lines: Array[Int], + start: Array[Int], end: Array[Int], y0: Int, line_height: Int) { - val document = model.recent_document() - def from_current(pos: Int) = model.from_current(document, pos) - def to_current(pos: Int) = model.to_current(document, pos) + Swing_Thread.now { + val document = model.recent_document() + def from_current(pos: Int) = model.from_current(document, pos) + def to_current(pos: Int) = model.to_current(document, pos) - val line_end = model.visible_line_end(start, end) - val line_height = text_area.getPainter.getFontMetrics.getHeight + val command_range: Iterable[(Command, Int)] = + { + val range = document.command_range(from_current(start(0))) + if (range.hasNext) { + val (cmd0, start0) = range.next + new Iterable[(Command, Int)] { + def iterator = Document.command_starts(document.commands.iterator(cmd0), start0) + } + } + else Iterable.empty + } - val saved_color = gfx.getColor - try { - for ((command, command_start) <- - document.command_range(from_current(start), from_current(line_end)) if !command.is_ignored) - { - val p = text_area.offsetToXY(start max to_current(command_start)) - val q = text_area.offsetToXY(line_end min to_current(command_start + command.length)) - assert(p.y == q.y) - gfx.setColor(Document_View.choose_color(document, command)) - gfx.fillRect(p.x, y, q.x - p.x, line_height) + val saved_color = gfx.getColor + try { + var y = y0 + for (i <- 0 until physical_lines.length) { + if (physical_lines(i) != -1) { + val line_start = start(i) + val line_end = model.visible_line_end(line_start, end(i)) + + val a = from_current(line_start) + val b = from_current(line_end) + val cmds = command_range.iterator. + dropWhile { case (cmd, c) => c + cmd.length <= a } . + takeWhile { case (_, c) => c < b } + + for ((command, command_start) <- cmds if !command.is_ignored) { + val p = text_area.offsetToXY(line_start max to_current(command_start)) + val q = text_area.offsetToXY(line_end min to_current(command_start + command.length)) + assert(p.y == q.y) + gfx.setColor(Document_View.choose_color(document, command)) + gfx.fillRect(p.x, y, q.x - p.x, line_height) + } + } + y += line_height + } } + finally { gfx.setColor(saved_color) } } - finally { gfx.setColor(saved_color) } } override def getToolTipText(x: Int, y: Int): String = @@ -186,30 +231,6 @@ } - /* (re)painting */ - - private val update_delay = Swing_Thread.delay_first(500) { model.buffer.propertiesChanged() } - // FIXME update_delay property - - private def update_syntax(document: Document, cmd: Command) - { - val (line1, line2) = model.lines_of_command(document, cmd) - if (line2 >= text_area.getFirstLine && - line1 <= text_area.getFirstLine + text_area.getVisibleLines) - update_delay() - } - - private def invalidate_line(document: Document, cmd: Command) = - { - val (start, stop) = model.lines_of_command(document, cmd) - text_area.invalidateLineRange(start, stop) - } - - private def invalidate_all() = - text_area.invalidateLineRange(text_area.getFirstPhysicalLine, - text_area.getLastPhysicalLine) - - /* overview of command status left of scrollbar */ private val overview = new JPanel(new BorderLayout) @@ -252,9 +273,9 @@ val buffer = model.buffer val document = model.recent_document() def to_current(pos: Int) = model.to_current(document, pos) - val saved_color = gfx.getColor + val saved_color = gfx.getColor // FIXME needed!? try { - for ((command, start) <- document.command_range(0) if !command.is_ignored) { + for ((command, start) <- document.command_starts if !command.is_ignored) { val line1 = buffer.getLineOfOffset(to_current(start)) val line2 = buffer.getLineOfOffset(to_current(start + command.length)) + 1 val y = line_to_y(line1)