| author | blanchet | 
| Thu, 12 May 2011 15:29:19 +0200 | |
| changeset 42738 | 2a9dcff63b80 | 
| parent 41549 | 2c65ad10bec8 | 
| child 44890 | 22f665a2e91c | 
| permissions | -rw-r--r-- | 
| 36098 
53992c639da5
added imperative SAT checker; improved headers of example files; adopted IsaMakefile
 bulwahn parents: 
35423diff
changeset | 1 | (* Title: HOL/Imperative_HOL/ex/Linked_Lists.thy | 
| 
53992c639da5
added imperative SAT checker; improved headers of example files; adopted IsaMakefile
 bulwahn parents: 
35423diff
changeset | 2 | Author: Lukas Bulwahn, TU Muenchen | 
| 
53992c639da5
added imperative SAT checker; improved headers of example files; adopted IsaMakefile
 bulwahn parents: 
35423diff
changeset | 3 | *) | 
| 
53992c639da5
added imperative SAT checker; improved headers of example files; adopted IsaMakefile
 bulwahn parents: 
35423diff
changeset | 4 | |
| 
53992c639da5
added imperative SAT checker; improved headers of example files; adopted IsaMakefile
 bulwahn parents: 
35423diff
changeset | 5 | header {* Linked Lists by ML references *}
 | 
| 
53992c639da5
added imperative SAT checker; improved headers of example files; adopted IsaMakefile
 bulwahn parents: 
