# HG changeset patch # User haftmann # Date 1278400875 -7200 # Node ID 6d28a2aea936081d51f05cb6f5596b2df273a391 # Parent 6607ccf779467cd0110099bbac2a97d2ebdc900c refactored reference operations diff -r 6607ccf77946 -r 6d28a2aea936 src/HOL/Imperative_HOL/Ref.thy --- a/src/HOL/Imperative_HOL/Ref.thy Tue Jul 06 09:21:13 2010 +0200 +++ b/src/HOL/Imperative_HOL/Ref.thy Tue Jul 06 09:21:15 2010 +0200 @@ -16,185 +16,175 @@ subsection {* Primitive layer *} -definition - ref_present :: "'a\heap ref \ heap \ bool" where - "ref_present r h \ addr_of_ref r < lim h" +definition present :: "heap \ 'a\heap ref \ bool" where + "present h r \ addr_of_ref r < 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 :: "heap \ 'a\heap ref \ 'a" where + "get h = from_nat \ refs h TYPEREP('a) \ addr_of_ref" -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 :: "'a\heap ref \ 'a \ heap \ heap" where + "set r x = refs_update + (\h. h(TYPEREP('a) := ((h (TYPEREP('a))) (addr_of_ref r := to_nat x))))" -definition ref :: "'a \ heap \ 'a\heap ref \ heap" where - "ref x h = (let +definition alloc :: "'a \ heap \ 'a\heap ref \ heap" where + "alloc x h = (let l = lim h; - r = Ref l; - h'' = set_ref r x (h\lim := l + 1\) - in (r, h''))" + r = Ref l + in (r, set r x (h\lim := l + 1\)))" -definition noteq_refs :: "('a\heap) ref \ ('b\heap) ref \ bool" (infix "=!=" 70) where +definition noteq :: "'a\heap ref \ 'b\heap ref \ bool" (infix "=!=" 70) where "r =!= s \ TYPEREP('a) \ TYPEREP('b) \ addr_of_ref r \ addr_of_ref s" -lemma noteq_refs_sym: "r =!= s \ s =!= r" - and unequal_refs [simp]: "r \ r' \ r =!= r'" -- "same types!" - unfolding noteq_refs_def by auto +lemma noteq_sym: "r =!= s \ s =!= r" + and unequal [simp]: "r \ r' \ r =!= r'" -- "same types!" + by (auto simp add: noteq_def) -lemma noteq_refs_irrefl: "r =!= r \ False" - unfolding noteq_refs_def by auto +lemma noteq_irrefl: "r =!= r \ False" + by (auto simp add: noteq_def) -lemma present_new_ref: "ref_present r h \ r =!= fst (ref v h)" - by (simp add: ref_present_def ref_def Let_def noteq_refs_def) +lemma present_alloc_neq: "present h r \ r =!= fst (alloc v h)" + by (simp add: present_def alloc_def noteq_def Let_def) -lemma next_ref_fresh [simp]: - assumes "(r, h') = ref x h" - shows "\ ref_present r h" - using assms by (cases h) (auto simp add: ref_def ref_present_def Let_def) +lemma next_fresh [simp]: + assumes "(r, h') = alloc x h" + shows "\ present h r" + using assms by (cases h) (auto simp add: alloc_def present_def Let_def) -lemma next_ref_present [simp]: - assumes "(r, h') = ref x h" - shows "ref_present r h'" - using assms by (cases h) (auto simp add: ref_def set_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 next_present [simp]: + assumes "(r, h') = alloc x h" + shows "present h' r" + using assms by (cases h) (auto simp add: alloc_def set_def present_def Let_def) -lemma 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) +lemma get_set_eq [simp]: + "get (set r x h) r = x" + by (simp add: get_def set_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 get_set_neq [simp]: + "r =!= s \ get (set s x h) r = get h r" + by (simp add: noteq_def get_def set_def) -lemma ref_set_get: "set_ref r (get_ref r h) h = h" -apply (simp add: set_ref_def get_ref_def) -oops +lemma set_same [simp]: + "set r x (set r y h) = set r x h" + by (simp add: set_def) -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 set_set_swap: + "r =!= r' \ set r x (set r' x' h) = set r' x' (set r x h)" + by (simp add: noteq_def set_def expand_fun_eq) -lemma 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 alloc_set: + "fst (alloc x (set r x' h)) = fst (alloc x h)" + by (simp add: alloc_def set_def Let_def) -lemma ref_new_set: "fst (ref v (set_ref r v' h)) = fst (ref v h)" - by (simp add: ref_def set_ref_def Let_def) +lemma get_alloc [simp]: + "get (snd (alloc x h)) (fst (alloc x' h)) = x" + by (simp add: alloc_def Let_def) -lemma 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 set_alloc [simp]: + "set (fst (alloc v h)) v' (snd (alloc v h)) = snd (alloc v' h)" + by (simp add: alloc_def Let_def) -lemma 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 noteq_refs_def) +lemma get_alloc_neq: "r =!= fst (alloc v h) \ + get (snd (alloc v h)) r = get h r" + by (simp add: get_def set_def alloc_def Let_def noteq_def) -lemma lim_set_ref [simp]: - "lim (set_ref r v h) = lim h" - by (simp add: set_ref_def) +lemma lim_set [simp]: + "lim (set r v h) = lim h" + by (simp add: set_def) -lemma ref_present_new_ref [simp]: - "ref_present r h \ ref_present r (snd (ref v h))" - by (simp add: ref_present_def ref_def Let_def) +lemma present_alloc [simp]: + "present h r \ present (snd (alloc v h)) r" + by (simp add: present_def alloc_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 present_set [simp]: + "present (set r v h) = present h" + by (simp add: present_def expand_fun_eq) -lemma noteq_refsI: "\ ref_present r h; \ref_present r' h \ \ r =!= r'" - unfolding noteq_refs_def ref_present_def - by auto +lemma noteq_I: + "present h r \ \ present h r' \ r =!= r'" + by (auto simp add: noteq_def present_def) subsection {* Primitives *} -definition - new :: "'a\heap \ 'a ref Heap" where - [code del]: "new v = Heap_Monad.heap (Ref.ref v)" +definition ref :: "'a\heap \ 'a ref Heap" where + [code del]: "ref v = Heap_Monad.heap (alloc v)" -definition - lookup :: "'a\heap ref \ 'a Heap" ("!_" 61) where - [code del]: "lookup r = Heap_Monad.heap (\h. (get_ref r h, h))" +definition lookup :: "'a\heap ref \ 'a Heap" ("!_" 61) where + [code del]: "lookup r = Heap_Monad.heap (\h. (get h r, h))" -definition - update :: "'a ref \ ('a\heap) \ unit Heap" ("_ := _" 62) where - [code del]: "update r e = Heap_Monad.heap (\h. ((), set_ref r e h))" +definition update :: "'a ref \ 'a\heap \ unit Heap" ("_ := _" 62) where + [code del]: "update r v = Heap_Monad.heap (\h. ((), set r v h))" subsection {* Derivates *} -definition - change :: "('a\heap \ 'a) \ 'a ref \ 'a Heap" -where - "change f r = (do x \ ! r; - let y = f x; - r := y; - return y - done)" - -hide_const (open) new lookup update change +definition change :: "('a\heap \ 'a) \ 'a ref \ 'a Heap" where + "change f r = (do + x \ ! r; + let y = f x; + r := y; + return y + done)" subsection {* Properties *} lemma lookup_chain: "(!r \ f) = f" - by (cases f) - (auto simp add: Let_def bindM_def lookup_def expand_fun_eq) + by (rule Heap_eqI) (simp add: lookup_def) lemma update_change [code]: - "r := e = Ref.change (\_. e) r \ return ()" - by (auto simp add: change_def lookup_chain) + "r := e = change (\_. e) r \ return ()" + by (rule Heap_eqI) (simp add: change_def lookup_chain) text {* Non-interaction between imperative array and imperative references *} -lemma get_array_set_ref [simp]: "get_array a (set_ref r v h) = get_array a h" - by (simp add: get_array_def set_ref_def) +lemma get_array_set [simp]: + "get_array a (set r v h) = get_array a h" + by (simp add: get_array_def set_def) -lemma nth_set_ref [simp]: "get_array a (set_ref r v h) ! i = get_array a h ! i" +lemma nth_set [simp]: + "get_array a (set r v h) ! i = get_array a h ! i" by simp -lemma get_ref_upd [simp]: "get_ref r (Array.change a i v h) = get_ref r h" - by (simp add: get_ref_def set_array_def Array.change_def) +lemma get_change [simp]: + "get (Array.change a i v h) r = get h r" + by (simp add: get_def Array.change_def set_array_def) -lemma new_ref_upd: "fst (ref v (Array.change a i v' h)) = fst (ref v h)" - by (simp add: set_array_def get_array_def Let_def ref_new_set Array.change_def ref_def) +lemma alloc_change: + "fst (alloc v (Array.change a i v' h)) = fst (alloc v h)" + by (simp add: Array.change_def get_array_def set_array_def alloc_def Let_def) -lemma upd_set_ref_swap: "Array.change a i v (set_ref r v' h) = set_ref r v' (Array.change a i v h)" - by (simp add: set_ref_def Array.change_def get_array_def set_array_def) +lemma change_set_swap: + "Array.change a i v (set r v' h) = set r v' (Array.change a i v h)" + by (simp add: Array.change_def get_array_def set_array_def set_def) -lemma length_new_ref[simp]: - "length a (snd (ref v h)) = length a h" - by (simp add: get_array_def set_ref_def length_def ref_def Let_def) +lemma length_alloc [simp]: + "Array.length a (snd (alloc v h)) = Array.length a h" + by (simp add: Array.length_def get_array_def alloc_def set_def Let_def) -lemma get_array_new_ref [simp]: - "get_array a (snd (ref v h)) = get_array a h" - by (simp add: ref_def set_ref_def get_array_def Let_def) +lemma get_array_alloc [simp]: + "get_array a (snd (alloc v h)) = get_array a h" + by (simp add: get_array_def alloc_def set_def Let_def) -lemma ref_present_upd [simp]: - "ref_present r (Array.change a i v h) = ref_present r h" - by (simp add: Array.change_def ref_present_def set_array_def get_array_def) +lemma present_change [simp]: + "present (Array.change a i v h) = present h" + by (simp add: Array.change_def set_array_def expand_fun_eq present_def) -lemma array_present_set_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_set [simp]: + "array_present a (set r v h) = array_present a h" + by (simp add: array_present_def set_def) -lemma array_present_new_ref [simp]: - "array_present a h \ array_present a (snd (ref v h))" - by (simp add: array_present_def ref_def Let_def) +lemma array_present_alloc [simp]: + "array_present a h \ array_present a (snd (alloc v h))" + by (simp add: array_present_def alloc_def Let_def) -lemma 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 set_array_set_swap: + "set_array a xs (set r x' h) = set r x' (set_array a xs h)" + by (simp add: set_array_def set_def) + +hide_const (open) present get set alloc lookup update change subsection {* Code generator setup *} @@ -203,7 +193,7 @@ code_type ref (SML "_/ Unsynchronized.ref") code_const Ref (SML "raise/ (Fail/ \"bare Ref\")") -code_const Ref.new (SML "(fn/ ()/ =>/ Unsynchronized.ref/ _)") +code_const ref (SML "(fn/ ()/ =>/ Unsynchronized.ref/ _)") code_const Ref.lookup (SML "(fn/ ()/ =>/ !/ _)") code_const Ref.update (SML "(fn/ ()/ =>/ _/ :=/ _)") @@ -214,7 +204,7 @@ code_type ref (OCaml "_/ ref") code_const Ref (OCaml "failwith/ \"bare Ref\")") -code_const Ref.new (OCaml "(fn/ ()/ =>/ ref/ _)") +code_const ref (OCaml "(fn/ ()/ =>/ ref/ _)") code_const Ref.lookup (OCaml "(fn/ ()/ =>/ !/ _)") code_const Ref.update (OCaml "(fn/ ()/ =>/ _/ :=/ _)") @@ -225,7 +215,7 @@ code_type ref (Haskell "Heap.STRef/ Heap.RealWorld/ _") code_const Ref (Haskell "error/ \"bare Ref\"") -code_const Ref.new (Haskell "Heap.newSTRef") +code_const ref (Haskell "Heap.newSTRef") code_const Ref.lookup (Haskell "Heap.readSTRef") code_const Ref.update (Haskell "Heap.writeSTRef") diff -r 6607ccf77946 -r 6d28a2aea936 src/HOL/Imperative_HOL/Relational.thy --- a/src/HOL/Imperative_HOL/Relational.thy Tue Jul 06 09:21:13 2010 +0200 +++ b/src/HOL/Imperative_HOL/Relational.thy Tue Jul 06 09:21:15 2010 +0200 @@ -311,42 +311,42 @@ 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'" +lemma crel_ref: + assumes "crel (ref v) h h' x" + obtains "Ref.get h' x = v" + and "\ Ref.present h x" + and "Ref.present h' x" + and "\y. Ref.present h y \ Ref.get h y = Ref.get h' y" (* and "lim h' = Suc (lim h)" *) - and "\y. ref_present y h \ ref_present y h'" + and "\y. Ref.present h y \ Ref.present h' y" using assms - unfolding Ref.new_def - apply (elim crel_heap) unfolding Ref.ref_def + apply (elim crel_heap) + unfolding Ref.alloc_def apply (simp add: Let_def) - unfolding ref_present_def + unfolding Ref.present_def apply auto - unfolding get_ref_def set_ref_def + unfolding Ref.get_def Ref.set_def apply auto done lemma crel_lookup: assumes "crel (!r') h h' r" - obtains "h = h'" "r = get_ref r' h" + obtains "h = h'" "r = Ref.get h r'" using assms unfolding Ref.lookup_def by (auto elim: crel_heap) lemma crel_update: assumes "crel (r' := v) h h' r" - obtains "h' = set_ref r' v h" "r = ()" + obtains "h' = Ref.set r' v h" "r = ()" using assms unfolding Ref.update_def by (auto elim: crel_heap) lemma crel_change: assumes "crel (Ref.change f r') h h' r" - obtains "h' = set_ref r' (f (get_ref r' h)) h" "r = f (get_ref r' h)" + obtains "h' = Ref.set r' (f (Ref.get h r')) h" "r = f (Ref.get h r')" using assms unfolding Ref.change_def Let_def by (auto elim!: crelE crel_lookup crel_update crel_return) @@ -467,15 +467,15 @@ subsubsection {* Introduction rules for reference commands *} lemma crel_lookupI: - shows "crel (!r) h h (get_ref r h)" + shows "crel (!r) h h (Ref.get h r)" unfolding lookup_def by (auto intro!: crel_heapI') lemma crel_updateI: - shows "crel (r := v) h (set_ref r v h) ()" + shows "crel (r := v) h (Ref.set r v h) ()" unfolding update_def by (auto intro!: crel_heapI') lemma crel_changeI: - shows "crel (Ref.change f r) h (set_ref r (f (get_ref r h)) h) (f (get_ref r h))" + shows "crel (Ref.change f r) h (Ref.set r (f (Ref.get h r)) h) (f (Ref.get h r))" unfolding change_def Let_def by (auto intro!: crelI crel_returnI crel_lookupI crel_updateI) subsubsection {* Introduction rules for the assert command *} @@ -667,8 +667,8 @@ subsection {* Introduction rules for the reference commands *} lemma noError_Ref_new: - shows "noError (Ref.new v) h" -unfolding Ref.new_def by (intro noError_heap) + shows "noError (ref v) h" + unfolding Ref.ref_def by (intro noError_heap) lemma noError_lookup: shows "noError (!r) h" @@ -696,7 +696,7 @@ lemmas crel_elim_all = crelE crelE' crel_return crel_raise crel_if crel_option_case crel_length crel_new_weak crel_nth crel_upd crel_of_list_weak crel_map_entry crel_swap crel_make_weak crel_freeze crel_map_weak - crel_Ref_new crel_lookup crel_update crel_change + crel_ref crel_lookup crel_update crel_change crel_assert lemmas crel_intro_all = diff -r 6607ccf77946 -r 6d28a2aea936 src/HOL/Imperative_HOL/ex/Linked_Lists.thy --- a/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Tue Jul 06 09:21:13 2010 +0200 +++ b/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Tue Jul 06 09:21:15 2010 +0200 @@ -13,7 +13,7 @@ setup {* Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \ 'a\type ref"}) *} datatype 'a node = Empty | Node 'a "('a node) ref" -fun +primrec node_encode :: "'a\countable node \ nat" where "node_encode Empty = 0" @@ -28,11 +28,11 @@ instance node :: (heap) heap .. -fun make_llist :: "'a\heap list \ 'a node Heap" +primrec make_llist :: "'a\heap list \ 'a node Heap" where [simp del]: "make_llist [] = return Empty" | "make_llist (x#xs) = do tl \ make_llist xs; - next \ Ref.new tl; + next \ ref tl; return (Node x next) done" @@ -63,24 +63,24 @@ subsection {* Definition of list_of, list_of', refs_of and refs_of' *} -fun list_of :: "heap \ ('a::heap) node \ 'a list \ bool" +primrec list_of :: "heap \ ('a::heap) node \ 'a list \ bool" where "list_of h r [] = (r = Empty)" -| "list_of h r (a#as) = (case r of Empty \ False | Node b bs \ (a = b \ list_of h (get_ref bs h) as))" +| "list_of h r (a#as) = (case r of Empty \ False | Node b bs \ (a = b \ list_of h (Ref.get h bs) as))" definition list_of' :: "heap \ ('a::heap) node ref \ 'a list \ bool" where - "list_of' h r xs = list_of h (get_ref r h) xs" + "list_of' h r xs = list_of h (Ref.get h r) xs" -fun refs_of :: "heap \ ('a::heap) node \ 'a node ref list \ bool" +primrec refs_of :: "heap \ ('a::heap) node \ 'a node ref list \ bool" where "refs_of h r [] = (r = Empty)" -| "refs_of h r (x#xs) = (case r of Empty \ False | Node b bs \ (x = bs) \ refs_of h (get_ref bs h) xs)" +| "refs_of h r (x#xs) = (case r of Empty \ False | Node b bs \ (x = bs) \ refs_of h (Ref.get h bs) xs)" -fun refs_of' :: "heap \ ('a::heap) node ref \ 'a node ref list \ bool" +primrec refs_of' :: "heap \ ('a::heap) node ref \ 'a node ref list \ bool" where "refs_of' h r [] = False" -| "refs_of' h r (x#xs) = ((x = r) \ refs_of h (get_ref x h) xs)" +| "refs_of' h r (x#xs) = ((x = r) \ refs_of h (Ref.get h x) xs)" subsection {* Properties of these definitions *} @@ -88,35 +88,35 @@ lemma list_of_Empty[simp]: "list_of h Empty xs = (xs = [])" by (cases xs, auto) -lemma list_of_Node[simp]: "list_of h (Node x ps) xs = (\xs'. (xs = x # xs') \ list_of h (get_ref ps h) xs')" +lemma list_of_Node[simp]: "list_of h (Node x ps) xs = (\xs'. (xs = x # xs') \ list_of h (Ref.get h ps) xs')" by (cases xs, auto) -lemma list_of'_Empty[simp]: "get_ref q h = Empty \ list_of' h q xs = (xs = [])" +lemma list_of'_Empty[simp]: "Ref.get h q = Empty \ list_of' h q xs = (xs = [])" unfolding list_of'_def by simp -lemma list_of'_Node[simp]: "get_ref q h = Node x ps \ list_of' h q xs = (\xs'. (xs = x # xs') \ list_of' h ps xs')" +lemma list_of'_Node[simp]: "Ref.get h q = Node x ps \ list_of' h q xs = (\xs'. (xs = x # xs') \ list_of' h ps xs')" unfolding list_of'_def by simp -lemma list_of'_Nil: "list_of' h q [] \ get_ref q h = Empty" +lemma list_of'_Nil: "list_of' h q [] \ Ref.get h q = Empty" unfolding list_of'_def by simp lemma list_of'_Cons: assumes "list_of' h q (x#xs)" -obtains n where "get_ref q h = Node x n" and "list_of' h n xs" +obtains n where "Ref.get h q = Node x n" and "list_of' h n xs" using assms unfolding list_of'_def by (auto split: node.split_asm) lemma refs_of_Empty[simp] : "refs_of h Empty xs = (xs = [])" by (cases xs, auto) -lemma refs_of_Node[simp]: "refs_of h (Node x ps) xs = (\prs. xs = ps # prs \ refs_of h (get_ref ps h) prs)" +lemma refs_of_Node[simp]: "refs_of h (Node x ps) xs = (\prs. xs = ps # prs \ refs_of h (Ref.get h ps) prs)" by (cases xs, auto) -lemma refs_of'_def': "refs_of' h p ps = (\prs. (ps = (p # prs)) \ refs_of h (get_ref p h) prs)" +lemma refs_of'_def': "refs_of' h p ps = (\prs. (ps = (p # prs)) \ refs_of h (Ref.get h p) prs)" by (cases ps, auto) lemma refs_of'_Node: assumes "refs_of' h p xs" - assumes "get_ref p h = Node x pn" + assumes "Ref.get h p = Node x pn" obtains pnrs where "xs = p # pnrs" and "refs_of' h pn pnrs" using assms @@ -166,7 +166,7 @@ assumes "list_of' h r xs" shows "\rs. refs_of' h r rs" proof - - from assms obtain rs' where "refs_of h (get_ref r h) rs'" + from assms obtain rs' where "refs_of h (Ref.get h r) rs'" unfolding list_of'_def by (rule list_of_refs_of) thus ?thesis unfolding refs_of'_def' by auto qed @@ -238,7 +238,7 @@ done lemma refs_of_next: -assumes "refs_of h (get_ref p h) rs" +assumes "refs_of h (Ref.get h p) rs" shows "p \ set rs" proof (rule ccontr) assume a: "\ (p \ set rs)" @@ -264,7 +264,7 @@ subsection {* Interaction of these predicates with our heap transitions *} -lemma list_of_set_ref: "refs_of h q rs \ p \ set rs \ list_of (set_ref p v h) q as = list_of h q as" +lemma list_of_set_ref: "refs_of h q rs \ p \ set rs \ list_of (Ref.set p v h) q as = list_of h q as" using assms proof (induct as arbitrary: q rs) case Nil thus ?case by simp @@ -275,15 +275,15 @@ case Empty thus ?thesis by auto next case (Node a ref) - from Cons(2) Node obtain rs' where 1: "refs_of h (get_ref ref h) rs'" and rs_rs': "rs = ref # rs'" by auto + from Cons(2) Node obtain rs' where 1: "refs_of h (Ref.get h ref) rs'" and rs_rs': "rs = ref # rs'" by auto from Cons(3) rs_rs' have "ref \ p" by fastsimp - hence ref_eq: "get_ref ref (set_ref p v h) = (get_ref ref h)" by (auto simp add: ref_get_set_neq) + hence ref_eq: "Ref.get (Ref.set p v h) ref = (Ref.get h ref)" by (auto simp add: Ref.get_set_neq) from rs_rs' Cons(3) have 2: "p \ set rs'" by simp from Cons.hyps[OF 1 2] Node ref_eq show ?thesis by simp qed qed -lemma refs_of_set_ref: "refs_of h q rs \ p \ set rs \ refs_of (set_ref p v h) q as = refs_of h q as" +lemma refs_of_set_ref: "refs_of h q rs \ p \ set rs \ refs_of (Ref.set p v h) q as = refs_of h q as" proof (induct as arbitrary: q rs) case Nil thus ?case by simp next @@ -293,15 +293,15 @@ case Empty thus ?thesis by auto next case (Node a ref) - from Cons(2) Node obtain rs' where 1: "refs_of h (get_ref ref h) rs'" and rs_rs': "rs = ref # rs'" by auto + from Cons(2) Node obtain rs' where 1: "refs_of h (Ref.get h ref) rs'" and rs_rs': "rs = ref # rs'" by auto from Cons(3) rs_rs' have "ref \ p" by fastsimp - hence ref_eq: "get_ref ref (set_ref p v h) = (get_ref ref h)" by (auto simp add: ref_get_set_neq) + hence ref_eq: "Ref.get (Ref.set p v h) ref = (Ref.get h ref)" by (auto simp add: Ref.get_set_neq) from rs_rs' Cons(3) have 2: "p \ set rs'" by simp from Cons.hyps[OF 1 2] Node ref_eq show ?thesis by auto qed qed -lemma refs_of_set_ref2: "refs_of (set_ref p v h) q rs \ p \ set rs \ refs_of (set_ref p v h) q rs = refs_of h q rs" +lemma refs_of_set_ref2: "refs_of (Ref.set p v h) q rs \ p \ set rs \ refs_of (Ref.set p v h) q rs = refs_of h q rs" proof (induct rs arbitrary: q) case Nil thus ?case by simp next @@ -311,9 +311,9 @@ case Empty thus ?thesis by auto next case (Node a ref) - from Cons(2) Node have 1:"refs_of (set_ref p v h) (get_ref ref (set_ref p v h)) xs" and x_ref: "x = ref" by auto + from Cons(2) Node have 1:"refs_of (Ref.set p v h) (Ref.get (Ref.set p v h) ref) xs" and x_ref: "x = ref" by auto from Cons(3) this have "ref \ p" by fastsimp - hence ref_eq: "get_ref ref (set_ref p v h) = (get_ref ref h)" by (auto simp add: ref_get_set_neq) + hence ref_eq: "Ref.get (Ref.set p v h) ref = (Ref.get h ref)" by (auto simp add: Ref.get_set_neq) from Cons(3) have 2: "p \ set xs" by simp with Cons.hyps 1 2 Node ref_eq show ?thesis by simp @@ -323,7 +323,7 @@ lemma list_of'_set_ref: assumes "refs_of' h q rs" assumes "p \ set rs" - shows "list_of' (set_ref p v h) q as = list_of' h q as" + shows "list_of' (Ref.set p v h) q as = list_of' h q as" proof - from assms have "q \ p" by (auto simp only: dest!: refs_of'E) with assms show ?thesis @@ -333,18 +333,18 @@ lemma list_of'_set_next_ref_Node[simp]: assumes "list_of' h r xs" - assumes "get_ref p h = Node x r'" + assumes "Ref.get h p = Node x r'" assumes "refs_of' h r rs" assumes "p \ set rs" - shows "list_of' (set_ref p (Node x r) h) p (x#xs) = list_of' h r xs" + shows "list_of' (Ref.set p (Node x r) h) p (x#xs) = list_of' h r xs" using assms unfolding list_of'_def refs_of'_def' -by (auto simp add: list_of_set_ref noteq_refs_sym) +by (auto simp add: list_of_set_ref Ref.noteq_sym) lemma refs_of'_set_ref: assumes "refs_of' h q rs" assumes "p \ set rs" - shows "refs_of' (set_ref p v h) q as = refs_of' h q as" + shows "refs_of' (Ref.set p v h) q as = refs_of' h q as" using assms proof - from assms have "q \ p" by (auto simp only: dest!: refs_of'E) @@ -354,9 +354,9 @@ qed lemma refs_of'_set_ref2: - assumes "refs_of' (set_ref p v h) q rs" + assumes "refs_of' (Ref.set p v h) q rs" assumes "p \ set rs" - shows "refs_of' (set_ref p v h) q as = refs_of' h q as" + shows "refs_of' (Ref.set p v h) q as = refs_of' h q as" using assms proof - from assms have "q \ p" by (auto simp only: dest!: refs_of'E) @@ -364,7 +364,7 @@ unfolding refs_of'_def' apply auto apply (subgoal_tac "prs = prsa") - apply (insert refs_of_set_ref2[of p v h "get_ref q h"]) + apply (insert refs_of_set_ref2[of p v h "Ref.get h q"]) apply (erule_tac x="prs" in meta_allE) apply auto apply (auto dest: refs_of_is_fun) @@ -372,15 +372,15 @@ qed lemma refs_of'_set_next_ref: -assumes "get_ref p h1 = Node x pn" -assumes "refs_of' (set_ref p (Node x r1) h1) p rs" +assumes "Ref.get h1 p = Node x pn" +assumes "refs_of' (Ref.set p (Node x r1) h1) p rs" obtains r1s where "rs = (p#r1s)" and "refs_of' h1 r1 r1s" using assms proof - from assms refs_of'_distinct[OF assms(2)] have "\ r1s. rs = (p # r1s) \ refs_of' h1 r1 r1s" apply - unfolding refs_of'_def'[of _ p] - apply (auto, frule refs_of_set_ref2) by (auto dest: noteq_refs_sym) + apply (auto, frule refs_of_set_ref2) by (auto dest: Ref.noteq_sym) with prems show thesis by auto qed @@ -388,7 +388,7 @@ lemma refs_of_invariant: assumes "refs_of h (r::('a::heap) node) xs" - assumes "\refs. refs_of h r refs \ (\ref \ set refs. ref_present ref h \ ref_present ref h' \ get_ref ref h = get_ref ref h')" + assumes "\refs. refs_of h r refs \ (\ref \ set refs. Ref.present h ref \ Ref.present h' ref \ Ref.get h ref = Ref.get h' ref)" shows "refs_of h' r xs" using assms proof (induct xs arbitrary: r) @@ -396,28 +396,28 @@ next case (Cons x xs') from Cons(2) obtain v where Node: "r = Node v x" by (cases r, auto) - from Cons(2) Node have refs_of_next: "refs_of h (get_ref x h) xs'" by simp - from Cons(2-3) Node have ref_eq: "get_ref x h = get_ref x h'" by auto - from ref_eq refs_of_next have 1: "refs_of h (get_ref x h') xs'" by simp - from Cons(2) Cons(3) have "\ref \ set xs'. ref_present ref h \ ref_present ref h' \ get_ref ref h = get_ref ref h'" + from Cons(2) Node have refs_of_next: "refs_of h (Ref.get h x) xs'" by simp + from Cons(2-3) Node have ref_eq: "Ref.get h x = Ref.get h' x" by auto + from ref_eq refs_of_next have 1: "refs_of h (Ref.get h' x) xs'" by simp + from Cons(2) Cons(3) have "\ref \ set xs'. Ref.present h ref \ Ref.present h' ref \ Ref.get h ref = Ref.get h' ref" by fastsimp - with Cons(3) 1 have 2: "\refs. refs_of h (get_ref x h') refs \ (\ref \ set refs. ref_present ref h \ ref_present ref h' \ get_ref ref h = get_ref ref h')" + with Cons(3) 1 have 2: "\refs. refs_of h (Ref.get h' x) refs \ (\ref \ set refs. Ref.present h ref \ Ref.present h' ref \ Ref.get h ref = Ref.get h' ref)" by (fastsimp dest: refs_of_is_fun) - from Cons.hyps[OF 1 2] have "refs_of h' (get_ref x h') xs'" . + from Cons.hyps[OF 1 2] have "refs_of h' (Ref.get h' x) xs'" . with Node show ?case by simp qed lemma refs_of'_invariant: assumes "refs_of' h r xs" - assumes "\refs. refs_of' h r refs \ (\ref \ set refs. ref_present ref h \ ref_present ref h' \ get_ref ref h = get_ref ref h')" + assumes "\refs. refs_of' h r refs \ (\ref \ set refs. Ref.present h ref \ Ref.present h' ref \ Ref.get h ref = Ref.get h' ref)" shows "refs_of' h' r xs" using assms proof - - from assms obtain prs where refs:"refs_of h (get_ref r h) prs" and xs_def: "xs = r # prs" + from assms obtain prs where refs:"refs_of h (Ref.get h r) prs" and xs_def: "xs = r # prs" unfolding refs_of'_def' by auto - from xs_def assms have x_eq: "get_ref r h = get_ref r h'" by fastsimp - from refs assms xs_def have 2: "\refs. refs_of h (get_ref r h) refs \ - (\ref\set refs. ref_present ref h \ ref_present ref h' \ get_ref ref h = get_ref ref h')" + from xs_def assms have x_eq: "Ref.get h r = Ref.get h' r" by fastsimp + from refs assms xs_def have 2: "\refs. refs_of h (Ref.get h r) refs \ + (\ref\set refs. Ref.present h ref \ Ref.present h' ref \ Ref.get h ref = Ref.get h' ref)" by (fastsimp dest: refs_of_is_fun) from refs_of_invariant [OF refs 2] xs_def x_eq show ?thesis unfolding refs_of'_def' by auto @@ -425,7 +425,7 @@ lemma list_of_invariant: assumes "list_of h (r::('a::heap) node) xs" - assumes "\refs. refs_of h r refs \ (\ref \ set refs. ref_present ref h \ ref_present ref h' \ get_ref ref h = get_ref ref h')" + assumes "\refs. refs_of h r refs \ (\ref \ set refs. Ref.present h ref \ Ref.present h' ref \ Ref.get h ref = Ref.get h' ref)" shows "list_of h' r xs" using assms proof (induct xs arbitrary: r) @@ -437,16 +437,16 @@ by (cases r, auto) from Cons(2) obtain rs where rs_def: "refs_of h r rs" by (rule list_of_refs_of) from Node rs_def obtain rss where refs_of: "refs_of h r (ref#rss)" and rss_def: "rs = ref#rss" by auto - from Cons(3) Node refs_of have ref_eq: "get_ref ref h = get_ref ref h'" + from Cons(3) Node refs_of have ref_eq: "Ref.get h ref = Ref.get h' ref" by auto - from Cons(2) ref_eq Node have 1: "list_of h (get_ref ref h') xs'" by simp - from refs_of Node ref_eq have refs_of_ref: "refs_of h (get_ref ref h') rss" by simp - from Cons(3) rs_def have rs_heap_eq: "\ref\set rs. ref_present ref h \ ref_present ref h' \ get_ref ref h = get_ref ref h'" by simp - from refs_of_ref rs_heap_eq rss_def have 2: "\refs. refs_of h (get_ref ref h') refs \ - (\ref\set refs. ref_present ref h \ ref_present ref h' \ get_ref ref h = get_ref ref h')" + from Cons(2) ref_eq Node have 1: "list_of h (Ref.get h' ref) xs'" by simp + from refs_of Node ref_eq have refs_of_ref: "refs_of h (Ref.get h' ref) rss" by simp + from Cons(3) rs_def have rs_heap_eq: "\ref\set rs. Ref.present h ref \ Ref.present h' ref \ Ref.get h ref = Ref.get h' ref" by simp + from refs_of_ref rs_heap_eq rss_def have 2: "\refs. refs_of h (Ref.get h' ref) refs \ + (\ref\set refs. Ref.present h ref \ Ref.present h' ref \ Ref.get h ref = Ref.get h' ref)" by (auto dest: refs_of_is_fun) from Cons(1)[OF 1 2] - have "list_of h' (get_ref ref h') xs'" . + have "list_of h' (Ref.get h' ref) xs'" . with Node show ?case unfolding list_of'_def by simp @@ -454,29 +454,29 @@ lemma make_llist: assumes "crel (make_llist xs) h h' r" -shows "list_of h' r xs \ (\rs. refs_of h' r rs \ (\ref \ (set rs). ref_present ref h'))" +shows "list_of h' r xs \ (\rs. refs_of h' r rs \ (\ref \ (set rs). Ref.present h' ref))" using assms proof (induct xs arbitrary: h h' r) case Nil thus ?case by (auto elim: crel_return simp add: make_llist.simps) next case (Cons x xs') from Cons.prems obtain h1 r1 r' where make_llist: "crel (make_llist xs') h h1 r1" - and crel_refnew:"crel (Ref.new r1) h1 h' r'" and Node: "r = Node x r'" + and crel_refnew:"crel (ref r1) h1 h' r'" and Node: "r = Node x r'" unfolding make_llist.simps by (auto elim!: crelE crel_return) from Cons.hyps[OF make_llist] have list_of_h1: "list_of h1 r1 xs'" .. from Cons.hyps[OF make_llist] obtain rs' where rs'_def: "refs_of h1 r1 rs'" by (auto intro: list_of_refs_of) - from Cons.hyps[OF make_llist] rs'_def have refs_present: "\ref\set rs'. ref_present ref h1" by simp + from Cons.hyps[OF make_llist] rs'_def have refs_present: "\ref\set rs'. Ref.present h1 ref" by simp from crel_refnew rs'_def refs_present have refs_unchanged: "\refs. refs_of h1 r1 refs \ - (\ref\set refs. ref_present ref h1 \ ref_present ref h' \ get_ref ref h1 = get_ref ref h')" - by (auto elim!: crel_Ref_new dest: refs_of_is_fun) + (\ref\set refs. Ref.present h1 ref \ Ref.present h' ref \ Ref.get h1 ref = Ref.get h' ref)" + by (auto elim!: crel_ref dest: refs_of_is_fun) with list_of_invariant[OF list_of_h1 refs_unchanged] Node crel_refnew have fstgoal: "list_of h' r (x # xs')" unfolding list_of.simps - by (auto elim!: crel_Ref_new) - from refs_unchanged rs'_def have refs_still_present: "\ref\set rs'. ref_present ref h'" by auto + by (auto elim!: crel_ref) + from refs_unchanged rs'_def have refs_still_present: "\ref\set rs'. Ref.present h' ref" by auto from refs_of_invariant[OF rs'_def refs_unchanged] refs_unchanged Node crel_refnew refs_still_present - have sndgoal: "\rs. refs_of h' r rs \ (\ref\set rs. ref_present ref h')" - by (fastsimp elim!: crel_Ref_new dest: refs_of_is_fun) + have sndgoal: "\rs. refs_of h' r rs \ (\ref\set rs. Ref.present h' ref)" + by (fastsimp elim!: crel_ref dest: refs_of_is_fun) from fstgoal sndgoal show ?case .. qed @@ -533,10 +533,10 @@ thm arg_cong2 by (auto simp add: expand_fun_eq intro: arg_cong2[where f = "op \="] split: node.split) -fun rev :: "('a:: heap) node \ 'a node Heap" +primrec rev :: "('a:: heap) node \ 'a node Heap" where "rev Empty = return Empty" -| "rev (Node x n) = (do q \ Ref.new Empty; p \ Ref.new (Node x n); v \ rev' (q, p); !v done)" +| "rev (Node x n) = (do q \ ref Empty; p \ ref (Node x n); v \ rev' (q, p); !v done)" subsection {* Correctness Proof *} @@ -556,17 +556,17 @@ case (Cons x xs) (*"LinkedList.list_of h' (get_ref v h') (List.rev xs @ x # qsa)"*) from Cons(4) obtain ref where - p_is_Node: "get_ref p h = Node x ref" + p_is_Node: "Ref.get h p = Node x ref" (*and "ref_present ref h"*) and list_of'_ref: "list_of' h ref xs" - unfolding list_of'_def by (cases "get_ref p h", auto) - from p_is_Node Cons(2) have crel_rev': "crel (rev' (p, ref)) (set_ref p (Node x q) h) h' v" + unfolding list_of'_def by (cases "Ref.get h p", auto) + from p_is_Node Cons(2) have crel_rev': "crel (rev' (p, ref)) (Ref.set p (Node x q) h) h' v" by (auto simp add: rev'_simps [of q p] elim!: crelE crel_lookup crel_update) from Cons(3) obtain qrs where qrs_def: "refs_of' h q qrs" by (elim list_of'_refs_of') from Cons(4) obtain prs where prs_def: "refs_of' h p prs" by (elim list_of'_refs_of') from qrs_def prs_def Cons(5) have distinct_pointers: "set qrs \ set prs = {}" by fastsimp from qrs_def prs_def distinct_pointers refs_of'E have p_notin_qrs: "p \ set qrs" by fastsimp - from Cons(3) qrs_def this have 1: "list_of' (set_ref p (Node x q) h) p (x#qs)" + from Cons(3) qrs_def this have 1: "list_of' (Ref.set p (Node x q) h) p (x#qs)" unfolding list_of'_def apply (simp) unfolding list_of'_def[symmetric] @@ -575,16 +575,16 @@ unfolding refs_of'_def' by auto from prs_refs prs_def have p_not_in_refs: "p \ set refs" by (fastsimp dest!: refs_of'_distinct) - with refs_def p_is_Node list_of'_ref have 2: "list_of' (set_ref p (Node x q) h) ref xs" + with refs_def p_is_Node list_of'_ref have 2: "list_of' (Ref.set p (Node x q) h) ref xs" by (auto simp add: list_of'_set_ref) - from p_notin_qrs qrs_def have refs_of1: "refs_of' (set_ref p (Node x q) h) p (p#qrs)" + from p_notin_qrs qrs_def have refs_of1: "refs_of' (Ref.set p (Node x q) h) p (p#qrs)" unfolding refs_of'_def' apply (simp) unfolding refs_of'_def'[symmetric] by (simp add: refs_of'_set_ref) - from p_not_in_refs p_is_Node refs_def have refs_of2: "refs_of' (set_ref p (Node x q) h) ref refs" + from p_not_in_refs p_is_Node refs_def have refs_of2: "refs_of' (Ref.set p (Node x q) h) ref refs" by (simp add: refs_of'_set_ref) - from p_not_in_refs refs_of1 refs_of2 distinct_pointers prs_refs have 3: "\qrs prs. refs_of' (set_ref p (Node x q) h) p qrs \ refs_of' (set_ref p (Node x q) h) ref prs \ set prs \ set qrs = {}" + from p_not_in_refs refs_of1 refs_of2 distinct_pointers prs_refs have 3: "\qrs prs. refs_of' (Ref.set p (Node x q) h) p qrs \ refs_of' (Ref.set p (Node x q) h) ref prs \ set prs \ set qrs = {}" apply - apply (rule allI)+ apply (rule impI) apply (erule conjE) apply (drule refs_of'_is_fun) back back apply assumption apply (drule refs_of'_is_fun) back back apply assumption @@ -595,7 +595,7 @@ lemma rev_correctness: assumes list_of_h: "list_of h r xs" - assumes validHeap: "\refs. refs_of h r refs \ (\r \ set refs. ref_present r h)" + assumes validHeap: "\refs. refs_of h r refs \ (\r \ set refs. Ref.present h r)" assumes crel_rev: "crel (rev r) h h' r'" shows "list_of h' r' (List.rev xs)" using assms @@ -606,39 +606,39 @@ next case (Node x ps) with crel_rev obtain p q h1 h2 h3 v where - init: "crel (Ref.new Empty) h h1 q" - "crel (Ref.new (Node x ps)) h1 h2 p" + init: "crel (ref Empty) h h1 q" + "crel (ref (Node x ps)) h1 h2 p" and crel_rev':"crel (rev' (q, p)) h2 h3 v" and lookup: "crel (!v) h3 h' r'" using rev.simps by (auto elim!: crelE) from init have a1:"list_of' h2 q []" unfolding list_of'_def - by (auto elim!: crel_Ref_new) + by (auto elim!: crel_ref) from list_of_h obtain refs where refs_def: "refs_of h r refs" by (rule list_of_refs_of) - from validHeap init refs_def have heap_eq: "\refs. refs_of h r refs \ (\ref\set refs. ref_present ref h \ ref_present ref h2 \ get_ref ref h = get_ref ref h2)" - by (fastsimp elim!: crel_Ref_new dest: refs_of_is_fun) + from validHeap init refs_def have heap_eq: "\refs. refs_of h r refs \ (\ref\set refs. Ref.present h ref \ Ref.present h2 ref \ Ref.get h ref = Ref.get h2 ref)" + by (fastsimp elim!: crel_ref dest: refs_of_is_fun) from list_of_invariant[OF list_of_h heap_eq] have "list_of h2 r xs" . from init this Node have a2: "list_of' h2 p xs" apply - unfolding list_of'_def - apply (auto elim!: crel_Ref_new) + apply (auto elim!: crel_ref) done from init have refs_of_q: "refs_of' h2 q [q]" - by (auto elim!: crel_Ref_new) + by (auto elim!: crel_ref) from refs_def Node have refs_of'_ps: "refs_of' h ps refs" by (auto simp add: refs_of'_def'[symmetric]) - from validHeap refs_def have all_ref_present: "\r\set refs. ref_present r h" by simp - from init refs_of'_ps Node this have heap_eq: "\refs. refs_of' h ps refs \ (\ref\set refs. ref_present ref h \ ref_present ref h2 \ get_ref ref h = get_ref ref h2)" - by (fastsimp elim!: crel_Ref_new dest: refs_of'_is_fun) + from validHeap refs_def have all_ref_present: "\r\set refs. Ref.present h r" by simp + from init refs_of'_ps Node this have heap_eq: "\refs. refs_of' h ps refs \ (\ref\set refs. Ref.present h ref \ Ref.present h2 ref \ Ref.get h ref = Ref.get h2 ref)" + by (fastsimp elim!: crel_ref dest: refs_of'_is_fun) from refs_of'_invariant[OF refs_of'_ps this] have "refs_of' h2 ps refs" . with init have refs_of_p: "refs_of' h2 p (p#refs)" - by (auto elim!: crel_Ref_new simp add: refs_of'_def') + by (auto elim!: crel_ref simp add: refs_of'_def') with init all_ref_present have q_is_new: "q \ set (p#refs)" - by (auto elim!: crel_Ref_new intro!: noteq_refsI) + by (auto elim!: crel_ref intro!: Ref.noteq_I) from refs_of_p refs_of_q q_is_new have a3: "\qrs prs. refs_of' h2 q qrs \ refs_of' h2 p prs \ set prs \ set qrs = {}" by (fastsimp simp only: set.simps dest: refs_of'_is_fun) - from rev'_invariant [OF crel_rev' a1 a2 a3] have "list_of h3 (get_ref v h3) (List.rev xs)" + from rev'_invariant [OF crel_rev' a1 a2 a3] have "list_of h3 (Ref.get h3 v) (List.rev xs)" unfolding list_of'_def by auto with lookup show ?thesis by (auto elim: crel_lookup) @@ -734,32 +734,32 @@ lemma merge_induct2: assumes "list_of' h (p::'a::{heap, ord} node ref) xs" assumes "list_of' h q ys" - assumes "\ ys p q. \ list_of' h p []; list_of' h q ys; get_ref p h = Empty \ \ P p q [] ys" - assumes "\ x xs' p q pn. \ list_of' h p (x#xs'); list_of' h q []; get_ref p h = Node x pn; get_ref q h = Empty \ \ P p q (x#xs') []" + assumes "\ ys p q. \ list_of' h p []; list_of' h q ys; Ref.get h p = Empty \ \ P p q [] ys" + assumes "\ x xs' p q pn. \ list_of' h p (x#xs'); list_of' h q []; Ref.get h p = Node x pn; Ref.get h q = Empty \ \ P p q (x#xs') []" assumes "\ x xs' y ys' p q pn qn. - \ list_of' h p (x#xs'); list_of' h q (y#ys'); get_ref p h = Node x pn; get_ref q h = Node y qn; + \ list_of' h p (x#xs'); list_of' h q (y#ys'); Ref.get h p = Node x pn; Ref.get h q = Node y qn; x \ y; P pn q xs' (y#ys') \ \ P p q (x#xs') (y#ys')" assumes "\ x xs' y ys' p q pn qn. - \ list_of' h p (x#xs'); list_of' h q (y#ys'); get_ref p h = Node x pn; get_ref q h = Node y qn; + \ list_of' h p (x#xs'); list_of' h q (y#ys'); Ref.get h p = Node x pn; Ref.get h q = Node y qn; \ x \ y; P p qn (x#xs') ys'\ \ P p q (x#xs') (y#ys')" shows "P p q xs ys" using assms(1-2) proof (induct xs ys arbitrary: p q rule: Lmerge.induct) case (2 ys) - from 2(1) have "get_ref p h = Empty" unfolding list_of'_def by simp + from 2(1) have "Ref.get h p = Empty" unfolding list_of'_def by simp with 2(1-2) assms(3) show ?case by blast next case (3 x xs') - from 3(1) obtain pn where Node: "get_ref p h = Node x pn" by (rule list_of'_Cons) - from 3(2) have "get_ref q h = Empty" unfolding list_of'_def by simp + from 3(1) obtain pn where Node: "Ref.get h p = Node x pn" by (rule list_of'_Cons) + from 3(2) have "Ref.get h q = Empty" unfolding list_of'_def by simp with Node 3(1-2) assms(4) show ?case by blast next case (1 x xs' y ys') - from 1(3) obtain pn where pNode:"get_ref p h = Node x pn" + from 1(3) obtain pn where pNode:"Ref.get h p = Node x pn" and list_of'_pn: "list_of' h pn xs'" by (rule list_of'_Cons) - from 1(4) obtain qn where qNode:"get_ref q h = Node y qn" + from 1(4) obtain qn where qNode:"Ref.get h q = Node y qn" and list_of'_qn: "list_of' h qn ys'" by (rule list_of'_Cons) show ?case proof (cases "x \ y") @@ -780,15 +780,15 @@ assumes "list_of' h p xs" assumes "list_of' h q ys" assumes "crel (merge p q) h h' r" -assumes "\ ys p q. \ list_of' h p []; list_of' h q ys; get_ref p h = Empty \ \ P p q h h q [] ys" -assumes "\ x xs' p q pn. \ list_of' h p (x#xs'); list_of' h q []; get_ref p h = Node x pn; get_ref q h = Empty \ \ P p q h h p (x#xs') []" +assumes "\ ys p q. \ list_of' h p []; list_of' h q ys; Ref.get h p = Empty \ \ P p q h h q [] ys" +assumes "\ x xs' p q pn. \ list_of' h p (x#xs'); list_of' h q []; Ref.get h p = Node x pn; Ref.get h q = Empty \ \ P p q h h p (x#xs') []" assumes "\ x xs' y ys' p q pn qn h1 r1 h'. - \ list_of' h p (x#xs'); list_of' h q (y#ys');get_ref p h = Node x pn; get_ref q h = Node y qn; - x \ y; crel (merge pn q) h h1 r1 ; P pn q h h1 r1 xs' (y#ys'); h' = set_ref p (Node x r1) h1 \ + \ list_of' h p (x#xs'); list_of' h q (y#ys');Ref.get h p = Node x pn; Ref.get h q = Node y qn; + x \ y; crel (merge pn q) h h1 r1 ; P pn q h h1 r1 xs' (y#ys'); h' = Ref.set p (Node x r1) h1 \ \ P p q h h' p (x#xs') (y#ys')" assumes "\ x xs' y ys' p q pn qn h1 r1 h'. - \ list_of' h p (x#xs'); list_of' h q (y#ys'); get_ref p h = Node x pn; get_ref q h = Node y qn; - \ x \ y; crel (merge p qn) h h1 r1; P p qn h h1 r1 (x#xs') ys'; h' = set_ref q (Node y r1) h1 \ + \ list_of' h p (x#xs'); list_of' h q (y#ys'); Ref.get h p = Node x pn; Ref.get h q = Node y qn; + \ x \ y; crel (merge p qn) h h1 r1; P p qn h h1 r1 (x#xs') ys'; h' = Ref.set q (Node y r1) h1 \ \ P p q h h' q (x#xs') (y#ys')" shows "P p q h h' r xs ys" using assms(3) @@ -808,7 +808,7 @@ case (3 x xs' y ys' p q pn qn) from 3(3-5) 3(7) obtain h1 r1 where 1: "crel (merge pn q) h h1 r1" - and 2: "h' = set_ref p (Node x r1) h1 \ r = p" + and 2: "h' = Ref.set p (Node x r1) h1 \ r = p" unfolding merge_simps[of p q] by (auto elim!: crel_lookup crelE crel_return crel_if crel_update) from 3(6)[OF 1] assms(6) [OF 3(1-5)] 1 2 show ?case by simp @@ -816,7 +816,7 @@ case (4 x xs' y ys' p q pn qn) from 4(3-5) 4(7) obtain h1 r1 where 1: "crel (merge p qn) h h1 r1" - and 2: "h' = set_ref q (Node y r1) h1 \ r = q" + and 2: "h' = Ref.set q (Node y r1) h1 \ r = q" unfolding merge_simps[of p q] by (auto elim!: crel_lookup crelE crel_return crel_if crel_update) from 4(6)[OF 1] assms(7) [OF 4(1-5)] 1 2 show ?case by simp @@ -834,7 +834,7 @@ assumes "crel (merge p q) h h' r'" assumes "set xs \ set ys = {}" assumes "r \ set xs \ set ys" - shows "get_ref r h = get_ref r h'" + shows "Ref.get h r = Ref.get h' r" proof - from assms(1) obtain ps where ps_def: "list_of' h p ps" by (rule refs_of'_list_of') from assms(2) obtain qs where qs_def: "list_of' h q qs" by (rule refs_of'_list_of') @@ -853,7 +853,8 @@ from pnrs_def 3(12) have "r \ p" by auto with 3(11) 3(12) pnrs_def refs_of'_distinct[OF 3(9)] have p_in: "p \ set pnrs \ set ys" by auto from 3(11) pnrs_def have no_inter: "set pnrs \ set ys = {}" by auto - from 3(7)[OF refs_of'_pn 3(10) this p_in] 3(3) have p_is_Node: "get_ref p h1 = Node x pn" by simp + from 3(7)[OF refs_of'_pn 3(10) this p_in] 3(3) have p_is_Node: "Ref.get h1 p = Node x pn" + by simp from 3(7)[OF refs_of'_pn 3(10) no_inter r_in] 3(8) `r \ p` show ?case by simp next @@ -866,7 +867,7 @@ from qnrs_def 4(12) have "r \ q" by auto with 4(11) 4(12) qnrs_def refs_of'_distinct[OF 4(10)] have q_in: "q \ set xs \ set qnrs" by auto from 4(11) qnrs_def have no_inter: "set xs \ set qnrs = {}" by auto - from 4(7)[OF 4(9) refs_of'_qn this q_in] 4(4) have q_is_Node: "get_ref q h1 = Node y qn" by simp + from 4(7)[OF 4(9) refs_of'_qn this q_in] 4(4) have q_is_Node: "Ref.get h1 q = Node y qn" by simp from 4(7)[OF 4(9) refs_of'_qn no_inter r_in] 4(8) `r \ q` show ?case by simp qed @@ -899,7 +900,7 @@ by (rule refs_of'_Node) from 3(10) 3(9) 3(11) pnrs_def refs_of'_distinct[OF 3(9)] have p_in: "p \ set pnrs \ set ys" by auto from 3(11) pnrs_def have no_inter: "set pnrs \ set ys = {}" by auto - from merge_unchanged[OF refs_of'_pn 3(10) 3(6) no_inter p_in] have p_stays: "get_ref p h1 = get_ref p h" .. + from merge_unchanged[OF refs_of'_pn 3(10) 3(6) no_inter p_in] have p_stays: "Ref.get h1 p = Ref.get h p" .. from 3 p_stays obtain r1s where rs_def: "rs = p#r1s" and refs_of'_r1:"refs_of' h1 r1 r1s" by (auto elim: refs_of'_set_next_ref) @@ -912,7 +913,7 @@ by (rule refs_of'_Node) from 4(10) 4(9) 4(11) qnrs_def refs_of'_distinct[OF 4(10)] have q_in: "q \ set xs \ set qnrs" by auto from 4(11) qnrs_def have no_inter: "set xs \ set qnrs = {}" by auto - from merge_unchanged[OF 4(9) refs_of'_qn 4(6) no_inter q_in] have q_stays: "get_ref q h1 = get_ref q h" .. + from merge_unchanged[OF 4(9) refs_of'_qn 4(6) no_inter q_in] have q_stays: "Ref.get h1 q = Ref.get h q" .. from 4 q_stays obtain r1s where rs_def: "rs = q#r1s" and refs_of'_r1:"refs_of' h1 r1 r1s" by (auto elim: refs_of'_set_next_ref) @@ -945,7 +946,7 @@ from prs_def qrs_def 3(9) pnrs_def have no_inter: "set pnrs \ set qrs = {}" by fastsimp from no_inter refs_of'_pn qrs_def have no_inter2: "\qrs prs. refs_of' h q qrs \ refs_of' h pn prs \ set prs \ set qrs = {}" by (fastsimp dest: refs_of'_is_fun) - from merge_unchanged[OF refs_of'_pn qrs_def 3(6) no_inter p_in] have p_stays: "get_ref p h1 = get_ref p h" .. + from merge_unchanged[OF refs_of'_pn qrs_def 3(6) no_inter p_in] have p_stays: "Ref.get h1 p = Ref.get h p" .. from 3(7)[OF no_inter2] obtain rs where rs_def: "refs_of' h1 r1 rs" by (rule list_of'_refs_of') from refs_of'_merge[OF refs_of'_pn qrs_def 3(6) no_inter this] p_in have p_rs: "p \ set rs" by auto with 3(7)[OF no_inter2] 3(1-5) 3(8) p_rs rs_def p_stays @@ -962,7 +963,7 @@ from prs_def qrs_def 4(9) qnrs_def have no_inter: "set prs \ set qnrs = {}" by fastsimp from no_inter refs_of'_qn prs_def have no_inter2: "\qrs prs. refs_of' h qn qrs \ refs_of' h p prs \ set prs \ set qrs = {}" by (fastsimp dest: refs_of'_is_fun) - from merge_unchanged[OF prs_def refs_of'_qn 4(6) no_inter q_in] have q_stays: "get_ref q h1 = get_ref q h" .. + from merge_unchanged[OF prs_def refs_of'_qn 4(6) no_inter q_in] have q_stays: "Ref.get h1 q = Ref.get h q" .. from 4(7)[OF no_inter2] obtain rs where rs_def: "refs_of' h1 r1 rs" by (rule list_of'_refs_of') from refs_of'_merge[OF prs_def refs_of'_qn 4(6) no_inter this] q_in have q_rs: "q \ set rs" by auto with 4(7)[OF no_inter2] 4(1-5) 4(8) q_rs rs_def q_stays @@ -984,8 +985,8 @@ (do ll_xs \ make_llist (filter (%n. n mod 2 = 0) [2..8]); ll_ys \ make_llist (filter (%n. n mod 2 = 1) [5..11]); - r \ Ref.new ll_xs; - q \ Ref.new ll_ys; + r \ ref ll_xs; + q \ ref ll_ys; p \ merge r q; ll_zs \ !p; zs \ traverse ll_zs; @@ -998,4 +999,4 @@ ML {* @{code test_2} () *} ML {* @{code test_3} () *} -end \ No newline at end of file +end