# HG changeset patch # User bulwahn # Date 1260442706 -3600 # Node ID 1a82e2e29d67ecdd035b4693715357d02a89debb # Parent 3d2acb18f2f22dc6050694ac02d6fb9547f6f052 added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas diff -r 3d2acb18f2f2 -r 1a82e2e29d67 src/HOL/Imperative_HOL/Heap.thy --- a/src/HOL/Imperative_HOL/Heap.thy Wed Dec 09 21:33:50 2009 +0100 +++ b/src/HOL/Imperative_HOL/Heap.thy Thu Dec 10 11:58:26 2009 +0100 @@ -365,6 +365,10 @@ "ref_present r (set_ref r' v h) = ref_present r h" by (simp add: set_ref_def ref_present_def) +lemma noteq_refsI: "\ ref_present r h; \ref_present r' h \ \ r =!= r'" + unfolding noteq_refs_def ref_present_def + by auto + lemma array_ranI: "\ Some b = get_array a h ! i; i < Heap.length a h \ \ b \ array_ran a h" unfolding array_ran_def Heap.length_def by simp diff -r 3d2acb18f2f2 -r 1a82e2e29d67 src/HOL/Imperative_HOL/Heap_Monad.thy --- a/src/HOL/Imperative_HOL/Heap_Monad.thy Wed Dec 09 21:33:50 2009 +0100 +++ b/src/HOL/Imperative_HOL/Heap_Monad.thy Thu Dec 10 11:58:26 2009 +0100 @@ -266,6 +266,81 @@ shows "(assert P x >>= f) = (assert P' x >>= f')" using assms by (auto simp add: assert_def return_bind raise_bind) +subsubsection {* A monadic combinator for simple recursive functions *} + +function (default "\(f,g,x,h). (Inr Exn, undefined)") + mrec +where + "mrec f g x h = + (case Heap_Monad.execute (f x) h of + (Inl (Inl r), h') \ (Inl r, h') + | (Inl (Inr s), h') \ + (case mrec f g s h' of + (Inl z, h'') \ Heap_Monad.execute (g x s z) h'' + | (Inr e, h'') \ (Inr e, h'')) + | (Inr e, h') \ (Inr e, h') + )" +by auto + +lemma graph_implies_dom: + "mrec_graph x y \ mrec_dom x" +apply (induct rule:mrec_graph.induct) +apply (rule accpI) +apply (erule mrec_rel.cases) +by simp + +lemma f_default: "\ mrec_dom (f, g, x, h) \ mrec f g x h = (Inr Exn, undefined)" + unfolding mrec_def + by (rule fundef_default_value[OF mrec_sumC_def graph_implies_dom, of _ _ "(f, g, x, h)", simplified]) + +lemma f_di_reverse: + assumes "\ mrec_dom (f, g, x, h)" + shows " + (case Heap_Monad.execute (f x) h of + (Inl (Inl r), h') \ mrecalse + | (Inl (Inr s), h') \ \ mrec_dom (f, g, s, h') + | (Inr e, h') \ mrecalse + )" +using assms +by (auto split:prod.splits sum.splits) + (erule notE, rule accpI, elim mrec_rel.cases, simp)+ + + +lemma mrec_rule: + "mrec f g x h = + (case Heap_Monad.execute (f x) h of + (Inl (Inl r), h') \ (Inl r, h') + | (Inl (Inr s), h') \ + (case mrec f g s h' of + (Inl z, h'') \ Heap_Monad.execute (g x s z) h'' + | (Inr e, h'') \ (Inr e, h'')) + | (Inr e, h') \ (Inr e, h') + )" +apply (cases "mrec_dom (f,g,x,h)", simp) +apply (frule f_default) +apply (frule f_di_reverse, simp) +by (auto split: sum.split prod.split simp: f_default) + + +definition + "MREC f g x = Heap (mrec f g x)" + +lemma MREC_rule: + "MREC f g x = + (do y \ f x; + (case y of + Inl r \ return r + | Inr s \ + do z \ MREC f g s ; + g x s z + done) done)" + unfolding MREC_def + unfolding bindM_def return_def + apply simp + apply (rule ext) + apply (unfold mrec_rule[of f g x]) + by (auto split:prod.splits sum.splits) + hide (open) const heap execute diff -r 3d2acb18f2f2 -r 1a82e2e29d67 src/HOL/Imperative_HOL/Imperative_HOL_ex.thy --- a/src/HOL/Imperative_HOL/Imperative_HOL_ex.thy Wed Dec 09 21:33:50 2009 +0100 +++ b/src/HOL/Imperative_HOL/Imperative_HOL_ex.thy Thu Dec 10 11:58:26 2009 +0100 @@ -6,7 +6,7 @@ header {* Monadic imperative HOL with examples *} theory Imperative_HOL_ex -imports Imperative_HOL "ex/Imperative_Quicksort" +imports Imperative_HOL "ex/Imperative_Quicksort" "ex/Imperative_Reverse" "ex/Linked_Lists" begin end diff -r 3d2acb18f2f2 -r 1a82e2e29d67 src/HOL/Imperative_HOL/Ref.thy --- a/src/HOL/Imperative_HOL/Ref.thy Wed Dec 09 21:33:50 2009 +0100 +++ b/src/HOL/Imperative_HOL/Ref.thy Thu Dec 10 11:58:26 2009 +0100 @@ -60,9 +60,9 @@ subsubsection {* SML *} -code_type ref (SML "_/ ref") +code_type ref (SML "_/ Unsynchronized.ref") code_const Ref (SML "raise/ (Fail/ \"bare Ref\")") -code_const Ref.new (SML "(fn/ ()/ =>/ ref/ _)") +code_const Ref.new (SML "(fn/ ()/ =>/ Unsynchronized.ref/ _)") code_const Ref.lookup (SML "(fn/ ()/ =>/ !/ _)") code_const Ref.update (SML "(fn/ ()/ =>/ _/ :=/ _)") diff -r 3d2acb18f2f2 -r 1a82e2e29d67 src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Imperative_HOL/ex/Imperative_Reverse.thy Thu Dec 10 11:58:26 2009 +0100 @@ -0,0 +1,112 @@ +theory Imperative_Reverse +imports "~~/src/HOL/Imperative_HOL/Imperative_HOL" Subarray +begin + +hide (open) const swap rev + +fun swap :: "'a\heap array \ nat \ nat \ unit Heap" where + "swap a i j = (do + x \ nth a i; + y \ nth a j; + upd i y a; + upd j x a; + return () + done)" + +fun rev :: "'a\heap array \ nat \ nat \ unit Heap" where + "rev a i j = (if (i < j) then (do + swap a i j; + rev a (i + 1) (j - 1) + done) + else return ())" + +notation (output) swap ("swap") +notation (output) rev ("rev") + +declare swap.simps [simp del] rev.simps [simp del] + +lemma swap_pointwise: assumes "crel (swap a i j) h h' r" + shows "get_array a h' ! k = (if k = i then get_array a h ! j + else if k = j then get_array a h ! i + else get_array a h ! k)" +using assms unfolding swap.simps +by (elim crel_elim_all) + (auto simp: Heap.length_def) + +lemma rev_pointwise: assumes "crel (rev a i j) h h' r" + shows "get_array a h' ! k = (if k < i then get_array a h ! k + else if j < k then get_array a h ! k + else get_array a h ! (j - (k - i)))" (is "?P a i j h h'") +using assms proof (induct a i j arbitrary: h h' rule: rev.induct) + case (1 a i j h h'') + thus ?case + proof (cases "i < j") + case True + with 1[unfolded rev.simps[of a i j]] + obtain h' where + swp: "crel (swap a i j) h h' ()" + and rev: "crel (rev a (i + 1) (j - 1)) h' h'' ()" + by (auto elim: crel_elim_all) + from rev 1 True + have eq: "?P a (i + 1) (j - 1) h' h''" by auto + + have "k < i \ i = k \ (i < k \ k < j) \ j = k \ j < k" by arith + with True show ?thesis + by (elim disjE) (auto simp: eq swap_pointwise[OF swp]) + next + case False + with 1[unfolded rev.simps[of a i j]] + show ?thesis + by (cases "k = j") (auto elim: crel_elim_all) + qed +qed + +lemma rev_length: + assumes "crel (rev a i j) h h' r" + shows "Heap.length a h = Heap.length a h'" +using assms +proof (induct a i j arbitrary: h h' rule: rev.induct) + case (1 a i j h h'') + thus ?case + proof (cases "i < j") + case True + with 1[unfolded rev.simps[of a i j]] + obtain h' where + swp: "crel (swap a i j) h h' ()" + and rev: "crel (rev a (i + 1) (j - 1)) h' h'' ()" + by (auto elim: crel_elim_all) + from swp rev 1 True show ?thesis + unfolding swap.simps + by (elim crel_elim_all) fastsimp + next + case False + with 1[unfolded rev.simps[of a i j]] + show ?thesis + by (auto elim: crel_elim_all) + qed +qed + +lemma rev2_rev': assumes "crel (rev a i j) h h' u" + assumes "j < Heap.length a h" + shows "subarray i (j + 1) a h' = List.rev (subarray i (j + 1) a h)" +proof - + { + fix k + assume "k < Suc j - i" + with rev_pointwise[OF assms(1)] have "get_array a h' ! (i + k) = get_array a h ! (j - k)" + by auto + } + with assms(2) rev_length[OF assms(1)] show ?thesis + unfolding subarray_def Heap.length_def + by (auto simp add: length_sublist' rev_nth min_def nth_sublist' intro!: nth_equalityI) +qed + +lemma rev2_rev: + assumes "crel (rev a 0 (Heap.length a h - 1)) h h' u" + shows "get_array a h' = List.rev (get_array a h)" + using rev2_rev'[OF assms] rev_length[OF assms] assms + by (cases "Heap.length a h = 0", auto simp add: Heap.length_def + subarray_def sublist'_all rev.simps[where j=0] elim!: crel_elim_all) + (drule sym[of "List.length (get_array a h)"], simp) + +end diff -r 3d2acb18f2f2 -r 1a82e2e29d67 src/HOL/Imperative_HOL/ex/Linked_Lists.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Thu Dec 10 11:58:26 2009 +0100 @@ -0,0 +1,993 @@ +theory Linked_Lists +imports "~~/src/HOL/Imperative_HOL/Imperative_HOL" Code_Integer +begin + +section {* Definition of Linked Lists *} + +setup {* Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \ 'a\type ref"}) *} +datatype 'a node = Empty | Node 'a "('a node) ref" + +fun + node_encode :: "'a\countable node \ nat" +where + "node_encode Empty = 0" + | "node_encode (Node x r) = Suc (to_nat (x, r))" + +instance node :: (countable) countable +proof (rule countable_classI [of "node_encode"]) + fix x y :: "'a\countable node" + show "node_encode x = node_encode y \ x = y" + by (induct x, auto, induct y, auto, induct y, auto) +qed + +instance node :: (heap) heap .. + +fun 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; + return (Node x next) + done" + + +text {* define traverse using the MREC combinator *} + +definition + traverse :: "'a\heap node \ 'a list Heap" +where +[code del]: "traverse = MREC (\n. case n of Empty \ return (Inl []) + | Node x r \ (do tl \ Ref.lookup r; + return (Inr tl) done)) + (\n tl xs. case n of Empty \ undefined + | Node x r \ return (x # xs))" + + +lemma traverse_simps[code, simp]: + "traverse Empty = return []" + "traverse (Node x r) = do tl \ Ref.lookup r; + xs \ traverse tl; + return (x#xs) + done" +unfolding traverse_def +by (auto simp: traverse_def monad_simp MREC_rule) + + +section {* Proving correctness with relational abstraction *} + +subsection {* Definition of list_of, list_of', refs_of and refs_of' *} + +fun list_of :: "heap \ ('a::heap) node \ 'a list \ bool" +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))" + +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" + +fun 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)" + +fun 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)" + + +subsection {* Properties of these definitions *} + +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')" +by (cases xs, auto) + +lemma list_of'_Empty[simp]: "get_ref q h = 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')" +unfolding list_of'_def by simp + +lemma list_of'_Nil: "list_of' h q [] \ get_ref q h = 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" +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)" + 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)" +by (cases ps, auto) + +lemma refs_of'_Node: + assumes "refs_of' h p xs" + assumes "get_ref p h = Node x pn" + obtains pnrs + where "xs = p # pnrs" and "refs_of' h pn pnrs" +using assms +unfolding refs_of'_def' by auto + +lemma list_of_is_fun: "\ list_of h n xs; list_of h n ys\ \ xs = ys" +proof (induct xs arbitrary: ys n) + case Nil thus ?case by auto +next + case (Cons x xs') + thus ?case + by (cases ys, auto split: node.split_asm) +qed + +lemma refs_of_is_fun: "\ refs_of h n xs; refs_of h n ys\ \ xs = ys" +proof (induct xs arbitrary: ys n) + case Nil thus ?case by auto +next + case (Cons x xs') + thus ?case + by (cases ys, auto split: node.split_asm) +qed + +lemma refs_of'_is_fun: "\ refs_of' h p as; refs_of' h p bs \ \ as = bs" +unfolding refs_of'_def' by (auto dest: refs_of_is_fun) + + +lemma list_of_refs_of_HOL: + assumes "list_of h r xs" + shows "\rs. refs_of h r rs" +using assms +proof (induct xs arbitrary: r) + case Nil thus ?case by auto +next + case (Cons x xs') + thus ?case + by (cases r, auto) +qed + +lemma list_of_refs_of: + assumes "list_of h r xs" + obtains rs where "refs_of h r rs" +using list_of_refs_of_HOL[OF assms] +by auto + +lemma list_of'_refs_of'_HOL: + 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'" + unfolding list_of'_def by (rule list_of_refs_of) + thus ?thesis unfolding refs_of'_def' by auto +qed + +lemma list_of'_refs_of': + assumes "list_of' h r xs" + obtains rs where "refs_of' h r rs" +using list_of'_refs_of'_HOL[OF assms] +by auto + +lemma refs_of_list_of_HOL: + assumes "refs_of h r rs" + shows "\xs. list_of h r xs" +using assms +proof (induct rs arbitrary: r) + case Nil thus ?case by auto +next + case (Cons r rs') + thus ?case + by (cases r, auto) +qed + +lemma refs_of_list_of: + assumes "refs_of h r rs" + obtains xs where "list_of h r xs" +using refs_of_list_of_HOL[OF assms] +by auto + +lemma refs_of'_list_of'_HOL: + assumes "refs_of' h r rs" + shows "\xs. list_of' h r xs" +using assms +unfolding list_of'_def refs_of'_def' +by (auto intro: refs_of_list_of) + + +lemma refs_of'_list_of': + assumes "refs_of' h r rs" + obtains xs where "list_of' h r xs" +using refs_of'_list_of'_HOL[OF assms] +by auto + +lemma refs_of'E: "refs_of' h q rs \ q \ set rs" +unfolding refs_of'_def' by auto + +lemma list_of'_refs_of'2: + assumes "list_of' h r xs" + shows "\rs'. refs_of' h r (r#rs')" +proof - + from assms obtain rs where "refs_of' h r rs" by (rule list_of'_refs_of') + thus ?thesis by (auto simp add: refs_of'_def') +qed + +subsection {* More complicated properties of these predicates *} + +lemma list_of_append: + "list_of h n (as @ bs) \ \m. list_of h m bs" +apply (induct as arbitrary: n) +apply auto +apply (case_tac n) +apply auto +done + +lemma refs_of_append: "refs_of h n (as @ bs) \ \m. refs_of h m bs" +apply (induct as arbitrary: n) +apply auto +apply (case_tac n) +apply auto +done + +lemma refs_of_next: +assumes "refs_of h (get_ref p h) rs" + shows "p \ set rs" +proof (rule ccontr) + assume a: "\ (p \ set rs)" + from this obtain as bs where split:"rs = as @ p # bs" by (fastsimp dest: split_list) + with assms obtain q where "refs_of h q (p # bs)" by (fast dest: refs_of_append) + with assms split show "False" + by (cases q,auto dest: refs_of_is_fun) +qed + +lemma refs_of_distinct: "refs_of h p rs \ distinct rs" +proof (induct rs arbitrary: p) + case Nil thus ?case by simp +next + case (Cons r rs') + thus ?case + by (cases p, auto simp add: refs_of_next) +qed + +lemma refs_of'_distinct: "refs_of' h p rs \ distinct rs" + unfolding refs_of'_def' + by (fastsimp simp add: refs_of_distinct refs_of_next) + + +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" +using assms +proof (induct as arbitrary: q rs) + case Nil thus ?case by simp +next + case (Cons x xs) + thus ?case + proof (cases q) + 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(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) + 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" +proof (induct as arbitrary: q rs) + case Nil thus ?case by simp +next + case (Cons x xs) + thus ?case + proof (cases q) + 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(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) + 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" +proof (induct rs arbitrary: q) + case Nil thus ?case by simp +next + case (Cons x xs) + thus ?case + proof (cases q) + 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(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) + from Cons(3) have 2: "p \ set xs" by simp + with Cons.hyps 1 2 Node ref_eq show ?thesis + by simp + qed +qed + +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" +proof - + from assms have "q \ p" by (auto simp only: dest!: refs_of'E) + with assms show ?thesis + unfolding list_of'_def refs_of'_def' + by (auto simp add: list_of_set_ref) +qed + +lemma list_of'_set_next_ref_Node[simp]: + assumes "list_of' h r xs" + assumes "get_ref p h = 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" +using assms +unfolding list_of'_def refs_of'_def' +by (auto simp add: list_of_set_ref noteq_refs_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" +using assms +proof - + from assms have "q \ p" by (auto simp only: dest!: refs_of'E) + with assms show ?thesis + unfolding refs_of'_def' + by (auto simp add: refs_of_set_ref) +qed + +lemma refs_of'_set_ref2: + assumes "refs_of' (set_ref p v h) q rs" + assumes "p \ set rs" + shows "refs_of' (set_ref 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) + with assms show ?thesis + 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 (erule_tac x="prs" in meta_allE) + apply auto + apply (auto dest: refs_of_is_fun) + done +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" +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) + with prems show thesis by auto +qed + +section {* Proving make_llist and traverse correct *} + +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')" + shows "refs_of h' r xs" +using assms +proof (induct xs arbitrary: r) + case Nil thus ?case by simp +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'" + 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')" + by (fastsimp dest: refs_of_is_fun) + from Cons.hyps[OF 1 2] have "refs_of h' (get_ref x h') 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')" + 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" + 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')" + 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 +qed + +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')" + shows "list_of h' r xs" +using assms +proof (induct xs arbitrary: r) + case Nil thus ?case by simp +next + case (Cons x xs') + + from Cons(2) obtain ref where Node: "r = Node x ref" + 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'" + 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')" + by (auto dest: refs_of_is_fun) + from Cons(1)[OF 1 2] + have "list_of h' (get_ref ref h') xs'" . + with Node show ?case + unfolding list_of'_def + by simp +qed + +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'))" +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'" + 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 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) + 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 + 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) + from fstgoal sndgoal show ?case .. +qed + +lemma traverse: "list_of h n r \ crel (traverse n) h h r" +proof (induct r arbitrary: n) + case Nil + thus ?case + by (auto intro: crel_returnI) +next + case (Cons x xs) + thus ?case + apply (cases n, auto) + by (auto intro!: crelI crel_returnI crel_lookupI) +qed + +lemma traverse_make_llist': + assumes crel: "crel (make_llist xs \= traverse) h h' r" + shows "r = xs" +proof - + from crel obtain h1 r1 + where makell: "crel (make_llist xs) h h1 r1" + and trav: "crel (traverse r1) h1 h' r" + by (auto elim!: crelE) + from make_llist[OF makell] have "list_of h1 r1 xs" .. + from traverse [OF this] trav show ?thesis + using crel_deterministic by fastsimp +qed + +section {* Proving correctness of in-place reversal *} + +subsection {* Definition of in-place reversal *} + +definition rev' :: "(('a::heap) node ref \ 'a node ref) \ 'a node ref Heap" +where "rev' = MREC (\(q, p). do v \ !p; (case v of Empty \ (return (Inl q)) + | Node x next \ do + p := Node x q; + return (Inr (p, next)) + done) done) + (\x s z. return z)" + +lemma rev'_simps [code]: + "rev' (q, p) = + do + v \ !p; + (case v of + Empty \ return q + | Node x next \ + do + p := Node x q; + rev' (p, next) + done) + done" + unfolding rev'_def MREC_rule[of _ _ "(q, p)"] unfolding rev'_def[symmetric] +thm arg_cong2 + by (auto simp add: monad_simp expand_fun_eq intro: arg_cong2[where f = "op \="] split: node.split) + +fun 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)" + +subsection {* Correctness Proof *} + +lemma rev'_invariant: + assumes "crel (rev' (q, p)) h h' v" + assumes "list_of' h q qs" + assumes "list_of' h p ps" + assumes "\qrs prs. refs_of' h q qrs \ refs_of' h p prs \ set prs \ set qrs = {}" + shows "\vs. list_of' h' v vs \ vs = (List.rev ps) @ qs" +using assms +proof (induct ps arbitrary: qs p q h) + case Nil + thus ?case + unfolding rev'_simps[of q p] list_of'_def + by (auto elim!: crelE crel_lookup crel_return) +next + 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" + (*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" + 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)" + unfolding list_of'_def + apply (simp) + unfolding list_of'_def[symmetric] + by (simp add: list_of'_set_ref) + from list_of'_refs_of'2[OF Cons(4)] p_is_Node prs_def obtain refs where refs_def: "refs_of' h ref refs" and prs_refs: "prs = p # refs" + 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" + 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)" + 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" + 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 = {}" + 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 + apply auto done + from Cons.hyps [OF crel_rev' 1 2 3] show ?case by simp +qed + + +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 crel_rev: "crel (rev r) h h' r'" + shows "list_of h' r' (List.rev xs)" +using assms +proof (cases r) + case Empty + with list_of_h crel_rev show ?thesis + by (auto simp add: list_of_Empty elim!: crel_return) +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" + 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) + 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 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) + done + from init have refs_of_q: "refs_of' h2 q [q]" + by (auto elim!: crel_Ref_new) + 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 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') + with init all_ref_present have q_is_new: "q \ set (p#refs)" + by (auto elim!: crel_Ref_new intro!: noteq_refsI) + 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)" + unfolding list_of'_def by auto + with lookup show ?thesis + by (auto elim: crel_lookup) +qed + + +section {* The merge function on Linked Lists *} +text {* We also prove merge correct *} + +text{* First, we define merge on lists in a natural way. *} + +fun Lmerge :: "('a::ord) list \ 'a list \ 'a list" +where + "Lmerge (x#xs) (y#ys) = + (if x \ y then x # Lmerge xs (y#ys) else y # Lmerge (x#xs) ys)" +| "Lmerge [] ys = ys" +| "Lmerge xs [] = xs" + +subsection {* Definition of merge function *} + +definition merge' :: "(('a::{heap, ord}) node ref * ('a::{heap, ord})) * ('a::{heap, ord}) node ref * ('a::{heap, ord}) node ref \ ('a::{heap, ord}) node ref Heap" +where +"merge' = MREC (\(_, p, q). (do v \ !p; w \ !q; + (case v of Empty \ return (Inl q) + | Node valp np \ + (case w of Empty \ return (Inl p) + | Node valq nq \ + if (valp \ valq) then + return (Inr ((p, valp), np, q)) + else + return (Inr ((q, valq), p, nq)))) done)) + (\ _ ((n, v), _, _) r. do n := Node v r; return n done)" + +definition merge where "merge p q = merge' (undefined, p, q)" + +lemma if_return: "(if P then return x else return y) = return (if P then x else y)" +by auto + +lemma if_distrib_App: "(if P then f else g) x = (if P then f x else g x)" +by auto +lemma redundant_if: "(if P then (if P then x else z) else y) = (if P then x else y)" + "(if P then x else (if P then z else y)) = (if P then x else y)" +by auto + + + +lemma sum_distrib: "sum_case fl fr (case x of Empty \ y | Node v n \ (z v n)) = (case x of Empty \ sum_case fl fr y | Node v n \ sum_case fl fr (z v n))" +by (cases x) auto + +lemma merge: "merge' (x, p, q) = merge p q" +unfolding merge'_def merge_def +apply (simp add: MREC_rule) done +term "Ref.change" +lemma merge_simps [code]: +shows "merge p q = +do v \ !p; + w \ !q; + (case v of node.Empty \ return q + | Node valp np \ + case w of node.Empty \ return p + | Node valq nq \ + if valp \ valq then do r \ merge np q; + p := (Node valp r); + return p + done + else do r \ merge p nq; + q := (Node valq r); + return q + done) +done" +proof - + {fix v x y + have case_return: "(case v of Empty \ return x | Node v n \ return (y v n)) = return (case v of Empty \ x | Node v n \ y v n)" by (cases v) auto + } note case_return = this +show ?thesis +unfolding merge_def[of p q] merge'_def +apply (simp add: MREC_rule[of _ _ "(undefined, p, q)"]) +unfolding bind_bind return_bind +unfolding merge'_def[symmetric] +unfolding if_return case_return bind_bind return_bind sum_distrib sum.cases +unfolding if_distrib[symmetric, where f="Inr"] +unfolding sum.cases +unfolding if_distrib +unfolding split_beta fst_conv snd_conv +unfolding if_distrib_App redundant_if merge +.. +qed + +subsection {* Induction refinement by applying the abstraction function to our induct rule *} + +text {* From our original induction rule Lmerge.induct, we derive a new rule with our list_of' predicate *} + +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 "\ 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; + 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; + \ 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 + 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 + 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" + 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" + and list_of'_qn: "list_of' h qn ys'" by (rule list_of'_Cons) + show ?case + proof (cases "x \ y") + case True + from 1(1)[OF True list_of'_pn 1(4)] assms(5) 1(3-4) pNode qNode True + show ?thesis by blast + next + case False + from 1(2)[OF False 1(3) list_of'_qn] assms(6) 1(3-4) pNode qNode False + show ?thesis by blast + qed +qed + + +text {* secondly, we add the crel statement in the premise, and derive the crel statements for the single cases which we then eliminate with our crel elim rules. *} + +lemma merge_induct3: +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 "\ 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 \ + \ 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 \ + \ P p q h h' q (x#xs') (y#ys')" +shows "P p q h h' r xs ys" +using assms(3) +proof (induct arbitrary: h' r rule: merge_induct2[OF assms(1) assms(2)]) + case (1 ys p q) + from 1(3-4) have "h = h' \ r = q" + unfolding merge_simps[of p q] + by (auto elim!: crel_lookup crelE crel_return) + with assms(4)[OF 1(1) 1(2) 1(3)] show ?case by simp +next + case (2 x xs' p q pn) + from 2(3-5) have "h = h' \ r = p" + unfolding merge_simps[of p q] + by (auto elim!: crel_lookup crelE crel_return) + with assms(5)[OF 2(1-4)] show ?case by simp +next + 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" + 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 +next + 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" + 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 +qed + + +subsection {* Proving merge correct *} + +text {* As many parts of the following three proofs are identical, we could actually move the +same reasoning into an extended induction rule *} + +lemma merge_unchanged: + assumes "refs_of' h p xs" + assumes "refs_of' h q ys" + 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'" +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') + show ?thesis using assms(1) assms(2) assms(4) assms(5) + proof (induct arbitrary: xs ys r rule: merge_induct3[OF ps_def qs_def assms(3)]) + case 1 thus ?case by simp + next + case 2 thus ?case by simp + next + case (3 x xs' y ys' p q pn qn h1 r1 h' xs ys r) + from 3(9) 3(3) obtain pnrs + where pnrs_def: "xs = p#pnrs" + and refs_of'_pn: "refs_of' h pn pnrs" + by (rule refs_of'_Node) + with 3(12) have r_in: "r \ set pnrs \ set ys" by auto + 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) no_inter r_in] 3(8) `r \ p` show ?case + by simp + next + case (4 x xs' y ys' p q pn qn h1 r1 h' xs ys r) + from 4(10) 4(4) obtain qnrs + where qnrs_def: "ys = q#qnrs" + and refs_of'_qn: "refs_of' h qn qnrs" + by (rule refs_of'_Node) + with 4(12) have r_in: "r \ set xs \ set qnrs" by auto + 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 no_inter r_in] 4(8) `r \ q` show ?case + by simp + qed +qed + +lemma refs_of'_merge: + assumes "refs_of' h p xs" + assumes "refs_of' h q ys" + assumes "crel (merge p q) h h' r" + assumes "set xs \ set ys = {}" + assumes "refs_of' h' r rs" + shows "set rs \ set xs \ set ys" +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') + show ?thesis using assms(1) assms(2) assms(4) assms(5) + proof (induct arbitrary: xs ys rs rule: merge_induct3[OF ps_def qs_def assms(3)]) + case 1 + from 1(5) 1(7) have "rs = ys" by (fastsimp simp add: refs_of'_is_fun) + thus ?case by auto + next + case 2 + from 2(5) 2(8) have "rs = xs" by (auto simp add: refs_of'_is_fun) + thus ?case by auto + next + case (3 x xs' y ys' p q pn qn h1 r1 h' xs ys rs) + from 3(9) 3(3) obtain pnrs + where pnrs_def: "xs = p#pnrs" + and refs_of'_pn: "refs_of' h pn pnrs" + 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 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) + from 3(7)[OF refs_of'_pn 3(10) no_inter refs_of'_r1] rs_def pnrs_def show ?case by auto + next + case (4 x xs' y ys' p q pn qn h1 r1 h' xs ys rs) + from 4(10) 4(4) obtain qnrs + where qnrs_def: "ys = q#qnrs" + and refs_of'_qn: "refs_of' h qn qnrs" + 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 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) + from 4(7)[OF 4(9) refs_of'_qn no_inter refs_of'_r1] rs_def qnrs_def show ?case by auto + qed +qed + +lemma + assumes "list_of' h p xs" + assumes "list_of' h q ys" + assumes "crel (merge p q) h h' r" + assumes "\qrs prs. refs_of' h q qrs \ refs_of' h p prs \ set prs \ set qrs = {}" + shows "list_of' h' r (Lmerge xs ys)" +using assms(4) +proof (induct rule: merge_induct3[OF assms(1-3)]) + case 1 + thus ?case by simp +next + case 2 + thus ?case by simp +next + case (3 x xs' y ys' p q pn qn h1 r1 h') + from 3(1) obtain prs where prs_def: "refs_of' h p prs" by (rule list_of'_refs_of') + from 3(2) obtain qrs where qrs_def: "refs_of' h q qrs" by (rule list_of'_refs_of') + from prs_def 3(3) obtain pnrs + where pnrs_def: "prs = p#pnrs" + and refs_of'_pn: "refs_of' h pn pnrs" + by (rule refs_of'_Node) + from prs_def qrs_def 3(9) pnrs_def refs_of'_distinct[OF prs_def] have p_in: "p \ set pnrs \ set qrs" by fastsimp + 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 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 + show ?case by auto +next + case (4 x xs' y ys' p q pn qn h1 r1 h') + from 4(1) obtain prs where prs_def: "refs_of' h p prs" by (rule list_of'_refs_of') + from 4(2) obtain qrs where qrs_def: "refs_of' h q qrs" by (rule list_of'_refs_of') + from qrs_def 4(4) obtain qnrs + where qnrs_def: "qrs = q#qnrs" + and refs_of'_qn: "refs_of' h qn qnrs" + by (rule refs_of'_Node) + from prs_def qrs_def 4(9) qnrs_def refs_of'_distinct[OF qrs_def] have q_in: "q \ set prs \ set qnrs" by fastsimp + 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 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 + show ?case by auto +qed + +section {* Code generation *} + +export_code merge in SML file - + +export_code rev in SML file - + +text {* A simple example program *} + +definition test_1 where "test_1 = (do ll_xs <- make_llist [1..(15::int)]; xs <- traverse ll_xs; return xs done)" +definition test_2 where "test_2 = (do ll_xs <- make_llist [1..(15::int)]; ll_ys <- rev ll_xs; ys <- traverse ll_ys; return ys done)" + +definition test_3 where "test_3 = + (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; + p \ merge r q; + ll_zs \ !p; + zs \ traverse ll_zs; + return zs + done)" + +ML {* @{code test_1} () *} +ML {* @{code test_2} () *} +ML {* @{code test_3} () *} + +end \ No newline at end of file