35423diff
changeset | 6 | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 7 | theory Linked_Lists | 
| 41549 | 8 | imports "../Imperative_HOL" Code_Integer | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 9 | begin | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 10 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 11 | section {* Definition of Linked Lists *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 12 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 13 | setup {* Sign.add_const_constraint (@{const_name Ref}, SOME @{typ "nat \<Rightarrow> 'a\<Colon>type ref"}) *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 14 | datatype 'a node = Empty | Node 'a "('a node) ref"
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 15 | |
| 37725 | 16 | primrec | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 17 | node_encode :: "'a\<Colon>countable node \<Rightarrow> nat" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 18 | where | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 19 | "node_encode Empty = 0" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 20 | | "node_encode (Node x r) = Suc (to_nat (x, r))" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 21 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 22 | instance node :: (countable) countable | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 23 | proof (rule countable_classI [of "node_encode"]) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 24 | fix x y :: "'a\<Colon>countable node" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 25 | show "node_encode x = node_encode y \<Longrightarrow> x = y" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 26 | by (induct x, auto, induct y, auto, induct y, auto) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 27 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 28 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 29 | instance node :: (heap) heap .. | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 30 | |
| 37725 | 31 | primrec make_llist :: "'a\<Colon>heap list \<Rightarrow> 'a node Heap" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 32 | where | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 33 | [simp del]: "make_llist [] = return Empty" | 
| 37792 | 34 |             | "make_llist (x#xs) = do { tl \<leftarrow> make_llist xs;
 | 
| 35 | next \<leftarrow> ref tl; | |
| 36 | return (Node x next) | |
| 37 | }" | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 38 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 39 | |
| 40174 
97b69fef5229
use partial_function instead of MREC combinator; curried rev'
 krauss parents: 
39302diff
changeset | 40 | partial_function (heap) traverse :: "'a\<Colon>heap node \<Rightarrow> 'a list Heap" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 41 | where | 
| 40174 
97b69fef5229
use partial_function instead of MREC combinator; curried rev'
 krauss parents: 
39302diff
changeset | 42 | [code del]: "traverse l = | 
| 
97b69fef5229
use partial_function instead of MREC combinator; curried rev'
 krauss parents: 
39302diff
changeset | 43 | (case l of Empty \<Rightarrow> return [] | 
| 
97b69fef5229
use partial_function instead of MREC combinator; curried rev'
 krauss parents: 
39302diff
changeset | 44 |      | Node x r \<Rightarrow> do { tl \<leftarrow> Ref.lookup r;
 | 
| 
97b69fef5229
use partial_function instead of MREC combinator; curried rev'
 krauss parents: 
39302diff
changeset | 45 | xs \<leftarrow> traverse tl; | 
| 
97b69fef5229
use partial_function instead of MREC combinator; curried rev'
 krauss parents: 
39302diff
changeset | 46 | return (x#xs) | 
| 
97b69fef5229
use partial_function instead of MREC combinator; curried rev'
 krauss parents: 
39302diff
changeset | 47 | })" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 48 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 49 | lemma traverse_simps[code, simp]: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 50 | "traverse Empty = return []" | 
| 37792 | 51 |   "traverse (Node x r) = do { tl \<leftarrow> Ref.lookup r;
 | 
| 52 | xs \<leftarrow> traverse tl; | |
| 53 | return (x#xs) | |
| 54 | }" | |
| 40174 
97b69fef5229
use partial_function instead of MREC combinator; curried rev'
 krauss parents: 
39302diff
changeset | 55 | by (simp_all add: traverse.simps[of "Empty"] traverse.simps[of "Node x r"]) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 56 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 57 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 58 | section {* Proving correctness with relational abstraction *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 59 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 60 | subsection {* Definition of list_of, list_of', refs_of and refs_of' *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 61 | |
| 37725 | 62 | primrec list_of :: "heap \<Rightarrow> ('a::heap) node \<Rightarrow> 'a list \<Rightarrow> bool"
 | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 63 | where | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 64 | "list_of h r [] = (r = Empty)" | 
| 37725 | 65 | | "list_of h r (a#as) = (case r of Empty \<Rightarrow> False | Node b bs \<Rightarrow> (a = b \<and> list_of h (Ref.get h bs) as))" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 66 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 67 | definition list_of' :: "heap \<Rightarrow> ('a::heap) node ref \<Rightarrow> 'a list \<Rightarrow> bool"
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 68 | where | 
| 37725 | 69 | "list_of' h r xs = list_of h (Ref.get h r) xs" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 70 | |
| 37725 | 71 | primrec refs_of :: "heap \<Rightarrow> ('a::heap) node \<Rightarrow> 'a node ref list \<Rightarrow> bool"
 | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 72 | where | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 73 | "refs_of h r [] = (r = Empty)" | 
| 37725 | 74 | | "refs_of h r (x#xs) = (case r of Empty \<Rightarrow> False | Node b bs \<Rightarrow> (x = bs) \<and> refs_of h (Ref.get h bs) xs)" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 75 | |
| 37725 | 76 | primrec refs_of' :: "heap \<Rightarrow> ('a::heap) node ref \<Rightarrow> 'a node ref list \<Rightarrow> bool"
 | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 77 | where | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 78 | "refs_of' h r [] = False" | 
| 37725 | 79 | | "refs_of' h r (x#xs) = ((x = r) \<and> refs_of h (Ref.get h x) xs)" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 80 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 81 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 82 | subsection {* Properties of these definitions *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 83 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 84 | lemma list_of_Empty[simp]: "list_of h Empty xs = (xs = [])" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 85 | by (cases xs, auto) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 86 | |
| 37725 | 87 | lemma list_of_Node[simp]: "list_of h (Node x ps) xs = (\<exists>xs'. (xs = x # xs') \<and> list_of h (Ref.get h ps) xs')" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 88 | by (cases xs, auto) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 89 | |
| 37725 | 90 | lemma list_of'_Empty[simp]: "Ref.get h q = Empty \<Longrightarrow> list_of' h q xs = (xs = [])" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 91 | unfolding list_of'_def by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 92 | |
| 37725 | 93 | lemma list_of'_Node[simp]: "Ref.get h q = Node x ps \<Longrightarrow> list_of' h q xs = (\<exists>xs'. (xs = x # xs') \<and> list_of' h ps xs')" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 94 | unfolding list_of'_def by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 95 | |
| 37725 | 96 | lemma list_of'_Nil: "list_of' h q [] \<Longrightarrow> Ref.get h q = Empty" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 97 | unfolding list_of'_def by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 98 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 99 | lemma list_of'_Cons: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 100 | assumes "list_of' h q (x#xs)" | 
| 37725 | 101 | obtains n where "Ref.get h q = Node x n" and "list_of' h n xs" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 102 | using assms unfolding list_of'_def by (auto split: node.split_asm) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 103 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 104 | lemma refs_of_Empty[simp] : "refs_of h Empty xs = (xs = [])" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 105 | by (cases xs, auto) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 106 | |
| 37725 | 107 | lemma refs_of_Node[simp]: "refs_of h (Node x ps) xs = (\<exists>prs. xs = ps # prs \<and> refs_of h (Ref.get h ps) prs)" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 108 | by (cases xs, auto) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 109 | |
| 37725 | 110 | lemma refs_of'_def': "refs_of' h p ps = (\<exists>prs. (ps = (p # prs)) \<and> refs_of h (Ref.get h p) prs)" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 111 | by (cases ps, auto) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 112 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 113 | lemma refs_of'_Node: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 114 | assumes "refs_of' h p xs" | 
| 37725 | 115 | assumes "Ref.get h p = Node x pn" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 116 | obtains pnrs | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 117 | where "xs = p # pnrs" and "refs_of' h pn pnrs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 118 | using assms | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 119 | unfolding refs_of'_def' by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 120 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 121 | lemma list_of_is_fun: "\<lbrakk> list_of h n xs; list_of h n ys\<rbrakk> \<Longrightarrow> xs = ys" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 122 | proof (induct xs arbitrary: ys n) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 123 | case Nil thus ?case by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 124 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 125 | case (Cons x xs') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 126 | thus ?case | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 127 | by (cases ys, auto split: node.split_asm) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 128 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 129 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 130 | lemma refs_of_is_fun: "\<lbrakk> refs_of h n xs; refs_of h n ys\<rbrakk> \<Longrightarrow> xs = ys" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 131 | proof (induct xs arbitrary: ys n) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 132 | case Nil thus ?case by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 133 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 134 | case (Cons x xs') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 135 | thus ?case | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 136 | by (cases ys, auto split: node.split_asm) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 137 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 138 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 139 | lemma refs_of'_is_fun: "\<lbrakk> refs_of' h p as; refs_of' h p bs \<rbrakk> \<Longrightarrow> as = bs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 140 | unfolding refs_of'_def' by (auto dest: refs_of_is_fun) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 141 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 142 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 143 | lemma list_of_refs_of_HOL: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 144 | assumes "list_of h r xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 145 | shows "\<exists>rs. refs_of h r rs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 146 | using assms | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 147 | proof (induct xs arbitrary: r) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 148 | case Nil thus ?case by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 149 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 150 | case (Cons x xs') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 151 | thus ?case | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 152 | by (cases r, auto) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 153 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 154 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 155 | lemma list_of_refs_of: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 156 | assumes "list_of h r xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 157 | obtains rs where "refs_of h r rs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 158 | using list_of_refs_of_HOL[OF assms] | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 159 | by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 160 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 161 | lemma list_of'_refs_of'_HOL: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 162 | assumes "list_of' h r xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 163 | shows "\<exists>rs. refs_of' h r rs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 164 | proof - | 
| 37725 | 165 | from assms obtain rs' where "refs_of h (Ref.get h r) rs'" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 166 | unfolding list_of'_def by (rule list_of_refs_of) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 167 | thus ?thesis unfolding refs_of'_def' by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 168 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 169 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 170 | lemma list_of'_refs_of': | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 171 | assumes "list_of' h r xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 172 | obtains rs where "refs_of' h r rs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 173 | using list_of'_refs_of'_HOL[OF assms] | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 174 | by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 175 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 176 | lemma refs_of_list_of_HOL: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 177 | assumes "refs_of h r rs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 178 | shows "\<exists>xs. list_of h r xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 179 | using assms | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 180 | proof (induct rs arbitrary: r) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 181 | case Nil thus ?case by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 182 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 183 | case (Cons r rs') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 184 | thus ?case | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 185 | by (cases r, auto) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 186 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 187 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 188 | lemma refs_of_list_of: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 189 | assumes "refs_of h r rs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 190 | obtains xs where "list_of h r xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 191 | using refs_of_list_of_HOL[OF assms] | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 192 | by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 193 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 194 | lemma refs_of'_list_of'_HOL: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 195 | assumes "refs_of' h r rs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 196 | shows "\<exists>xs. list_of' h r xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 197 | using assms | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 198 | unfolding list_of'_def refs_of'_def' | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 199 | by (auto intro: refs_of_list_of) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 200 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 201 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 202 | lemma refs_of'_list_of': | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 203 | assumes "refs_of' h r rs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 204 | obtains xs where "list_of' h r xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 205 | using refs_of'_list_of'_HOL[OF assms] | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 206 | by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 207 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 208 | lemma refs_of'E: "refs_of' h q rs \<Longrightarrow> q \<in> set rs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 209 | unfolding refs_of'_def' by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 210 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 211 | lemma list_of'_refs_of'2: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 212 | assumes "list_of' h r xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 213 | shows "\<exists>rs'. refs_of' h r (r#rs')" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 214 | proof - | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 215 | from assms obtain rs where "refs_of' h r rs" by (rule list_of'_refs_of') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 216 | thus ?thesis by (auto simp add: refs_of'_def') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 217 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 218 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 219 | subsection {* More complicated properties of these predicates *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 220 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 221 | lemma list_of_append: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 222 | "list_of h n (as @ bs) \<Longrightarrow> \<exists>m. list_of h m bs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 223 | apply (induct as arbitrary: n) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 224 | apply auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 225 | apply (case_tac n) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 226 | apply auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 227 | done | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 228 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 229 | lemma refs_of_append: "refs_of h n (as @ bs) \<Longrightarrow> \<exists>m. refs_of h m bs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 230 | apply (induct as arbitrary: n) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 231 | apply auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 232 | apply (case_tac n) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 233 | apply auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 234 | done | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 235 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 236 | lemma refs_of_next: | 
| 37725 | 237 | assumes "refs_of h (Ref.get h p) rs" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 238 | shows "p \<notin> set rs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 239 | proof (rule ccontr) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 240 | assume a: "\<not> (p \<notin> set rs)" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 241 | from this obtain as bs where split:"rs = as @ p # bs" by (fastsimp dest: split_list) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 242 | with assms obtain q where "refs_of h q (p # bs)" by (fast dest: refs_of_append) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 243 | with assms split show "False" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 244 | by (cases q,auto dest: refs_of_is_fun) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 245 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 246 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 247 | lemma refs_of_distinct: "refs_of h p rs \<Longrightarrow> distinct rs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 248 | proof (induct rs arbitrary: p) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 249 | case Nil thus ?case by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 250 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 251 | case (Cons r rs') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 252 | thus ?case | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 253 | by (cases p, auto simp add: refs_of_next) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 254 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 255 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 256 | lemma refs_of'_distinct: "refs_of' h p rs \<Longrightarrow> distinct rs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 257 | unfolding refs_of'_def' | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 258 | by (fastsimp simp add: refs_of_distinct refs_of_next) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 259 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 260 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 261 | subsection {* Interaction of these predicates with our heap transitions *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 262 | |
| 37725 | 263 | lemma list_of_set_ref: "refs_of h q rs \<Longrightarrow> p \<notin> set rs \<Longrightarrow> list_of (Ref.set p v h) q as = list_of h q as" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 264 | using assms | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 265 | proof (induct as arbitrary: q rs) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 266 | case Nil thus ?case by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 267 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 268 | case (Cons x xs) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 269 | thus ?case | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 270 | proof (cases q) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 271 | case Empty thus ?thesis by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 272 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 273 | case (Node a ref) | 
| 37725 | 274 | from Cons(2) Node obtain rs' where 1: "refs_of h (Ref.get h ref) rs'" and rs_rs': "rs = ref # rs'" by auto | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 275 | from Cons(3) rs_rs' have "ref \<noteq> p" by fastsimp | 
| 37725 | 276 | hence ref_eq: "Ref.get (Ref.set p v h) ref = (Ref.get h ref)" by (auto simp add: Ref.get_set_neq) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 277 | from rs_rs' Cons(3) have 2: "p \<notin> set rs'" by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 278 | from Cons.hyps[OF 1 2] Node ref_eq show ?thesis by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 279 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 280 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 281 | |
| 37725 | 282 | lemma refs_of_set_ref: "refs_of h q rs \<Longrightarrow> p \<notin> set rs \<Longrightarrow> refs_of (Ref.set p v h) q as = refs_of h q as" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 283 | proof (induct as arbitrary: q rs) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 284 | case Nil thus ?case by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 285 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 286 | case (Cons x xs) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 287 | thus ?case | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 288 | proof (cases q) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 289 | case Empty thus ?thesis by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 290 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 291 | case (Node a ref) | 
| 37725 | 292 | from Cons(2) Node obtain rs' where 1: "refs_of h (Ref.get h ref) rs'" and rs_rs': "rs = ref # rs'" by auto | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 293 | from Cons(3) rs_rs' have "ref \<noteq> p" by fastsimp | 
| 37725 | 294 | hence ref_eq: "Ref.get (Ref.set p v h) ref = (Ref.get h ref)" by (auto simp add: Ref.get_set_neq) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 295 | from rs_rs' Cons(3) have 2: "p \<notin> set rs'" by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 296 | from Cons.hyps[OF 1 2] Node ref_eq show ?thesis by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 297 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 298 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 299 | |
| 37725 | 300 | lemma refs_of_set_ref2: "refs_of (Ref.set p v h) q rs \<Longrightarrow> p \<notin> set rs \<Longrightarrow> refs_of (Ref.set p v h) q rs = refs_of h q rs" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 301 | proof (induct rs arbitrary: q) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 302 | case Nil thus ?case by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 303 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 304 | case (Cons x xs) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 305 | thus ?case | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 306 | proof (cases q) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 307 | case Empty thus ?thesis by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 308 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 309 | case (Node a ref) | 
| 37725 | 310 | 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 | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 311 | from Cons(3) this have "ref \<noteq> p" by fastsimp | 
| 37725 | 312 | hence ref_eq: "Ref.get (Ref.set p v h) ref = (Ref.get h ref)" by (auto simp add: Ref.get_set_neq) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 313 | from Cons(3) have 2: "p \<notin> set xs" by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 314 | with Cons.hyps 1 2 Node ref_eq show ?thesis | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 315 | by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 316 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 317 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 318 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 319 | lemma list_of'_set_ref: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 320 | assumes "refs_of' h q rs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 321 | assumes "p \<notin> set rs" | 
| 37725 | 322 | shows "list_of' (Ref.set p v h) q as = list_of' h q as" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 323 | proof - | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 324 | from assms have "q \<noteq> p" by (auto simp only: dest!: refs_of'E) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 325 | with assms show ?thesis | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 326 | unfolding list_of'_def refs_of'_def' | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 327 | by (auto simp add: list_of_set_ref) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 328 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 329 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 330 | lemma list_of'_set_next_ref_Node[simp]: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 331 | assumes "list_of' h r xs" | 
| 37725 | 332 | assumes "Ref.get h p = Node x r'" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 333 | assumes "refs_of' h r rs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 334 | assumes "p \<notin> set rs" | 
| 37725 | 335 | shows "list_of' (Ref.set p (Node x r) h) p (x#xs) = list_of' h r xs" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 336 | using assms | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 337 | unfolding list_of'_def refs_of'_def' | 
| 37725 | 338 | by (auto simp add: list_of_set_ref Ref.noteq_sym) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 339 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 340 | lemma refs_of'_set_ref: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 341 | assumes "refs_of' h q rs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 342 | assumes "p \<notin> set rs" | 
| 37725 | 343 | shows "refs_of' (Ref.set p v h) q as = refs_of' h q as" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 344 | using assms | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 345 | proof - | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 346 | from assms have "q \<noteq> p" by (auto simp only: dest!: refs_of'E) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 347 | with assms show ?thesis | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 348 | unfolding refs_of'_def' | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 349 | by (auto simp add: refs_of_set_ref) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 350 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 351 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 352 | lemma refs_of'_set_ref2: | 
| 37725 | 353 | assumes "refs_of' (Ref.set p v h) q rs" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 354 | assumes "p \<notin> set rs" | 
| 37725 | 355 | shows "refs_of' (Ref.set p v h) q as = refs_of' h q as" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 356 | using assms | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 357 | proof - | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 358 | from assms have "q \<noteq> p" by (auto simp only: dest!: refs_of'E) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 359 | with assms show ?thesis | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 360 | unfolding refs_of'_def' | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 361 | apply auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 362 | apply (subgoal_tac "prs = prsa") | 
| 37725 | 363 | apply (insert refs_of_set_ref2[of p v h "Ref.get h q"]) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 364 | apply (erule_tac x="prs" in meta_allE) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 365 | apply auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 366 | apply (auto dest: refs_of_is_fun) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 367 | done | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 368 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 369 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 370 | lemma refs_of'_set_next_ref: | 
| 37725 | 371 | assumes "Ref.get h1 p = Node x pn" | 
| 372 | assumes "refs_of' (Ref.set p (Node x r1) h1) p rs" | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 373 | obtains r1s where "rs = (p#r1s)" and "refs_of' h1 r1 r1s" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 374 | proof - | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 375 | from assms refs_of'_distinct[OF assms(2)] have "\<exists> r1s. rs = (p # r1s) \<and> refs_of' h1 r1 r1s" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 376 | apply - | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 377 | unfolding refs_of'_def'[of _ p] | 
| 37725 | 378 | apply (auto, frule refs_of_set_ref2) by (auto dest: Ref.noteq_sym) | 
| 41549 | 379 | with assms that show thesis by auto | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 380 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 381 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 382 | section {* Proving make_llist and traverse correct *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 383 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 384 | lemma refs_of_invariant: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 385 |   assumes "refs_of h (r::('a::heap) node) xs"
 | 
| 37725 | 386 | assumes "\<forall>refs. refs_of h r refs \<longrightarrow> (\<forall>ref \<in> set refs. Ref.present h ref \<and> Ref.present h' ref \<and> Ref.get h ref = Ref.get h' ref)" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 387 | shows "refs_of h' r xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 388 | using assms | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 389 | proof (induct xs arbitrary: r) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 390 | case Nil thus ?case by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 391 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 392 | case (Cons x xs') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 393 | from Cons(2) obtain v where Node: "r = Node v x" by (cases r, auto) | 
| 37725 | 394 | from Cons(2) Node have refs_of_next: "refs_of h (Ref.get h x) xs'" by simp | 
| 395 | from Cons(2-3) Node have ref_eq: "Ref.get h x = Ref.get h' x" by auto | |
| 396 | from ref_eq refs_of_next have 1: "refs_of h (Ref.get h' x) xs'" by simp | |
| 397 | from Cons(2) Cons(3) have "\<forall>ref \<in> set xs'. Ref.present h ref \<and> Ref.present h' ref \<and> Ref.get h ref = Ref.get h' ref" | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 398 | by fastsimp | 
| 37725 | 399 | with Cons(3) 1 have 2: "\<forall>refs. refs_of h (Ref.get h' x) refs \<longrightarrow> (\<forall>ref \<in> set refs. Ref.present h ref \<and> Ref.present h' ref \<and> Ref.get h ref = Ref.get h' ref)" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 400 | by (fastsimp dest: refs_of_is_fun) | 
| 37725 | 401 | from Cons.hyps[OF 1 2] have "refs_of h' (Ref.get h' x) xs'" . | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 402 | with Node show ?case by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 403 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 404 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 405 | lemma refs_of'_invariant: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 406 | assumes "refs_of' h r xs" | 
| 37725 | 407 | assumes "\<forall>refs. refs_of' h r refs \<longrightarrow> (\<forall>ref \<in> set refs. Ref.present h ref \<and> Ref.present h' ref \<and> Ref.get h ref = Ref.get h' ref)" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 408 | shows "refs_of' h' r xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 409 | using assms | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 410 | proof - | 
| 37725 | 411 | from assms obtain prs where refs:"refs_of h (Ref.get h r) prs" and xs_def: "xs = r # prs" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 412 | unfolding refs_of'_def' by auto | 
| 37725 | 413 | from xs_def assms have x_eq: "Ref.get h r = Ref.get h' r" by fastsimp | 
| 414 | from refs assms xs_def have 2: "\<forall>refs. refs_of h (Ref.get h r) refs \<longrightarrow> | |
| 415 | (\<forall>ref\<in>set refs. Ref.present h ref \<and> Ref.present h' ref \<and> Ref.get h ref = Ref.get h' ref)" | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 416 | by (fastsimp dest: refs_of_is_fun) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 417 | from refs_of_invariant [OF refs 2] xs_def x_eq show ?thesis | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 418 | unfolding refs_of'_def' by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 419 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 420 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 421 | lemma list_of_invariant: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 422 |   assumes "list_of h (r::('a::heap) node) xs"
 | 
| 37725 | 423 | assumes "\<forall>refs. refs_of h r refs \<longrightarrow> (\<forall>ref \<in> set refs. Ref.present h ref \<and> Ref.present h' ref \<and> Ref.get h ref = Ref.get h' ref)" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 424 | shows "list_of h' r xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 425 | using assms | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 426 | proof (induct xs arbitrary: r) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 427 | case Nil thus ?case by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 428 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 429 | case (Cons x xs') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 430 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 431 | from Cons(2) obtain ref where Node: "r = Node x ref" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 432 | by (cases r, auto) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 433 | from Cons(2) obtain rs where rs_def: "refs_of h r rs" by (rule list_of_refs_of) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 434 | from Node rs_def obtain rss where refs_of: "refs_of h r (ref#rss)" and rss_def: "rs = ref#rss" by auto | 
| 37725 | 435 | from Cons(3) Node refs_of have ref_eq: "Ref.get h ref = Ref.get h' ref" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 436 | by auto | 
| 37725 | 437 | from Cons(2) ref_eq Node have 1: "list_of h (Ref.get h' ref) xs'" by simp | 
| 438 | from refs_of Node ref_eq have refs_of_ref: "refs_of h (Ref.get h' ref) rss" by simp | |
| 439 | from Cons(3) rs_def have rs_heap_eq: "\<forall>ref\<in>set rs. Ref.present h ref \<and> Ref.present h' ref \<and> Ref.get h ref = Ref.get h' ref" by simp | |
| 440 | from refs_of_ref rs_heap_eq rss_def have 2: "\<forall>refs. refs_of h (Ref.get h' ref) refs \<longrightarrow> | |
| 441 | (\<forall>ref\<in>set refs. Ref.present h ref \<and> Ref.present h' ref \<and> Ref.get h ref = Ref.get h' ref)" | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 442 | by (auto dest: refs_of_is_fun) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 443 | from Cons(1)[OF 1 2] | 
| 37725 | 444 | have "list_of h' (Ref.get h' ref) xs'" . | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 445 | with Node show ?case | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 446 | unfolding list_of'_def | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 447 | by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 448 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 449 | |
| 40671 | 450 | lemma effect_ref: | 
| 451 | assumes "effect (ref v) h h' x" | |
| 37771 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37765diff
changeset | 452 | obtains "Ref.get h' x = v" | 
| 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37765diff
changeset | 453 | and "\<not> Ref.present h x" | 
| 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37765diff
changeset | 454 | and "Ref.present h' x" | 
| 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37765diff
changeset | 455 | and "\<forall>y. Ref.present h y \<longrightarrow> Ref.get h y = Ref.get h' y" | 
| 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37765diff
changeset | 456 | (* and "lim h' = Suc (lim h)" *) | 
| 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37765diff
changeset | 457 | and "\<forall>y. Ref.present h y \<longrightarrow> Ref.present h' y" | 
| 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37765diff
changeset | 458 | using assms | 
| 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37765diff
changeset | 459 | unfolding Ref.ref_def | 
| 40671 | 460 | apply (elim effect_heapE) | 
| 37771 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37765diff
changeset | 461 | unfolding Ref.alloc_def | 
| 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37765diff
changeset | 462 | apply (simp add: Let_def) | 
| 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37765diff
changeset | 463 | unfolding Ref.present_def | 
| 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37765diff
changeset | 464 | apply auto | 
| 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37765diff
changeset | 465 | unfolding Ref.get_def Ref.set_def | 
| 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37765diff
changeset | 466 | apply auto | 
| 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37765diff
changeset | 467 | done | 
| 
1bec64044b5e
spelt out relational framework in a consistent way
 haftmann parents: 
37765diff
changeset | 468 | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 469 | lemma make_llist: | 
| 40671 | 470 | assumes "effect (make_llist xs) h h' r" | 
| 37725 | 471 | shows "list_of h' r xs \<and> (\<forall>rs. refs_of h' r rs \<longrightarrow> (\<forall>ref \<in> (set rs). Ref.present h' ref))" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 472 | using assms | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 473 | proof (induct xs arbitrary: h h' r) | 
| 40671 | 474 | case Nil thus ?case by (auto elim: effect_returnE simp add: make_llist.simps) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 475 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 476 | case (Cons x xs') | 
| 40671 | 477 | from Cons.prems obtain h1 r1 r' where make_llist: "effect (make_llist xs') h h1 r1" | 
| 478 | and effect_refnew:"effect (ref r1) h1 h' r'" and Node: "r = Node x r'" | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 479 | unfolding make_llist.simps | 
| 40671 | 480 | by (auto elim!: effect_bindE effect_returnE) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 481 | from Cons.hyps[OF make_llist] have list_of_h1: "list_of h1 r1 xs'" .. | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 482 | from Cons.hyps[OF make_llist] obtain rs' where rs'_def: "refs_of h1 r1 rs'" by (auto intro: list_of_refs_of) | 
| 37725 | 483 | from Cons.hyps[OF make_llist] rs'_def have refs_present: "\<forall>ref\<in>set rs'. Ref.present h1 ref" by simp | 
| 40671 | 484 | from effect_refnew rs'_def refs_present have refs_unchanged: "\<forall>refs. refs_of h1 r1 refs \<longrightarrow> | 
| 37725 | 485 | (\<forall>ref\<in>set refs. Ref.present h1 ref \<and> Ref.present h' ref \<and> Ref.get h1 ref = Ref.get h' ref)" | 
| 40671 | 486 | by (auto elim!: effect_ref dest: refs_of_is_fun) | 
| 487 | with list_of_invariant[OF list_of_h1 refs_unchanged] Node effect_refnew have fstgoal: "list_of h' r (x # xs')" | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 488 | unfolding list_of.simps | 
| 40671 | 489 | by (auto elim!: effect_refE) | 
| 37725 | 490 | from refs_unchanged rs'_def have refs_still_present: "\<forall>ref\<in>set rs'. Ref.present h' ref" by auto | 
| 40671 | 491 | from refs_of_invariant[OF rs'_def refs_unchanged] refs_unchanged Node effect_refnew refs_still_present | 
| 37725 | 492 | have sndgoal: "\<forall>rs. refs_of h' r rs \<longrightarrow> (\<forall>ref\<in>set rs. Ref.present h' ref)" | 
| 40671 | 493 | by (fastsimp elim!: effect_refE dest: refs_of_is_fun) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 494 | from fstgoal sndgoal show ?case .. | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 495 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 496 | |
| 40671 | 497 | lemma traverse: "list_of h n r \<Longrightarrow> effect (traverse n) h h r" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 498 | proof (induct r arbitrary: n) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 499 | case Nil | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 500 | thus ?case | 
| 40671 | 501 | by (auto intro: effect_returnI) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 502 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 503 | case (Cons x xs) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 504 | thus ?case | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 505 | apply (cases n, auto) | 
| 40671 | 506 | by (auto intro!: effect_bindI effect_returnI effect_lookupI) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 507 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 508 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 509 | lemma traverse_make_llist': | 
| 40671 | 510 | assumes effect: "effect (make_llist xs \<guillemotright>= traverse) h h' r" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 511 | shows "r = xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 512 | proof - | 
| 40671 | 513 | from effect obtain h1 r1 | 
| 514 | where makell: "effect (make_llist xs) h h1 r1" | |
| 515 | and trav: "effect (traverse r1) h1 h' r" | |
| 516 | by (auto elim!: effect_bindE) | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 517 | from make_llist[OF makell] have "list_of h1 r1 xs" .. | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 518 | from traverse [OF this] trav show ?thesis | 
| 40671 | 519 | using effect_deterministic by fastsimp | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 520 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 521 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 522 | section {* Proving correctness of in-place reversal *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 523 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 524 | subsection {* Definition of in-place reversal *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 525 | |
| 40174 
97b69fef5229
use partial_function instead of MREC combinator; curried rev'
 krauss parents: 
39302diff
changeset | 526 | partial_function (heap) rev' :: "('a::heap) node ref \<Rightarrow> 'a node ref \<Rightarrow> 'a node ref Heap"
 | 
| 
97b69fef5229
use partial_function instead of MREC combinator; curried rev'
 krauss parents: 
39302diff
changeset | 527 | where | 
| 
97b69fef5229
use partial_function instead of MREC combinator; curried rev'
 krauss parents: 
39302diff
changeset | 528 | [code]: "rev' q p = | 
| 37792 | 529 |    do {
 | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 530 | v \<leftarrow> !p; | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 531 | (case v of | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 532 | Empty \<Rightarrow> return q | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 533 | | Node x next \<Rightarrow> | 
| 37792 | 534 |         do {
 | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 535 | p := Node x q; | 
| 40174 
97b69fef5229
use partial_function instead of MREC combinator; curried rev'
 krauss parents: 
39302diff
changeset | 536 | rev' p next | 
| 37792 | 537 | }) | 
| 40174 
97b69fef5229
use partial_function instead of MREC combinator; curried rev'
 krauss parents: 
39302diff
changeset | 538 | }" | 
| 
97b69fef5229
use partial_function instead of MREC combinator; curried rev'
 krauss parents: 
39302diff
changeset | 539 | |
| 37725 | 540 | primrec rev :: "('a:: heap) node \<Rightarrow> 'a node Heap" 
 | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 541 | where | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 542 | "rev Empty = return Empty" | 
| 40174 
97b69fef5229
use partial_function instead of MREC combinator; curried rev'
 krauss parents: 
39302diff
changeset | 543 | | "rev (Node x n) = do { q \<leftarrow> ref Empty; p \<leftarrow> ref (Node x n); v \<leftarrow> rev' q p; !v }"
 | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 544 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 545 | subsection {* Correctness Proof *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 546 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 547 | lemma rev'_invariant: | 
| 40671 | 548 | assumes "effect (rev' q p) h h' v" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 549 | assumes "list_of' h q qs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 550 | assumes "list_of' h p ps" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 551 |   assumes "\<forall>qrs prs. refs_of' h q qrs \<and> refs_of' h p prs \<longrightarrow> set prs \<inter> set qrs = {}"
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 552 | shows "\<exists>vs. list_of' h' v vs \<and> vs = (List.rev ps) @ qs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 553 | using assms | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 554 | proof (induct ps arbitrary: qs p q h) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 555 | case Nil | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 556 | thus ?case | 
| 40174 
97b69fef5229
use partial_function instead of MREC combinator; curried rev'
 krauss parents: 
39302diff
changeset | 557 | unfolding rev'.simps[of q p] list_of'_def | 
| 40671 | 558 | by (auto elim!: effect_bindE effect_lookupE effect_returnE) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 559 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 560 | case (Cons x xs) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 561 | (*"LinkedList.list_of h' (get_ref v h') (List.rev xs @ x # qsa)"*) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 562 | from Cons(4) obtain ref where | 
| 37725 | 563 | p_is_Node: "Ref.get h p = Node x ref" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 564 | (*and "ref_present ref h"*) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 565 | and list_of'_ref: "list_of' h ref xs" | 
| 37725 | 566 | unfolding list_of'_def by (cases "Ref.get h p", auto) | 
| 40671 | 567 | from p_is_Node Cons(2) have effect_rev': "effect (rev' p ref) (Ref.set p (Node x q) h) h' v" | 
| 568 | by (auto simp add: rev'.simps [of q p] elim!: effect_bindE effect_lookupE effect_updateE) | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 569 | from Cons(3) obtain qrs where qrs_def: "refs_of' h q qrs" by (elim list_of'_refs_of') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 570 | from Cons(4) obtain prs where prs_def: "refs_of' h p prs" by (elim list_of'_refs_of') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 571 |   from qrs_def prs_def Cons(5) have distinct_pointers: "set qrs \<inter> set prs = {}" by fastsimp
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 572 | from qrs_def prs_def distinct_pointers refs_of'E have p_notin_qrs: "p \<notin> set qrs" by fastsimp | 
| 37725 | 573 | from Cons(3) qrs_def this have 1: "list_of' (Ref.set p (Node x q) h) p (x#qs)" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 574 | unfolding list_of'_def | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 575 | apply (simp) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 576 | unfolding list_of'_def[symmetric] | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 577 | by (simp add: list_of'_set_ref) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 578 | 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" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 579 | unfolding refs_of'_def' by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 580 | from prs_refs prs_def have p_not_in_refs: "p \<notin> set refs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 581 | by (fastsimp dest!: refs_of'_distinct) | 
| 37725 | 582 | with refs_def p_is_Node list_of'_ref have 2: "list_of' (Ref.set p (Node x q) h) ref xs" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 583 | by (auto simp add: list_of'_set_ref) | 
| 37725 | 584 | from p_notin_qrs qrs_def have refs_of1: "refs_of' (Ref.set p (Node x q) h) p (p#qrs)" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 585 | unfolding refs_of'_def' | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 586 | apply (simp) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 587 | unfolding refs_of'_def'[symmetric] | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 588 | by (simp add: refs_of'_set_ref) | 
| 37725 | 589 | from p_not_in_refs p_is_Node refs_def have refs_of2: "refs_of' (Ref.set p (Node x q) h) ref refs" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 590 | by (simp add: refs_of'_set_ref) | 
| 37725 | 591 |   from p_not_in_refs refs_of1 refs_of2 distinct_pointers prs_refs have 3: "\<forall>qrs prs. refs_of' (Ref.set p (Node x q) h) p qrs \<and> refs_of' (Ref.set p (Node x q) h) ref prs \<longrightarrow> set prs \<inter> set qrs = {}"
 | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 592 | apply - apply (rule allI)+ apply (rule impI) apply (erule conjE) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 593 | apply (drule refs_of'_is_fun) back back apply assumption | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 594 | apply (drule refs_of'_is_fun) back back apply assumption | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 595 | apply auto done | 
| 40671 | 596 | from Cons.hyps [OF effect_rev' 1 2 3] show ?case by simp | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 597 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 598 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 599 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 600 | lemma rev_correctness: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 601 | assumes list_of_h: "list_of h r xs" | 
| 37725 | 602 | assumes validHeap: "\<forall>refs. refs_of h r refs \<longrightarrow> (\<forall>r \<in> set refs. Ref.present h r)" | 
| 40671 | 603 | assumes effect_rev: "effect (rev r) h h' r'" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 604 | shows "list_of h' r' (List.rev xs)" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 605 | using assms | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 606 | proof (cases r) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 607 | case Empty | 
| 40671 | 608 | with list_of_h effect_rev show ?thesis | 
| 609 | by (auto simp add: list_of_Empty elim!: effect_returnE) | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 610 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 611 | case (Node x ps) | 
| 40671 | 612 | with effect_rev obtain p q h1 h2 h3 v where | 
| 613 | init: "effect (ref Empty) h h1 q" | |
| 614 | "effect (ref (Node x ps)) h1 h2 p" | |
| 615 | and effect_rev':"effect (rev' q p) h2 h3 v" | |
| 616 | and lookup: "effect (!v) h3 h' r'" | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 617 | using rev.simps | 
| 40671 | 618 | by (auto elim!: effect_bindE) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 619 | from init have a1:"list_of' h2 q []" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 620 | unfolding list_of'_def | 
| 40671 | 621 | by (auto elim!: effect_ref) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 622 | from list_of_h obtain refs where refs_def: "refs_of h r refs" by (rule list_of_refs_of) | 
| 37725 | 623 | from validHeap init refs_def have heap_eq: "\<forall>refs. refs_of h r refs \<longrightarrow> (\<forall>ref\<in>set refs. Ref.present h ref \<and> Ref.present h2 ref \<and> Ref.get h ref = Ref.get h2 ref)" | 
| 40671 | 624 | by (fastsimp elim!: effect_ref dest: refs_of_is_fun) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 625 | from list_of_invariant[OF list_of_h heap_eq] have "list_of h2 r xs" . | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 626 | from init this Node have a2: "list_of' h2 p xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 627 | apply - | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 628 | unfolding list_of'_def | 
| 40671 | 629 | apply (auto elim!: effect_refE) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 630 | done | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 631 | from init have refs_of_q: "refs_of' h2 q [q]" | 
| 40671 | 632 | by (auto elim!: effect_ref) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 633 | from refs_def Node have refs_of'_ps: "refs_of' h ps refs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 634 | by (auto simp add: refs_of'_def'[symmetric]) | 
| 37725 | 635 | from validHeap refs_def have all_ref_present: "\<forall>r\<in>set refs. Ref.present h r" by simp | 
| 38410 | 636 | from init refs_of'_ps this | 
| 637 | have heap_eq: "\<forall>refs. refs_of' h ps refs \<longrightarrow> (\<forall>ref\<in>set refs. Ref.present h ref \<and> Ref.present h2 ref \<and> Ref.get h ref = Ref.get h2 ref)" | |
| 40671 | 638 | by (auto elim!: effect_ref [where ?'a="'a node", where ?'b="'a node", where ?'c="'a node"] dest: refs_of'_is_fun) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 639 | from refs_of'_invariant[OF refs_of'_ps this] have "refs_of' h2 ps refs" . | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 640 | with init have refs_of_p: "refs_of' h2 p (p#refs)" | 
| 40671 | 641 | by (auto elim!: effect_refE simp add: refs_of'_def') | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 642 | with init all_ref_present have q_is_new: "q \<notin> set (p#refs)" | 
| 40671 | 643 | by (auto elim!: effect_refE intro!: Ref.noteq_I) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 644 |   from refs_of_p refs_of_q q_is_new have a3: "\<forall>qrs prs. refs_of' h2 q qrs \<and> refs_of' h2 p prs \<longrightarrow> set prs \<inter> set qrs = {}"
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 645 | by (fastsimp simp only: set.simps dest: refs_of'_is_fun) | 
| 40671 | 646 | from rev'_invariant [OF effect_rev' a1 a2 a3] have "list_of h3 (Ref.get h3 v) (List.rev xs)" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 647 | unfolding list_of'_def by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 648 | with lookup show ?thesis | 
| 40671 | 649 | by (auto elim: effect_lookupE) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 650 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 651 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 652 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 653 | section {* The merge function on Linked Lists *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 654 | text {* We also prove merge correct *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 655 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 656 | text{* First, we define merge on lists in a natural way. *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 657 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 658 | fun Lmerge :: "('a::ord) list \<Rightarrow> 'a list \<Rightarrow> 'a list"
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 659 | where | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 660 | "Lmerge (x#xs) (y#ys) = | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 661 | (if x \<le> y then x # Lmerge xs (y#ys) else y # Lmerge (x#xs) ys)" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 662 | | "Lmerge [] ys = ys" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 663 | | "Lmerge xs [] = xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 664 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 665 | subsection {* Definition of merge function *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 666 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 667 | definition merge' :: "(('a::{heap, ord}) node ref * ('a::{heap, ord})) * ('a::{heap, ord}) node ref * ('a::{heap, ord}) node ref \<Rightarrow> ('a::{heap, ord}) node ref Heap"
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 668 | where | 
| 37792 | 669 | "merge' = MREC (\<lambda>(_, p, q). do { v \<leftarrow> !p; w \<leftarrow> !q;
 | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 670 | (case v of Empty \<Rightarrow> return (Inl q) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 671 | | Node valp np \<Rightarrow> | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 672 | (case w of Empty \<Rightarrow> return (Inl p) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 673 | | Node valq nq \<Rightarrow> | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 674 | if (valp \<le> valq) then | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 675 | return (Inr ((p, valp), np, q)) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 676 | else | 
| 37792 | 677 | return (Inr ((q, valq), p, nq)))) }) | 
| 678 |  (\<lambda> _ ((n, v), _, _) r. do { n := Node v r; return n })"
 | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 679 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 680 | definition merge where "merge p q = merge' (undefined, p, q)" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 681 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 682 | lemma if_return: "(if P then return x else return y) = return (if P then x else y)" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 683 | by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 684 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 685 | lemma if_distrib_App: "(if P then f else g) x = (if P then f x else g x)" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 686 | by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 687 | lemma redundant_if: "(if P then (if P then x else z) else y) = (if P then x else y)" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 688 | "(if P then x else (if P then z else y)) = (if P then x else y)" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 689 | by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 690 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 691 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 692 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 693 | lemma sum_distrib: "sum_case fl fr (case x of Empty \<Rightarrow> y | Node v n \<Rightarrow> (z v n)) = (case x of Empty \<Rightarrow> sum_case fl fr y | Node v n \<Rightarrow> sum_case fl fr (z v n))" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 694 | by (cases x) auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 695 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 696 | lemma merge: "merge' (x, p, q) = merge p q" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 697 | unfolding merge'_def merge_def | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 698 | apply (simp add: MREC_rule) done | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 699 | term "Ref.change" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 700 | lemma merge_simps [code]: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 701 | shows "merge p q = | 
| 37792 | 702 | do { v \<leftarrow> !p;
 | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 703 | w \<leftarrow> !q; | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 704 | (case v of node.Empty \<Rightarrow> return q | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 705 | | Node valp np \<Rightarrow> | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 706 | case w of node.Empty \<Rightarrow> return p | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 707 | | Node valq nq \<Rightarrow> | 
| 37792 | 708 |             if valp \<le> valq then do { r \<leftarrow> merge np q;
 | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 709 | p := (Node valp r); | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 710 | return p | 
| 37792 | 711 | } | 
| 712 |             else do { r \<leftarrow> merge p nq;
 | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 713 | q := (Node valq r); | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 714 | return q | 
| 37792 | 715 | }) | 
| 716 | }" | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 717 | proof - | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 718 |   {fix v x y
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 719 | have case_return: "(case v of Empty \<Rightarrow> return x | Node v n \<Rightarrow> return (y v n)) = return (case v of Empty \<Rightarrow> x | Node v n \<Rightarrow> y v n)" by (cases v) auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 720 | } note case_return = this | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 721 | show ?thesis | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 722 | unfolding merge_def[of p q] merge'_def | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 723 | apply (simp add: MREC_rule[of _ _ "(undefined, p, q)"]) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 724 | unfolding bind_bind return_bind | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 725 | unfolding merge'_def[symmetric] | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 726 | unfolding if_return case_return bind_bind return_bind sum_distrib sum.cases | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 727 | unfolding if_distrib[symmetric, where f="Inr"] | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 728 | unfolding sum.cases | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 729 | unfolding if_distrib | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 730 | unfolding split_beta fst_conv snd_conv | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 731 | unfolding if_distrib_App redundant_if merge | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 732 | .. | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 733 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 734 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 735 | subsection {* Induction refinement by applying the abstraction function to our induct rule *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 736 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 737 | text {* From our original induction rule Lmerge.induct, we derive a new rule with our list_of' predicate *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 738 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 739 | lemma merge_induct2: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 740 |   assumes "list_of' h (p::'a::{heap, ord} node ref) xs"
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 741 | assumes "list_of' h q ys" | 
| 37725 | 742 | assumes "\<And> ys p q. \<lbrakk> list_of' h p []; list_of' h q ys; Ref.get h p = Empty \<rbrakk> \<Longrightarrow> P p q [] ys" | 
| 743 | assumes "\<And> x xs' p q pn. \<lbrakk> list_of' h p (x#xs'); list_of' h q []; Ref.get h p = Node x pn; Ref.get h q = Empty \<rbrakk> \<Longrightarrow> P p q (x#xs') []" | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 744 | assumes "\<And> x xs' y ys' p q pn qn. | 
| 37725 | 745 | \<lbrakk> 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; | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 746 | x \<le> y; P pn q xs' (y#ys') \<rbrakk> | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 747 | \<Longrightarrow> P p q (x#xs') (y#ys')" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 748 | assumes "\<And> x xs' y ys' p q pn qn. | 
| 37725 | 749 | \<lbrakk> 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; | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 750 | \<not> x \<le> y; P p qn (x#xs') ys'\<rbrakk> | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 751 | \<Longrightarrow> P p q (x#xs') (y#ys')" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 752 | shows "P p q xs ys" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 753 | using assms(1-2) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 754 | proof (induct xs ys arbitrary: p q rule: Lmerge.induct) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 755 | case (2 ys) | 
| 37725 | 756 | from 2(1) have "Ref.get h p = Empty" unfolding list_of'_def by simp | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 757 | with 2(1-2) assms(3) show ?case by blast | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 758 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 759 | case (3 x xs') | 
| 37725 | 760 | from 3(1) obtain pn where Node: "Ref.get h p = Node x pn" by (rule list_of'_Cons) | 
| 761 | from 3(2) have "Ref.get h q = Empty" unfolding list_of'_def by simp | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 762 | with Node 3(1-2) assms(4) show ?case by blast | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 763 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 764 | case (1 x xs' y ys') | 
| 37725 | 765 | from 1(3) obtain pn where pNode:"Ref.get h p = Node x pn" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 766 | and list_of'_pn: "list_of' h pn xs'" by (rule list_of'_Cons) | 
| 37725 | 767 | from 1(4) obtain qn where qNode:"Ref.get h q = Node y qn" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 768 | and list_of'_qn: "list_of' h qn ys'" by (rule list_of'_Cons) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 769 | show ?case | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 770 | proof (cases "x \<le> y") | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 771 | case True | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 772 | from 1(1)[OF True list_of'_pn 1(4)] assms(5) 1(3-4) pNode qNode True | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 773 | show ?thesis by blast | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 774 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 775 | case False | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 776 | from 1(2)[OF False 1(3) list_of'_qn] assms(6) 1(3-4) pNode qNode False | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 777 | show ?thesis by blast | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 778 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 779 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 780 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 781 | |
| 40671 | 782 | text {* secondly, we add the effect statement in the premise, and derive the effect statements for the single cases which we then eliminate with our effect elim rules. *}
 | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 783 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 784 | lemma merge_induct3: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 785 | assumes "list_of' h p xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 786 | assumes "list_of' h q ys" | 
| 40671 | 787 | assumes "effect (merge p q) h h' r" | 
| 37725 | 788 | assumes "\<And> ys p q. \<lbrakk> list_of' h p []; list_of' h q ys; Ref.get h p = Empty \<rbrakk> \<Longrightarrow> P p q h h q [] ys" | 
| 789 | assumes "\<And> x xs' p q pn. \<lbrakk> list_of' h p (x#xs'); list_of' h q []; Ref.get h p = Node x pn; Ref.get h q = Empty \<rbrakk> \<Longrightarrow> P p q h h p (x#xs') []" | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 790 | assumes "\<And> x xs' y ys' p q pn qn h1 r1 h'. | 
| 37725 | 791 | \<lbrakk> 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; | 
| 40671 | 792 | x \<le> y; effect (merge pn q) h h1 r1 ; P pn q h h1 r1 xs' (y#ys'); h' = Ref.set p (Node x r1) h1 \<rbrakk> | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 793 | \<Longrightarrow> P p q h h' p (x#xs') (y#ys')" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 794 | assumes "\<And> x xs' y ys' p q pn qn h1 r1 h'. | 
| 37725 | 795 | \<lbrakk> 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; | 
| 40671 | 796 | \<not> x \<le> y; effect (merge p qn) h h1 r1; P p qn h h1 r1 (x#xs') ys'; h' = Ref.set q (Node y r1) h1 \<rbrakk> | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 797 | \<Longrightarrow> P p q h h' q (x#xs') (y#ys')" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 798 | shows "P p q h h' r xs ys" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 799 | using assms(3) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 800 | proof (induct arbitrary: h' r rule: merge_induct2[OF assms(1) assms(2)]) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 801 | case (1 ys p q) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 802 | from 1(3-4) have "h = h' \<and> r = q" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 803 | unfolding merge_simps[of p q] | 
| 40671 | 804 | by (auto elim!: effect_lookupE effect_bindE effect_returnE) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 805 | with assms(4)[OF 1(1) 1(2) 1(3)] show ?case by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 806 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 807 | case (2 x xs' p q pn) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 808 | from 2(3-5) have "h = h' \<and> r = p" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 809 | unfolding merge_simps[of p q] | 
| 40671 | 810 | by (auto elim!: effect_lookupE effect_bindE effect_returnE) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 811 | with assms(5)[OF 2(1-4)] show ?case by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 812 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 813 | case (3 x xs' y ys' p q pn qn) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 814 | from 3(3-5) 3(7) obtain h1 r1 where | 
| 40671 | 815 | 1: "effect (merge pn q) h h1 r1" | 
| 37725 | 816 | and 2: "h' = Ref.set p (Node x r1) h1 \<and> r = p" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 817 | unfolding merge_simps[of p q] | 
| 40671 | 818 | by (auto elim!: effect_lookupE effect_bindE effect_returnE effect_ifE effect_updateE) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 819 | from 3(6)[OF 1] assms(6) [OF 3(1-5)] 1 2 show ?case by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 820 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 821 | case (4 x xs' y ys' p q pn qn) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 822 | from 4(3-5) 4(7) obtain h1 r1 where | 
| 40671 | 823 | 1: "effect (merge p qn) h h1 r1" | 
| 37725 | 824 | and 2: "h' = Ref.set q (Node y r1) h1 \<and> r = q" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 825 | unfolding merge_simps[of p q] | 
| 40671 | 826 | by (auto elim!: effect_lookupE effect_bindE effect_returnE effect_ifE effect_updateE) | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 827 | from 4(6)[OF 1] assms(7) [OF 4(1-5)] 1 2 show ?case by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 828 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 829 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 830 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 831 | subsection {* Proving merge correct *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 832 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 833 | text {* As many parts of the following three proofs are identical, we could actually move the
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 834 | same reasoning into an extended induction rule *} | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 835 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 836 | lemma merge_unchanged: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 837 | assumes "refs_of' h p xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 838 | assumes "refs_of' h q ys" | 
| 40671 | 839 | assumes "effect (merge p q) h h' r'" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 840 |   assumes "set xs \<inter> set ys = {}"
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 841 | assumes "r \<notin> set xs \<union> set ys" | 
| 37725 | 842 | shows "Ref.get h r = Ref.get h' r" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 843 | proof - | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 844 | from assms(1) obtain ps where ps_def: "list_of' h p ps" by (rule refs_of'_list_of') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 845 | from assms(2) obtain qs where qs_def: "list_of' h q qs" by (rule refs_of'_list_of') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 846 | show ?thesis using assms(1) assms(2) assms(4) assms(5) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 847 | proof (induct arbitrary: xs ys r rule: merge_induct3[OF ps_def qs_def assms(3)]) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 848 | case 1 thus ?case by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 849 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 850 | case 2 thus ?case by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 851 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 852 | case (3 x xs' y ys' p q pn qn h1 r1 h' xs ys r) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 853 | from 3(9) 3(3) obtain pnrs | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 854 | where pnrs_def: "xs = p#pnrs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 855 | and refs_of'_pn: "refs_of' h pn pnrs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 856 | by (rule refs_of'_Node) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 857 | with 3(12) have r_in: "r \<notin> set pnrs \<union> set ys" by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 858 | from pnrs_def 3(12) have "r \<noteq> p" by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 859 | with 3(11) 3(12) pnrs_def refs_of'_distinct[OF 3(9)] have p_in: "p \<notin> set pnrs \<union> set ys" by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 860 |     from 3(11) pnrs_def have no_inter: "set pnrs \<inter> set ys = {}" by auto
 | 
| 37725 | 861 | 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" | 
| 862 | by simp | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 863 | from 3(7)[OF refs_of'_pn 3(10) no_inter r_in] 3(8) `r \<noteq> p` show ?case | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 864 | by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 865 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 866 | case (4 x xs' y ys' p q pn qn h1 r1 h' xs ys r) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 867 | from 4(10) 4(4) obtain qnrs | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 868 | where qnrs_def: "ys = q#qnrs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 869 | and refs_of'_qn: "refs_of' h qn qnrs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 870 | by (rule refs_of'_Node) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 871 | with 4(12) have r_in: "r \<notin> set xs \<union> set qnrs" by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 872 | from qnrs_def 4(12) have "r \<noteq> q" by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 873 | with 4(11) 4(12) qnrs_def refs_of'_distinct[OF 4(10)] have q_in: "q \<notin> set xs \<union> set qnrs" by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 874 |     from 4(11) qnrs_def have no_inter: "set xs \<inter> set qnrs = {}" by auto
 | 
| 37725 | 875 | 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 | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 876 | from 4(7)[OF 4(9) refs_of'_qn no_inter r_in] 4(8) `r \<noteq> q` show ?case | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 877 | by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 878 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 879 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 880 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 881 | lemma refs_of'_merge: | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 882 | assumes "refs_of' h p xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 883 | assumes "refs_of' h q ys" | 
| 40671 | 884 | assumes "effect (merge p q) h h' r" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 885 |   assumes "set xs \<inter> set ys = {}"
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 886 | assumes "refs_of' h' r rs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 887 | shows "set rs \<subseteq> set xs \<union> set ys" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 888 | proof - | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 889 | from assms(1) obtain ps where ps_def: "list_of' h p ps" by (rule refs_of'_list_of') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 890 | from assms(2) obtain qs where qs_def: "list_of' h q qs" by (rule refs_of'_list_of') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 891 | show ?thesis using assms(1) assms(2) assms(4) assms(5) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 892 | proof (induct arbitrary: xs ys rs rule: merge_induct3[OF ps_def qs_def assms(3)]) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 893 | case 1 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 894 | from 1(5) 1(7) have "rs = ys" by (fastsimp simp add: refs_of'_is_fun) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 895 | thus ?case by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 896 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 897 | case 2 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 898 | from 2(5) 2(8) have "rs = xs" by (auto simp add: refs_of'_is_fun) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 899 | thus ?case by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 900 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 901 | case (3 x xs' y ys' p q pn qn h1 r1 h' xs ys rs) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 902 | from 3(9) 3(3) obtain pnrs | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 903 | where pnrs_def: "xs = p#pnrs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 904 | and refs_of'_pn: "refs_of' h pn pnrs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 905 | by (rule refs_of'_Node) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 906 | from 3(10) 3(9) 3(11) pnrs_def refs_of'_distinct[OF 3(9)] have p_in: "p \<notin> set pnrs \<union> set ys" by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 907 |     from 3(11) pnrs_def have no_inter: "set pnrs \<inter> set ys = {}" by auto
 | 
| 37725 | 908 | 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" .. | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 909 | from 3 p_stays obtain r1s | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 910 | where rs_def: "rs = p#r1s" and refs_of'_r1:"refs_of' h1 r1 r1s" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 911 | by (auto elim: refs_of'_set_next_ref) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 912 | from 3(7)[OF refs_of'_pn 3(10) no_inter refs_of'_r1] rs_def pnrs_def show ?case by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 913 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 914 | case (4 x xs' y ys' p q pn qn h1 r1 h' xs ys rs) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 915 | from 4(10) 4(4) obtain qnrs | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 916 | where qnrs_def: "ys = q#qnrs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 917 | and refs_of'_qn: "refs_of' h qn qnrs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 918 | by (rule refs_of'_Node) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 919 | from 4(10) 4(9) 4(11) qnrs_def refs_of'_distinct[OF 4(10)] have q_in: "q \<notin> set xs \<union> set qnrs" by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 920 |     from 4(11) qnrs_def have no_inter: "set xs \<inter> set qnrs = {}" by auto
 | 
| 37725 | 921 | 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" .. | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 922 | from 4 q_stays obtain r1s | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 923 | where rs_def: "rs = q#r1s" and refs_of'_r1:"refs_of' h1 r1 r1s" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 924 | by (auto elim: refs_of'_set_next_ref) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 925 | from 4(7)[OF 4(9) refs_of'_qn no_inter refs_of'_r1] rs_def qnrs_def show ?case by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 926 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 927 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 928 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 929 | lemma | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 930 | assumes "list_of' h p xs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 931 | assumes "list_of' h q ys" | 
| 40671 | 932 | assumes "effect (merge p q) h h' r" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 933 |   assumes "\<forall>qrs prs. refs_of' h q qrs \<and> refs_of' h p prs \<longrightarrow> set prs \<inter> set qrs = {}"
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 934 | shows "list_of' h' r (Lmerge xs ys)" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 935 | using assms(4) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 936 | proof (induct rule: merge_induct3[OF assms(1-3)]) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 937 | case 1 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 938 | thus ?case by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 939 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 940 | case 2 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 941 | thus ?case by simp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 942 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 943 | case (3 x xs' y ys' p q pn qn h1 r1 h') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 944 | from 3(1) obtain prs where prs_def: "refs_of' h p prs" by (rule list_of'_refs_of') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 945 | from 3(2) obtain qrs where qrs_def: "refs_of' h q qrs" by (rule list_of'_refs_of') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 946 | from prs_def 3(3) obtain pnrs | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 947 | where pnrs_def: "prs = p#pnrs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 948 | and refs_of'_pn: "refs_of' h pn pnrs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 949 | by (rule refs_of'_Node) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 950 | from prs_def qrs_def 3(9) pnrs_def refs_of'_distinct[OF prs_def] have p_in: "p \<notin> set pnrs \<union> set qrs" by fastsimp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 951 |   from prs_def qrs_def 3(9) pnrs_def have no_inter: "set pnrs \<inter> set qrs = {}" by fastsimp
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 952 |   from no_inter refs_of'_pn qrs_def have no_inter2: "\<forall>qrs prs. refs_of' h q qrs \<and> refs_of' h pn prs \<longrightarrow> set prs \<inter> set qrs = {}"
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 953 | by (fastsimp dest: refs_of'_is_fun) | 
| 37725 | 954 | 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" .. | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 955 | from 3(7)[OF no_inter2] obtain rs where rs_def: "refs_of' h1 r1 rs" by (rule list_of'_refs_of') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 956 | from refs_of'_merge[OF refs_of'_pn qrs_def 3(6) no_inter this] p_in have p_rs: "p \<notin> set rs" by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 957 | with 3(7)[OF no_inter2] 3(1-5) 3(8) p_rs rs_def p_stays | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 958 | show ?case by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 959 | next | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 960 | case (4 x xs' y ys' p q pn qn h1 r1 h') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 961 | from 4(1) obtain prs where prs_def: "refs_of' h p prs" by (rule list_of'_refs_of') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 962 | from 4(2) obtain qrs where qrs_def: "refs_of' h q qrs" by (rule list_of'_refs_of') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 963 | from qrs_def 4(4) obtain qnrs | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 964 | where qnrs_def: "qrs = q#qnrs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 965 | and refs_of'_qn: "refs_of' h qn qnrs" | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 966 | by (rule refs_of'_Node) | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 967 | from prs_def qrs_def 4(9) qnrs_def refs_of'_distinct[OF qrs_def] have q_in: "q \<notin> set prs \<union> set qnrs" by fastsimp | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 968 |   from prs_def qrs_def 4(9) qnrs_def have no_inter: "set prs \<inter> set qnrs = {}" by fastsimp
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 969 |   from no_inter refs_of'_qn prs_def have no_inter2: "\<forall>qrs prs. refs_of' h qn qrs \<and> refs_of' h p prs \<longrightarrow> set prs \<inter> set qrs = {}"
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 970 | by (fastsimp dest: refs_of'_is_fun) | 
| 37725 | 971 | 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" .. | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 972 | from 4(7)[OF no_inter2] obtain rs where rs_def: "refs_of' h1 r1 rs" by (rule list_of'_refs_of') | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 973 | from refs_of'_merge[OF prs_def refs_of'_qn 4(6) no_inter this] q_in have q_rs: "q \<notin> set rs" by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 974 | with 4(7)[OF no_inter2] 4(1-5) 4(8) q_rs rs_def q_stays | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 975 | show ?case by auto | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 976 | qed | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 977 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 978 | section {* Code generation *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 979 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 980 | text {* A simple example program *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 981 | |
| 37792 | 982 | definition test_1 where "test_1 = (do { ll_xs <- make_llist [1..(15::int)]; xs <- traverse ll_xs; return xs })" 
 | 
| 983 | 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 })"
 | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 984 | |
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 985 | definition test_3 where "test_3 = | 
| 37792 | 986 |   (do {
 | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 987 | ll_xs \<leftarrow> make_llist (filter (%n. n mod 2 = 0) [2..8]); | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 988 | ll_ys \<leftarrow> make_llist (filter (%n. n mod 2 = 1) [5..11]); | 
| 37725 | 989 | r \<leftarrow> ref ll_xs; | 
| 990 | q \<leftarrow> ref ll_ys; | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 991 | p \<leftarrow> merge r q; | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 992 | ll_zs \<leftarrow> !p; | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 993 | zs \<leftarrow> traverse ll_zs; | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 994 | return zs | 
| 37792 | 995 | })" | 
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 996 | |
| 35041 
6eb917794a5c
avoid upto in generated code (is infix operator in library.ML)
 haftmann parents: 
34051diff
changeset | 997 | code_reserved SML upto | 
| 
6eb917794a5c
avoid upto in generated code (is infix operator in library.ML)
 haftmann parents: 
34051diff
changeset | 998 | |
| 34051 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 999 | ML {* @{code test_1} () *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 1000 | ML {* @{code test_2} () *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 1001 | ML {* @{code test_3} () *}
 | 
| 
1a82e2e29d67
added Imperative_HOL examples; added tail-recursive combinator for monadic heap functions; adopted code generation of references; added lemmas
 bulwahn parents: diff
changeset | 1002 | |
| 38068 
00042fd999a8
intermediate operation avoids invariance problem in Scala
 haftmann parents: 
38058diff
changeset | 1003 | export_code test_1 test_2 test_3 checking SML SML_imp OCaml? OCaml_imp? Haskell? Scala? Scala_imp? | 
| 37750 
82e0fe8b07eb
dropped ancient in-place compilation of SML; more tests
 haftmann parents: 
37725diff
changeset | 1004 | |
| 37725 | 1005 | end |