# HG changeset patch # User haftmann # Date 1278333300 -7200 # Node ID 3f499df0751f71d7665fc3d017a845e6bd03916e # Parent 694815d76240030122f153d43fc2de7c6d61a33c# Parent 70fafefbcc9802a1e54e60571a9cd69c28a2cbb2 merged diff -r 694815d76240 -r 3f499df0751f src/HOL/Imperative_HOL/Array.thy --- a/src/HOL/Imperative_HOL/Array.thy Mon Jul 05 11:25:06 2010 +0200 +++ b/src/HOL/Imperative_HOL/Array.thy Mon Jul 05 14:35:00 2010 +0200 @@ -28,7 +28,7 @@ [code del]: "nth a i = (do len \ length a; (if i < len then Heap_Monad.heap (\h. (get_array a h ! i, h)) - else raise (''array lookup: index out of range'')) + else raise ''array lookup: index out of range'') done)" definition @@ -37,18 +37,12 @@ [code del]: "upd i x a = (do len \ length a; (if i < len then Heap_Monad.heap (\h. (a, Heap.upd a i x h)) - else raise (''array update: index out of range'')) + else raise ''array update: index out of range'') done)" lemma upd_return: "upd i x a \ return a = upd i x a" -proof (rule Heap_eqI) - fix h - obtain len h' where "Heap_Monad.execute (Array.length a) h = (len, h')" - by (cases "Heap_Monad.execute (Array.length a) h") - then show "Heap_Monad.execute (upd i x a \ return a) h = Heap_Monad.execute (upd i x a) h" - by (auto simp add: upd_def bindM_def split: sum.split) -qed + by (rule Heap_eqI) (simp add: upd_def bindM_def split: option.split) subsection {* Derivates *} @@ -99,14 +93,11 @@ lemma array_make [code]: "Array.new n x = make n (\_. x)" - by (induct n) (simp_all add: make_def new_def Heap_Monad.heap_def - monad_simp array_of_list_replicate [symmetric] - map_replicate_trivial replicate_append_same - of_list_def) + by (rule Heap_eqI) (simp add: make_def new_def array_of_list_replicate map_replicate_trivial of_list_def) lemma array_of_list_make [code]: "of_list xs = make (List.length xs) (\n. xs ! n)" - unfolding make_def map_nth .. + by (rule Heap_eqI) (simp add: make_def map_nth) subsection {* Code generator setup *} @@ -135,13 +126,11 @@ by (simp add: make'_def o_def) definition length' where - [code del]: "length' = Array.length \== liftM Code_Numeral.of_nat" + [code del]: "length' a = Array.length a \= (\n. return (Code_Numeral.of_nat n))" hide_const (open) length' lemma [code]: - "Array.length = Array.length' \== liftM Code_Numeral.nat_of" - by (simp add: length'_def monad_simp', - simp add: liftM_def comp_def monad_simp, - simp add: monad_simp') + "Array.length a = Array.length' a \= (\i. return (Code_Numeral.nat_of i))" + by (simp add: length'_def) definition nth' where [code del]: "nth' a = Array.nth a o Code_Numeral.nat_of" @@ -155,7 +144,7 @@ hide_const (open) upd' lemma [code]: "Array.upd i x a = Array.upd' a (Code_Numeral.of_nat i) x \ return a" - by (simp add: upd'_def monad_simp upd_return) + by (simp add: upd'_def upd_return) subsubsection {* SML *} diff -r 694815d76240 -r 3f499df0751f src/HOL/Imperative_HOL/Heap_Monad.thy --- a/src/HOL/Imperative_HOL/Heap_Monad.thy Mon Jul 05 11:25:06 2010 +0200 +++ b/src/HOL/Imperative_HOL/Heap_Monad.thy Mon Jul 05 14:35:00 2010 +0200 @@ -12,16 +12,12 @@ subsubsection {* Monad combinators *} -datatype exception = Exn - text {* Monadic heap actions either produce values and transform the heap, or fail *} -datatype 'a Heap = Heap "heap \ ('a + exception) \ heap" +datatype 'a Heap = Heap "heap \ ('a \ heap) option" -primrec - execute :: "'a Heap \ heap \ ('a + exception) \ heap" where - "execute (Heap f) = f" -lemmas [code del] = execute.simps +primrec execute :: "'a Heap \ heap \ ('a \ heap) option" where + [code del]: "execute (Heap f) = f" lemma Heap_execute [simp]: "Heap (execute f) = f" by (cases f) simp_all @@ -34,58 +30,67 @@ "(\h. (\x. execute (f x) h) = (\y. execute (g y) h)) \ f = g" by (auto simp: expand_fun_eq intro: Heap_eqI) -lemma Heap_strip: "(\f. PROP P f) \ (\g. PROP P (Heap g))" -proof - fix g :: "heap \ ('a + exception) \ heap" - assume "\f. PROP P f" - then show "PROP P (Heap g)" . -next - fix f :: "'a Heap" - assume assm: "\g. PROP P (Heap g)" - then have "PROP P (Heap (execute f))" . - then show "PROP P f" by simp -qed - -definition - heap :: "(heap \ 'a \ heap) \ 'a Heap" where - [code del]: "heap f = Heap (\h. apfst Inl (f h))" +definition heap :: "(heap \ 'a \ heap) \ 'a Heap" where + [code del]: "heap f = Heap (Some \ f)" lemma execute_heap [simp]: - "execute (heap f) h = apfst Inl (f h)" + "execute (heap f) = Some \ f" by (simp add: heap_def) -definition - bindM :: "'a Heap \ ('a \ 'b Heap) \ 'b Heap" (infixl ">>=" 54) where - [code del]: "f >>= g = Heap (\h. case execute f h of - (Inl x, h') \ execute (g x) h' - | r \ r)" +lemma heap_cases [case_names succeed fail]: + fixes f and h + assumes succeed: "\x h'. execute f h = Some (x, h') \ P" + assumes fail: "execute f h = None \ P" + shows P + using assms by (cases "execute f h") auto -notation - bindM (infixl "\=" 54) - -abbreviation - chainM :: "'a Heap \ 'b Heap \ 'b Heap" (infixl ">>" 54) where - "f >> g \ f >>= (\_. g)" - -notation - chainM (infixl "\" 54) - -definition - return :: "'a \ 'a Heap" where +definition return :: "'a \ 'a Heap" where [code del]: "return x = heap (Pair x)" lemma execute_return [simp]: - "execute (return x) h = apfst Inl (x, h)" + "execute (return x) = Some \ Pair x" by (simp add: return_def) -definition - raise :: "string \ 'a Heap" where -- {* the string is just decoration *} - [code del]: "raise s = Heap (Pair (Inr Exn))" +definition raise :: "string \ 'a Heap" where -- {* the string is just decoration *} + [code del]: "raise s = Heap (\_. None)" lemma execute_raise [simp]: - "execute (raise s) h = (Inr Exn, h)" + "execute (raise s) = (\_. None)" by (simp add: raise_def) +definition bindM :: "'a Heap \ ('a \ 'b Heap) \ 'b Heap" (infixl ">>=" 54) where + [code del]: "f >>= g = Heap (\h. case execute f h of + Some (x, h') \ execute (g x) h' + | None \ None)" + +notation bindM (infixl "\=" 54) + +lemma execute_bind [simp]: + "execute f h = Some (x, h') \ execute (f \= g) h = execute (g x) h'" + "execute f h = None \ execute (f \= g) h = None" + by (simp_all add: bindM_def) + +lemma execute_bind_heap [simp]: + "execute (heap f \= g) h = execute (g (fst (f h))) (snd (f h))" + by (simp add: bindM_def split_def) + +lemma return_bind [simp]: "return x \= f = f x" + by (rule Heap_eqI) simp + +lemma bind_return [simp]: "f \= return = f" + by (rule Heap_eqI) (simp add: bindM_def split: option.splits) + +lemma bind_bind [simp]: "(f \= g) \= k = f \= (\x. g x \= k)" + by (rule Heap_eqI) (simp add: bindM_def split: option.splits) + +lemma raise_bind [simp]: "raise e \= f = raise e" + by (rule Heap_eqI) simp + +abbreviation chainM :: "'a Heap \ 'b Heap \ 'b Heap" (infixl ">>" 54) where + "f >> g \ f >>= (\_. g)" + +notation chainM (infixl "\" 54) + subsubsection {* do-syntax *} @@ -170,88 +175,10 @@ subsection {* Monad properties *} -subsubsection {* Monad laws *} - -lemma return_bind: "return x \= f = f x" - by (simp add: bindM_def return_def) - -lemma bind_return: "f \= return = f" -proof (rule Heap_eqI) - fix h - show "execute (f \= return) h = execute f h" - by (auto simp add: bindM_def return_def split: sum.splits prod.splits) -qed - -lemma bind_bind: "(f \= g) \= h = f \= (\x. g x \= h)" - by (rule Heap_eqI) (auto simp add: bindM_def split: split: sum.splits prod.splits) - -lemma bind_bind': "f \= (\x. g x \= h x) = f \= (\x. g x \= (\y. return (x, y))) \= (\(x, y). h x y)" - by (rule Heap_eqI) (auto simp add: bindM_def split: split: sum.splits prod.splits) - -lemma raise_bind: "raise e \= f = raise e" - by (simp add: raise_def bindM_def) - - -lemmas monad_simp = return_bind bind_return bind_bind raise_bind - - subsection {* Generic combinators *} -definition - liftM :: "('a \ 'b) \ 'a \ 'b Heap" -where - "liftM f = return o f" - -definition - compM :: "('a \ 'b Heap) \ ('b \ 'c Heap) \ 'a \ 'c Heap" (infixl ">>==" 54) -where - "(f >>== g) = (\x. f x \= g)" - -notation - compM (infixl "\==" 54) - -lemma liftM_collapse: "liftM f x = return (f x)" - by (simp add: liftM_def) - -lemma liftM_compM: "liftM f \== g = g o f" - by (auto intro: Heap_eqI' simp add: expand_fun_eq liftM_def compM_def bindM_def) - -lemma compM_return: "f \== return = f" - by (simp add: compM_def monad_simp) - -lemma compM_compM: "(f \== g) \== h = f \== (g \== h)" - by (simp add: compM_def monad_simp) - -lemma liftM_bind: - "(\x. liftM f x \= liftM g) = liftM (\x. g (f x))" - by (rule Heap_eqI') (simp add: monad_simp liftM_def bindM_def) - -lemma liftM_comp: - "liftM f o g = liftM (f o g)" - by (rule Heap_eqI') (simp add: liftM_def) - -lemmas monad_simp' = monad_simp liftM_compM compM_return - compM_compM liftM_bind liftM_comp - -primrec - mapM :: "('a \ 'b Heap) \ 'a list \ 'b list Heap" -where - "mapM f [] = return []" - | "mapM f (x#xs) = do y \ f x; - ys \ mapM f xs; - return (y # ys) - done" - -primrec - foldM :: "('a \ 'b \ 'b Heap) \ 'a list \ 'b \ 'b Heap" -where - "foldM f [] s = return s" - | "foldM f (x#xs) s = f x s \= foldM f xs" - -definition - assert :: "('a \ bool) \ 'a \ 'a Heap" -where - "assert P x = (if P x then return x else raise (''assert''))" +definition assert :: "('a \ bool) \ 'a \ 'a Heap" where + "assert P x = (if P x then return x else raise ''assert'')" lemma assert_cong [fundef_cong]: assumes "P = P'" @@ -259,28 +186,42 @@ shows "(assert P x >>= f) = (assert P' x >>= f')" using assms by (auto simp add: assert_def return_bind raise_bind) +definition liftM :: "('a \ 'b) \ 'a \ 'b Heap" where + "liftM f = return o f" + +lemma liftM_collapse [simp]: + "liftM f x = return (f x)" + by (simp add: liftM_def) + +lemma bind_liftM: + "(f \= liftM g) = (f \= (\x. return (g x)))" + by (simp add: liftM_def comp_def) + +primrec mapM :: "('a \ 'b Heap) \ 'a list \ 'b list Heap" where + "mapM f [] = return []" +| "mapM f (x#xs) = do + y \ f x; + ys \ mapM f xs; + return (y # ys) + done" + + subsubsection {* A monadic combinator for simple recursive functions *} text {* Using a locale to fix arguments f and g of MREC *} locale mrec = -fixes - f :: "'a => ('b + 'a) Heap" - and g :: "'a => 'a => 'b => 'b Heap" + fixes f :: "'a \ ('b + 'a) Heap" + and g :: "'a \ 'a \ 'b \ 'b Heap" begin -function (default "\(x,h). (Inr Exn, undefined)") - mrec -where - "mrec x h = - (case Heap_Monad.execute (f x) h of - (Inl (Inl r), h') \ (Inl r, h') - | (Inl (Inr s), h') \ - (case mrec s h' of - (Inl z, h'') \ Heap_Monad.execute (g x s z) h'' - | (Inr e, h'') \ (Inr e, h'')) - | (Inr e, h') \ (Inr e, h') - )" +function (default "\(x, h). None") mrec :: "'a \ heap \ ('b \ heap) option" where + "mrec x h = (case execute (f x) h of + Some (Inl r, h') \ Some (r, h') + | Some (Inr s, h') \ (case mrec s h' of + Some (z, h'') \ execute (g x s z) h'' + | None \ None) + | None \ None)" by auto lemma graph_implies_dom: @@ -290,38 +231,37 @@ apply (erule mrec_rel.cases) by simp -lemma mrec_default: "\ mrec_dom (x, h) \ mrec x h = (Inr Exn, undefined)" +lemma mrec_default: "\ mrec_dom (x, h) \ mrec x h = None" unfolding mrec_def by (rule fundef_default_value[OF mrec_sumC_def graph_implies_dom, of _ _ "(x, h)", simplified]) lemma mrec_di_reverse: assumes "\ mrec_dom (x, h)" shows " - (case Heap_Monad.execute (f x) h of - (Inl (Inl r), h') \ False - | (Inl (Inr s), h') \ \ mrec_dom (s, h') - | (Inr e, h') \ False + (case execute (f x) h of + Some (Inl r, h') \ False + | Some (Inr s, h') \ \ mrec_dom (s, h') + | None \ False )" -using assms -by (auto split:prod.splits sum.splits) - (erule notE, rule accpI, elim mrec_rel.cases, simp)+ - +using assms apply (auto split: option.split sum.split) +apply (rule ccontr) +apply (erule notE, rule accpI, elim mrec_rel.cases, auto)+ +done lemma mrec_rule: "mrec x h = - (case Heap_Monad.execute (f x) h of - (Inl (Inl r), h') \ (Inl r, h') - | (Inl (Inr s), h') \ + (case execute (f x) h of + Some (Inl r, h') \ Some (r, h') + | Some (Inr s, h') \ (case mrec s h' of - (Inl z, h'') \ Heap_Monad.execute (g x s z) h'' - | (Inr e, h'') \ (Inr e, h'')) - | (Inr e, h') \ (Inr e, h') + Some (z, h'') \ execute (g x s z) h'' + | None \ None) + | None \ None )" apply (cases "mrec_dom (x,h)", simp) apply (frule mrec_default) apply (frule mrec_di_reverse, simp) -by (auto split: sum.split prod.split simp: mrec_default) - +by (auto split: sum.split option.split simp: mrec_default) definition "MREC x = Heap (mrec x)" @@ -340,32 +280,31 @@ apply simp apply (rule ext) apply (unfold mrec_rule[of x]) - by (auto split:prod.splits sum.splits) - + by (auto split: option.splits prod.splits sum.splits) lemma MREC_pinduct: - assumes "Heap_Monad.execute (MREC x) h = (Inl r, h')" - assumes non_rec_case: "\ x h h' r. Heap_Monad.execute (f x) h = (Inl (Inl r), h') \ P x h h' r" - assumes rec_case: "\ x h h1 h2 h' s z r. Heap_Monad.execute (f x) h = (Inl (Inr s), h1) \ Heap_Monad.execute (MREC s) h1 = (Inl z, h2) \ P s h1 h2 z - \ Heap_Monad.execute (g x s z) h2 = (Inl r, h') \ P x h h' r" + assumes "execute (MREC x) h = Some (r, h')" + assumes non_rec_case: "\ x h h' r. execute (f x) h = Some (Inl r, h') \ P x h h' r" + assumes rec_case: "\ x h h1 h2 h' s z r. execute (f x) h = Some (Inr s, h1) \ execute (MREC s) h1 = Some (z, h2) \ P s h1 h2 z + \ execute (g x s z) h2 = Some (r, h') \ P x h h' r" shows "P x h h' r" proof - - from assms(1) have mrec: "mrec x h = (Inl r, h')" + from assms(1) have mrec: "mrec x h = Some (r, h')" unfolding MREC_def execute.simps . from mrec have dom: "mrec_dom (x, h)" apply - apply (rule ccontr) apply (drule mrec_default) by auto - from mrec have h'_r: "h' = (snd (mrec x h))" "r = (Sum_Type.Projl (fst (mrec x h)))" + from mrec have h'_r: "h' = snd (the (mrec x h))" "r = fst (the (mrec x h))" by auto - from mrec have "P x h (snd (mrec x h)) (Sum_Type.Projl (fst (mrec x h)))" + from mrec have "P x h (snd (the (mrec x h))) (fst (the (mrec x h)))" proof (induct arbitrary: r h' rule: mrec.pinduct[OF dom]) case (1 x h) - obtain rr h' where "mrec x h = (rr, h')" by fastsimp - obtain fret h1 where exec_f: "Heap_Monad.execute (f x) h = (fret, h1)" by fastsimp + obtain rr h' where "the (mrec x h) = (rr, h')" by fastsimp show ?case - proof (cases fret) - case (Inl a) + proof (cases "execute (f x) h") + case (Some result) + then obtain a h1 where exec_f: "execute (f x) h = Some (a, h1)" by fastsimp note Inl' = this show ?thesis proof (cases a) @@ -375,23 +314,28 @@ next case (Inr b) note Inr' = this - obtain ret_mrec h2 where mrec_rec: "mrec b h1 = (ret_mrec, h2)" by fastsimp - from this Inl 1(1) exec_f mrec show ?thesis - proof (cases "ret_mrec") - case (Inl aaa) - from this mrec exec_f Inl' Inr' 1(1) mrec_rec 1(2) [OF exec_f [symmetric] Inl' Inr', of "aaa" "h2"] 1(3) - show ?thesis - apply auto - apply (rule rec_case) - unfolding MREC_def by auto + show ?thesis + proof (cases "mrec b h1") + case (Some result) + then obtain aaa h2 where mrec_rec: "mrec b h1 = Some (aaa, h2)" by fastsimp + moreover from this have "P b h1 (snd (the (mrec b h1))) (fst (the (mrec b h1)))" + apply (intro 1(2)) + apply (auto simp add: Inr Inl') + done + moreover note mrec mrec_rec exec_f Inl' Inr' 1(1) 1(3) + ultimately show ?thesis + apply auto + apply (rule rec_case) + apply auto + unfolding MREC_def by auto next - case (Inr b) - from this Inl 1(1) exec_f mrec Inr' mrec_rec 1(3) show ?thesis by auto + case None + from this 1(1) exec_f mrec Inr' 1(3) show ?thesis by auto qed qed next - case (Inr b) - from this 1(1) mrec exec_f 1(3) show ?thesis by simp + case None + from this 1(1) mrec 1(3) show ?thesis by simp qed qed from this h'_r show ?thesis by simp @@ -412,38 +356,29 @@ subsubsection {* Logical intermediate layer *} -definition - Fail :: "String.literal \ exception" -where - [code del]: "Fail s = Exn" +primrec raise' :: "String.literal \ 'a Heap" where + [code del, code_post]: "raise' (STR s) = raise s" -definition - raise_exc :: "exception \ 'a Heap" -where - [code del]: "raise_exc e = raise []" +lemma raise_raise' [code_inline]: + "raise s = raise' (STR s)" + by simp -lemma raise_raise_exc [code, code_unfold]: - "raise s = raise_exc (Fail (STR s))" - unfolding Fail_def raise_exc_def raise_def .. +code_datatype raise' -- {* avoid @{const "Heap"} formally *} -hide_const (open) Fail raise_exc +hide_const (open) raise' subsubsection {* SML and OCaml *} code_type Heap (SML "unit/ ->/ _") -code_const Heap (SML "raise/ (Fail/ \"bare Heap\")") code_const "op \=" (SML "!(fn/ f'_/ =>/ fn/ ()/ =>/ f'_/ (_/ ())/ ())") code_const return (SML "!(fn/ ()/ =>/ _)") -code_const "Heap_Monad.Fail" (SML "Fail") -code_const "Heap_Monad.raise_exc" (SML "!(fn/ ()/ =>/ raise/ _)") +code_const Heap_Monad.raise' (SML "!(raise/ Fail/ _)") code_type Heap (OCaml "_") -code_const Heap (OCaml "failwith/ \"bare Heap\"") code_const "op \=" (OCaml "!(fun/ f'_/ ()/ ->/ f'_/ (_/ ())/ ())") code_const return (OCaml "!(fun/ ()/ ->/ _)") -code_const "Heap_Monad.Fail" (OCaml "Failure") -code_const "Heap_Monad.raise_exc" (OCaml "!(fun/ ()/ ->/ raise/ _)") +code_const Heap_Monad.raise' (OCaml "failwith/ _") setup {* @@ -514,8 +449,6 @@ *} -code_reserved OCaml Failure raise - subsubsection {* Haskell *} @@ -556,10 +489,8 @@ text {* Monad *} code_type Heap (Haskell "Heap.ST/ Heap.RealWorld/ _") -code_const Heap (Haskell "error/ \"bare Heap\"") code_monad "op \=" Haskell code_const return (Haskell "return") -code_const "Heap_Monad.Fail" (Haskell "_") -code_const "Heap_Monad.raise_exc" (Haskell "error") +code_const Heap_Monad.raise' (Haskell "error/ _") end diff -r 694815d76240 -r 3f499df0751f src/HOL/Imperative_HOL/Ref.thy --- a/src/HOL/Imperative_HOL/Ref.thy Mon Jul 05 11:25:06 2010 +0200 +++ b/src/HOL/Imperative_HOL/Ref.thy Mon Jul 05 14:35:00 2010 +0200 @@ -1,5 +1,4 @@ (* Title: HOL/Library/Ref.thy - ID: $Id$ Author: John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen *) @@ -53,7 +52,7 @@ lemma update_change [code]: "r := e = Ref.change (\_. e) r \ return ()" - by (auto simp add: monad_simp change_def lookup_chain) + by (auto simp add: change_def lookup_chain) subsection {* Code generator setup *} diff -r 694815d76240 -r 3f499df0751f src/HOL/Imperative_HOL/Relational.thy --- a/src/HOL/Imperative_HOL/Relational.thy Mon Jul 05 11:25:06 2010 +0200 +++ b/src/HOL/Imperative_HOL/Relational.thy Mon Jul 05 14:35:00 2010 +0200 @@ -9,10 +9,10 @@ definition crel :: "'a Heap \ heap \ heap \ 'a \ bool" where - crel_def': "crel c h h' r \ Heap_Monad.execute c h = (Inl r, h')" + crel_def': "crel c h h' r \ Heap_Monad.execute c h = Some (r, h')" lemma crel_def: -- FIXME - "crel c h h' r \ (Inl r, h') = Heap_Monad.execute c h" + "crel c h h' r \ Some (r, h') = Heap_Monad.execute c h" unfolding crel_def' by auto lemma crel_deterministic: "\ crel f h h' a; crel f h h'' b \ \ (a = b) \ (h' = h'')" @@ -28,8 +28,7 @@ lemma crelE[consumes 1]: assumes "crel (f >>= g) h h'' r'" obtains h' r where "crel f h h' r" "crel (g r) h' h'' r'" - using assms - by (auto simp add: crel_def bindM_def Let_def prod_case_beta split_def Pair_fst_snd_eq split add: sum.split_asm) + using assms by (auto simp add: crel_def bindM_def split: option.split_asm) lemma crelE'[consumes 1]: assumes "crel (f >> g) h h'' r'" @@ -86,8 +85,8 @@ lemma crel_heap: assumes "crel (Heap_Monad.heap f) h h' r" obtains "h' = snd (f h)" "r = fst (f h)" - using assms - unfolding heap_def crel_def apfst_def split_def prod_fun_def by simp_all + using assms by (cases "f h") (simp add: crel_def) + subsection {* Elimination rules for array commands *} @@ -369,11 +368,9 @@ apply (cases f) apply simp apply (simp add: expand_fun_eq split_def) +apply (auto split: option.split) +apply (erule_tac x="x" in meta_allE) apply auto -apply (case_tac "fst (fun x)") -apply (simp_all add: Pair_fst_snd_eq) -apply (erule_tac x="x" in meta_allE) -apply fastsimp done section {* Introduction rules *} @@ -502,10 +499,10 @@ shows "P x h h' r" proof (rule MREC_pinduct[OF assms(1)[unfolded crel_def, symmetric]]) fix x h h1 h2 h' s z r - assume "Heap_Monad.execute (f x) h = (Inl (Inr s), h1)" - "Heap_Monad.execute (MREC f g s) h1 = (Inl z, h2)" + assume "Heap_Monad.execute (f x) h = Some (Inr s, h1)" + "Heap_Monad.execute (MREC f g s) h1 = Some (z, h2)" "P s h1 h2 z" - "Heap_Monad.execute (g x s z) h2 = (Inl r, h')" + "Heap_Monad.execute (g x s z) h2 = Some (r, h')" from assms(3)[unfolded crel_def, OF this(1)[symmetric] this(2)[symmetric] this(3) this(4)[symmetric]] show "P x h h' r" . next @@ -519,15 +516,15 @@ definition noError :: "'a Heap \ heap \ bool" where - "noError c h \ (\r h'. (Inl r, h') = Heap_Monad.execute c h)" + "noError c h \ (\r h'. Some (r, h') = Heap_Monad.execute c h)" lemma noError_def': -- FIXME - "noError c h \ (\r h'. Heap_Monad.execute c h = (Inl r, h'))" + "noError c h \ (\r h'. Heap_Monad.execute c h = Some (r, h'))" unfolding noError_def apply auto proof - fix r h' - assume "(Inl r, h') = Heap_Monad.execute c h" - then have "Heap_Monad.execute c h = (Inl r, h')" .. - then show "\r h'. Heap_Monad.execute c h = (Inl r, h')" by blast + assume "Some (r, h') = Heap_Monad.execute c h" + then have "Heap_Monad.execute c h = Some (r, h')" .. + then show "\r h'. Heap_Monad.execute c h = Some (r, h')" by blast qed subsection {* Introduction rules for basic monadic commands *} @@ -640,7 +637,7 @@ (*TODO: move to HeapMonad *) lemma mapM_append: "mapM f (xs @ ys) = mapM f xs \= (\xs. mapM f ys \= (\ys. return (xs @ ys)))" - by (induct xs) (simp_all add: monad_simp) + by (induct xs) simp_all lemma noError_freeze: shows "noError (freeze a) h" diff -r 694815d76240 -r 3f499df0751f src/HOL/Imperative_HOL/ex/Linked_Lists.thy --- a/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Mon Jul 05 11:25:06 2010 +0200 +++ b/src/HOL/Imperative_HOL/ex/Linked_Lists.thy Mon Jul 05 14:35:00 2010 +0200 @@ -56,7 +56,7 @@ return (x#xs) done" unfolding traverse_def -by (auto simp: traverse_def monad_simp MREC_rule) +by (auto simp: traverse_def MREC_rule) section {* Proving correctness with relational abstraction *} @@ -531,7 +531,7 @@ done" unfolding rev'_def MREC_rule[of _ _ "(q, p)"] unfolding rev'_def[symmetric] thm arg_cong2 - by (auto simp add: monad_simp expand_fun_eq intro: arg_cong2[where f = "op \="] split: node.split) + by (auto simp add: expand_fun_eq intro: arg_cong2[where f = "op \="] split: node.split) fun rev :: "('a:: heap) node \ 'a node Heap" where diff -r 694815d76240 -r 3f499df0751f src/HOL/Imperative_HOL/ex/SatChecker.thy --- a/src/HOL/Imperative_HOL/ex/SatChecker.thy Mon Jul 05 11:25:06 2010 +0200 +++ b/src/HOL/Imperative_HOL/ex/SatChecker.thy Mon Jul 05 14:35:00 2010 +0200 @@ -402,6 +402,12 @@ res_thm' l cli clj done))" +primrec + foldM :: "('a \ 'b \ 'b Heap) \ 'a list \ 'b \ 'b Heap" +where + "foldM f [] s = return s" + | "foldM f (x#xs) s = f x s \= foldM f xs" + fun doProofStep2 :: "Clause option array \ ProofStep \ Clause list \ Clause list Heap" where "doProofStep2 a (Conflict saveTo (i, rs)) rcs =