# HG changeset patch # User haftmann # Date 1231431041 -3600 # Node ID ebcd69a00872df7df555696525c41ec0ccd56f57 # Parent aab26a65e80f04183056a359c0c862aa94f9cd8e split of Imperative_HOL theories from HOL-Library diff -r aab26a65e80f -r ebcd69a00872 src/HOL/Imperative_HOL/Array.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Imperative_HOL/Array.thy Thu Jan 08 17:10:41 2009 +0100 @@ -0,0 +1,209 @@ +(* Title: HOL/Library/Array.thy + ID: $Id$ + Author: John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen +*) + +header {* Monadic arrays *} + +theory Array +imports Heap_Monad Code_Index +begin + +subsection {* Primitives *} + +definition + new :: "nat \ 'a\heap \ 'a array Heap" where + [code del]: "new n x = Heap_Monad.heap (Heap.array 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)" + +definition + length :: "'a\heap array \ nat Heap" where + [code del]: "length arr = Heap_Monad.heap (\h. (Heap.length arr h, h))" + +definition + nth :: "'a\heap array \ nat \ 'a Heap" +where + [code del]: "nth a i = (do len \ length a; + (if i < len + then Heap_Monad.heap (\h. (get_array a h ! i, h)) + else raise (''array lookup: index out of range'')) + done)" + +-- {* FIXME adjustion for List theory *} +no_syntax + nth :: "'a list \ nat \ 'a" (infixl "!" 100) + +abbreviation + nth_list :: "'a list \ nat \ 'a" (infixl "!" 100) +where + "nth_list \ List.nth" + +definition + upd :: "nat \ 'a \ 'a\heap array \ 'a\heap array Heap" +where + [code del]: "upd i x a = (do len \ length a; + (if i < len + then Heap_Monad.heap (\h. (a, Heap.upd 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 + + +subsection {* Derivates *} + +definition + map_entry :: "nat \ ('a\heap \ 'a) \ 'a array \ 'a array Heap" +where + "map_entry i f a = (do + x \ nth a i; + upd i (f x) a + done)" + +definition + swap :: "nat \ 'a \ 'a\heap array \ 'a Heap" +where + "swap i x a = (do + y \ nth a i; + upd i x a; + return y + done)" + +definition + make :: "nat \ (nat \ 'a\heap) \ 'a array Heap" +where + "make n f = of_list (map f [0 ..< n])" + +definition + freeze :: "'a\heap array \ 'a list Heap" +where + "freeze a = (do + n \ length a; + mapM (nth a) [0..heap \ 'a) \ 'a array \ 'a array Heap" +where + "map f a = (do + n \ length 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) + +lemma array_of_list_make [code]: + "of_list xs = make (List.length xs) (\n. xs ! n)" + unfolding make_def map_nth .. + + +subsection {* Code generator setup *} + +subsubsection {* Logical intermediate layer *} + +definition new' where + [code del]: "new' = Array.new o nat_of_index" +hide (open) const new' +lemma [code]: + "Array.new = Array.new' o index_of_nat" + by (simp add: new'_def o_def) + +definition of_list' where + [code del]: "of_list' i xs = Array.of_list (take (nat_of_index i) xs)" +hide (open) const of_list' +lemma [code]: + "Array.of_list xs = Array.of_list' (index_of_nat (List.length xs)) xs" + by (simp add: of_list'_def) + +definition make' where + [code del]: "make' i f = Array.make (nat_of_index i) (f o index_of_nat)" +hide (open) const make' +lemma [code]: + "Array.make n f = Array.make' (index_of_nat n) (f o nat_of_index)" + by (simp add: make'_def o_def) + +definition length' where + [code del]: "length' = Array.length \== liftM index_of_nat" +hide (open) const length' +lemma [code]: + "Array.length = Array.length' \== liftM nat_of_index" + by (simp add: length'_def monad_simp', + simp add: liftM_def comp_def monad_simp, + simp add: monad_simp') + +definition nth' where + [code del]: "nth' a = Array.nth a o nat_of_index" +hide (open) const nth' +lemma [code]: + "Array.nth a n = Array.nth' a (index_of_nat n)" + by (simp add: nth'_def) + +definition upd' where + [code del]: "upd' a i x = Array.upd (nat_of_index i) x a \ return ()" +hide (open) const upd' +lemma [code]: + "Array.upd i x a = Array.upd' a (index_of_nat i) x \ return a" + by (simp add: upd'_def monad_simp upd_return) + + +subsubsection {* SML *} + +code_type array (SML "_/ array") +code_const Array (SML "raise/ (Fail/ \"bare Array\")") +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.nth' (SML "(fn/ ()/ =>/ Array.sub/ ((_),/ (_)))") +code_const Array.upd' (SML "(fn/ ()/ =>/ Array.update/ ((_),/ (_),/ (_)))") + +code_reserved SML Array + + +subsubsection {* OCaml *} + +code_type array (OCaml "_/ array") +code_const Array (OCaml "failwith/ \"bare Array\"") +code_const Array.new' (OCaml "(fun/ ()/ ->/ Array.make/ _/ _)") +code_const Array.of_list (OCaml "(fun/ ()/ ->/ Array.of'_list/ _)") +code_const Array.make' (OCaml "(fun/ ()/ ->/ Array.init/ _/ _)") +code_const Array.length' (OCaml "(fun/ ()/ ->/ Array.length/ _)") +code_const Array.nth' (OCaml "(fun/ ()/ ->/ Array.get/ _/ _)") +code_const Array.upd' (OCaml "(fun/ ()/ ->/ Array.set/ _/ _/ _)") + +code_reserved OCaml Array + + +subsubsection {* Haskell *} + +code_type array (Haskell "STArray/ RealWorld/ _") +code_const Array (Haskell "error/ \"bare Array\"") +code_const Array.new' (Haskell "newArray/ (0,/ _)") +code_const Array.of_list' (Haskell "newListArray/ (0,/ _)") +code_const Array.length' (Haskell "lengthArray") +code_const Array.nth' (Haskell "readArray") +code_const Array.upd' (Haskell "writeArray") + +end diff -r aab26a65e80f -r ebcd69a00872 src/HOL/Imperative_HOL/Heap.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Imperative_HOL/Heap.thy Thu Jan 08 17:10:41 2009 +0100 @@ -0,0 +1,434 @@ +(* Title: HOL/Library/Heap.thy + ID: $Id$ + Author: John Matthews, Galois Connections; Alexander Krauss, TU Muenchen +*) + +header {* A polymorphic heap based on cantor encodings *} + +theory Heap +imports Plain "~~/src/HOL/List" Countable Typerep +begin + +subsection {* Representable types *} + +text {* The type class of representable types *} + +class heap = typerep + countable + +text {* Instances for common HOL types *} + +instance nat :: heap .. + +instance "*" :: (heap, heap) heap .. + +instance "+" :: (heap, heap) heap .. + +instance list :: (heap) heap .. + +instance option :: (heap) heap .. + +instance int :: heap .. + +instance message_string :: countable + by (rule countable_classI [of "message_string_case to_nat"]) + (auto split: message_string.splits) + +instance message_string :: 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 *} + +types addr = nat -- "untyped heap references" + +datatype 'a array = Array addr +datatype 'a ref = Ref addr -- "note the phantom type 'a " + +primrec addr_of_array :: "'a array \ addr" where + "addr_of_array (Array x) = x" + +primrec addr_of_ref :: "'a ref \ addr" where + "addr_of_ref (Ref x) = x" + +lemma addr_of_array_inj [simp]: + "addr_of_array a = addr_of_array a' \ a = a'" + by (cases a, cases a') simp_all + +lemma addr_of_ref_inj [simp]: + "addr_of_ref r = addr_of_ref r' \ r = r'" + by (cases r, cases r') simp_all + +instance array :: (type) countable + by (rule countable_classI [of addr_of_array]) simp + +instance ref :: (type) countable + by (rule countable_classI [of addr_of_ref]) simp + +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"}) + #> Sign.add_const_constraint (@{const_name addr_of_array}, SOME @{typ "'a\heap array \ nat"}) + #> 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 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) + +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 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 (open) const empty array array_of_list upd length ref + +end diff -r aab26a65e80f -r ebcd69a00872 src/HOL/Imperative_HOL/Heap_Monad.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Imperative_HOL/Heap_Monad.thy Thu Jan 08 17:10:41 2009 +0100 @@ -0,0 +1,425 @@ +(* Title: HOL/Library/Heap_Monad.thy + ID: $Id$ + Author: John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen +*) + +header {* A monad with a polymorphic heap *} + +theory Heap_Monad +imports Heap +begin + +subsection {* The monad *} + +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" + +primrec + execute :: "'a Heap \ heap \ ('a + exception) \ heap" where + "execute (Heap f) = f" +lemmas [code del] = execute.simps + +lemma Heap_execute [simp]: + "Heap (execute f) = f" by (cases f) simp_all + +lemma Heap_eqI: + "(\h. execute f h = execute g h) \ f = g" + by (cases f, cases g) (auto simp: expand_fun_eq) + +lemma Heap_eqI': + "(\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))" + +lemma execute_heap [simp]: + "execute (heap f) h = apfst Inl (f h)" + 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) + +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 + [code del]: "return x = heap (Pair x)" + +lemma execute_return [simp]: + "execute (return x) h = apfst Inl (x, h)" + by (simp add: return_def) + +definition + raise :: "string \ 'a Heap" where -- {* the string is just decoration *} + [code del]: "raise s = Heap (Pair (Inr Exn))" + +notation (latex output) + "raise" ("\<^raw:{\textsf{raise}}>") + +lemma execute_raise [simp]: + "execute (raise s) h = (Inr Exn, h)" + by (simp add: raise_def) + + +subsubsection {* do-syntax *} + +text {* + We provide a convenient do-notation for monadic expressions + well-known from Haskell. @{const Let} is printed + specially in do-expressions. +*} + +nonterminals do_expr + +syntax + "_do" :: "do_expr \ 'a" + ("(do (_)//done)" [12] 100) + "_bindM" :: "pttrn \ 'a \ do_expr \ do_expr" + ("_ <- _;//_" [1000, 13, 12] 12) + "_chainM" :: "'a \ do_expr \ do_expr" + ("_;//_" [13, 12] 12) + "_let" :: "pttrn \ 'a \ do_expr \ do_expr" + ("let _ = _;//_" [1000, 13, 12] 12) + "_nil" :: "'a \ do_expr" + ("_" [12] 12) + +syntax (xsymbols) + "_bindM" :: "pttrn \ 'a \ do_expr \ do_expr" + ("_ \ _;//_" [1000, 13, 12] 12) +syntax (latex output) + "_do" :: "do_expr \ 'a" + ("(\<^raw:{\textsf{do}}> (_))" [12] 100) + "_let" :: "pttrn \ 'a \ do_expr \ do_expr" + ("\<^raw:\textsf{let}> _ = _;//_" [1000, 13, 12] 12) +notation (latex output) + "return" ("\<^raw:{\textsf{return}}>") + +translations + "_do f" => "f" + "_bindM x f g" => "f \= (\x. g)" + "_chainM f g" => "f \ g" + "_let x t f" => "CONST Let t (\x. f)" + "_nil f" => "f" + +print_translation {* +let + fun dest_abs_eta (Abs (abs as (_, ty, _))) = + let + val (v, t) = Syntax.variant_abs abs; + in (Free (v, ty), t) end + | dest_abs_eta t = + let + val (v, t) = Syntax.variant_abs ("", dummyT, t $ Bound 0); + in (Free (v, dummyT), t) end; + fun unfold_monad (Const (@{const_syntax bindM}, _) $ f $ g) = + let + val (v, g') = dest_abs_eta g; + val vs = fold_aterms (fn Free (v, _) => insert (op =) v | _ => I) v []; + val v_used = fold_aterms + (fn Free (w, _) => (fn s => s orelse member (op =) vs w) | _ => I) g' false; + in if v_used then + Const ("_bindM", dummyT) $ v $ f $ unfold_monad g' + else + Const ("_chainM", dummyT) $ f $ unfold_monad g' + end + | unfold_monad (Const (@{const_syntax chainM}, _) $ f $ g) = + Const ("_chainM", dummyT) $ f $ unfold_monad g + | unfold_monad (Const (@{const_syntax Let}, _) $ f $ g) = + let + val (v, g') = dest_abs_eta g; + in Const ("_let", dummyT) $ v $ f $ unfold_monad g' end + | unfold_monad (Const (@{const_syntax Pair}, _) $ f) = + Const (@{const_syntax return}, dummyT) $ f + | unfold_monad f = f; + fun contains_bindM (Const (@{const_syntax bindM}, _) $ _ $ _) = true + | contains_bindM (Const (@{const_syntax Let}, _) $ _ $ Abs (_, _, t)) = + contains_bindM t; + fun bindM_monad_tr' (f::g::ts) = list_comb + (Const ("_do", dummyT) $ unfold_monad (Const (@{const_syntax bindM}, dummyT) $ f $ g), ts); + fun Let_monad_tr' (f :: (g as Abs (_, _, g')) :: ts) = if contains_bindM g' then list_comb + (Const ("_do", dummyT) $ unfold_monad (Const (@{const_syntax Let}, dummyT) $ f $ g), ts) + else raise Match; +in [ + (@{const_syntax bindM}, bindM_monad_tr'), + (@{const_syntax Let}, Let_monad_tr') +] end; +*} + + +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''))" + +lemma assert_cong [fundef_cong]: + assumes "P = P'" + assumes "\x. P' x \ f x = f' x" + shows "(assert P x >>= f) = (assert P' x >>= f')" + using assms by (auto simp add: assert_def return_bind raise_bind) + +hide (open) const heap execute + + +subsection {* Code generator setup *} + +subsubsection {* Logical intermediate layer *} + +definition + Fail :: "message_string \ exception" +where + [code del]: "Fail s = Exn" + +definition + raise_exc :: "exception \ 'a Heap" +where + [code del]: "raise_exc e = raise []" + +lemma raise_raise_exc [code, code inline]: + "raise s = raise_exc (Fail (STR s))" + unfolding Fail_def raise_exc_def raise_def .. + +hide (open) const Fail raise_exc + + +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_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/ _)") + +setup {* let + open Code_Thingol; + + fun lookup naming = the o Code_Thingol.lookup_const naming; + + fun imp_monad_bind'' bind' return' unit' ts = + let + val dummy_name = ""; + val dummy_type = ITyVar dummy_name; + val dummy_case_term = IVar dummy_name; + (*assumption: dummy values are not relevant for serialization*) + val unitt = IConst (unit', ([], [])); + fun dest_abs ((v, ty) `|-> t, _) = ((v, ty), t) + | dest_abs (t, ty) = + let + val vs = Code_Thingol.fold_varnames cons t []; + val v = Name.variant vs "x"; + val ty' = (hd o fst o Code_Thingol.unfold_fun) ty; + in ((v, ty'), t `$ IVar v) end; + fun force (t as IConst (c, _) `$ t') = if c = return' + then t' else t `$ unitt + | force t = t `$ unitt; + fun tr_bind' [(t1, _), (t2, ty2)] = + let + val ((v, ty), t) = dest_abs (t2, ty2); + in ICase (((force t1, ty), [(IVar v, tr_bind'' t)]), dummy_case_term) end + and tr_bind'' t = case Code_Thingol.unfold_app t + of (IConst (c, (_, ty1 :: ty2 :: _)), [x1, x2]) => if c = bind' + then tr_bind' [(x1, ty1), (x2, ty2)] + else force t + | _ => force t; + in (dummy_name, dummy_type) `|-> ICase (((IVar dummy_name, dummy_type), + [(unitt, tr_bind' ts)]), dummy_case_term) end + and imp_monad_bind' bind' return' unit' (const as (c, (_, tys))) ts = if c = bind' then case (ts, tys) + of ([t1, t2], ty1 :: ty2 :: _) => imp_monad_bind'' bind' return' unit' [(t1, ty1), (t2, ty2)] + | ([t1, t2, t3], ty1 :: ty2 :: _) => imp_monad_bind'' bind' return' unit' [(t1, ty1), (t2, ty2)] `$ t3 + | (ts, _) => imp_monad_bind bind' return' unit' (eta_expand 2 (const, ts)) + else IConst const `$$ map (imp_monad_bind bind' return' unit') ts + and imp_monad_bind bind' return' unit' (IConst const) = imp_monad_bind' bind' return' unit' const [] + | imp_monad_bind bind' return' unit' (t as IVar _) = t + | imp_monad_bind bind' return' unit' (t as _ `$ _) = (case unfold_app t + of (IConst const, ts) => imp_monad_bind' bind' return' unit' const ts + | (t, ts) => imp_monad_bind bind' return' unit' t `$$ map (imp_monad_bind bind' return' unit') ts) + | imp_monad_bind bind' return' unit' (v_ty `|-> t) = v_ty `|-> imp_monad_bind bind' return' unit' t + | imp_monad_bind bind' return' unit' (ICase (((t, ty), pats), t0)) = ICase + (((imp_monad_bind bind' return' unit' t, ty), (map o pairself) (imp_monad_bind bind' return' unit') pats), imp_monad_bind bind' return' unit' t0); + + fun imp_program naming = (Graph.map_nodes o map_terms_stmt) + (imp_monad_bind (lookup naming @{const_name bindM}) + (lookup naming @{const_name return}) + (lookup naming @{const_name Unity})); + +in + + Code_Target.extend_target ("SML_imp", ("SML", imp_program)) + #> Code_Target.extend_target ("OCaml_imp", ("OCaml", imp_program)) + +end +*} + + +code_reserved OCaml Failure raise + + +subsubsection {* Haskell *} + +text {* Adaption layer *} + +code_include Haskell "STMonad" +{*import qualified Control.Monad; +import qualified Control.Monad.ST; +import qualified Data.STRef; +import qualified Data.Array.ST; + +type RealWorld = Control.Monad.ST.RealWorld; +type ST s a = Control.Monad.ST.ST s a; +type STRef s a = Data.STRef.STRef s a; +type STArray s a = Data.Array.ST.STArray s Int a; + +runST :: (forall s. ST s a) -> a; +runST s = Control.Monad.ST.runST s; + +newSTRef = Data.STRef.newSTRef; +readSTRef = Data.STRef.readSTRef; +writeSTRef = Data.STRef.writeSTRef; + +newArray :: (Int, Int) -> a -> ST s (STArray s a); +newArray = Data.Array.ST.newArray; + +newListArray :: (Int, Int) -> [a] -> ST s (STArray s a); +newListArray = Data.Array.ST.newListArray; + +lengthArray :: STArray s a -> ST s Int; +lengthArray a = Control.Monad.liftM snd (Data.Array.ST.getBounds a); + +readArray :: STArray s a -> Int -> ST s a; +readArray = Data.Array.ST.readArray; + +writeArray :: STArray s a -> Int -> a -> ST s (); +writeArray = Data.Array.ST.writeArray;*} + +code_reserved Haskell RealWorld ST STRef Array + runST + newSTRef reasSTRef writeSTRef + newArray newListArray lengthArray readArray writeArray + +text {* Monad *} + +code_type Heap (Haskell "ST/ 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") + +end diff -r aab26a65e80f -r ebcd69a00872 src/HOL/Imperative_HOL/Imperative_HOL.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Imperative_HOL/Imperative_HOL.thy Thu Jan 08 17:10:41 2009 +0100 @@ -0,0 +1,12 @@ +(* Title: HOL/Library/Imperative_HOL.thy + ID: $Id$ + Author: John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen +*) + +header {* Entry point into monadic imperative HOL *} + +theory Imperative_HOL +imports Array Ref Relational +begin + +end diff -r aab26a65e80f -r ebcd69a00872 src/HOL/Imperative_HOL/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Imperative_HOL/ROOT.ML Thu Jan 08 17:10:41 2009 +0100 @@ -0,0 +1,2 @@ + +use_thy "Imperative_HOL"; diff -r aab26a65e80f -r ebcd69a00872 src/HOL/Imperative_HOL/Ref.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Imperative_HOL/Ref.thy Thu Jan 08 17:10:41 2009 +0100 @@ -0,0 +1,91 @@ +(* 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 +begin + +text {* + Imperative reference operations; modeled after their ML counterparts. + See http://caml.inria.fr/pub/docs/manual-caml-light/node14.15.html + and http://www.smlnj.org/doc/Conversion/top-level-comparison.html +*} + +subsection {* Primitives *} + +definition + new :: "'a\heap \ 'a ref Heap" where + [code del]: "new v = Heap_Monad.heap (Heap.ref v)" + +definition + lookup :: "'a\heap ref \ 'a Heap" ("!_" 61) where + [code del]: "lookup r = Heap_Monad.heap (\h. (get_ref r h, 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))" + + +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 (open) const new lookup update change + + +subsection {* Properties *} + +lemma lookup_chain: + "(!r \ f) = f" + by (cases f) + (auto simp add: Let_def bindM_def lookup_def expand_fun_eq) + +lemma update_change [code]: + "r := e = Ref.change (\_. e) r \ return ()" + by (auto simp add: monad_simp change_def lookup_chain) + + +subsection {* Code generator setup *} + +subsubsection {* SML *} + +code_type ref (SML "_/ ref") +code_const Ref (SML "raise/ (Fail/ \"bare Ref\")") +code_const Ref.new (SML "(fn/ ()/ =>/ ref/ _)") +code_const Ref.lookup (SML "(fn/ ()/ =>/ !/ _)") +code_const Ref.update (SML "(fn/ ()/ =>/ _/ :=/ _)") + +code_reserved SML ref + + +subsubsection {* OCaml *} + +code_type ref (OCaml "_/ ref") +code_const Ref (OCaml "failwith/ \"bare Ref\")") +code_const Ref.new (OCaml "(fn/ ()/ =>/ ref/ _)") +code_const Ref.lookup (OCaml "(fn/ ()/ =>/ !/ _)") +code_const Ref.update (OCaml "(fn/ ()/ =>/ _/ :=/ _)") + +code_reserved OCaml ref + + +subsubsection {* Haskell *} + +code_type ref (Haskell "STRef/ RealWorld/ _") +code_const Ref (Haskell "error/ \"bare Ref\"") +code_const Ref.new (Haskell "newSTRef") +code_const Ref.lookup (Haskell "readSTRef") +code_const Ref.update (Haskell "writeSTRef") + +end \ No newline at end of file diff -r aab26a65e80f -r ebcd69a00872 src/HOL/Imperative_HOL/Relational.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Imperative_HOL/Relational.thy Thu Jan 08 17:10:41 2009 +0100 @@ -0,0 +1,700 @@ +theory Relational +imports Array Ref +begin + +section{* Definition of the Relational framework *} + +text {* The crel predicate states that when a computation c runs with the heap h + will result in return value r and a heap h' (if no exception occurs). *} + +definition crel :: "'a Heap \ heap \ heap \ 'a \ bool" +where + crel_def': "crel c h h' r \ Heap_Monad.execute c h = (Inl r, h')" + +lemma crel_def: -- FIXME + "crel c h h' r \ (Inl 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'')" +unfolding crel_def' by auto + +section {* Elimination rules *} + +text {* For all commands, we define simple elimination rules. *} +(* FIXME: consumes 1 necessary ?? *) + +subsection {* Elimination rules for basic monadic commands *} + +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) + +lemma crelE'[consumes 1]: + assumes "crel (f >> g) h h'' r'" + obtains h' r where "crel f h h' r" "crel g h' h'' r'" + using assms + by (elim crelE) auto + +lemma crel_return[consumes 1]: + assumes "crel (return x) h h' r" + obtains "r = x" "h = h'" + using assms + unfolding crel_def return_def by simp + +lemma crel_raise[consumes 1]: + assumes "crel (raise x) h h' r" + obtains "False" + using assms + unfolding crel_def raise_def by simp + +lemma crel_if: + assumes "crel (if c then t else e) h h' r" + obtains "c" "crel t h h' r" + | "\c" "crel e h h' r" + using assms + unfolding crel_def by auto + +lemma crel_option_case: + assumes "crel (case x of None \ n | Some y \ s y) h h' r" + obtains "x = None" "crel n h h' r" + | y where "x = Some y" "crel (s y) h h' r" + using assms + unfolding crel_def by auto + +lemma crel_mapM: + assumes "crel (mapM f xs) h h' r" + assumes "\h h'. P f [] h h' []" + assumes "\h h1 h' x xs y ys. \ crel (f x) h h1 y; crel (mapM f xs) h1 h' ys; P f xs h1 h' ys \ \ P f (x#xs) h h' (y#ys)" + shows "P f xs h h' r" +using assms(1) +proof (induct xs arbitrary: h h' r) + case Nil with assms(2) show ?case + by (auto elim: crel_return) +next + case (Cons x xs) + from Cons(2) obtain h1 y ys where crel_f: "crel (f x) h h1 y" + and crel_mapM: "crel (mapM f xs) h1 h' ys" + and r_def: "r = y#ys" + unfolding mapM.simps + by (auto elim!: crelE crel_return) + from Cons(1)[OF crel_mapM] crel_mapM crel_f assms(3) r_def + show ?case by auto +qed + +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 + +subsection {* Elimination rules for array commands *} + +lemma crel_length: + assumes "crel (length a) h h' r" + obtains "h = h'" "r = Heap.length a h'" + using assms + unfolding length_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) + +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" + 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" + using assms + unfolding upd_def + by (elim crelE crel_if crel_return crel_raise + crel_length crel_heap) auto + +(* Strong version of the lemma for this operation is missing *) +lemma crel_of_list_weak: + assumes "crel (Array.of_list xs) h h' r" + obtains "get_array r h' = xs" + using assms + unfolding of_list_def + by (elim crel_heap) (simp add:get_array_init_array_list) + +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" + 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" + using assms + unfolding Array.swap_def + by (elim crelE crel_upd crel_nth crel_return) auto + +(* Strong version of the lemma for this operation is missing *) +lemma crel_make_weak: + assumes "crel (Array.make n f) h h' r" + obtains "i < n \ get_array r h' ! i = f i" + using assms + unfolding Array.make_def + by (elim crel_of_list_weak) auto + +lemma upt_conv_Cons': + assumes "Suc a \ b" + shows "[b - Suc a.. Heap.length a h" + shows "h = h' \ xs = drop (Heap.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) +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) [Heap.length a h - n..n. map_entry n f a) [Heap.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" + shows "get_array a h' ! i = f (get_array a h ! i)" +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.. 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..n. map_entry n f a) [Heap.length a h - n.. Heap.length a h" + shows "Heap.length a h' = Heap.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) [Heap.length a ?h1 - n..n. map_entry n f a) [0..n. map_entry n f a) [Heap.length a h - Heap.length a h.. \ ref_present x h +extends h' h x \ ref_present x h' +extends h' h x \ ref_present y h \ ref_present y h' +extends h' h x \ ref_present y h \ get_ref y h = get_ref y h' +extends h' h x \ 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'" + (* and "lim h' = Suc (lim h)" *) + and "\y. ref_present y h \ ref_present y h'" + using assms + unfolding Ref.new_def + apply (elim crel_heap) + unfolding Heap.ref_def + apply (simp add: Let_def) + unfolding Heap.new_ref_def + apply (simp add: Let_def) + unfolding ref_present_def + apply auto + unfolding get_ref_def set_ref_def + apply auto + done + +lemma crel_lookup: + assumes "crel (!ref) h h' r" + obtains "h = h'" "r = get_ref ref h" +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 = ()" +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)" +using assms +unfolding Ref.change_def Let_def +by (auto elim!: crelE crel_lookup crel_update crel_return) + +subsection {* Elimination rules for the assert command *} + +lemma crel_assert[consumes 1]: + assumes "crel (assert P x) h h' r" + obtains "P x" "r = x" "h = h'" + using assms + unfolding assert_def + by (elim crel_if crel_return crel_raise) auto + +lemma crel_assert_eq: "(\h h' r. crel f h h' r \ P r) \ f \= assert P = f" +unfolding crel_def bindM_def Let_def assert_def + raise_def return_def prod_case_beta +apply (cases f) +apply simp +apply (simp add: expand_fun_eq split_def) +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 *} + +subsection {* Introduction rules for basic monadic commands *} + +lemma crelI: + assumes "crel f h h' r" "crel (g r) h' h'' r'" + shows "crel (f >>= g) h h'' r'" + using assms by (simp add: crel_def' bindM_def) + +lemma crelI': + assumes "crel f h h' r" "crel g h' h'' r'" + shows "crel (f >> g) h h'' r'" + using assms by (intro crelI) auto + +lemma crel_returnI: + shows "crel (return x) h h x" + unfolding crel_def return_def by simp + +lemma crel_raiseI: + shows "\ (crel (raise x) h h' r)" + unfolding crel_def raise_def by simp + +lemma crel_ifI: + assumes "c \ crel t h h' r" + "\c \ crel e h h' r" + shows "crel (if c then t else e) h h' r" + using assms + unfolding crel_def by auto + +lemma crel_option_caseI: + assumes "\y. x = Some y \ crel (s y) h h' r" + assumes "x = None \ crel n h h' r" + shows "crel (case x of None \ n | Some y \ s y) h h' r" +using assms +by (auto split: option.split) + +lemma crel_heapI: + shows "crel (Heap_Monad.heap f) h (snd (f h)) (fst (f h))" + by (simp add: crel_def apfst_def split_def prod_fun_def) + +lemma crel_heapI': + assumes "h' = snd (f h)" "r = fst (f h)" + shows "crel (Heap_Monad.heap f) h h' r" + using assms + by (simp add: crel_def split_def apfst_def prod_fun_def) + +lemma crelI2: + assumes "\h' rs'. crel f h h' rs' \ (\h'' rs. crel (g rs') h' h'' rs)" + shows "\h'' rs. crel (f\= g) h h'' rs" + oops + +lemma crel_ifI2: + assumes "c \ \h' r. crel t h h' r" + "\ c \ \h' r. crel e h h' r" + shows "\ h' r. crel (if c then t else e) h h' r" + oops + +subsection {* Introduction rules for array commands *} + +lemma crel_lengthI: + shows "crel (length a) h h (Heap.length a h)" + unfolding length_def + by (rule crel_heapI') auto + +(* thm crel_newI for Array.new is missing *) + +lemma crel_nthI: + assumes "i < Heap.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" + using assms + unfolding upd_def + by (auto intro!: crelI crel_ifI crel_returnI crel_raiseI + crel_lengthI crel_heapI') + +(* thm crel_of_listI is missing *) + +(* thm crel_map_entryI is missing *) + +(* thm crel_swapI is missing *) + +(* thm crel_makeI is missing *) + +(* thm crel_freezeI is missing *) + +(* thm crel_mapI is missing *) + +subsection {* Introduction rules for reference commands *} + +lemma crel_lookupI: + shows "crel (!ref) h h (get_ref ref h)" + unfolding lookup_def by (auto intro!: crel_heapI') + +lemma crel_updateI: + shows "crel (ref := v) h (set_ref ref 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))" +unfolding change_def Let_def by (auto intro!: crelI crel_returnI crel_lookupI crel_updateI) + +subsection {* Introduction rules for the assert command *} + +lemma crel_assertI: + assumes "P x" + shows "crel (assert P x) h h x" + using assms + unfolding assert_def + by (auto intro!: crel_ifI crel_returnI crel_raiseI) + +section {* Defintion of the noError predicate *} + +text {* We add a simple definitional setting for crel intro rules + where we only would like to show that the computation does not result in a exception for heap h, + but we do not care about statements about the resulting heap and return value.*} + +definition noError :: "'a Heap \ heap \ bool" +where + "noError c h \ (\r h'. (Inl r, h') = Heap_Monad.execute c h)" + +lemma noError_def': -- FIXME + "noError c h \ (\r h'. Heap_Monad.execute c h = (Inl 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 +qed + +subsection {* Introduction rules for basic monadic commands *} + +lemma noErrorI: + assumes "noError f h" + assumes "\h' r. crel f h h' r \ noError (g r) h'" + shows "noError (f \= g) h" + using assms + by (auto simp add: noError_def' crel_def' bindM_def) + +lemma noErrorI': + assumes "noError f h" + assumes "\h' r. crel f h h' r \ noError g h'" + shows "noError (f \ g) h" + using assms + by (auto simp add: noError_def' crel_def' bindM_def) + +lemma noErrorI2: +"\crel f h h' r ; noError f h; noError (g r) h'\ +\ noError (f \= g) h" +by (auto simp add: noError_def' crel_def' bindM_def) + +lemma noError_return: + shows "noError (return x) h" + unfolding noError_def return_def + by auto + +lemma noError_if: + assumes "c \ noError t h" "\ c \ noError e h" + shows "noError (if c then t else e) h" + using assms + unfolding noError_def + by auto + +lemma noError_option_case: + assumes "\y. x = Some y \ noError (s y) h" + assumes "noError n h" + shows "noError (case x of None \ n | Some y \ s y) h" +using assms +by (auto split: option.split) + +lemma noError_mapM: +assumes "\x \ set xs. noError (f x) h \ crel (f x) h h (r x)" +shows "noError (mapM f xs) h" +using assms +proof (induct xs) + case Nil + thus ?case + unfolding mapM.simps by (intro noError_return) +next + case (Cons x xs) + thus ?case + unfolding mapM.simps + by (auto intro: noErrorI2[of "f x"] noErrorI noError_return) +qed + +lemma noError_heap: + shows "noError (Heap_Monad.heap f) h" + by (simp add: noError_def' apfst_def prod_fun_def split_def) + +subsection {* Introduction rules for array commands *} + +lemma noError_length: + shows "noError (Array.length a) h" + unfolding length_def + by (intro noError_heap) + +lemma noError_new: + shows "noError (Array.new n v) h" +unfolding Array.new_def by (intro noError_heap) + +lemma noError_upd: + assumes "i < Heap.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" + shows "noError (Array.nth a i) h" + using assms + unfolding nth_def + by (auto intro!: noErrorI noError_if noError_return noError_length noError_heap) (auto elim: crel_length) + +lemma noError_of_list: + shows "noError (of_list ls) h" + unfolding of_list_def by (rule noError_heap) + +lemma noError_map_entry: + assumes "i < Heap.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" + shows "noError (swap i x a) h" + using assms + unfolding swap_def + by (auto elim: crel_nth intro!: noErrorI noError_return noError_nth noError_upd) + +lemma noError_make: + shows "noError (make n f) h" + unfolding make_def + by (auto intro: noError_of_list) + +(*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) + +lemma noError_freeze: + shows "noError (freeze a) h" +unfolding freeze_def +by (auto intro!: noErrorI noError_length noError_mapM[of _ _ _ "\x. get_array a h ! x"] + 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.. 'a\heap \ 'a array Heap" where - [code del]: "new n x = Heap_Monad.heap (Heap.array 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)" - -definition - length :: "'a\heap array \ nat Heap" where - [code del]: "length arr = Heap_Monad.heap (\h. (Heap.length arr h, h))" - -definition - nth :: "'a\heap array \ nat \ 'a Heap" -where - [code del]: "nth a i = (do len \ length a; - (if i < len - then Heap_Monad.heap (\h. (get_array a h ! i, h)) - else raise (''array lookup: index out of range'')) - done)" - --- {* FIXME adjustion for List theory *} -no_syntax - nth :: "'a list \ nat \ 'a" (infixl "!" 100) - -abbreviation - nth_list :: "'a list \ nat \ 'a" (infixl "!" 100) -where - "nth_list \ List.nth" - -definition - upd :: "nat \ 'a \ 'a\heap array \ 'a\heap array Heap" -where - [code del]: "upd i x a = (do len \ length a; - (if i < len - then Heap_Monad.heap (\h. (a, Heap.upd 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 - - -subsection {* Derivates *} - -definition - map_entry :: "nat \ ('a\heap \ 'a) \ 'a array \ 'a array Heap" -where - "map_entry i f a = (do - x \ nth a i; - upd i (f x) a - done)" - -definition - swap :: "nat \ 'a \ 'a\heap array \ 'a Heap" -where - "swap i x a = (do - y \ nth a i; - upd i x a; - return y - done)" - -definition - make :: "nat \ (nat \ 'a\heap) \ 'a array Heap" -where - "make n f = of_list (map f [0 ..< n])" - -definition - freeze :: "'a\heap array \ 'a list Heap" -where - "freeze a = (do - n \ length a; - mapM (nth a) [0..heap \ 'a) \ 'a array \ 'a array Heap" -where - "map f a = (do - n \ length 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) - -lemma array_of_list_make [code]: - "of_list xs = make (List.length xs) (\n. xs ! n)" - unfolding make_def map_nth .. - - -subsection {* Code generator setup *} - -subsubsection {* Logical intermediate layer *} - -definition new' where - [code del]: "new' = Array.new o nat_of_index" -hide (open) const new' -lemma [code]: - "Array.new = Array.new' o index_of_nat" - by (simp add: new'_def o_def) - -definition of_list' where - [code del]: "of_list' i xs = Array.of_list (take (nat_of_index i) xs)" -hide (open) const of_list' -lemma [code]: - "Array.of_list xs = Array.of_list' (index_of_nat (List.length xs)) xs" - by (simp add: of_list'_def) - -definition make' where - [code del]: "make' i f = Array.make (nat_of_index i) (f o index_of_nat)" -hide (open) const make' -lemma [code]: - "Array.make n f = Array.make' (index_of_nat n) (f o nat_of_index)" - by (simp add: make'_def o_def) - -definition length' where - [code del]: "length' = Array.length \== liftM index_of_nat" -hide (open) const length' -lemma [code]: - "Array.length = Array.length' \== liftM nat_of_index" - by (simp add: length'_def monad_simp', - simp add: liftM_def comp_def monad_simp, - simp add: monad_simp') - -definition nth' where - [code del]: "nth' a = Array.nth a o nat_of_index" -hide (open) const nth' -lemma [code]: - "Array.nth a n = Array.nth' a (index_of_nat n)" - by (simp add: nth'_def) - -definition upd' where - [code del]: "upd' a i x = Array.upd (nat_of_index i) x a \ return ()" -hide (open) const upd' -lemma [code]: - "Array.upd i x a = Array.upd' a (index_of_nat i) x \ return a" - by (simp add: upd'_def monad_simp upd_return) - - -subsubsection {* SML *} - -code_type array (SML "_/ array") -code_const Array (SML "raise/ (Fail/ \"bare Array\")") -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.nth' (SML "(fn/ ()/ =>/ Array.sub/ ((_),/ (_)))") -code_const Array.upd' (SML "(fn/ ()/ =>/ Array.update/ ((_),/ (_),/ (_)))") - -code_reserved SML Array - - -subsubsection {* OCaml *} - -code_type array (OCaml "_/ array") -code_const Array (OCaml "failwith/ \"bare Array\"") -code_const Array.new' (OCaml "(fun/ ()/ ->/ Array.make/ _/ _)") -code_const Array.of_list (OCaml "(fun/ ()/ ->/ Array.of'_list/ _)") -code_const Array.make' (OCaml "(fun/ ()/ ->/ Array.init/ _/ _)") -code_const Array.length' (OCaml "(fun/ ()/ ->/ Array.length/ _)") -code_const Array.nth' (OCaml "(fun/ ()/ ->/ Array.get/ _/ _)") -code_const Array.upd' (OCaml "(fun/ ()/ ->/ Array.set/ _/ _/ _)") - -code_reserved OCaml Array - - -subsubsection {* Haskell *} - -code_type array (Haskell "STArray/ RealWorld/ _") -code_const Array (Haskell "error/ \"bare Array\"") -code_const Array.new' (Haskell "newArray/ (0,/ _)") -code_const Array.of_list' (Haskell "newListArray/ (0,/ _)") -code_const Array.length' (Haskell "lengthArray") -code_const Array.nth' (Haskell "readArray") -code_const Array.upd' (Haskell "writeArray") - -end diff -r aab26a65e80f -r ebcd69a00872 src/HOL/Library/Heap.thy --- a/src/HOL/Library/Heap.thy Thu Jan 08 10:53:48 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,434 +0,0 @@ -(* Title: HOL/Library/Heap.thy - ID: $Id$ - Author: John Matthews, Galois Connections; Alexander Krauss, TU Muenchen -*) - -header {* A polymorphic heap based on cantor encodings *} - -theory Heap -imports Plain "~~/src/HOL/List" Countable Typerep -begin - -subsection {* Representable types *} - -text {* The type class of representable types *} - -class heap = typerep + countable - -text {* Instances for common HOL types *} - -instance nat :: heap .. - -instance "*" :: (heap, heap) heap .. - -instance "+" :: (heap, heap) heap .. - -instance list :: (heap) heap .. - -instance option :: (heap) heap .. - -instance int :: heap .. - -instance message_string :: countable - by (rule countable_classI [of "message_string_case to_nat"]) - (auto split: message_string.splits) - -instance message_string :: 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 *} - -types addr = nat -- "untyped heap references" - -datatype 'a array = Array addr -datatype 'a ref = Ref addr -- "note the phantom type 'a " - -primrec addr_of_array :: "'a array \ addr" where - "addr_of_array (Array x) = x" - -primrec addr_of_ref :: "'a ref \ addr" where - "addr_of_ref (Ref x) = x" - -lemma addr_of_array_inj [simp]: - "addr_of_array a = addr_of_array a' \ a = a'" - by (cases a, cases a') simp_all - -lemma addr_of_ref_inj [simp]: - "addr_of_ref r = addr_of_ref r' \ r = r'" - by (cases r, cases r') simp_all - -instance array :: (type) countable - by (rule countable_classI [of addr_of_array]) simp - -instance ref :: (type) countable - by (rule countable_classI [of addr_of_ref]) simp - -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"}) - #> Sign.add_const_constraint (@{const_name addr_of_array}, SOME @{typ "'a\heap array \ nat"}) - #> 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 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) - -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 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 (open) const empty array array_of_list upd length ref - -end diff -r aab26a65e80f -r ebcd69a00872 src/HOL/Library/Heap_Monad.thy --- a/src/HOL/Library/Heap_Monad.thy Thu Jan 08 10:53:48 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,425 +0,0 @@ -(* Title: HOL/Library/Heap_Monad.thy - ID: $Id$ - Author: John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen -*) - -header {* A monad with a polymorphic heap *} - -theory Heap_Monad -imports Heap -begin - -subsection {* The monad *} - -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" - -primrec - execute :: "'a Heap \ heap \ ('a + exception) \ heap" where - "execute (Heap f) = f" -lemmas [code del] = execute.simps - -lemma Heap_execute [simp]: - "Heap (execute f) = f" by (cases f) simp_all - -lemma Heap_eqI: - "(\h. execute f h = execute g h) \ f = g" - by (cases f, cases g) (auto simp: expand_fun_eq) - -lemma Heap_eqI': - "(\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))" - -lemma execute_heap [simp]: - "execute (heap f) h = apfst Inl (f h)" - 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) - -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 - [code del]: "return x = heap (Pair x)" - -lemma execute_return [simp]: - "execute (return x) h = apfst Inl (x, h)" - by (simp add: return_def) - -definition - raise :: "string \ 'a Heap" where -- {* the string is just decoration *} - [code del]: "raise s = Heap (Pair (Inr Exn))" - -notation (latex output) - "raise" ("\<^raw:{\textsf{raise}}>") - -lemma execute_raise [simp]: - "execute (raise s) h = (Inr Exn, h)" - by (simp add: raise_def) - - -subsubsection {* do-syntax *} - -text {* - We provide a convenient do-notation for monadic expressions - well-known from Haskell. @{const Let} is printed - specially in do-expressions. -*} - -nonterminals do_expr - -syntax - "_do" :: "do_expr \ 'a" - ("(do (_)//done)" [12] 100) - "_bindM" :: "pttrn \ 'a \ do_expr \ do_expr" - ("_ <- _;//_" [1000, 13, 12] 12) - "_chainM" :: "'a \ do_expr \ do_expr" - ("_;//_" [13, 12] 12) - "_let" :: "pttrn \ 'a \ do_expr \ do_expr" - ("let _ = _;//_" [1000, 13, 12] 12) - "_nil" :: "'a \ do_expr" - ("_" [12] 12) - -syntax (xsymbols) - "_bindM" :: "pttrn \ 'a \ do_expr \ do_expr" - ("_ \ _;//_" [1000, 13, 12] 12) -syntax (latex output) - "_do" :: "do_expr \ 'a" - ("(\<^raw:{\textsf{do}}> (_))" [12] 100) - "_let" :: "pttrn \ 'a \ do_expr \ do_expr" - ("\<^raw:\textsf{let}> _ = _;//_" [1000, 13, 12] 12) -notation (latex output) - "return" ("\<^raw:{\textsf{return}}>") - -translations - "_do f" => "f" - "_bindM x f g" => "f \= (\x. g)" - "_chainM f g" => "f \ g" - "_let x t f" => "CONST Let t (\x. f)" - "_nil f" => "f" - -print_translation {* -let - fun dest_abs_eta (Abs (abs as (_, ty, _))) = - let - val (v, t) = Syntax.variant_abs abs; - in (Free (v, ty), t) end - | dest_abs_eta t = - let - val (v, t) = Syntax.variant_abs ("", dummyT, t $ Bound 0); - in (Free (v, dummyT), t) end; - fun unfold_monad (Const (@{const_syntax bindM}, _) $ f $ g) = - let - val (v, g') = dest_abs_eta g; - val vs = fold_aterms (fn Free (v, _) => insert (op =) v | _ => I) v []; - val v_used = fold_aterms - (fn Free (w, _) => (fn s => s orelse member (op =) vs w) | _ => I) g' false; - in if v_used then - Const ("_bindM", dummyT) $ v $ f $ unfold_monad g' - else - Const ("_chainM", dummyT) $ f $ unfold_monad g' - end - | unfold_monad (Const (@{const_syntax chainM}, _) $ f $ g) = - Const ("_chainM", dummyT) $ f $ unfold_monad g - | unfold_monad (Const (@{const_syntax Let}, _) $ f $ g) = - let - val (v, g') = dest_abs_eta g; - in Const ("_let", dummyT) $ v $ f $ unfold_monad g' end - | unfold_monad (Const (@{const_syntax Pair}, _) $ f) = - Const (@{const_syntax return}, dummyT) $ f - | unfold_monad f = f; - fun contains_bindM (Const (@{const_syntax bindM}, _) $ _ $ _) = true - | contains_bindM (Const (@{const_syntax Let}, _) $ _ $ Abs (_, _, t)) = - contains_bindM t; - fun bindM_monad_tr' (f::g::ts) = list_comb - (Const ("_do", dummyT) $ unfold_monad (Const (@{const_syntax bindM}, dummyT) $ f $ g), ts); - fun Let_monad_tr' (f :: (g as Abs (_, _, g')) :: ts) = if contains_bindM g' then list_comb - (Const ("_do", dummyT) $ unfold_monad (Const (@{const_syntax Let}, dummyT) $ f $ g), ts) - else raise Match; -in [ - (@{const_syntax bindM}, bindM_monad_tr'), - (@{const_syntax Let}, Let_monad_tr') -] end; -*} - - -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''))" - -lemma assert_cong [fundef_cong]: - assumes "P = P'" - assumes "\x. P' x \ f x = f' x" - shows "(assert P x >>= f) = (assert P' x >>= f')" - using assms by (auto simp add: assert_def return_bind raise_bind) - -hide (open) const heap execute - - -subsection {* Code generator setup *} - -subsubsection {* Logical intermediate layer *} - -definition - Fail :: "message_string \ exception" -where - [code del]: "Fail s = Exn" - -definition - raise_exc :: "exception \ 'a Heap" -where - [code del]: "raise_exc e = raise []" - -lemma raise_raise_exc [code, code inline]: - "raise s = raise_exc (Fail (STR s))" - unfolding Fail_def raise_exc_def raise_def .. - -hide (open) const Fail raise_exc - - -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_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/ _)") - -setup {* let - open Code_Thingol; - - fun lookup naming = the o Code_Thingol.lookup_const naming; - - fun imp_monad_bind'' bind' return' unit' ts = - let - val dummy_name = ""; - val dummy_type = ITyVar dummy_name; - val dummy_case_term = IVar dummy_name; - (*assumption: dummy values are not relevant for serialization*) - val unitt = IConst (unit', ([], [])); - fun dest_abs ((v, ty) `|-> t, _) = ((v, ty), t) - | dest_abs (t, ty) = - let - val vs = Code_Thingol.fold_varnames cons t []; - val v = Name.variant vs "x"; - val ty' = (hd o fst o Code_Thingol.unfold_fun) ty; - in ((v, ty'), t `$ IVar v) end; - fun force (t as IConst (c, _) `$ t') = if c = return' - then t' else t `$ unitt - | force t = t `$ unitt; - fun tr_bind' [(t1, _), (t2, ty2)] = - let - val ((v, ty), t) = dest_abs (t2, ty2); - in ICase (((force t1, ty), [(IVar v, tr_bind'' t)]), dummy_case_term) end - and tr_bind'' t = case Code_Thingol.unfold_app t - of (IConst (c, (_, ty1 :: ty2 :: _)), [x1, x2]) => if c = bind' - then tr_bind' [(x1, ty1), (x2, ty2)] - else force t - | _ => force t; - in (dummy_name, dummy_type) `|-> ICase (((IVar dummy_name, dummy_type), - [(unitt, tr_bind' ts)]), dummy_case_term) end - and imp_monad_bind' bind' return' unit' (const as (c, (_, tys))) ts = if c = bind' then case (ts, tys) - of ([t1, t2], ty1 :: ty2 :: _) => imp_monad_bind'' bind' return' unit' [(t1, ty1), (t2, ty2)] - | ([t1, t2, t3], ty1 :: ty2 :: _) => imp_monad_bind'' bind' return' unit' [(t1, ty1), (t2, ty2)] `$ t3 - | (ts, _) => imp_monad_bind bind' return' unit' (eta_expand 2 (const, ts)) - else IConst const `$$ map (imp_monad_bind bind' return' unit') ts - and imp_monad_bind bind' return' unit' (IConst const) = imp_monad_bind' bind' return' unit' const [] - | imp_monad_bind bind' return' unit' (t as IVar _) = t - | imp_monad_bind bind' return' unit' (t as _ `$ _) = (case unfold_app t - of (IConst const, ts) => imp_monad_bind' bind' return' unit' const ts - | (t, ts) => imp_monad_bind bind' return' unit' t `$$ map (imp_monad_bind bind' return' unit') ts) - | imp_monad_bind bind' return' unit' (v_ty `|-> t) = v_ty `|-> imp_monad_bind bind' return' unit' t - | imp_monad_bind bind' return' unit' (ICase (((t, ty), pats), t0)) = ICase - (((imp_monad_bind bind' return' unit' t, ty), (map o pairself) (imp_monad_bind bind' return' unit') pats), imp_monad_bind bind' return' unit' t0); - - fun imp_program naming = (Graph.map_nodes o map_terms_stmt) - (imp_monad_bind (lookup naming @{const_name bindM}) - (lookup naming @{const_name return}) - (lookup naming @{const_name Unity})); - -in - - Code_Target.extend_target ("SML_imp", ("SML", imp_program)) - #> Code_Target.extend_target ("OCaml_imp", ("OCaml", imp_program)) - -end -*} - - -code_reserved OCaml Failure raise - - -subsubsection {* Haskell *} - -text {* Adaption layer *} - -code_include Haskell "STMonad" -{*import qualified Control.Monad; -import qualified Control.Monad.ST; -import qualified Data.STRef; -import qualified Data.Array.ST; - -type RealWorld = Control.Monad.ST.RealWorld; -type ST s a = Control.Monad.ST.ST s a; -type STRef s a = Data.STRef.STRef s a; -type STArray s a = Data.Array.ST.STArray s Int a; - -runST :: (forall s. ST s a) -> a; -runST s = Control.Monad.ST.runST s; - -newSTRef = Data.STRef.newSTRef; -readSTRef = Data.STRef.readSTRef; -writeSTRef = Data.STRef.writeSTRef; - -newArray :: (Int, Int) -> a -> ST s (STArray s a); -newArray = Data.Array.ST.newArray; - -newListArray :: (Int, Int) -> [a] -> ST s (STArray s a); -newListArray = Data.Array.ST.newListArray; - -lengthArray :: STArray s a -> ST s Int; -lengthArray a = Control.Monad.liftM snd (Data.Array.ST.getBounds a); - -readArray :: STArray s a -> Int -> ST s a; -readArray = Data.Array.ST.readArray; - -writeArray :: STArray s a -> Int -> a -> ST s (); -writeArray = Data.Array.ST.writeArray;*} - -code_reserved Haskell RealWorld ST STRef Array - runST - newSTRef reasSTRef writeSTRef - newArray newListArray lengthArray readArray writeArray - -text {* Monad *} - -code_type Heap (Haskell "ST/ 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") - -end diff -r aab26a65e80f -r ebcd69a00872 src/HOL/Library/Imperative_HOL.thy --- a/src/HOL/Library/Imperative_HOL.thy Thu Jan 08 10:53:48 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,12 +0,0 @@ -(* Title: HOL/Library/Imperative_HOL.thy - ID: $Id$ - Author: John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen -*) - -header {* Entry point into monadic imperative HOL *} - -theory Imperative_HOL -imports Array Ref Relational -begin - -end diff -r aab26a65e80f -r ebcd69a00872 src/HOL/Library/Library.thy --- a/src/HOL/Library/Library.thy Thu Jan 08 10:53:48 2009 +0100 +++ b/src/HOL/Library/Library.thy Thu Jan 08 17:10:41 2009 +0100 @@ -22,7 +22,6 @@ Executable_Set Float FuncSet - Imperative_HOL Infinite_Set ListVector Multiset diff -r aab26a65e80f -r ebcd69a00872 src/HOL/Library/Ref.thy --- a/src/HOL/Library/Ref.thy Thu Jan 08 10:53:48 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,91 +0,0 @@ -(* 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 -begin - -text {* - Imperative reference operations; modeled after their ML counterparts. - See http://caml.inria.fr/pub/docs/manual-caml-light/node14.15.html - and http://www.smlnj.org/doc/Conversion/top-level-comparison.html -*} - -subsection {* Primitives *} - -definition - new :: "'a\heap \ 'a ref Heap" where - [code del]: "new v = Heap_Monad.heap (Heap.ref v)" - -definition - lookup :: "'a\heap ref \ 'a Heap" ("!_" 61) where - [code del]: "lookup r = Heap_Monad.heap (\h. (get_ref r h, 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))" - - -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 (open) const new lookup update change - - -subsection {* Properties *} - -lemma lookup_chain: - "(!r \ f) = f" - by (cases f) - (auto simp add: Let_def bindM_def lookup_def expand_fun_eq) - -lemma update_change [code]: - "r := e = Ref.change (\_. e) r \ return ()" - by (auto simp add: monad_simp change_def lookup_chain) - - -subsection {* Code generator setup *} - -subsubsection {* SML *} - -code_type ref (SML "_/ ref") -code_const Ref (SML "raise/ (Fail/ \"bare Ref\")") -code_const Ref.new (SML "(fn/ ()/ =>/ ref/ _)") -code_const Ref.lookup (SML "(fn/ ()/ =>/ !/ _)") -code_const Ref.update (SML "(fn/ ()/ =>/ _/ :=/ _)") - -code_reserved SML ref - - -subsubsection {* OCaml *} - -code_type ref (OCaml "_/ ref") -code_const Ref (OCaml "failwith/ \"bare Ref\")") -code_const Ref.new (OCaml "(fn/ ()/ =>/ ref/ _)") -code_const Ref.lookup (OCaml "(fn/ ()/ =>/ !/ _)") -code_const Ref.update (OCaml "(fn/ ()/ =>/ _/ :=/ _)") - -code_reserved OCaml ref - - -subsubsection {* Haskell *} - -code_type ref (Haskell "STRef/ RealWorld/ _") -code_const Ref (Haskell "error/ \"bare Ref\"") -code_const Ref.new (Haskell "newSTRef") -code_const Ref.lookup (Haskell "readSTRef") -code_const Ref.update (Haskell "writeSTRef") - -end \ No newline at end of file diff -r aab26a65e80f -r ebcd69a00872 src/HOL/Library/Relational.thy --- a/src/HOL/Library/Relational.thy Thu Jan 08 10:53:48 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,700 +0,0 @@ -theory Relational -imports Array Ref -begin - -section{* Definition of the Relational framework *} - -text {* The crel predicate states that when a computation c runs with the heap h - will result in return value r and a heap h' (if no exception occurs). *} - -definition crel :: "'a Heap \ heap \ heap \ 'a \ bool" -where - crel_def': "crel c h h' r \ Heap_Monad.execute c h = (Inl r, h')" - -lemma crel_def: -- FIXME - "crel c h h' r \ (Inl 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'')" -unfolding crel_def' by auto - -section {* Elimination rules *} - -text {* For all commands, we define simple elimination rules. *} -(* FIXME: consumes 1 necessary ?? *) - -subsection {* Elimination rules for basic monadic commands *} - -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) - -lemma crelE'[consumes 1]: - assumes "crel (f >> g) h h'' r'" - obtains h' r where "crel f h h' r" "crel g h' h'' r'" - using assms - by (elim crelE) auto - -lemma crel_return[consumes 1]: - assumes "crel (return x) h h' r" - obtains "r = x" "h = h'" - using assms - unfolding crel_def return_def by simp - -lemma crel_raise[consumes 1]: - assumes "crel (raise x) h h' r" - obtains "False" - using assms - unfolding crel_def raise_def by simp - -lemma crel_if: - assumes "crel (if c then t else e) h h' r" - obtains "c" "crel t h h' r" - | "\c" "crel e h h' r" - using assms - unfolding crel_def by auto - -lemma crel_option_case: - assumes "crel (case x of None \ n | Some y \ s y) h h' r" - obtains "x = None" "crel n h h' r" - | y where "x = Some y" "crel (s y) h h' r" - using assms - unfolding crel_def by auto - -lemma crel_mapM: - assumes "crel (mapM f xs) h h' r" - assumes "\h h'. P f [] h h' []" - assumes "\h h1 h' x xs y ys. \ crel (f x) h h1 y; crel (mapM f xs) h1 h' ys; P f xs h1 h' ys \ \ P f (x#xs) h h' (y#ys)" - shows "P f xs h h' r" -using assms(1) -proof (induct xs arbitrary: h h' r) - case Nil with assms(2) show ?case - by (auto elim: crel_return) -next - case (Cons x xs) - from Cons(2) obtain h1 y ys where crel_f: "crel (f x) h h1 y" - and crel_mapM: "crel (mapM f xs) h1 h' ys" - and r_def: "r = y#ys" - unfolding mapM.simps - by (auto elim!: crelE crel_return) - from Cons(1)[OF crel_mapM] crel_mapM crel_f assms(3) r_def - show ?case by auto -qed - -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 - -subsection {* Elimination rules for array commands *} - -lemma crel_length: - assumes "crel (length a) h h' r" - obtains "h = h'" "r = Heap.length a h'" - using assms - unfolding length_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) - -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" - 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" - using assms - unfolding upd_def - by (elim crelE crel_if crel_return crel_raise - crel_length crel_heap) auto - -(* Strong version of the lemma for this operation is missing *) -lemma crel_of_list_weak: - assumes "crel (Array.of_list xs) h h' r" - obtains "get_array r h' = xs" - using assms - unfolding of_list_def - by (elim crel_heap) (simp add:get_array_init_array_list) - -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" - 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" - using assms - unfolding Array.swap_def - by (elim crelE crel_upd crel_nth crel_return) auto - -(* Strong version of the lemma for this operation is missing *) -lemma crel_make_weak: - assumes "crel (Array.make n f) h h' r" - obtains "i < n \ get_array r h' ! i = f i" - using assms - unfolding Array.make_def - by (elim crel_of_list_weak) auto - -lemma upt_conv_Cons': - assumes "Suc a \ b" - shows "[b - Suc a.. Heap.length a h" - shows "h = h' \ xs = drop (Heap.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) -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) [Heap.length a h - n..n. map_entry n f a) [Heap.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" - shows "get_array a h' ! i = f (get_array a h ! i)" -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.. 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..n. map_entry n f a) [Heap.length a h - n.. Heap.length a h" - shows "Heap.length a h' = Heap.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) [Heap.length a ?h1 - n..n. map_entry n f a) [0..n. map_entry n f a) [Heap.length a h - Heap.length a h.. \ ref_present x h -extends h' h x \ ref_present x h' -extends h' h x \ ref_present y h \ ref_present y h' -extends h' h x \ ref_present y h \ get_ref y h = get_ref y h' -extends h' h x \ 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'" - (* and "lim h' = Suc (lim h)" *) - and "\y. ref_present y h \ ref_present y h'" - using assms - unfolding Ref.new_def - apply (elim crel_heap) - unfolding Heap.ref_def - apply (simp add: Let_def) - unfolding Heap.new_ref_def - apply (simp add: Let_def) - unfolding ref_present_def - apply auto - unfolding get_ref_def set_ref_def - apply auto - done - -lemma crel_lookup: - assumes "crel (!ref) h h' r" - obtains "h = h'" "r = get_ref ref h" -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 = ()" -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)" -using assms -unfolding Ref.change_def Let_def -by (auto elim!: crelE crel_lookup crel_update crel_return) - -subsection {* Elimination rules for the assert command *} - -lemma crel_assert[consumes 1]: - assumes "crel (assert P x) h h' r" - obtains "P x" "r = x" "h = h'" - using assms - unfolding assert_def - by (elim crel_if crel_return crel_raise) auto - -lemma crel_assert_eq: "(\h h' r. crel f h h' r \ P r) \ f \= assert P = f" -unfolding crel_def bindM_def Let_def assert_def - raise_def return_def prod_case_beta -apply (cases f) -apply simp -apply (simp add: expand_fun_eq split_def) -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 *} - -subsection {* Introduction rules for basic monadic commands *} - -lemma crelI: - assumes "crel f h h' r" "crel (g r) h' h'' r'" - shows "crel (f >>= g) h h'' r'" - using assms by (simp add: crel_def' bindM_def) - -lemma crelI': - assumes "crel f h h' r" "crel g h' h'' r'" - shows "crel (f >> g) h h'' r'" - using assms by (intro crelI) auto - -lemma crel_returnI: - shows "crel (return x) h h x" - unfolding crel_def return_def by simp - -lemma crel_raiseI: - shows "\ (crel (raise x) h h' r)" - unfolding crel_def raise_def by simp - -lemma crel_ifI: - assumes "c \ crel t h h' r" - "\c \ crel e h h' r" - shows "crel (if c then t else e) h h' r" - using assms - unfolding crel_def by auto - -lemma crel_option_caseI: - assumes "\y. x = Some y \ crel (s y) h h' r" - assumes "x = None \ crel n h h' r" - shows "crel (case x of None \ n | Some y \ s y) h h' r" -using assms -by (auto split: option.split) - -lemma crel_heapI: - shows "crel (Heap_Monad.heap f) h (snd (f h)) (fst (f h))" - by (simp add: crel_def apfst_def split_def prod_fun_def) - -lemma crel_heapI': - assumes "h' = snd (f h)" "r = fst (f h)" - shows "crel (Heap_Monad.heap f) h h' r" - using assms - by (simp add: crel_def split_def apfst_def prod_fun_def) - -lemma crelI2: - assumes "\h' rs'. crel f h h' rs' \ (\h'' rs. crel (g rs') h' h'' rs)" - shows "\h'' rs. crel (f\= g) h h'' rs" - oops - -lemma crel_ifI2: - assumes "c \ \h' r. crel t h h' r" - "\ c \ \h' r. crel e h h' r" - shows "\ h' r. crel (if c then t else e) h h' r" - oops - -subsection {* Introduction rules for array commands *} - -lemma crel_lengthI: - shows "crel (length a) h h (Heap.length a h)" - unfolding length_def - by (rule crel_heapI') auto - -(* thm crel_newI for Array.new is missing *) - -lemma crel_nthI: - assumes "i < Heap.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" - using assms - unfolding upd_def - by (auto intro!: crelI crel_ifI crel_returnI crel_raiseI - crel_lengthI crel_heapI') - -(* thm crel_of_listI is missing *) - -(* thm crel_map_entryI is missing *) - -(* thm crel_swapI is missing *) - -(* thm crel_makeI is missing *) - -(* thm crel_freezeI is missing *) - -(* thm crel_mapI is missing *) - -subsection {* Introduction rules for reference commands *} - -lemma crel_lookupI: - shows "crel (!ref) h h (get_ref ref h)" - unfolding lookup_def by (auto intro!: crel_heapI') - -lemma crel_updateI: - shows "crel (ref := v) h (set_ref ref 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))" -unfolding change_def Let_def by (auto intro!: crelI crel_returnI crel_lookupI crel_updateI) - -subsection {* Introduction rules for the assert command *} - -lemma crel_assertI: - assumes "P x" - shows "crel (assert P x) h h x" - using assms - unfolding assert_def - by (auto intro!: crel_ifI crel_returnI crel_raiseI) - -section {* Defintion of the noError predicate *} - -text {* We add a simple definitional setting for crel intro rules - where we only would like to show that the computation does not result in a exception for heap h, - but we do not care about statements about the resulting heap and return value.*} - -definition noError :: "'a Heap \ heap \ bool" -where - "noError c h \ (\r h'. (Inl r, h') = Heap_Monad.execute c h)" - -lemma noError_def': -- FIXME - "noError c h \ (\r h'. Heap_Monad.execute c h = (Inl 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 -qed - -subsection {* Introduction rules for basic monadic commands *} - -lemma noErrorI: - assumes "noError f h" - assumes "\h' r. crel f h h' r \ noError (g r) h'" - shows "noError (f \= g) h" - using assms - by (auto simp add: noError_def' crel_def' bindM_def) - -lemma noErrorI': - assumes "noError f h" - assumes "\h' r. crel f h h' r \ noError g h'" - shows "noError (f \ g) h" - using assms - by (auto simp add: noError_def' crel_def' bindM_def) - -lemma noErrorI2: -"\crel f h h' r ; noError f h; noError (g r) h'\ -\ noError (f \= g) h" -by (auto simp add: noError_def' crel_def' bindM_def) - -lemma noError_return: - shows "noError (return x) h" - unfolding noError_def return_def - by auto - -lemma noError_if: - assumes "c \ noError t h" "\ c \ noError e h" - shows "noError (if c then t else e) h" - using assms - unfolding noError_def - by auto - -lemma noError_option_case: - assumes "\y. x = Some y \ noError (s y) h" - assumes "noError n h" - shows "noError (case x of None \ n | Some y \ s y) h" -using assms -by (auto split: option.split) - -lemma noError_mapM: -assumes "\x \ set xs. noError (f x) h \ crel (f x) h h (r x)" -shows "noError (mapM f xs) h" -using assms -proof (induct xs) - case Nil - thus ?case - unfolding mapM.simps by (intro noError_return) -next - case (Cons x xs) - thus ?case - unfolding mapM.simps - by (auto intro: noErrorI2[of "f x"] noErrorI noError_return) -qed - -lemma noError_heap: - shows "noError (Heap_Monad.heap f) h" - by (simp add: noError_def' apfst_def prod_fun_def split_def) - -subsection {* Introduction rules for array commands *} - -lemma noError_length: - shows "noError (Array.length a) h" - unfolding length_def - by (intro noError_heap) - -lemma noError_new: - shows "noError (Array.new n v) h" -unfolding Array.new_def by (intro noError_heap) - -lemma noError_upd: - assumes "i < Heap.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" - shows "noError (Array.nth a i) h" - using assms - unfolding nth_def - by (auto intro!: noErrorI noError_if noError_return noError_length noError_heap) (auto elim: crel_length) - -lemma noError_of_list: - shows "noError (of_list ls) h" - unfolding of_list_def by (rule noError_heap) - -lemma noError_map_entry: - assumes "i < Heap.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" - shows "noError (swap i x a) h" - using assms - unfolding swap_def - by (auto elim: crel_nth intro!: noErrorI noError_return noError_nth noError_upd) - -lemma noError_make: - shows "noError (make n f) h" - unfolding make_def - by (auto intro: noError_of_list) - -(*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) - -lemma noError_freeze: - shows "noError (freeze a) h" -unfolding freeze_def -by (auto intro!: noErrorI noError_length noError_mapM[of _ _ _ "\x. get_array a h ! x"] - 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.. 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) -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) -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 -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 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_0: "m \ Heap.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 -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 -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 -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 -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 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 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) - -end \ No newline at end of file diff -r aab26a65e80f -r ebcd69a00872 src/HOL/Library/Sublist.thy --- a/src/HOL/Library/Sublist.thy Thu Jan 08 10:53:48 2009 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,507 +0,0 @@ -(* $Id$ *) - -header {* Slices of lists *} - -theory Sublist -imports Multiset -begin - - -lemma sublist_split: "i \ j \ j \ k \ sublist xs {i.. j - 1 \ j - 1 \ k - 1") -apply simp -apply (subgoal_tac "{ja. Suc ja < j} = {0.. Suc ja \ Suc ja < k} = {j - Suc 0.. Suc ja \ Suc ja < j} = {i - 1 .. Suc ja \ Suc ja < k} = {j - 1.. Suc j \ Suc j < k} = {i - 1.. j - 1 \ j - 1 \ k - 1") -apply simp -apply fastsimp -apply fastsimp -apply fastsimp -apply fastsimp -done - -lemma sublist_update1: "i \ inds \ sublist (xs[i := v]) inds = sublist xs inds" -apply (induct xs arbitrary: i inds) -apply simp -apply (case_tac i) -apply (simp add: sublist_Cons) -apply (simp add: sublist_Cons) -done - -lemma sublist_update2: "i \ inds \ sublist (xs[i := v]) inds = (sublist xs inds)[(card {k \ inds. k < i}):= v]" -proof (induct xs arbitrary: i inds) - case Nil thus ?case by simp -next - case (Cons x xs) - thus ?case - proof (cases i) - case 0 with Cons show ?thesis by (simp add: sublist_Cons) - next - case (Suc i') - with Cons show ?thesis - apply simp - apply (simp add: sublist_Cons) - apply auto - apply (auto simp add: nat.split) - apply (simp add: card_less) - apply (simp add: card_less) - apply (simp add: card_less_Suc[symmetric]) - apply (simp add: card_less_Suc2) - done - qed -qed - -lemma sublist_update: "sublist (xs[i := v]) inds = (if i \ inds then (sublist xs inds)[(card {k \ inds. k < i}) := v] else sublist xs inds)" -by (simp add: sublist_update1 sublist_update2) - -lemma sublist_take: "sublist xs {j. j < m} = take m xs" -apply (induct xs arbitrary: m) -apply simp -apply (case_tac m) -apply simp -apply (simp add: sublist_Cons) -done - -lemma sublist_take': "sublist xs {0.. sublist xs {a} = [xs ! a]" -apply (induct xs arbitrary: a) -apply simp -apply(case_tac aa) -apply simp -apply (simp add: sublist_Cons) -apply simp -apply (simp add: sublist_Cons) -done - -lemma sublist_is_Nil: "\i \ inds. i \ length xs \ sublist xs inds = []" -apply (induct xs arbitrary: inds) -apply simp -apply (simp add: sublist_Cons) -apply auto -apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) -apply auto -done - -lemma sublist_Nil': "sublist xs inds = [] \ \i \ inds. i \ length xs" -apply (induct xs arbitrary: inds) -apply simp -apply (simp add: sublist_Cons) -apply (auto split: if_splits) -apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) -apply (case_tac x, auto) -done - -lemma sublist_Nil[simp]: "(sublist xs inds = []) = (\i \ inds. i \ length xs)" -apply (induct xs arbitrary: inds) -apply simp -apply (simp add: sublist_Cons) -apply auto -apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) -apply (case_tac x, auto) -done - -lemma sublist_eq_subseteq: " \ inds' \ inds; sublist xs inds = sublist ys inds \ \ sublist xs inds' = sublist ys inds'" -apply (induct xs arbitrary: ys inds inds') -apply simp -apply (drule sym, rule sym) -apply (simp add: sublist_Nil, fastsimp) -apply (case_tac ys) -apply (simp add: sublist_Nil, fastsimp) -apply (auto simp add: sublist_Cons) -apply (erule_tac x="list" in meta_allE) -apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) -apply (erule_tac x="{j. Suc j \ inds'}" in meta_allE) -apply fastsimp -apply (erule_tac x="list" in meta_allE) -apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) -apply (erule_tac x="{j. Suc j \ inds'}" in meta_allE) -apply fastsimp -done - -lemma sublist_eq: "\ \i \ inds. ((i < length xs) \ (i < length ys)) \ ((i \ length xs ) \ (i \ length ys)); \i \ inds. xs ! i = ys ! i \ \ sublist xs inds = sublist ys inds" -apply (induct xs arbitrary: ys inds) -apply simp -apply (rule sym, simp add: sublist_Nil) -apply (case_tac ys) -apply (simp add: sublist_Nil) -apply (auto simp add: sublist_Cons) -apply (erule_tac x="list" in meta_allE) -apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) -apply fastsimp -apply (erule_tac x="list" in meta_allE) -apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) -apply fastsimp -done - -lemma sublist_eq_samelength: "\ length xs = length ys; \i \ inds. xs ! i = ys ! i \ \ sublist xs inds = sublist ys inds" -by (rule sublist_eq, auto) - -lemma sublist_eq_samelength_iff: "length xs = length ys \ (sublist xs inds = sublist ys inds) = (\i \ inds. xs ! i = ys ! i)" -apply (induct xs arbitrary: ys inds) -apply simp -apply (rule sym, simp add: sublist_Nil) -apply (case_tac ys) -apply (simp add: sublist_Nil) -apply (auto simp add: sublist_Cons) -apply (case_tac i) -apply auto -apply (case_tac i) -apply auto -done - -section {* Another sublist function *} - -function sublist' :: "nat \ nat \ 'a list \ 'a list" -where - "sublist' n m [] = []" -| "sublist' n 0 xs = []" -| "sublist' 0 (Suc m) (x#xs) = (x#sublist' 0 m xs)" -| "sublist' (Suc n) (Suc m) (x#xs) = sublist' n m xs" -by pat_completeness auto -termination by lexicographic_order - -subsection {* Proving equivalence to the other sublist command *} - -lemma sublist'_sublist: "sublist' n m xs = sublist xs {j. n \ j \ j < m}" -apply (induct xs arbitrary: n m) -apply simp -apply (case_tac n) -apply (case_tac m) -apply simp -apply (simp add: sublist_Cons) -apply (case_tac m) -apply simp -apply (simp add: sublist_Cons) -done - - -lemma "sublist' n m xs = sublist xs {n.. (x # sublist' 0 j xs) | Suc i' \ sublist' i' j xs)" -by (cases i) auto - -lemma sublist'_Cons2[simp]: "sublist' i j (x#xs) = (if (j = 0) then [] else ((if (i = 0) then [x] else []) @ sublist' (i - 1) (j - 1) xs))" -apply (cases j) -apply auto -apply (cases i) -apply auto -done - -lemma sublist_n_0: "sublist' n 0 xs = []" -by (induct xs, auto) - -lemma sublist'_Nil': "n \ m \ sublist' n m xs = []" -apply (induct xs arbitrary: n m) -apply simp -apply (case_tac m) -apply simp -apply (case_tac n) -apply simp -apply simp -done - -lemma sublist'_Nil2: "n \ length xs \ sublist' n m xs = []" -apply (induct xs arbitrary: n m) -apply simp -apply (case_tac m) -apply simp -apply (case_tac n) -apply simp -apply simp -done - -lemma sublist'_Nil3: "(sublist' n m xs = []) = ((n \ m) \ (n \ length xs))" -apply (induct xs arbitrary: n m) -apply simp -apply (case_tac m) -apply simp -apply (case_tac n) -apply simp -apply simp -done - -lemma sublist'_notNil: "\ n < length xs; n < m \ \ sublist' n m xs \ []" -apply (induct xs arbitrary: n m) -apply simp -apply (case_tac m) -apply simp -apply (case_tac n) -apply simp -apply simp -done - -lemma sublist'_single: "n < length xs \ sublist' n (Suc n) xs = [xs ! n]" -apply (induct xs arbitrary: n) -apply simp -apply simp -apply (case_tac n) -apply (simp add: sublist_n_0) -apply simp -done - -lemma sublist'_update1: "i \ m \ sublist' n m (xs[i:=v]) = sublist' n m xs" -apply (induct xs arbitrary: n m i) -apply simp -apply simp -apply (case_tac i) -apply simp -apply simp -done - -lemma sublist'_update2: "i < n \ sublist' n m (xs[i:=v]) = sublist' n m xs" -apply (induct xs arbitrary: n m i) -apply simp -apply simp -apply (case_tac i) -apply simp -apply simp -done - -lemma sublist'_update3: "\n \ i; i < m\ \ sublist' n m (xs[i := v]) = (sublist' n m xs)[i - n := v]" -proof (induct xs arbitrary: n m i) - case Nil thus ?case by auto -next - case (Cons x xs) - thus ?case - apply - - apply auto - apply (cases i) - apply auto - apply (cases i) - apply auto - done -qed - -lemma "\ sublist' i j xs = sublist' i j ys; n \ i; m \ j \ \ sublist' n m xs = sublist' n m ys" -proof (induct xs arbitrary: i j ys n m) - case Nil - thus ?case - apply - - apply (rule sym, drule sym) - apply (simp add: sublist'_Nil) - apply (simp add: sublist'_Nil3) - apply arith - done -next - case (Cons x xs i j ys n m) - note c = this - thus ?case - proof (cases m) - case 0 thus ?thesis by (simp add: sublist_n_0) - next - case (Suc m') - note a = this - thus ?thesis - proof (cases n) - case 0 note b = this - show ?thesis - proof (cases ys) - case Nil with a b Cons.prems show ?thesis by (simp add: sublist'_Nil3) - next - case (Cons y ys) - show ?thesis - proof (cases j) - case 0 with a b Cons.prems show ?thesis by simp - next - case (Suc j') with a b Cons.prems Cons show ?thesis - apply - - apply (simp, rule Cons.hyps [of "0" "j'" "ys" "0" "m'"], auto) - done - qed - qed - next - case (Suc n') - show ?thesis - proof (cases ys) - case Nil with Suc a Cons.prems show ?thesis by (auto simp add: sublist'_Nil3) - next - case (Cons y ys) with Suc a Cons.prems show ?thesis - apply - - apply simp - apply (cases j) - apply simp - apply (cases i) - apply simp - apply (rule_tac j="nat" in Cons.hyps [of "0" _ "ys" "n'" "m'"]) - apply simp - apply simp - apply simp - apply simp - apply (rule_tac i="nata" and j="nat" in Cons.hyps [of _ _ "ys" "n'" "m'"]) - apply simp - apply simp - apply simp - done - qed - qed - qed -qed - -lemma length_sublist': "j \ length xs \ length (sublist' i j xs) = j - i" -by (induct xs arbitrary: i j, auto) - -lemma sublist'_front: "\ i < j; i < length xs \ \ sublist' i j xs = xs ! i # sublist' (Suc i) j xs" -apply (induct xs arbitrary: a i j) -apply simp -apply (case_tac j) -apply simp -apply (case_tac i) -apply simp -apply simp -done - -lemma sublist'_back: "\ i < j; j \ length xs \ \ sublist' i j xs = sublist' i (j - 1) xs @ [xs ! (j - 1)]" -apply (induct xs arbitrary: a i j) -apply simp -apply simp -apply (case_tac j) -apply simp -apply auto -apply (case_tac nat) -apply auto -done - -(* suffices that j \ length xs and length ys *) -lemma sublist'_eq_samelength_iff: "length xs = length ys \ (sublist' i j xs = sublist' i j ys) = (\i'. i \ i' \ i' < j \ xs ! i' = ys ! i')" -proof (induct xs arbitrary: ys i j) - case Nil thus ?case by simp -next - case (Cons x xs) - thus ?case - apply - - apply (cases ys) - apply simp - apply simp - apply auto - apply (case_tac i', auto) - apply (erule_tac x="Suc i'" in allE, auto) - apply (erule_tac x="i' - 1" in allE, auto) - apply (case_tac i', auto) - apply (erule_tac x="Suc i'" in allE, auto) - done -qed - -lemma sublist'_all[simp]: "sublist' 0 (length xs) xs = xs" -by (induct xs, auto) - -lemma sublist'_sublist': "sublist' n m (sublist' i j xs) = sublist' (i + n) (min (i + m) j) xs" -by (induct xs arbitrary: i j n m) (auto simp add: min_diff) - -lemma sublist'_append: "\ i \ j; j \ k \ \(sublist' i j xs) @ (sublist' j k xs) = sublist' i k xs" -by (induct xs arbitrary: i j k) auto - -lemma nth_sublist': "\ k < j - i; j \ length xs \ \ (sublist' i j xs) ! k = xs ! (i + k)" -apply (induct xs arbitrary: i j k) -apply auto -apply (case_tac k) -apply auto -apply (case_tac i) -apply auto -done - -lemma set_sublist': "set (sublist' i j xs) = {x. \k. i \ k \ k < j \ k < List.length xs \ x = xs ! k}" -apply (simp add: sublist'_sublist) -apply (simp add: set_sublist) -apply auto -done - -lemma all_in_set_sublist'_conv: "(\j. j \ set (sublist' l r xs) \ P j) = (\k. l \ k \ k < r \ k < List.length xs \ P (xs ! k))" -unfolding set_sublist' by blast - -lemma ball_in_set_sublist'_conv: "(\j \ set (sublist' l r xs). P j) = (\k. l \ k \ k < r \ k < List.length xs \ P (xs ! k))" -unfolding set_sublist' by blast - - -lemma multiset_of_sublist: -assumes l_r: "l \ r \ r \ List.length xs" -assumes left: "\ i. i < l \ (xs::'a list) ! i = ys ! i" -assumes right: "\ i. i \ r \ (xs::'a list) ! i = ys ! i" -assumes multiset: "multiset_of xs = multiset_of ys" - shows "multiset_of (sublist' l r xs) = multiset_of (sublist' l r ys)" -proof - - from l_r have xs_def: "xs = (sublist' 0 l xs) @ (sublist' l r xs) @ (sublist' r (List.length xs) xs)" (is "_ = ?xs_long") - by (simp add: sublist'_append) - from multiset have length_eq: "List.length xs = List.length ys" by (rule multiset_of_eq_length) - with l_r have ys_def: "ys = (sublist' 0 l ys) @ (sublist' l r ys) @ (sublist' r (List.length ys) ys)" (is "_ = ?ys_long") - by (simp add: sublist'_append) - from xs_def ys_def multiset have "multiset_of ?xs_long = multiset_of ?ys_long" by simp - moreover - from left l_r length_eq have "sublist' 0 l xs = sublist' 0 l ys" - by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI) - moreover - from right l_r length_eq have "sublist' r (List.length xs) xs = sublist' r (List.length ys) ys" - by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI) - moreover - ultimately show ?thesis by (simp add: multiset_of_append) -qed - - -end diff -r aab26a65e80f -r ebcd69a00872 src/HOL/ex/ImperativeQuicksort.thy --- a/src/HOL/ex/ImperativeQuicksort.thy Thu Jan 08 10:53:48 2009 +0100 +++ b/src/HOL/ex/ImperativeQuicksort.thy Thu Jan 08 17:10:41 2009 +0100 @@ -1,5 +1,5 @@ theory ImperativeQuicksort -imports Imperative_HOL Subarray Multiset Efficient_Nat +imports "~~/src/HOL/Imperative_HOL/Imperative_HOL" Subarray Multiset Efficient_Nat begin text {* We prove QuickSort correct in the Relational Calculus. *} diff -r aab26a65e80f -r ebcd69a00872 src/HOL/ex/Subarray.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Subarray.thy Thu Jan 08 17:10:41 2009 +0100 @@ -0,0 +1,66 @@ +theory Subarray +imports Array Sublist +begin + +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) +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) +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 +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 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_0: "m \ Heap.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 +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 +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 +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 +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 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 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) + +end \ No newline at end of file diff -r aab26a65e80f -r ebcd69a00872 src/HOL/ex/Sublist.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Sublist.thy Thu Jan 08 17:10:41 2009 +0100 @@ -0,0 +1,507 @@ +(* $Id$ *) + +header {* Slices of lists *} + +theory Sublist +imports Multiset +begin + + +lemma sublist_split: "i \ j \ j \ k \ sublist xs {i.. j - 1 \ j - 1 \ k - 1") +apply simp +apply (subgoal_tac "{ja. Suc ja < j} = {0.. Suc ja \ Suc ja < k} = {j - Suc 0.. Suc ja \ Suc ja < j} = {i - 1 .. Suc ja \ Suc ja < k} = {j - 1.. Suc j \ Suc j < k} = {i - 1.. j - 1 \ j - 1 \ k - 1") +apply simp +apply fastsimp +apply fastsimp +apply fastsimp +apply fastsimp +done + +lemma sublist_update1: "i \ inds \ sublist (xs[i := v]) inds = sublist xs inds" +apply (induct xs arbitrary: i inds) +apply simp +apply (case_tac i) +apply (simp add: sublist_Cons) +apply (simp add: sublist_Cons) +done + +lemma sublist_update2: "i \ inds \ sublist (xs[i := v]) inds = (sublist xs inds)[(card {k \ inds. k < i}):= v]" +proof (induct xs arbitrary: i inds) + case Nil thus ?case by simp +next + case (Cons x xs) + thus ?case + proof (cases i) + case 0 with Cons show ?thesis by (simp add: sublist_Cons) + next + case (Suc i') + with Cons show ?thesis + apply simp + apply (simp add: sublist_Cons) + apply auto + apply (auto simp add: nat.split) + apply (simp add: card_less) + apply (simp add: card_less) + apply (simp add: card_less_Suc[symmetric]) + apply (simp add: card_less_Suc2) + done + qed +qed + +lemma sublist_update: "sublist (xs[i := v]) inds = (if i \ inds then (sublist xs inds)[(card {k \ inds. k < i}) := v] else sublist xs inds)" +by (simp add: sublist_update1 sublist_update2) + +lemma sublist_take: "sublist xs {j. j < m} = take m xs" +apply (induct xs arbitrary: m) +apply simp +apply (case_tac m) +apply simp +apply (simp add: sublist_Cons) +done + +lemma sublist_take': "sublist xs {0.. sublist xs {a} = [xs ! a]" +apply (induct xs arbitrary: a) +apply simp +apply(case_tac aa) +apply simp +apply (simp add: sublist_Cons) +apply simp +apply (simp add: sublist_Cons) +done + +lemma sublist_is_Nil: "\i \ inds. i \ length xs \ sublist xs inds = []" +apply (induct xs arbitrary: inds) +apply simp +apply (simp add: sublist_Cons) +apply auto +apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) +apply auto +done + +lemma sublist_Nil': "sublist xs inds = [] \ \i \ inds. i \ length xs" +apply (induct xs arbitrary: inds) +apply simp +apply (simp add: sublist_Cons) +apply (auto split: if_splits) +apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) +apply (case_tac x, auto) +done + +lemma sublist_Nil[simp]: "(sublist xs inds = []) = (\i \ inds. i \ length xs)" +apply (induct xs arbitrary: inds) +apply simp +apply (simp add: sublist_Cons) +apply auto +apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) +apply (case_tac x, auto) +done + +lemma sublist_eq_subseteq: " \ inds' \ inds; sublist xs inds = sublist ys inds \ \ sublist xs inds' = sublist ys inds'" +apply (induct xs arbitrary: ys inds inds') +apply simp +apply (drule sym, rule sym) +apply (simp add: sublist_Nil, fastsimp) +apply (case_tac ys) +apply (simp add: sublist_Nil, fastsimp) +apply (auto simp add: sublist_Cons) +apply (erule_tac x="list" in meta_allE) +apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) +apply (erule_tac x="{j. Suc j \ inds'}" in meta_allE) +apply fastsimp +apply (erule_tac x="list" in meta_allE) +apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) +apply (erule_tac x="{j. Suc j \ inds'}" in meta_allE) +apply fastsimp +done + +lemma sublist_eq: "\ \i \ inds. ((i < length xs) \ (i < length ys)) \ ((i \ length xs ) \ (i \ length ys)); \i \ inds. xs ! i = ys ! i \ \ sublist xs inds = sublist ys inds" +apply (induct xs arbitrary: ys inds) +apply simp +apply (rule sym, simp add: sublist_Nil) +apply (case_tac ys) +apply (simp add: sublist_Nil) +apply (auto simp add: sublist_Cons) +apply (erule_tac x="list" in meta_allE) +apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) +apply fastsimp +apply (erule_tac x="list" in meta_allE) +apply (erule_tac x="{j. Suc j \ inds}" in meta_allE) +apply fastsimp +done + +lemma sublist_eq_samelength: "\ length xs = length ys; \i \ inds. xs ! i = ys ! i \ \ sublist xs inds = sublist ys inds" +by (rule sublist_eq, auto) + +lemma sublist_eq_samelength_iff: "length xs = length ys \ (sublist xs inds = sublist ys inds) = (\i \ inds. xs ! i = ys ! i)" +apply (induct xs arbitrary: ys inds) +apply simp +apply (rule sym, simp add: sublist_Nil) +apply (case_tac ys) +apply (simp add: sublist_Nil) +apply (auto simp add: sublist_Cons) +apply (case_tac i) +apply auto +apply (case_tac i) +apply auto +done + +section {* Another sublist function *} + +function sublist' :: "nat \ nat \ 'a list \ 'a list" +where + "sublist' n m [] = []" +| "sublist' n 0 xs = []" +| "sublist' 0 (Suc m) (x#xs) = (x#sublist' 0 m xs)" +| "sublist' (Suc n) (Suc m) (x#xs) = sublist' n m xs" +by pat_completeness auto +termination by lexicographic_order + +subsection {* Proving equivalence to the other sublist command *} + +lemma sublist'_sublist: "sublist' n m xs = sublist xs {j. n \ j \ j < m}" +apply (induct xs arbitrary: n m) +apply simp +apply (case_tac n) +apply (case_tac m) +apply simp +apply (simp add: sublist_Cons) +apply (case_tac m) +apply simp +apply (simp add: sublist_Cons) +done + + +lemma "sublist' n m xs = sublist xs {n.. (x # sublist' 0 j xs) | Suc i' \ sublist' i' j xs)" +by (cases i) auto + +lemma sublist'_Cons2[simp]: "sublist' i j (x#xs) = (if (j = 0) then [] else ((if (i = 0) then [x] else []) @ sublist' (i - 1) (j - 1) xs))" +apply (cases j) +apply auto +apply (cases i) +apply auto +done + +lemma sublist_n_0: "sublist' n 0 xs = []" +by (induct xs, auto) + +lemma sublist'_Nil': "n \ m \ sublist' n m xs = []" +apply (induct xs arbitrary: n m) +apply simp +apply (case_tac m) +apply simp +apply (case_tac n) +apply simp +apply simp +done + +lemma sublist'_Nil2: "n \ length xs \ sublist' n m xs = []" +apply (induct xs arbitrary: n m) +apply simp +apply (case_tac m) +apply simp +apply (case_tac n) +apply simp +apply simp +done + +lemma sublist'_Nil3: "(sublist' n m xs = []) = ((n \ m) \ (n \ length xs))" +apply (induct xs arbitrary: n m) +apply simp +apply (case_tac m) +apply simp +apply (case_tac n) +apply simp +apply simp +done + +lemma sublist'_notNil: "\ n < length xs; n < m \ \ sublist' n m xs \ []" +apply (induct xs arbitrary: n m) +apply simp +apply (case_tac m) +apply simp +apply (case_tac n) +apply simp +apply simp +done + +lemma sublist'_single: "n < length xs \ sublist' n (Suc n) xs = [xs ! n]" +apply (induct xs arbitrary: n) +apply simp +apply simp +apply (case_tac n) +apply (simp add: sublist_n_0) +apply simp +done + +lemma sublist'_update1: "i \ m \ sublist' n m (xs[i:=v]) = sublist' n m xs" +apply (induct xs arbitrary: n m i) +apply simp +apply simp +apply (case_tac i) +apply simp +apply simp +done + +lemma sublist'_update2: "i < n \ sublist' n m (xs[i:=v]) = sublist' n m xs" +apply (induct xs arbitrary: n m i) +apply simp +apply simp +apply (case_tac i) +apply simp +apply simp +done + +lemma sublist'_update3: "\n \ i; i < m\ \ sublist' n m (xs[i := v]) = (sublist' n m xs)[i - n := v]" +proof (induct xs arbitrary: n m i) + case Nil thus ?case by auto +next + case (Cons x xs) + thus ?case + apply - + apply auto + apply (cases i) + apply auto + apply (cases i) + apply auto + done +qed + +lemma "\ sublist' i j xs = sublist' i j ys; n \ i; m \ j \ \ sublist' n m xs = sublist' n m ys" +proof (induct xs arbitrary: i j ys n m) + case Nil + thus ?case + apply - + apply (rule sym, drule sym) + apply (simp add: sublist'_Nil) + apply (simp add: sublist'_Nil3) + apply arith + done +next + case (Cons x xs i j ys n m) + note c = this + thus ?case + proof (cases m) + case 0 thus ?thesis by (simp add: sublist_n_0) + next + case (Suc m') + note a = this + thus ?thesis + proof (cases n) + case 0 note b = this + show ?thesis + proof (cases ys) + case Nil with a b Cons.prems show ?thesis by (simp add: sublist'_Nil3) + next + case (Cons y ys) + show ?thesis + proof (cases j) + case 0 with a b Cons.prems show ?thesis by simp + next + case (Suc j') with a b Cons.prems Cons show ?thesis + apply - + apply (simp, rule Cons.hyps [of "0" "j'" "ys" "0" "m'"], auto) + done + qed + qed + next + case (Suc n') + show ?thesis + proof (cases ys) + case Nil with Suc a Cons.prems show ?thesis by (auto simp add: sublist'_Nil3) + next + case (Cons y ys) with Suc a Cons.prems show ?thesis + apply - + apply simp + apply (cases j) + apply simp + apply (cases i) + apply simp + apply (rule_tac j="nat" in Cons.hyps [of "0" _ "ys" "n'" "m'"]) + apply simp + apply simp + apply simp + apply simp + apply (rule_tac i="nata" and j="nat" in Cons.hyps [of _ _ "ys" "n'" "m'"]) + apply simp + apply simp + apply simp + done + qed + qed + qed +qed + +lemma length_sublist': "j \ length xs \ length (sublist' i j xs) = j - i" +by (induct xs arbitrary: i j, auto) + +lemma sublist'_front: "\ i < j; i < length xs \ \ sublist' i j xs = xs ! i # sublist' (Suc i) j xs" +apply (induct xs arbitrary: a i j) +apply simp +apply (case_tac j) +apply simp +apply (case_tac i) +apply simp +apply simp +done + +lemma sublist'_back: "\ i < j; j \ length xs \ \ sublist' i j xs = sublist' i (j - 1) xs @ [xs ! (j - 1)]" +apply (induct xs arbitrary: a i j) +apply simp +apply simp +apply (case_tac j) +apply simp +apply auto +apply (case_tac nat) +apply auto +done + +(* suffices that j \ length xs and length ys *) +lemma sublist'_eq_samelength_iff: "length xs = length ys \ (sublist' i j xs = sublist' i j ys) = (\i'. i \ i' \ i' < j \ xs ! i' = ys ! i')" +proof (induct xs arbitrary: ys i j) + case Nil thus ?case by simp +next + case (Cons x xs) + thus ?case + apply - + apply (cases ys) + apply simp + apply simp + apply auto + apply (case_tac i', auto) + apply (erule_tac x="Suc i'" in allE, auto) + apply (erule_tac x="i' - 1" in allE, auto) + apply (case_tac i', auto) + apply (erule_tac x="Suc i'" in allE, auto) + done +qed + +lemma sublist'_all[simp]: "sublist' 0 (length xs) xs = xs" +by (induct xs, auto) + +lemma sublist'_sublist': "sublist' n m (sublist' i j xs) = sublist' (i + n) (min (i + m) j) xs" +by (induct xs arbitrary: i j n m) (auto simp add: min_diff) + +lemma sublist'_append: "\ i \ j; j \ k \ \(sublist' i j xs) @ (sublist' j k xs) = sublist' i k xs" +by (induct xs arbitrary: i j k) auto + +lemma nth_sublist': "\ k < j - i; j \ length xs \ \ (sublist' i j xs) ! k = xs ! (i + k)" +apply (induct xs arbitrary: i j k) +apply auto +apply (case_tac k) +apply auto +apply (case_tac i) +apply auto +done + +lemma set_sublist': "set (sublist' i j xs) = {x. \k. i \ k \ k < j \ k < List.length xs \ x = xs ! k}" +apply (simp add: sublist'_sublist) +apply (simp add: set_sublist) +apply auto +done + +lemma all_in_set_sublist'_conv: "(\j. j \ set (sublist' l r xs) \ P j) = (\k. l \ k \ k < r \ k < List.length xs \ P (xs ! k))" +unfolding set_sublist' by blast + +lemma ball_in_set_sublist'_conv: "(\j \ set (sublist' l r xs). P j) = (\k. l \ k \ k < r \ k < List.length xs \ P (xs ! k))" +unfolding set_sublist' by blast + + +lemma multiset_of_sublist: +assumes l_r: "l \ r \ r \ List.length xs" +assumes left: "\ i. i < l \ (xs::'a list) ! i = ys ! i" +assumes right: "\ i. i \ r \ (xs::'a list) ! i = ys ! i" +assumes multiset: "multiset_of xs = multiset_of ys" + shows "multiset_of (sublist' l r xs) = multiset_of (sublist' l r ys)" +proof - + from l_r have xs_def: "xs = (sublist' 0 l xs) @ (sublist' l r xs) @ (sublist' r (List.length xs) xs)" (is "_ = ?xs_long") + by (simp add: sublist'_append) + from multiset have length_eq: "List.length xs = List.length ys" by (rule multiset_of_eq_length) + with l_r have ys_def: "ys = (sublist' 0 l ys) @ (sublist' l r ys) @ (sublist' r (List.length ys) ys)" (is "_ = ?ys_long") + by (simp add: sublist'_append) + from xs_def ys_def multiset have "multiset_of ?xs_long = multiset_of ?ys_long" by simp + moreover + from left l_r length_eq have "sublist' 0 l xs = sublist' 0 l ys" + by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI) + moreover + from right l_r length_eq have "sublist' r (List.length xs) xs = sublist' r (List.length ys) ys" + by (auto simp add: length_sublist' nth_sublist' intro!: nth_equalityI) + moreover + ultimately show ?thesis by (simp add: multiset_of_append) +qed + + +end