diff --git a/thys/Van_Emde_Boas_Trees/Imperative_HOL_Time/Array_Time.thy b/thys/Van_Emde_Boas_Trees/Imperative_HOL_Time/Array_Time.thy --- a/thys/Van_Emde_Boas_Trees/Imperative_HOL_Time/Array_Time.thy +++ b/thys/Van_Emde_Boas_Trees/Imperative_HOL_Time/Array_Time.thy @@ -1,511 +1,511 @@ (* Title: Imperative_HOL_Time/Array_Time.thy Author: Maximilian P. L. Haslbeck & Bohua Zhan, TU Muenchen *) section \Monadic arrays\ text \This theory is an adaptation of \HOL/Imperative_HOL/Array_Time.thy\, adding time bookkeeping.\ theory Array_Time imports Heap_Time_Monad begin subsection \Primitives\ definition present :: "heap \ 'a::heap array \ bool" where "present h a \ addr_of_array a < lim h" definition get :: "heap \ 'a::heap array \ 'a list" where "get h a = map from_nat (arrays h (TYPEREP('a)) (addr_of_array a))" definition set :: "'a::heap array \ 'a list \ heap \ heap" where "set a x = arrays_update (\h. h(TYPEREP('a) := ((h(TYPEREP('a))) (addr_of_array a:=map to_nat x))))" definition alloc :: "'a list \ heap \ 'a::heap array \ heap" where "alloc xs h = (let l = lim h; r = Array l; h'' = set r xs (h\lim := l + 1\) in (r, h''))" definition length :: "heap \ 'a::heap array \ nat" where "length h a = List.length (get h a)" definition update :: "'a::heap array \ nat \ 'a \ heap \ heap" where "update a i x h = set a ((get h a)[i:=x]) h" definition noteq :: "'a::heap array \ 'b::heap array \ bool" (infix "=!!=" 70) where "r =!!= s \ TYPEREP('a) \ TYPEREP('b) \ addr_of_array r \ addr_of_array s" subsection \Monad operations\ definition new :: "nat \ 'a::heap \ 'a array Heap" where [code del]: "new n x = Heap_Time_Monad.heap (%h. let (r,h') = alloc (replicate n x) h in (r,h',n+1))" definition of_list :: "'a::heap list \ 'a array Heap" where [code del]: "of_list xs = Heap_Time_Monad.heap (%h. let (r,h') = alloc xs h in (r,h',1+List.length xs))" definition make :: "nat \ (nat \ 'a::heap) \ 'a array Heap" where [code del]: "make n f = Heap_Time_Monad.heap (%h. let (r,h') = alloc (map f [0 ..< n]) h in (r,h',n+1))" definition len :: "'a::heap array \ nat Heap" where [code del]: "len a = Heap_Time_Monad.tap (\h. length h a)" definition nth :: "'a::heap array \ nat \ 'a Heap" where [code del]: "nth a i = Heap_Time_Monad.guard (\h. i < length h a) (\h. (get h a ! i, h, 1))" definition upd :: "nat \ 'a \ 'a::heap array \ 'a::heap array Heap" where [code del]: "upd i x a = Heap_Time_Monad.guard (\h. i < length h a) (\h. (a, update a i x h, 1))" definition map_entry :: "nat \ ('a::heap \ 'a) \ 'a array \ 'a array Heap" where [code del]: "map_entry i f a = Heap_Time_Monad.guard (\h. i < length h a) (\h. (a, update a i (f (get h a ! i)) h, 2))" definition swap :: "nat \ 'a \ 'a::heap array \ 'a Heap" where [code del]: "swap i x a = Heap_Time_Monad.guard (\h. i < length h a) (\h. (get h a ! i, update a i x h, 2 ))" (* questionable *) definition freeze :: "'a::heap array \ 'a list Heap" where [code del]: "freeze a = Heap_Time_Monad.heap (\h. (get h a, h, 1+length h a)) " subsection \Properties\ text \FIXME: Does there exist a "canonical" array axiomatisation in the literature?\ text \Primitives\ lemma noteq_sym: "a =!!= b \ b =!!= a" and unequal [simp]: "a \ a' \ a =!!= a'" unfolding noteq_def by auto lemma noteq_irrefl: "r =!!= r \ False" unfolding noteq_def by auto lemma present_alloc_noteq: "present h a \ a =!!= fst (alloc xs h)" by (simp add: present_def noteq_def alloc_def Let_def) lemma get_set_eq [simp]: "get (set r x h) r = x" by (simp add: get_def set_def o_def) lemma get_set_neq [simp]: "r =!!= s \ get (set s x h) r = get h r" by (simp add: noteq_def get_def set_def) lemma set_same [simp]: "set r x (set r y h) = set r x h" by (simp add: set_def) lemma set_set_swap: "r =!!= r' \ set r x (set r' x' h) = set r' x' (set r x h)" by (simp add: Let_def fun_eq_iff noteq_def set_def) lemma get_update_eq [simp]: "get (update a i v h) a = (get h a) [i := v]" by (simp add: update_def) lemma nth_update_neq [simp]: "a =!!= b \ get (update b j v h) a ! i = get h a ! i" by (simp add: update_def noteq_def) lemma get_update_elem_neqIndex [simp]: "i \ j \ get (update a j v h) a ! i = get h a ! i" by simp lemma length_update [simp]: "length (update b i v h) = length h" by (simp add: update_def length_def set_def get_def fun_eq_iff) lemma update_swap_neq: "a =!!= a' \ update a i v (update a' i' v' h) = update a' i' v' (update a i v h)" apply (unfold update_def) apply simp apply (subst set_set_swap, assumption) apply (subst get_set_neq) apply (erule noteq_sym) apply simp done lemma update_swap_neqIndex: "\ i \ i' \ \ update a i v (update a i' v' h) = update a i' v' (update a i v h)" by (auto simp add: update_def set_set_swap list_update_swap) lemma get_alloc: "get (snd (alloc xs h)) (fst (alloc ys h)) = xs" by (simp add: Let_def split_def alloc_def) lemma length_alloc: "length (snd (alloc (xs :: 'a::heap list) h)) (fst (alloc (ys :: 'a list) h)) = List.length xs" by (simp add: Array_Time.length_def get_alloc) lemma set: "set (fst (alloc ls h)) new_ls (snd (alloc ls h)) = snd (alloc new_ls h)" by (simp add: Let_def split_def alloc_def) lemma present_update [simp]: "present (update b i v h) = present h" by (simp add: update_def present_def set_def get_def fun_eq_iff) lemma present_alloc [simp]: "present (snd (alloc xs h)) (fst (alloc xs h))" by (simp add: present_def alloc_def set_def Let_def) lemma not_present_alloc [simp]: "\ present h (fst (alloc xs h))" by (simp add: present_def alloc_def Let_def) text \Monad operations\ lemma execute_new [execute_simps]: "execute (new n x) h = Some (let (r,h') = alloc (replicate n x) h in (r,h',n+1))" by (simp add: new_def execute_simps) lemma success_newI [success_intros]: "success (new n x) h" by (auto intro: success_intros simp add: new_def) lemma effect_newI [effect_intros]: assumes "(a, h') = alloc (replicate n x) h" shows "effect (new n x) h h' a (n+1)" apply (rule effectI) apply (simp add: assms execute_simps) by (metis assms case_prod_conv) lemma effect_newE [effect_elims]: assumes "effect (new n x) h h' r n'" obtains "r = fst (alloc (replicate n x) h)" "h' = snd (alloc (replicate n x) h)" "get h' r = replicate n x" "present h' r" "\ present h r" "n+1=n'" using assms apply (rule effectE) using case_prod_beta get_alloc execute_new by (metis (mono_tags, lifting) fst_conv not_present_alloc option.sel present_alloc sndI) (* apply (si mp add: case_prod_beta get_alloc execute_simps) refactor proof *) lemma execute_of_list [execute_simps]: "execute (of_list xs) h = Some (let (r,h') = alloc xs h in (r,h',1 + List.length xs))" by (simp add: of_list_def execute_simps) lemma success_of_listI [success_intros]: "success (of_list xs) h" by (auto intro: success_intros simp add: of_list_def) lemma effect_of_listI [effect_intros]: assumes "(a, h') = alloc xs h" shows "effect (of_list xs) h h' a (1 + List.length xs)" by (rule effectI, simp add: assms execute_simps, metis assms case_prod_conv) lemma effect_of_listE [effect_elims]: assumes "effect (of_list xs) h h' r n'" obtains "r = fst (alloc xs h)" "h' = snd (alloc xs h)" "get h' r = xs" "present h' r" "\ present h r" "n' = 1 + List.length xs" using assms apply (rule effectE) apply (simp add: get_alloc execute_of_list) by (simp add: case_prod_unfold) lemma execute_make [execute_simps]: "execute (make n f) h = Some (let (r,h') = alloc (map f [0 ..< n]) h in (r,h',n+1))" by (simp add: make_def execute_simps) lemma success_makeI [success_intros]: "success (make n f) h" by (auto intro: success_intros simp add: make_def) lemma effect_makeI [effect_intros]: assumes "(a, h') = alloc (map f [0 ..< n]) h" shows "effect (make n f) h h' a (n+1)" by (rule effectI) (simp add: assms execute_simps, metis assms case_prod_conv) lemma effect_makeE [effect_elims]: assumes "effect (make n f) h h' r n'" obtains "r = fst (alloc (map f [0 ..< n]) h)" "h' = snd (alloc (map f [0 ..< n]) h)" "get h' r = map f [0 ..< n]" "present h' r" "\ present h r" "n+1=n'" using assms apply (rule effectE) using get_alloc by (metis (mono_tags, opaque_lifting) effectE effect_makeI not_present_alloc present_alloc prod.collapse) (* apply (si mp add: get_alloc execute_make) by (s imp add: case_prod_unfold) *) lemma execute_len [execute_simps]: "execute (len a) h = Some (length h a, h, 1)" by (simp add: len_def execute_simps) lemma success_lenI [success_intros]: "success (len a) h" by (auto intro: success_intros simp add: len_def) lemma effect_lengthI [effect_intros]: assumes "h' = h" "r = length h a" "n=1" shows "effect (len a) h h' r n" by (rule effectI) (simp add: assms execute_simps) lemma effect_lengthE [effect_elims]: assumes "effect (len a) h h' r n" obtains "r = length h' a" "h' = h" "n=1" using assms by (rule effectE) (simp add: execute_simps) lemma execute_nth [execute_simps]: "i < length h a \ execute (nth a i) h = Some (get h a ! i, h,1)" "i \ length h a \ execute (nth a i) h = None" by (simp_all add: nth_def execute_simps) lemma success_nthI [success_intros]: "i < length h a \ success (nth a i) h" by (auto intro: success_intros simp add: nth_def) lemma effect_nthI [effect_intros]: assumes "i < length h a" "h' = h" "r = get h a ! i" "n=1" shows "effect (nth a i) h h' r n" by (rule effectI) (insert assms, simp add: execute_simps) lemma effect_nthE [effect_elims]: assumes "effect (nth a i) h h' r n" obtains "i < length h a" "r = get h a ! i" "h' = h" "n=1" using assms by (rule effectE) (cases "i < length h a", auto simp: execute_simps elim: successE) lemma execute_upd [execute_simps]: "i < length h a \ execute (upd i x a) h = Some (a, update a i x h, 1)" "i \ length h a \ execute (upd i x a) h = None" by (simp_all add: upd_def execute_simps) lemma success_updI [success_intros]: "i < length h a \ success (upd i x a) h" by (auto intro: success_intros simp add: upd_def) lemma effect_updI [effect_intros]: assumes "i < length h a" "h' = update a i v h" "n=1" shows "effect (upd i v a) h h' a n" by (rule effectI) (insert assms, simp add: execute_simps) lemma effect_updE [effect_elims]: assumes "effect (upd i v a) h h' r n" obtains "r = a" "h' = update a i v h" "i < length h a" "n=1" using assms by (rule effectE) (cases "i < length h a", auto simp: execute_simps elim: successE) lemma execute_map_entry [execute_simps]: "i < length h a \ execute (map_entry i f a) h = Some (a, update a i (f (get h a ! i)) h, 2)" "i \ length h a \ execute (map_entry i f a) h = None" by (simp_all add: map_entry_def execute_simps) lemma success_map_entryI [success_intros]: "i < length h a \ success (map_entry i f a) h" by (auto intro: success_intros simp add: map_entry_def) lemma effect_map_entryI [effect_intros]: assumes "i < length h a" "h' = update a i (f (get h a ! i)) h" "r = a" "n=2" shows "effect (map_entry i f a) h h' r n" by (rule effectI) (insert assms, simp add: execute_simps) lemma effect_map_entryE [effect_elims]: assumes "effect (map_entry i f a) h h' r n" obtains "r = a" "h' = update a i (f (get h a ! i)) h" "i < length h a" "n=2" using assms by (rule effectE) (cases "i < length h a", auto simp: execute_simps elim: successE) lemma execute_swap [execute_simps]: "i < length h a \ execute (swap i x a) h = Some (get h a ! i, update a i x h, 2)" "i \ length h a \ execute (swap i x a) h = None" by (simp_all add: swap_def execute_simps) lemma success_swapI [success_intros]: "i < length h a \ success (swap i x a) h" by (auto intro: success_intros simp add: swap_def) lemma effect_swapI [effect_intros]: assumes "i < length h a" "h' = update a i x h" "r = get h a ! i" "n=2" shows "effect (swap i x a) h h' r n" by (rule effectI) (insert assms, simp add: execute_simps) lemma effect_swapE [effect_elims]: assumes "effect (swap i x a) h h' r n" obtains "r = get h a ! i" "h' = update a i x h" "i < length h a" "n=2" using assms by (rule effectE) (cases "i < length h a", auto simp: execute_simps elim: successE) lemma execute_freeze [execute_simps]: "execute (freeze a) h = Some (get h a, h, 1+length h a)" by (simp add: freeze_def execute_simps) lemma success_freezeI [success_intros]: "success (freeze a) h" by (auto intro: success_intros simp add: freeze_def) lemma effect_freezeI [effect_intros]: assumes "h' = h" "r = get h a" "n=length h a" shows "effect (freeze a) h h' r (n+1)" by (rule effectI) (insert assms, simp add: execute_simps) lemma effect_freezeE [effect_elims]: assumes "effect (freeze a) h h' r n" obtains "h' = h" "r = get h a" "n=length h a+1" using assms by (rule effectE) (simp add: execute_simps) lemma upd_ureturn: "upd i x a \ ureturn a = upd i x a " by (rule Heap_eqI) (simp add: bind_def guard_def upd_def execute_simps) lemma array_make: "new n x = make n (\_. x)" by (rule Heap_eqI) (simp add: map_replicate_trivial execute_simps) lemma array_of_list_make [code]: "of_list xs = make (List.length xs) (\n. xs ! n)" by (rule Heap_eqI) (simp add: map_nth execute_simps) hide_const (open) present get set alloc length update noteq new of_list make len nth upd map_entry swap freeze subsection \Code generator setup\ subsubsection \Logical intermediate layer\ definition new' where [code del]: "new' = Array_Time.new o nat_of_integer" lemma [code]: "Array_Time.new = new' o of_nat" by (simp add: new'_def o_def) definition make' where [code del]: "make' i f = Array_Time.make (nat_of_integer i) (f o of_nat)" lemma [code]: "Array_Time.make n f = make' (of_nat n) (f o nat_of_integer)" by (simp add: make'_def o_def) definition len' where [code del]: "len' a = Array_Time.len a \ (\n. ureturn (of_nat n))" lemma [code]: "Array_Time.len a = len' a \ (\i. ureturn (nat_of_integer i))" by (simp add: len'_def execute_simps) definition nth' where [code del]: "nth' a = Array_Time.nth a o nat_of_integer" lemma [code]: "Array_Time.nth a n = nth' a (of_nat n)" by (simp add: nth'_def) definition upd' where [code del]: "upd' a i x = Array_Time.upd (nat_of_integer i) x a \ ureturn ()" lemma [code]: "Array_Time.upd i x a = upd' a (of_nat i) x \ ureturn a" by (simp add: upd'_def upd_ureturn execute_simps) lemma [code]: "Array_Time.map_entry i f a = do { x \ Array_Time.nth a i; Array_Time.upd i (f x) a }" by (rule Heap_eqI) (simp add: bind_def guard_def map_entry_def execute_simps) lemma [code]: "Array_Time.swap i x a = do { y \ Array_Time.nth a i; Array_Time.upd i x a; ureturn y }" by (rule Heap_eqI) (simp add: bind_def guard_def swap_def execute_simps) (* lemma [code]: "Array_Time.freeze a = do { n \ Array_Time.len a; Heap_Monad.fold_map (\i. Array_Time.nth a i) [0..x. fst (the (if x < Array_Time.length h a then Some (Array_Time.get h a ! x, h) else None))) [0.. Array_Time.len a; Heap_Monad.fold_map (Array_Time.nth a) [0.. Array_Time.len a; Heap_Monad.fold_map (Array_Time.nth a) [0..SML\ code_printing type_constructor array \ (SML) "_/ array" code_printing constant Array \ (SML) "raise/ (Fail/ \"bare Array\")" code_printing constant Array_Time.new' \ (SML) "(fn/ ()/ =>/ Array.array/ ((_),/ (_)))" code_printing constant Array_Time.of_list \ (SML) "(fn/ ()/ =>/ Array.fromList/ _)" code_printing constant Array_Time.make' \ (SML) "(fn/ ()/ =>/ Array.tabulate/ ((_),/ (_)))" code_printing constant Array_Time.len' \ (SML) "(fn/ ()/ =>/ Array.length/ _)" code_printing constant Array_Time.nth' \ (SML) "(fn/ ()/ =>/ Array.sub/ ((_),/ (_)))" code_printing constant Array_Time.upd' \ (SML) "(fn/ ()/ =>/ Array.update/ ((_),/ (_),/ (_)))" code_printing constant "HOL.equal :: 'a array \ 'a array \ bool" \ (SML) infixl 6 "=" code_reserved SML Array text \OCaml\ code_printing type_constructor array \ (OCaml) "_/ array" code_printing constant Array \ (OCaml) "failwith/ \"bare Array\"" code_printing constant Array_Time.new' \ (OCaml) "(fun/ ()/ ->/ Array.make/ (Big'_int.int'_of'_big'_int/ _)/ _)" code_printing constant Array_Time.of_list \ (OCaml) "(fun/ ()/ ->/ Array.of'_list/ _)" code_printing constant Array_Time.make' \ (OCaml) "(fun/ ()/ ->/ Array.init/ (Big'_int.int'_of'_big'_int/ _)/ (fun k'_ ->/ _/ (Big'_int.big'_int'_of'_int/ k'_)))" code_printing constant Array_Time.len' \ (OCaml) "(fun/ ()/ ->/ Big'_int.big'_int'_of'_int/ (Array.length/ _))" code_printing constant Array_Time.nth' \ (OCaml) "(fun/ ()/ ->/ Array.get/ _/ (Big'_int.int'_of'_big'_int/ _))" code_printing constant Array_Time.upd' \ (OCaml) "(fun/ ()/ ->/ Array.set/ _/ (Big'_int.int'_of'_big'_int/ _)/ _)" code_printing constant "HOL.equal :: 'a array \ 'a array \ bool" \ (OCaml) infixl 4 "=" code_reserved OCaml Array text \Haskell\ code_printing type_constructor array \ (Haskell) "Heap.STArray/ Heap.RealWorld/ _" code_printing constant Array \ (Haskell) "error/ \"bare Array\"" code_printing constant Array_Time.new' \ (Haskell) "Heap.newArray" code_printing constant Array_Time.of_list \ (Haskell) "Heap.newListArray" code_printing constant Array_Time.make' \ (Haskell) "Heap.newFunArray" code_printing constant Array_Time.len' \ (Haskell) "Heap.lengthArray" code_printing constant Array_Time.nth' \ (Haskell) "Heap.readArray" code_printing constant Array_Time.upd' \ (Haskell) "Heap.writeArray" code_printing constant "HOL.equal :: 'a array \ 'a array \ bool" \ (Haskell) infix 4 "==" code_printing class_instance array :: HOL.equal \ (Haskell) - text \Scala\ -code_printing type_constructor array \ (Scala) "!collection.mutable.ArraySeq[_]" +code_printing type_constructor array \ (Scala) "!Array.T[_]" code_printing constant Array \ (Scala) "!sys.error(\"bare Array\")" code_printing constant Array_Time.new' \ (Scala) "('_: Unit)/ => / Array.alloc((_))((_))" code_printing constant Array_Time.make' \ (Scala) "('_: Unit)/ =>/ Array.make((_))((_))" code_printing constant Array_Time.len' \ (Scala) "('_: Unit)/ =>/ Array.len((_))" code_printing constant Array_Time.nth' \ (Scala) "('_: Unit)/ =>/ Array.nth((_), (_))" code_printing constant Array_Time.upd' \ (Scala) "('_: Unit)/ =>/ Array.upd((_), (_), (_))" code_printing constant Array_Time.freeze \ (Scala) "('_: Unit)/ =>/ Array.freeze((_))" code_printing constant "HOL.equal :: 'a array \ 'a array \ bool" \ (Scala) infixl 5 "==" end diff --git a/thys/Van_Emde_Boas_Trees/Imperative_HOL_Time/Heap_Time_Monad.thy b/thys/Van_Emde_Boas_Trees/Imperative_HOL_Time/Heap_Time_Monad.thy --- a/thys/Van_Emde_Boas_Trees/Imperative_HOL_Time/Heap_Time_Monad.thy +++ b/thys/Van_Emde_Boas_Trees/Imperative_HOL_Time/Heap_Time_Monad.thy @@ -1,769 +1,777 @@ (* Title: Imperative_HOL_Time/Heap_Time_Monad.thy Author: Maximilian P. L. Haslbeck & Bohua Zhan, TU Muenchen *) section \A monad with a polymorphic heap and time and primitive reasoning infrastructure\ text \This theory is an adapted version of \Imperative_HOL/Heap_Time_Monad\, where the heap is extended by time bookkeeping.\ theory Heap_Time_Monad imports "HOL-Imperative_HOL.Heap" "HOL-Library.Monad_Syntax" begin subsection \The monad\ subsubsection \Monad construction\ text \Monadic heap actions either produce values and transform the heap, or fail\ datatype 'a Heap = Heap "heap \ ('a \ heap \ nat) option" lemma [code, code del]: "(Code_Evaluation.term_of :: 'a::typerep Heap \ Code_Evaluation.term) = Code_Evaluation.term_of" .. primrec execute :: "'a Heap \ heap \ ('a \ heap \ nat) option" where [code del]: "execute (Heap f) = f" 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 lemma Heap_execute [simp]: "Heap (execute f) = f" by (cases f) simp_all lemma Heap_eqI: "(\h. execute f h = execute g h) \ f = g" by (cases f, cases g) (auto simp: fun_eq_iff) named_theorems execute_simps "simplification rules for execute" lemma execute_Let [execute_simps]: "execute (let x = t in f x) = (let x = t in execute (f x))" by (simp add: Let_def) subsubsection \Specialised lifters\ definition tap :: "(heap \ 'a) \ 'a Heap" where [code del]: "tap f = Heap (\h. Some (f h, h, 1))" lemma execute_tap [execute_simps]: "execute (tap f) h = Some (f h, h, 1)" by (simp add: tap_def) definition heap :: "(heap \ 'a \ heap \ nat) \ 'a Heap" where [code del]: "heap f = Heap (Some \ f)" lemma execute_heap [execute_simps]: "execute (heap f) = Some \ f" by (simp add: heap_def) definition guard :: "(heap \ bool) \ (heap \ 'a \ heap \ nat) \ 'a Heap" where [code del]: "guard P f = Heap (\h. if P h then Some (f h) else None)" lemma execute_guard [execute_simps]: "\ P h \ execute (guard P f) h = None" "P h \ execute (guard P f) h = Some (f h)" by (simp_all add: guard_def) subsubsection \Predicate classifying successful computations\ definition success :: "'a Heap \ heap \ bool" where "success f h \ execute f h \ None" lemma successI: "execute f h \ None \ success f h" by (simp add: success_def) lemma successE: assumes "success f h" obtains r h' where "execute f h = Some (r, h')" using assms by (auto simp: success_def) named_theorems success_intros "introduction rules for success" lemma success_tapI [success_intros]: "success (tap f) h" by (rule successI) (simp add: execute_simps) lemma success_heapI [success_intros]: "success (heap f) h" by (rule successI) (simp add: execute_simps) lemma success_guardI [success_intros]: "P h \ success (guard P f) h" by (rule successI) (simp add: execute_guard) lemma success_LetI [success_intros]: "x = t \ success (f x) h \ success (let x = t in f x) h" by (simp add: Let_def) lemma success_ifI: "(c \ success t h) \ (\ c \ success e h) \ success (if c then t else e) h" by (simp add: success_def) subsubsection \Predicate for a simple relational calculus\ text \ The \effect\ vebt_predicate states that when a computation \c\ runs with the heap \h\ will result in return value \r\ and a heap \h'\, i.e.~no exception occurs. AND consume time \n\ \ definition effect :: "'a Heap \ heap \ heap \ 'a \ nat \ bool" where effect_def: "effect c h h' r n \ execute c h = Some (r, h', n)" lemma effectI: "execute c h = Some (r, h',n) \ effect c h h' r n" by (simp add: effect_def) lemma effectE: assumes "effect c h h' r n" obtains "r = fst (the (execute c h))" and "h' = fst (snd (the (execute c h)))" and "n = snd (snd (the (execute c h)))" and "success c h" proof (rule that) from assms have *: "execute c h = Some (r, h',n)" by (simp add: effect_def) then show "success c h" by (simp add: success_def) from * have "fst (the (execute c h)) = r" and "fst (snd (the (execute c h))) = h'" and "snd (snd (the (execute c h))) = n" by simp_all then show "r = fst (the (execute c h))" and "h' = fst (snd (the (execute c h)))" and "n = snd (snd (the (execute c h)))" by simp_all qed lemma effect_success: "effect c h h' r n \ success c h" by (simp add: effect_def success_def) lemma success_effectE: assumes "success c h" obtains r h' n where "effect c h h' r n" using assms by (auto simp add: effect_def success_def) lemma effect_deterministic: assumes "effect f h h' a n" and "effect f h h'' b n'" shows "a = b" and "h' = h''" and "n = n'" using assms unfolding effect_def by auto named_theorems effect_intros "introduction rules for effect" and effect_elims "elimination rules for effect" lemma effect_LetI [effect_intros]: assumes "x = t" "effect (f x) h h' r n" shows "effect (let x = t in f x) h h' r n" using assms by simp lemma effect_LetE [effect_elims]: assumes "effect (let x = t in f x) h h' r n" obtains "effect (f t) h h' r n" using assms by simp lemma effect_ifI: assumes "c \ effect t h h' r n" and "\ c \ effect e h h' r n" shows "effect (if c then t else e) h h' r n" by (cases c) (simp_all add: assms) lemma effect_ifE: assumes "effect (if c then t else e) h h' r n" obtains "c" "effect t h h' r n" | "\ c" "effect e h h' r n" using assms by (cases c) simp_all lemma effect_tapI [effect_intros]: assumes "h' = h" "r = f h" shows "effect (tap f) h h' r 1" by (rule effectI) (simp add: assms execute_simps) lemma effect_tapE [effect_elims]: assumes "effect (tap f) h h' r n" obtains "h' = h" and "r = f h" and "n=1" using assms by (rule effectE) (auto simp add: execute_simps) lemma effect_heapI [effect_intros]: assumes "n = snd (snd (f h))" "h' = fst (snd (f h))" "r = fst (f h)" shows "effect (heap f) h h' r n" by (rule effectI) (simp add: assms execute_simps) lemma effect_heapE [effect_elims]: assumes "effect (heap f) h h' r n" obtains "h' = fst (snd (f h))" and "n = snd (snd (f h))" and "r = fst (f h)" using assms by (rule effectE) (simp add: execute_simps) lemma effect_guardI [effect_intros]: assumes "P h" "h' = fst (snd (f h))" "n = snd (snd (f h))" "r = fst (f h)" shows "effect (guard P f) h h' r n" by (rule effectI) (simp add: assms execute_simps) lemma effect_guardE [effect_elims]: assumes "effect (guard P f) h h' r n" obtains "h' = fst (snd (f h))" "n = snd (snd (f h))" "r = fst (f h)" "P h" using assms by (rule effectE) (auto simp add: execute_simps elim!: successE, cases "P h", auto simp add: execute_simps) subsubsection \Monad combinators\ definition return :: "'a \ 'a Heap" where [code del]: "return x = heap (\h. (x,h,1))" lemma execute_return [execute_simps]: "execute (return x) = Some \ (\h. (x,h,1))" by (simp add: return_def execute_simps) lemma success_returnI [success_intros]: "success (return x) h" by (rule successI) (simp add: execute_simps) lemma effect_returnI [effect_intros]: "h = h' \ effect (return x) h h' x 1" by (rule effectI) (simp add: execute_simps) lemma effect_returnE [effect_elims]: assumes "effect (return x) h h' r n" obtains "r = x" "h' = h" "n=1" using assms by (rule effectE) (simp add: execute_simps) definition ureturn :: "'a \ 'a Heap" where [code del]: "ureturn x = heap (\h. (x,h,0))" lemma execute_ureturn [execute_simps]: "execute (ureturn x) = Some \ (\h. (x,h,0))" by (simp add: ureturn_def execute_simps) lemma success_ureturnI [success_intros]: "success (ureturn x) h" by (rule successI) (simp add: execute_simps) lemma effect_ureturnI [effect_intros]: "h = h' \ effect (ureturn x) h h' x 0" by (rule effectI) (simp add: execute_simps) lemma effect_ureturnE [effect_elims]: assumes "effect (ureturn x) h h' r n" obtains "r = x" "h' = h" "n=0" using assms by (rule effectE) (simp add: execute_simps) definition raise :: "string \ 'a Heap" where \ \the string is just decoration\ [code del]: "raise s = Heap (\_. None)" lemma execute_raise [execute_simps]: "execute (raise s) = (\_. None)" by (simp add: raise_def) lemma effect_raiseE [effect_elims]: assumes "effect (raise x) h h' r n" obtains "False" using assms by (rule effectE) (simp add: success_def execute_simps) fun timeFrame :: "nat \ ('a \ heap \ nat) option \ ('a \ heap \ nat) option" where "timeFrame n (Some (r,h,n')) = Some (r, h, n+n')" | "timeFrame n None = None" lemma timeFrame_zero[simp]: "timeFrame 0 h = h" apply(cases h) by auto lemma timeFrame_assoc[simp]: "timeFrame n (timeFrame n' f) = timeFrame (n+n') f" by (metis (no_types, lifting) ab_semigroup_add_class.add_ac(1) timeFrame.elims timeFrame.simps(1)) (* refactor proof *) definition bind :: "'a Heap \ ('a \ 'b Heap) \ 'b Heap" where [code del]: "bind f g = Heap (\h. case execute f h of Some (r, h', n) \ timeFrame n (execute (g r) h') | None \ None)" adhoc_overloading Monad_Syntax.bind Heap_Time_Monad.bind lemma execute_bind [execute_simps]: "execute f h = Some (x, h',n) \ execute (f \ g) h = timeFrame n (execute (g x) h')" "execute f h = None \ execute (f \ g) h = None" by (simp_all add: bind_def) lemma execute_bind_case: "execute (f \ g) h = (case (execute f h) of Some (x, h',n) \ timeFrame n (execute (g x) h') | None \ None)" by (simp add: bind_def) lemma execute_bind_success: "success f h \ execute (f \ g) h = timeFrame (snd (snd (the (execute f h)))) (execute (g (fst (the (execute f h)))) (fst (snd (the (execute f h)))))" by (cases f h rule: Heap_cases) (auto elim: successE simp add: bind_def) lemma success_bind_executeI: "execute f h = Some (x, h',n) \ success (g x) h' \ success (f \ g) h" by (auto intro!: successI elim: successE simp add: bind_def) lemma success_bind_effectI [success_intros]: "effect f h h' x n \ success (g x) h' \ success (f \ g) h" by (auto simp add: effect_def success_def bind_def) lemma effect_bindI [effect_intros]: assumes "effect f h h' r n" "effect (g r) h' h'' r' n'" shows "effect (f \ g) h h'' r' (n+n')" using assms apply (auto intro!: effectI elim!: effectE successE) apply (subst execute_bind, simp_all) apply auto done lemma effect_bindE [effect_elims]: assumes "effect (f \ g) h h'' r' n" obtains h' r n1 n2 where "effect f h h' r n1" "effect (g r) h' h'' r' n2" "n = n1 + n2" using assms apply (auto simp add: effect_def bind_def split: option.split_asm) (* To cleanup *) by (smt Pair_inject option.distinct(1) option.inject timeFrame.elims) lemma execute_bind_eq_SomeI: assumes "execute f h = Some (x, h',n)" and "execute (g x) h' = Some (y, h'',n')" shows "execute (f \ g) h = Some (y, h'',n+n')" using assms by (simp add: bind_def) term "return x \ f" definition wait :: "nat \ unit Heap" where [execute_simps]: "wait n = Heap (\h. Some ((),h,n))" lemma [simp]: "wait n \ (%_. wait m) = wait (n+m)" unfolding wait_def bind_def by auto term "wait 1 \ (%_. f x)" term "return x \ f" lemma ureturn_bind [execute_simps]: "ureturn x \ f = f x" apply (rule Heap_eqI) by (simp add: execute_simps) lemma return_bind [execute_simps]: "return x \ f = (wait 1) \ f x" apply (rule Heap_eqI) by (simp add: execute_simps) lemma bind_return [execute_simps]: "f \ return = wait 1 \ f" by (rule Heap_eqI) (simp add: bind_def execute_simps split: option.splits) lemma bind_ureturn [execute_simps]: "f \ ureturn = f" by (rule Heap_eqI) (simp add: bind_def execute_simps split: option.splits) lemma bind_bind [simp]: "(f \ g) \ k = (f :: 'a Heap) \ (\x. g x \ k)" by (rule Heap_eqI) (simp add: bind_def execute_simps split: option.splits) lemma raise_bind [simp]: "raise e \ f = raise e" by (rule Heap_eqI) (simp add: execute_simps) subsection \Generic combinators\ subsubsection \Assertions\ definition assert :: "('a \ bool) \ 'a \ 'a Heap" where "assert P x = (if P x then return x else raise ''assert'')" lemma execute_assert [execute_simps]: "P x \ execute (assert P x) h = Some (x, h, 1)" "\ P x \ execute (assert P x) h = None" by (simp_all add: assert_def execute_simps) lemma success_assertI [success_intros]: "P x \ success (assert P x) h" by (rule successI) (simp add: execute_assert) lemma effect_assertI [effect_intros]: "P x \ h' = h \ r = x \ n = 1 \ effect (assert P x) h h' r n" by (rule effectI) (simp add: execute_assert) lemma effect_assertE [effect_elims]: assumes "effect (assert P x) h h' r n" obtains "P x" "r = x" "h' = h" "n=1" using assms by (rule effectE) (cases "P x", simp_all add: execute_assert success_def) lemma assert_cong [fundef_cong]: assumes "P = P'" assumes "\x. P' x \ f x = f' x" shows "(assert P x \ f) = (assert P' x \ f')" by (rule Heap_eqI) (insert assms, simp add: assert_def execute_simps) subsubsection \Plain lifting\ definition lift :: "('a \ 'b) \ 'a \ 'b Heap" where "lift f = return o f" lemma lift_collapse [simp]: "lift f x = return (f x)" by (simp add: lift_def) lemma bind_lift: "(f \ lift g) = (f \ (\x. return (g x)))" by (simp add: lift_def comp_def) (* subsubsection \Iteration -- warning: this is rarely useful!\ primrec fold_map :: "('a \ 'b Heap) \ 'a list \ 'b list Heap" where "fold_map f [] = return []" | "fold_map f (x # xs) = do { y \ f x; ys \ fold_map f xs; return (y # ys) }" lemma fold_map_append: "fold_map f (xs @ ys) = fold_map f xs \ (\xs. fold_map f ys \ (\ys. return (xs @ ys)))" by (induct xs) simp_all lemma execute_fold_map_unchanged_heap [execute_simps]: assumes "\x. x \ set xs \ \y. execute (f x) h = Some (y, h)" shows "execute (fold_map f xs) h = Some (List.map (\x. fst (the (execute (f x) h))) xs, h)" using assms proof (induct xs) case Nil show ?case by (simp add: execute_simps) next case (Cons x xs) from Cons.prems obtain y where y: "execute (f x) h = Some (y, h)" by auto moreover from Cons.prems Cons.hyps have "execute (fold_map f xs) h = Some (map (\x. fst (the (execute (f x) h))) xs, h)" by auto ultimately show ?case by (simp, simp only: execute_bind(1), simp add: execute_simps) qed *) subsection \Partial function definition setup\ definition Heap_ord :: "'a Heap \ 'a Heap \ bool" where "Heap_ord = img_ord execute (fun_ord option_ord)" definition Heap_lub :: "'a Heap set \ 'a Heap" where "Heap_lub = img_lub execute Heap (fun_lub (flat_lub None))" lemma Heap_lub_empty: "Heap_lub {} = Heap Map.empty" by(simp add: Heap_lub_def img_lub_def fun_lub_def flat_lub_def) lemma heap_interpretation: "partial_function_definitions Heap_ord Heap_lub" proof - have "partial_function_definitions (fun_ord option_ord) (fun_lub (flat_lub None))" by (rule partial_function_lift) (rule flat_interpretation) then have "partial_function_definitions (img_ord execute (fun_ord option_ord)) (img_lub execute Heap (fun_lub (flat_lub None)))" by (rule partial_function_image) (auto intro: Heap_eqI) then show "partial_function_definitions Heap_ord Heap_lub" by (simp only: Heap_ord_def Heap_lub_def) qed interpretation heap: partial_function_definitions Heap_ord Heap_lub rewrites "Heap_lub {} \ Heap Map.empty" by (fact heap_interpretation)(simp add: Heap_lub_empty) lemma heap_step_admissible: "option.admissible (\f:: 'a => ('b * 'c* 'd) option. \h h' r n. f h = Some (r, h',n) \ P x h h' r n)" proof (rule ccpo.admissibleI) fix A :: "('a \ ('b * 'c * 'd) option) set" assume ch: "Complete_Partial_Order.chain option.le_fun A" and IH: "\f\A. \h h' r n. f h = Some (r, h', n) \ P x h h' r n" from ch have ch': "\x. Complete_Partial_Order.chain option_ord {y. \f\A. y = f x}" by (rule chain_fun) show "\h h' r n. option.lub_fun A h = Some (r, h', n) \ P x h h' r n" proof (intro allI impI) fix h h' r n assume "option.lub_fun A h = Some (r, h', n)" from flat_lub_in_chain[OF ch' this[unfolded fun_lub_def]] have "Some (r, h', n) \ {y. \f\A. y = f h}" by simp then have "\f\A. f h = Some (r, h', n)" by auto with IH show "P x h h' r n" by auto qed qed lemma admissible_heap: "heap.admissible (\f. \x h h' r n. effect (f x) h h' r n \ P x h h' r n)" proof (rule admissible_fun[OF heap_interpretation]) fix x show "ccpo.admissible Heap_lub Heap_ord (\a. \h h' r n. effect a h h' r n \ P x h h' r n)" unfolding Heap_ord_def Heap_lub_def proof (intro admissible_image partial_function_lift flat_interpretation) show "option.admissible ((\a. \h h' r n. effect a h h' r n \ P x h h' r n) \ Heap)" unfolding comp_def effect_def execute.simps by (rule heap_step_admissible) qed (auto simp add: Heap_eqI) qed lemma fixp_induct_heap: fixes F :: "'c \ 'c" and U :: "'c \ 'b \ 'a Heap" and C :: "('b \ 'a Heap) \ 'c" and P :: "'b \ heap \ heap \ 'a \ nat \ bool" assumes mono: "\x. monotone (fun_ord Heap_ord) Heap_ord (\f. U (F (C f)) x)" assumes eq: "f \ C (ccpo.fixp (fun_lub Heap_lub) (fun_ord Heap_ord) (\f. U (F (C f))))" assumes inverse2: "\f. U (C f) = f" assumes step: "\f x h h' r n. (\x h h' r n. effect (U f x) h h' r n \ P x h h' r n) \ effect (U (F f) x) h h' r n \ P x h h' r n" assumes defined: "effect (U f x) h h' r n" shows "P x h h' r n" using step defined heap.fixp_induct_uc[of U F C, OF mono eq inverse2 admissible_heap, of P] unfolding effect_def execute.simps by blast declaration \Partial_Function.init "heap_time" @{term heap.fixp_fun} @{term heap.mono_body} @{thm heap.fixp_rule_uc} @{thm heap.fixp_induct_uc} (SOME @{thm fixp_induct_heap})\ abbreviation "mono_Heap \ monotone (fun_ord Heap_ord) Heap_ord" lemma Heap_ordI: assumes "\h. execute x h = None \ execute x h = execute y h" shows "Heap_ord x y" using assms unfolding Heap_ord_def img_ord_def fun_ord_def flat_ord_def by blast lemma Heap_ordE: assumes "Heap_ord x y" obtains "execute x h = None" | "execute x h = execute y h" using assms unfolding Heap_ord_def img_ord_def fun_ord_def flat_ord_def by atomize_elim blast lemma bind_mono [partial_function_mono]: assumes mf: "mono_Heap B" and mg: "\y. mono_Heap (\f. C y f)" shows "mono_Heap (\f. B f \ (\y. C y f))" proof (rule monotoneI) fix f g :: "'a \ 'b Heap" assume fg: "fun_ord Heap_ord f g" from mf have 1: "Heap_ord (B f) (B g)" by (rule monotoneD) (rule fg) from mg have 2: "\y'. Heap_ord (C y' f) (C y' g)" by (rule monotoneD) (rule fg) have "Heap_ord (B f \ (\y. C y f)) (B g \ (\y. C y f))" (is "Heap_ord ?L ?R") proof (rule Heap_ordI) fix h from 1 show "execute ?L h = None \ execute ?L h = execute ?R h" by (rule Heap_ordE[where h = h]) (auto simp: execute_bind_case) qed also have "Heap_ord (B g \ (\y'. C y' f)) (B g \ (\y'. C y' g))" (is "Heap_ord ?L ?R") proof (rule Heap_ordI) fix h show "execute ?L h = None \ execute ?L h = execute ?R h" proof (cases "execute (B g) h") case None then have "execute ?L h = None" by (auto simp: execute_bind_case) thus ?thesis .. next case Some then obtain r h' n where "execute (B g) h = Some (r, h', n)" by (metis surjective_pairing) then have "execute ?L h = timeFrame n (execute (C r f) h')" "execute ?R h = timeFrame n (execute (C r g) h')" by (auto simp: execute_bind_case) with 2[of r] show ?thesis apply (auto elim: Heap_ordE) by (metis Heap_ordE timeFrame.simps(2)) qed qed finally (heap.leq_trans) show "Heap_ord (B f \ (\y. C y f)) (B g \ (\y'. C y' g))" . qed subsection \Code generator setup\ subsubsection \SML and OCaml\ code_printing type_constructor Heap \ (SML) "(unit/ ->/ _)" code_printing constant bind \ (SML) "!(fn/ f'_/ =>/ fn/ ()/ =>/ f'_/ (_/ ())/ ())" code_printing constant return \ (SML) "!(fn/ ()/ =>/ _)" code_printing constant Heap_Time_Monad.raise \ (SML) "!(raise/ Fail/ _)" code_printing type_constructor Heap \ (OCaml) "(unit/ ->/ _)" code_printing constant bind \ (OCaml) "!(fun/ f'_/ ()/ ->/ f'_/ (_/ ())/ ())" code_printing constant return \ (OCaml) "!(fun/ ()/ ->/ _)" code_printing constant Heap_Time_Monad.raise \ (OCaml) "failwith" subsubsection \Haskell\ text \Adaption layer\ code_printing code_module "Heap" \ (Haskell) \import qualified Control.Monad; import qualified Control.Monad.ST; import qualified Data.STRef; import qualified Data.Array.ST; type RealWorld = Control.Monad.ST.RealWorld; type ST s a = Control.Monad.ST.ST s a; type STRef s a = Data.STRef.STRef s a; type STArray s a = Data.Array.ST.STArray s Integer a; newSTRef = Data.STRef.newSTRef; readSTRef = Data.STRef.readSTRef; writeSTRef = Data.STRef.writeSTRef; newArray :: Integer -> a -> ST s (STArray s a); newArray k = Data.Array.ST.newArray (0, k - 1); newListArray :: [a] -> ST s (STArray s a); newListArray xs = Data.Array.ST.newListArray (0, (fromInteger . toInteger . length) xs - 1) xs; newFunArray :: Integer -> (Integer -> a) -> ST s (STArray s a); newFunArray k f = Data.Array.ST.newListArray (0, k - 1) (map f [0..k-1]); lengthArray :: STArray s a -> ST s Integer; lengthArray a = Control.Monad.liftM (\(_, l) -> l + 1) (Data.Array.ST.getBounds a); readArray :: STArray s a -> Integer -> ST s a; readArray = Data.Array.ST.readArray; writeArray :: STArray s a -> Integer -> a -> ST s (); writeArray = Data.Array.ST.writeArray;\ code_reserved Haskell Heap text \Monad\ code_printing type_constructor Heap \ (Haskell) "Heap.ST/ Heap.RealWorld/ _" code_monad bind Haskell code_printing constant return \ (Haskell) "return" code_printing constant Heap_Time_Monad.raise \ (Haskell) "error" subsubsection \Scala\ code_printing code_module "Heap" \ (Scala) \object Heap { - def bind[A, B](f: Unit => A, g: A => Unit => B): Unit => B = (_: Unit) => g (f ()) () + def bind[A, B](f: Unit => A, g: A => Unit => B): Unit => B = (_: Unit) => g(f(()))(()) } class Ref[A](x: A) { var value = x } object Ref { def apply[A](x: A): Ref[A] = new Ref[A](x) def lookup[A](r: Ref[A]): A = r.value def update[A](r: Ref[A], x: A): Unit = { r.value = x } } object Array { - import collection.mutable.ArraySeq - def alloc[A](n: BigInt)(x: A): ArraySeq[A] = - ArraySeq.fill(n.toInt)(x) - def make[A](n: BigInt)(f: BigInt => A): ArraySeq[A] = - ArraySeq.tabulate(n.toInt)((k: Int) => f(BigInt(k))) - def len[A](a: ArraySeq[A]): BigInt = - BigInt(a.length) - def nth[A](a: ArraySeq[A], n: BigInt): A = - a(n.toInt) - def upd[A](a: ArraySeq[A], n: BigInt, x: A): Unit = - a.update(n.toInt, x) - def freeze[A](a: ArraySeq[A]): List[A] = - a.toList + class T[A](n: Int) + { + val array: Array[AnyRef] = new Array[AnyRef](n) + def apply(i: Int): A = array(i).asInstanceOf[A] + def update(i: Int, x: A): Unit = array(i) = x.asInstanceOf[AnyRef] + def length: Int = array.length + def toList: List[A] = array.toList.asInstanceOf[List[A]] + override def toString: String = array.mkString("Array.T(", ",", ")") + } + def make[A](n: BigInt)(f: BigInt => A): T[A] = + { + val m = n.toInt + val a = new T[A](m) + for (i <- 0 until m) a(i) = f(i) + a + } + def alloc[A](n: BigInt)(x: A): T[A] = make(n)(_ => x) + def len[A](a: T[A]): BigInt = BigInt(a.length) + def nth[A](a: T[A], n: BigInt): A = a(n.toInt) + def upd[A](a: T[A], n: BigInt, x: A): Unit = a.update(n.toInt, x) + def freeze[A](a: T[A]): List[A] = a.toList } \ code_reserved Scala Heap Ref Array code_printing type_constructor Heap \ (Scala) "(Unit/ =>/ _)" code_printing constant bind \ (Scala) "Heap.bind" code_printing constant return \ (Scala) "('_: Unit)/ =>/ _" code_printing constant Heap_Time_Monad.raise \ (Scala) "!sys.error((_))" subsubsection \Target variants with less units\ setup \ let open Code_Thingol; val imp_program = let val is_bind = curry (op =) @{const_name bind}; val is_return = curry (op =) @{const_name return}; val dummy_name = ""; val dummy_case_term = IVar NONE; (*assumption: dummy values are not relevant for serialization*) val unitT = @{type_name unit} `%% []; val unitt = IConst { sym = Code_Symbol.Constant @{const_name Unity}, typargs = [], dicts = [], dom = [], annotation = NONE }; fun dest_abs ((v, ty) `|=> t, _) = ((v, ty), t) | dest_abs (t, ty) = let val vs = fold_varnames cons t []; val v = singleton (Name.variant_list vs) "x"; val ty' = (hd o fst o unfold_fun) ty; in ((SOME v, ty'), t `$ IVar (SOME v)) end; fun force (t as IConst { sym = Code_Symbol.Constant c, ... } `$ t') = if is_return c then t' else t `$ unitt | force t = t `$ unitt; fun tr_bind'' [(t1, _), (t2, ty2)] = let val ((v, ty), t) = dest_abs (t2, ty2); in ICase { term = force t1, typ = ty, clauses = [(IVar v, tr_bind' t)], primitive = dummy_case_term } end and tr_bind' t = case unfold_app t of (IConst { sym = Code_Symbol.Constant c, dom = ty1 :: ty2 :: _, ... }, [x1, x2]) => if is_bind c then tr_bind'' [(x1, ty1), (x2, ty2)] else force t | _ => force t; fun imp_monad_bind'' ts = (SOME dummy_name, unitT) `|=> ICase { term = IVar (SOME dummy_name), typ = unitT, clauses = [(unitt, tr_bind'' ts)], primitive = dummy_case_term } fun imp_monad_bind' (const as { sym = Code_Symbol.Constant c, dom = dom, ... }) ts = if is_bind c then case (ts, dom) of ([t1, t2], ty1 :: ty2 :: _) => imp_monad_bind'' [(t1, ty1), (t2, ty2)] | ([t1, t2, t3], ty1 :: ty2 :: _) => imp_monad_bind'' [(t1, ty1), (t2, ty2)] `$ t3 | (ts, _) => imp_monad_bind (eta_expand 2 (const, ts)) else IConst const `$$ map imp_monad_bind ts and imp_monad_bind (IConst const) = imp_monad_bind' const [] | imp_monad_bind (t as IVar _) = t | imp_monad_bind (t as _ `$ _) = (case unfold_app t of (IConst const, ts) => imp_monad_bind' const ts | (t, ts) => imp_monad_bind t `$$ map imp_monad_bind ts) | imp_monad_bind (v_ty `|=> t) = v_ty `|=> imp_monad_bind t | imp_monad_bind (ICase { term = t, typ = ty, clauses = clauses, primitive = t0 }) = ICase { term = imp_monad_bind t, typ = ty, clauses = (map o apply2) imp_monad_bind clauses, primitive = imp_monad_bind t0 }; in (Code_Symbol.Graph.map o K o map_terms_stmt) imp_monad_bind end; in Code_Target.add_derived_target ("SML_imp", [("SML", imp_program)]) #> Code_Target.add_derived_target ("OCaml_imp", [("OCaml", imp_program)]) #> Code_Target.add_derived_target ("Scala_imp", [("Scala", imp_program)]) end \ hide_const (open) Heap heap guard (* fold_map TODO *) (*two additional lemmas, not by Haslbeck and Zhan, lemmas just added for this project*) lemma fold_if_return: "(if b then return c else return d) = return (if b then c else d)" by simp lemma distrib_if_bind: "do { x \ if b then (c::_ Heap) else d; f x } = (if b then do {x \ c; f x} else do { x\d; f x })" by simp lemmas heap_monad_laws = bind_return return_bind bind_bind end