(* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) theory TraceMonad imports Lib Strengthen begin text \ The ``Interference Trace Monad''. This nondeterministic monad records the state at every interference point, permitting nondeterministic interference by the environment at interference points. The trace set initially includes all possible environment behaviours. Trace steps are tagged as environment or self actions, and can then be constrained to a smaller set where the environment acts according to a rely constraint (i.e. rely-guarantee reasoning), or to set the environment actions to be the self actions of another program (parallel composition). \ section "The Trace Monad" text \Trace monad identifier. Me corresponds to the current thread running and Env to the environment.\ datatype tmid = Me | Env text \Results associated with traces. Traces may correspond to incomplete, failed, or completed executions.\ datatype ('s, 'a) tmres = Failed | Incomplete | Result "('a \ 's)" abbreviation map_tmres_rv :: "('a \ 'b) \ ('s, 'a) tmres \ ('s, 'b) tmres" where "map_tmres_rv f \ map_tmres id f" section "The Monad" text \ tmonad returns a set of non-deterministic computations, including a trace as a list of "thread identifier" \ state, and an optional pair result, state when the computation did not fail. \ type_synonym ('s, 'a) tmonad = "'s \ ((tmid \ 's) list \ ('s, 'a) tmres) set" text \Returns monad results, ignoring failures and traces.\ definition mres :: "((tmid \ 's) list \ ('s, 'a) tmres) set \ ('a \ 's) set" where "mres r = Result -` (snd ` r)" text \ The definition of fundamental monad functions @{text return} and @{text bind}. The monad function @{text "return x"} does not change the state, does not fail, and returns @{text "x"}. \ definition return :: "'a \ ('s,'a) tmonad" where "return a \ \s. ({([], Result (a, s))})" text \ The monad function @{text "bind f g"}, also written @{text "f >>= g"}, is the execution of @{term f} followed by the execution of @{text g}. The function @{text g} takes the result value \emph{and} the result state of @{text f} as parameter. The definition says that the result of the combined operation is the union of the set of sets that is created by @{text g} applied to the result sets of @{text f}. The combined operation may have failed, if @{text f} may have failed or @{text g} may have failed on any of the results of @{text f}. \ abbreviation (input) fst_upd :: "('a \ 'c) \ 'a \ 'b \ 'c \ 'b" where "fst_upd f \ \(a,b). (f a, b)" abbreviation (input) snd_upd :: "('b \ 'c) \ 'a \ 'b \ 'a \ 'c" where "snd_upd f \ \(a,b). (a, f b)" definition bind :: "('s, 'a) tmonad \ ('a \ ('s, 'b) tmonad) \ ('s, 'b) tmonad" (infixl ">>=" 60) where "bind f g \ \s. \(xs, r) \ (f s). case r of Failed \ {(xs, Failed)} | Incomplete \ {(xs, Incomplete)} | Result (rv, s) \ fst_upd (\ys. ys @ xs) ` g rv s" text \Sometimes it is convenient to write @{text bind} in reverse order.\ abbreviation(input) bind_rev :: "('c \ ('a, 'b) tmonad) \ ('a, 'c) tmonad \ ('a, 'b) tmonad" (infixl "=<<" 60) where "g =<< f \ f >>= g" text \ The basic accessor functions of the state monad. @{text get} returns the current state as result, does not fail, and does not change the state. @{text "put s"} returns nothing (@{typ unit}), changes the current state to @{text s} and does not fail. \ definition get :: "('s,'s) tmonad" where "get \ \s. {([], Result (s, s))}" definition put :: "'s \ ('s, unit) tmonad" where "put s \ \st. {([], Result ((), s))}" definition put_trace_elem :: "(tmid \ 's) \ ('s, unit) tmonad" where "put_trace_elem x = (\s. {([], Incomplete), ([x], Result ((), s))})" primrec put_trace :: "(tmid \ 's) list \ ('s, unit) tmonad" where "put_trace [] = return ()" | "put_trace (x # xs) = (put_trace xs >>= (\_. put_trace_elem x))" subsection "Nondeterminism" text \ Basic nondeterministic functions. @{text "select A"} chooses an element of the set @{text A}, does not change the state, and does not fail (even if the set is empty). @{text "f OR g"} executes @{text f} or executes @{text g}. It retuns the union of results of @{text f} and @{text g}, and may have failed if either may have failed. \ definition select :: "'a set \ ('s, 'a) tmonad" where (* Should we have Failed when A = {} ? *) "select A \ \s. (Pair [] ` Result ` (A \ {s}))" definition alternative :: "('s,'a) tmonad \ ('s,'a) tmonad \ ('s,'a) tmonad" (infixl "OR" 20) where "f OR g \ \s. (f s \ g s)" text \ Alternative notation for @{text OR} \ notation (xsymbols) alternative (infixl "\" 20) text \ The @{text selet_f} function was left out here until we figure out what variant we actually need. \ subsection "Failure" text \ The monad function that always fails. Returns an empty set of results and sets the failure flag. \ definition fail :: "('s, 'a) tmonad" where "fail \ \s. {([], Failed)}" text \ Assertions: fail if the property @{text P} is not true \ definition assert :: "bool \ ('a, unit) tmonad" where "assert P \ if P then return () else fail" text \ Fail if the value is @{const None}, return result @{text v} for @{term "Some v"} \ definition assert_opt :: "'a option \ ('b, 'a) tmonad" where "assert_opt v \ case v of None \ fail | Some v \ return v" text \ An assertion that also can introspect the current state. \ definition state_assert :: "('s \ bool) \ ('s, unit) tmonad" where "state_assert P \ get >>= (\s. assert (P s))" subsection "Generic functions on top of the state monad" text \ Apply a function to the current state and return the result without changing the state. \ definition gets :: "('s \ 'a) \ ('s, 'a) tmonad" where "gets f \ get >>= (\s. return (f s))" text \ Modify the current state using the function passed in. \ definition modify :: "('s \ 's) \ ('s, unit) tmonad" where "modify f \ get >>= (\s. put (f s))" lemma simpler_gets_def: "gets f = (\s. {([], Result ((f s), s))})" by (simp add: fun_eq_iff gets_def return_def bind_def get_def split_def) lemma simpler_modify_def: "modify f = (\s. {([], Result ((),(f s)))})" by (simp add: fun_eq_iff modify_def bind_def get_def put_def split_def) text \ Execute the given monad when the condition is true, return @{text "()"} otherwise. \ definition "when" :: "bool \ ('s, unit) tmonad \ ('s, unit) tmonad" where "when P m \ if P then m else return ()" text \ Execute the given monad unless the condition is true, return @{text "()"} otherwise. \ definition unless :: "bool \ ('s, unit) tmonad \ ('s, unit) tmonad" where "unless P m \ when (\P) m" text \ Perform a test on the current state, performing the left monad if the result is true or the right monad if the result is false. \ definition condition :: "('s \ bool) \ ('s, 'r) tmonad \ ('s, 'r) tmonad \ ('s, 'r) tmonad" where "condition P L R \ \s. if (P s) then (L s) else (R s)" notation (output) condition ("(condition (_)// (_)// (_))" [1000,1000,1000] 1000) text \ Apply an option valued function to the current state, fail if it returns @{const None}, return @{text v} if it returns @{term "Some v"}. \ definition gets_the :: "('s \ 'a option) \ ('s, 'a) tmonad" where "gets_the f \ gets f >>= assert_opt" subsection \ The Monad Laws \ text \An alternative definition of bind, sometimes more convenient.\ lemma bind_def2: "bind f g \ (\s. ((\xs. (xs, Failed)) ` {xs. (xs, Failed) \ f s}) \ ((\xs. (xs, Incomplete)) ` {xs. (xs, Incomplete) \ f s}) \ (\(xs, rv, s) \ {(xs, rv, s'). (xs, Result (rv, s')) \ f s}. fst_upd (\ys. ys @ xs) ` g rv s))" apply (clarsimp simp add: bind_def fun_eq_iff Un_Union_image split_def intro!: eq_reflection) apply (auto split: tmres.splits elim!:rev_bexI[where A="f x" for x]) apply (fastforce intro: image_eqI[rotated]) done lemma elem_bindE: "(tr, res) \ bind f (\x. g x) s \ ((res = Incomplete | res = Failed) \ (tr, map_tmres undefined undefined res) \ f s \ P) \ (\tr' tr'' x s'. (tr', Result (x, s')) \ f s \ (tr'', res) \ g x s' \ tr = tr'' @ tr' \ P) \ P" by (auto simp: bind_def2) text \ Each monad satisfies at least the following three laws. \ text \ @{term return} is absorbed at the left of a @{term bind}, applying the return value directly: \ declare map_option.identity[simp] lemma return_bind [simp]: "(return x >>= f) = f x" by (auto simp add: return_def bind_def split_def split:if_splits) text \ @{term return} is absorbed on the right of a @{term bind} \ lemma bind_return [simp]: "(m >>= return) = m" by (auto simp add: fun_eq_iff bind_def return_def split: tmres.splits) text \ @{term bind} is associative \ lemma bind_assoc: fixes m :: "('a,'b) tmonad" fixes f :: "'b \ ('a,'c) tmonad" fixes g :: "'c \ ('a,'d) tmonad" shows "(m >>= f) >>= g = m >>= (\x. f x >>= g)" apply (unfold bind_def Let_def split_def) apply (rule ext) apply clarsimp apply (rule SUP_cong[OF refl], clarsimp) apply (split tmres.split; intro conjI impI; clarsimp) apply (simp add: image_Union) apply (rule SUP_cong[OF refl], clarsimp) apply (split tmres.split; intro conjI impI; clarsimp) apply (simp add: image_image) done section \ Adding Exceptions \ text \ The type @{typ "('s,'a) tmonad"} gives us nondeterminism and failure. We now extend this monad with exceptional return values that abort normal execution, but can be handled explicitly. We use the sum type to indicate exceptions. In @{typ "('s, 'e + 'a) tmonad"}, @{typ "'s"} is the state, @{typ 'e} is an exception, and @{typ 'a} is a normal return value. This new type itself forms a monad again. Since type classes in Isabelle are not powerful enough to express the class of monads, we provide new names for the @{term return} and @{term bind} functions in this monad. We call them @{text returnOk} (for normal return values) and @{text bindE} (for composition). We also define @{text throwError} to return an exceptional value. \ definition returnOk :: "'a \ ('s, 'e + 'a) tmonad" where "returnOk \ return o Inr" definition throwError :: "'e \ ('s, 'e + 'a) tmonad" where "throwError \ return o Inl" text \ Lifting a function over the exception type: if the input is an exception, return that exception; otherwise continue execution. \ definition lift :: "('a \ ('s, 'e + 'b) tmonad) \ 'e +'a \ ('s, 'e + 'b) tmonad" where "lift f v \ case v of Inl e \ throwError e | Inr v' \ f v'" text \ The definition of @{term bind} in the exception monad (new name @{text bindE}): the same as normal @{term bind}, but the right-hand side is skipped if the left-hand side produced an exception. \ definition bindE :: "('s, 'e + 'a) tmonad \ ('a \ ('s, 'e + 'b) tmonad) \ ('s, 'e + 'b) tmonad" (infixl ">>=E" 60) where "bindE f g \ bind f (lift g)" text \ Lifting a normal nondeterministic monad into the exception monad is achieved by always returning its result as normal result and never throwing an exception. \ definition liftE :: "('s,'a) tmonad \ ('s, 'e+'a) tmonad" where "liftE f \ f >>= (\r. return (Inr r))" text \ Since the underlying type and @{text return} function changed, we need new definitions for when and unless: \ definition whenE :: "bool \ ('s, 'e + unit) tmonad \ ('s, 'e + unit) tmonad" where "whenE P f \ if P then f else returnOk ()" definition unlessE :: "bool \ ('s, 'e + unit) tmonad \ ('s, 'e + unit) tmonad" where "unlessE P f \ if P then returnOk () else f" text \ Throwing an exception when the parameter is @{term None}, otherwise returning @{term "v"} for @{term "Some v"}. \ definition throw_opt :: "'e \ 'a option \ ('s, 'e + 'a) tmonad" where "throw_opt ex x \ case x of None \ throwError ex | Some v \ returnOk v" text \ Failure in the exception monad is redefined in the same way as @{const whenE} and @{const unlessE}, with @{term returnOk} instead of @{term return}. \ definition assertE :: "bool \ ('a, 'e + unit) tmonad" where "assertE P \ if P then returnOk () else fail" subsection "Monad Laws for the Exception Monad" text \ More direct definition of @{const liftE}: \ lemma liftE_def2: "liftE f = (\s. snd_upd (map_tmres_rv Inr) ` (f s))" apply (clarsimp simp: fun_eq_iff liftE_def return_def split_def bind_def image_def) apply (rule set_eqI) apply (rule iffI) apply clarsimp apply (erule rev_bexI[where A="f s" for s]) apply (clarsimp split: tmres.splits) apply clarsimp apply (rule exI) apply (rule conjI) apply (erule rev_bexI[where A="f s" for s]) apply (rule refl) apply (clarsimp split: tmres.splits) done text \ Left @{const returnOk} absorbtion over @{term bindE}: \ lemma returnOk_bindE [simp]: "(returnOk x >>=E f) = f x" apply (unfold bindE_def returnOk_def) apply (clarsimp simp: lift_def) done lemma lift_return [simp]: "lift (return \ Inr) = return" by (simp add: fun_eq_iff lift_def throwError_def split: sum.splits) text \ Right @{const returnOk} absorbtion over @{term bindE}: \ lemma bindE_returnOk [simp]: "(m >>=E returnOk) = m" by (simp add: bindE_def returnOk_def) text \ Associativity of @{const bindE}: \ lemma bindE_assoc: "(m >>=E f) >>=E g = m >>=E (\x. f x >>=E g)" apply (simp add: bindE_def bind_assoc) apply (rule arg_cong [where f="\x. m >>= x"]) apply (rule ext) apply (case_tac x, simp_all add: lift_def throwError_def) done text \ @{const returnOk} could also be defined via @{const liftE}: \ lemma returnOk_liftE: "returnOk x = liftE (return x)" by (simp add: liftE_def returnOk_def) text \ Execution after throwing an exception is skipped: \ lemma throwError_bindE [simp]: "(throwError E >>=E f) = throwError E" by (simp add: fun_eq_iff bindE_def bind_def throwError_def lift_def return_def split_def) section "Syntax" text \ This section defines traditional Haskell-like do-syntax for the state monad in Isabelle. \ subsection "Syntax for the Nondeterministic State Monad" text \ We use @{text K_bind} to syntactically indicate the case where the return argument of the left side of a @{term bind} is ignored \ definition K_bind_def [iff]: "K_bind \ \x y. x" nonterminal dobinds and dobind and nobind syntax "_dobind" :: "[pttrn, 'a] => dobind" ("(_ <-/ _)" 10) "" :: "dobind => dobinds" ("_") "_nobind" :: "'a => dobind" ("_") "_dobinds" :: "[dobind, dobinds] => dobinds" ("(_);//(_)") "_do" :: "[dobinds, 'a] => 'a" ("(do ((_);//(_))//od)" 100) syntax (xsymbols) "_dobind" :: "[pttrn, 'a] => dobind" ("(_ \/ _)" 10) translations "_do (_dobinds b bs) e" == "_do b (_do bs e)" "_do (_nobind b) e" == "b >>= (CONST K_bind e)" "do x <- a; e od" == "a >>= (\x. e)" text \ Syntax examples: \ lemma "do x \ return 1; return (2::nat); return x od = return 1 >>= (\x. return (2::nat) >>= K_bind (return x))" by (rule refl) lemma "do x \ return 1; return 2; return x od = return 1" by simp subsection "Interference command" text \Interference commands must be inserted in between actions that can be interfered with commands running in other threads. \ definition last_st_tr :: "(tmid * 's) list \ 's \ 's" where "last_st_tr tr s0 = (hd (map snd tr @ [s0]))" definition env_steps :: "('s,unit) tmonad" where "env_steps \ do s \ get; \ \Add unfiltered environment events to the trace\ xs \ select UNIV; tr \ return (map (Pair Env) xs); put_trace tr; \ \Pick the last event of the trace as the final state\ put (last_st_tr tr s) od" definition commit_step :: "('s, unit) tmonad" where "commit_step \ do s \ get; put_trace [(Me,s)] od" definition interference :: "('s,unit) tmonad" where "interference \ do commit_step; env_steps od" subsection "Syntax for the Exception Monad" text \ Since the exception monad is a different type, we need to syntactically distinguish it in the syntax. We use @{text doE}/@{text odE} for this, but can re-use most of the productions from @{text do}/@{text od} above. \ syntax "_doE" :: "[dobinds, 'a] => 'a" ("(doE ((_);//(_))//odE)" 100) translations "_doE (_dobinds b bs) e" == "_doE b (_doE bs e)" "_doE (_nobind b) e" == "b >>=E (CONST K_bind e)" "doE x <- a; e odE" == "a >>=E (\x. e)" text \ Syntax examples: \ lemma "doE x \ returnOk 1; returnOk (2::nat); returnOk x odE = returnOk 1 >>=E (\x. returnOk (2::nat) >>=E K_bind (returnOk x))" by (rule refl) lemma "doE x \ returnOk 1; returnOk 2; returnOk x odE = returnOk 1" by simp section "Library of additional Monadic Functions and Combinators" text \ Lifting a normal function into the monad type: \ definition liftM :: "('a \ 'b) \ ('s,'a) tmonad \ ('s, 'b) tmonad" where "liftM f m \ do x \ m; return (f x) od" text \ The same for the exception monad: \ definition liftME :: "('a \ 'b) \ ('s,'e+'a) tmonad \ ('s,'e+'b) tmonad" where "liftME f m \ doE x \ m; returnOk (f x) odE" text \ Run a sequence of monads from left to right, ignoring return values. \ definition sequence_x :: "('s, 'a) tmonad list \ ('s, unit) tmonad" where "sequence_x xs \ foldr (\x y. x >>= (\_. y)) xs (return ())" text \ Map a monadic function over a list by applying it to each element of the list from left to right, ignoring return values. \ definition mapM_x :: "('a \ ('s,'b) tmonad) \ 'a list \ ('s, unit) tmonad" where "mapM_x f xs \ sequence_x (map f xs)" text \ Map a monadic function with two parameters over two lists, going through both lists simultaneously, left to right, ignoring return values. \ definition zipWithM_x :: "('a \ 'b \ ('s,'c) tmonad) \ 'a list \ 'b list \ ('s, unit) tmonad" where "zipWithM_x f xs ys \ sequence_x (zipWith f xs ys)" text \ The same three functions as above, but returning a list of return values instead of @{text unit} \ definition sequence :: "('s, 'a) tmonad list \ ('s, 'a list) tmonad" where "sequence xs \ let mcons = (\p q. p >>= (\x. q >>= (\y. return (x#y)))) in foldr mcons xs (return [])" definition mapM :: "('a \ ('s,'b) tmonad) \ 'a list \ ('s, 'b list) tmonad" where "mapM f xs \ sequence (map f xs)" definition zipWithM :: "('a \ 'b \ ('s,'c) tmonad) \ 'a list \ 'b list \ ('s, 'c list) tmonad" where "zipWithM f xs ys \ sequence (zipWith f xs ys)" definition foldM :: "('b \ 'a \ ('s, 'a) tmonad) \ 'b list \ 'a \ ('s, 'a) tmonad" where "foldM m xs a \ foldr (\p q. q >>= m p) xs (return a) " definition foldME ::"('b \ 'a \ ('s,('e + 'b)) tmonad) \ 'b \ 'a list \ ('s, ('e + 'b)) tmonad" where "foldME m a xs \ foldr (\p q. q >>=E swp m p) xs (returnOk a)" text \ The sequence and map functions above for the exception monad, with and without lists of return value \ definition sequenceE_x :: "('s, 'e+'a) tmonad list \ ('s, 'e+unit) tmonad" where "sequenceE_x xs \ foldr (\x y. doE _ <- x; y odE) xs (returnOk ())" definition mapME_x :: "('a \ ('s,'e+'b) tmonad) \ 'a list \ ('s,'e+unit) tmonad" where "mapME_x f xs \ sequenceE_x (map f xs)" definition sequenceE :: "('s, 'e+'a) tmonad list \ ('s, 'e+'a list) tmonad" where "sequenceE xs \ let mcons = (\p q. p >>=E (\x. q >>=E (\y. returnOk (x#y)))) in foldr mcons xs (returnOk [])" definition mapME :: "('a \ ('s,'e+'b) tmonad) \ 'a list \ ('s,'e+'b list) tmonad" where "mapME f xs \ sequenceE (map f xs)" text \ Filtering a list using a monadic function as predicate: \ primrec filterM :: "('a \ ('s, bool) tmonad) \ 'a list \ ('s, 'a list) tmonad" where "filterM P [] = return []" | "filterM P (x # xs) = do b <- P x; ys <- filterM P xs; return (if b then (x # ys) else ys) od" text \ @{text select_state} takes a relationship between states, and outputs nondeterministically a state related to the input state. \ definition state_select :: "('s \ 's) set \ ('s, unit) tmonad" where "state_select r \ (do s \ get; S \ return {s'. (s, s') \ r}; assert (S \ {}); s' \ select S; put s' od)" section "Catching and Handling Exceptions" text \ Turning an exception monad into a normal state monad by catching and handling any potential exceptions: \ definition catch :: "('s, 'e + 'a) tmonad \ ('e \ ('s, 'a) tmonad) \ ('s, 'a) tmonad" (infix "" 10) where "f handler \ do x \ f; case x of Inr b \ return b | Inl e \ handler e od" text \ Handling exceptions, but staying in the exception monad. The handler may throw a type of exceptions different from the left side. \ definition handleE' :: "('s, 'e1 + 'a) tmonad \ ('e1 \ ('s, 'e2 + 'a) tmonad) \ ('s, 'e2 + 'a) tmonad" (infix "" 10) where "f handler \ do v \ f; case v of Inl e \ handler e | Inr v' \ return (Inr v') od" text \ A type restriction of the above that is used more commonly in practice: the exception handle (potentially) throws exception of the same type as the left-hand side. \ definition handleE :: "('s, 'x + 'a) tmonad \ ('x \ ('s, 'x + 'a) tmonad) \ ('s, 'x + 'a) tmonad" (infix "" 10) where "handleE \ handleE'" text \ Handling exceptions, and additionally providing a continuation if the left-hand side throws no exception: \ definition handle_elseE :: "('s, 'e + 'a) tmonad \ ('e \ ('s, 'ee + 'b) tmonad) \ ('a \ ('s, 'ee + 'b) tmonad) \ ('s, 'ee + 'b) tmonad" ("_ _ _" 10) where "f handler continue \ do v \ f; case v of Inl e \ handler e | Inr v' \ continue v' od" subsection "Loops" text \ Loops are handled using the following inductive predicate; non-termination is represented using the failure flag of the monad. \ inductive_set whileLoop_results :: "('r \ 's \ bool) \ ('r \ ('s, 'r) tmonad) \ (('r \ 's) \ ((tmid \ 's) list \ ('s, 'r) tmres)) set" for C B where "\ \ C r s \ \ ((r, s), ([], Result (r, s))) \ whileLoop_results C B" | "\ C r s; (ts, Failed) \ B r s \ \ ((r, s), (ts, Failed)) \ whileLoop_results C B" | "\ C r s; (ts, Incomplete) \ B r s \ \ ((r, s), (ts, Incomplete)) \ whileLoop_results C B" | "\ C r s; (ts, Result (r', s')) \ B r s; ((r', s'), (ts',z)) \ whileLoop_results C B \ \ ((r, s), (ts'@ts,z)) \ whileLoop_results C B" inductive_cases whileLoop_results_cases_result_end: "((x,y), ([],Result r)) \ whileLoop_results C B" inductive_cases whileLoop_results_cases_fail: "((x,y), (ts, Failed)) \ whileLoop_results C B" inductive_cases whileLoop_results_cases_incomplete: "((x,y), (ts, Incomplete)) \ whileLoop_results C B" inductive_simps whileLoop_results_simps_valid: "((x,y), ([], Result z)) \ whileLoop_results C B" inductive whileLoop_terminates :: "('r \ 's \ bool) \ ('r \ ('s, 'r) tmonad) \ 'r \ 's \ bool" for C B where "\ C r s \ whileLoop_terminates C B r s" | "\ C r s; \(r', s') \ Result -` snd ` (B r s). whileLoop_terminates C B r' s' \ \ whileLoop_terminates C B r s" inductive_cases whileLoop_terminates_cases: "whileLoop_terminates C B r s" inductive_simps whileLoop_terminates_simps: "whileLoop_terminates C B r s" definition whileLoop :: "('r \ 's \ bool) \ ('r \ ('s, 'r) tmonad) \ 'r \ ('s, 'r) tmonad" where "whileLoop C B \ (\r s. {(ts, res). ((r,s), ts,res) \ whileLoop_results C B})" notation (output) whileLoop ("(whileLoop (_)// (_))" [1000, 1000] 1000) definition whileLoopT :: "('r \ 's \ bool) \ ('r \ ('s, 'r) tmonad) \ 'r \ ('s, 'r) tmonad" where "whileLoopT C B \ (\r s. {(ts, res). ((r,s), ts,res) \ whileLoop_results C B \ whileLoop_terminates C B r s})" notation (output) whileLoopT ("(whileLoopT (_)// (_))" [1000, 1000] 1000) definition whileLoopE :: "('r \ 's \ bool) \ ('r \ ('s, 'e + 'r) tmonad) \ 'r \ ('s, ('e + 'r)) tmonad" where "whileLoopE C body \ \r. whileLoop (\r s. (case r of Inr v \ C v s | _ \ False)) (lift body) (Inr r)" notation (output) whileLoopE ("(whileLoopE (_)// (_))" [1000, 1000] 1000) subsection "Await command" text \ @{term "Await c f"} blocks the execution until the @{term "c"} is true, and atomically executes @{term "f"}. \ definition Await :: "('s \ bool) \ ('s,unit) tmonad" where "Await c \ do s \ get; \ \Add unfiltered environment events, with the last one satisfying the `c' state predicate\ xs \ select {xs. c (last_st_tr (map (Pair Env) xs) s)}; tr \ return (map (Pair Env) xs); put_trace tr; \ \Pick the last event of the trace\ put (last_st_tr tr s) od" section "Hoare Logic" subsection "Validity" text \ This section defines a Hoare logic for partial correctness for the nondeterministic state monad as well as the exception monad. The logic talks only about the behaviour part of the monad and ignores the failure flag. The logic is defined semantically. Rules work directly on the validity predicate. In the nondeterministic state monad, validity is a triple of precondition, monad, and postcondition. The precondition is a function from state to bool (a state predicate), the postcondition is a function from return value to state to bool. A triple is valid if for all states that satisfy the precondition, all result values and result states that are returned by the monad satisfy the postcondition. Note that if the computation returns the empty set, the triple is trivially valid. This means @{term "assert P"} does not require us to prove that @{term P} holds, but rather allows us to assume @{term P}! Proving non-failure is done via separate predicate and calculus (see below). \ definition valid :: "('s \ bool) \ ('s,'a) tmonad \ ('a \ 's \ bool) \ bool" ("\_\/ _ /\_\") where "\P\ f \Q\ \ \s. P s \ (\(r,s') \ mres (f s). Q r s')" text \ We often reason about invariant predicates. The following provides shorthand syntax that avoids repeating potentially long predicates. \ abbreviation (input) invariant :: "('s,'a) tmonad \ ('s \ bool) \ bool" ("_ \_\" [59,0] 60) where "invariant f P \ \P\ f \\_. P\" text \rg_pred type: Rely-Guaranty predicates (state before => state after => bool)\ type_synonym 's rg_pred = "'s \ 's \ bool" text \ Validity for the exception monad is similar and build on the standard validity above. Instead of one postcondition, we have two: one for normal and one for exceptional results. \ definition validE :: "('s \ bool) \ ('s, 'a + 'b) tmonad \ ('b \ 's \ bool) \ ('a \ 's \ bool) \ bool" ("\_\/ _ /(\_\,/ \_\)" ) where "\P\ f \Q\,\E\ \ \P\ f \ \v s. case v of Inr r \ Q r s | Inl e \ E e s \" (* text \ Validity for exception monad with interferences. Not as easy to phrase as we need to \ definition validIE :: "('s, 'a + 'b) tmonad \ 's rg_pred \ 's rg_pred \ 's rg_pred \ ('b \ 's rg_pred) \ ('a \ 's rg_pred) \ bool" ("_ //PRE _//RELY _//GUAR _//POST _//EXC _" [59,0,0,0,0,0] 60) where "validIE f P R G Q E \ f SAT [P,R,G,\v. case v of Inr r \ Q r | Inl e \ E e]" abbreviation (input) validIEsat :: "('s, 'a + 'b) tmonad \ 's rg_pred \ 's rg_pred \ 's rg_pred \ ('b \ 's rg_pred) \ ('a \ 's rg_pred) \ bool" ("_ //SAT [_, _, _, _, _]" [59,0,0,0,0,0] 60) where "validIEsat f P R G Q E \ validIE f P R G Q E" *) text \ The following two instantiations are convenient to separate reasoning for exceptional and normal case. \ definition validE_R :: "('s \ bool) \ ('s, 'e + 'a) tmonad \ ('a \ 's \ bool) \ bool" ("\_\/ _ /\_\, -") where "\P\ f \Q\,- \ validE P f Q (\x y. True)" definition validE_E :: "('s \ bool) \ ('s, 'e + 'a) tmonad \ ('e \ 's \ bool) \ bool" ("\_\/ _ /-, \_\") where "\P\ f -,\Q\ \ validE P f (\x y. True) Q" text \ Abbreviations for trivial preconditions: \ abbreviation(input) top :: "'a \ bool" ("\") where "\ \ \_. True" abbreviation(input) bottom :: "'a \ bool" ("\") where "\ \ \_. False" text \ Abbreviations for trivial postconditions (taking two arguments): \ abbreviation(input) toptop :: "'a \ 'b \ bool" ("\\") where "\\ \ \_ _. True" abbreviation(input) toptoptop :: "'a \ 'b \ 'b \ bool" ("\\\") where "\\\ \ \_ _ _. True" abbreviation(input) botbot :: "'a \ 'b \ bool" ("\\") where "\\ \ \_ _. False" abbreviation(input) botbotbot :: "'a \ 'b \ 'b \ bool" ("\\\") where "\\\ \ \_ _ _. False" text \ Lifting @{text "\"} and @{text "\"} over two arguments. Lifting @{text "\"} and @{text "\"} over one argument is already defined (written @{text "and"} and @{text "or"}). \ definition bipred_conj :: "('a \ 'b \ bool) \ ('a \ 'b \ bool) \ ('a \ 'b \ bool)" (infixl "And" 96) where "bipred_conj P Q \ \x y. P x y \ Q x y" definition bipred_disj :: "('a \ 'b \ bool) \ ('a \ 'b \ bool) \ ('a \ 'b \ bool)" (infixl "Or" 91) where "bipred_disj P Q \ \x y. P x y \ Q x y" subsection "Determinism" text \ A monad of type @{text tmonad} is deterministic iff it returns an empty trace, exactly one state and result and does not fail \ definition det :: "('a,'s) tmonad \ bool" where "det f \ \s. \r. f s = {([], Result r)}" text \ A deterministic @{text tmonad} can be turned into a normal state monad: \ definition the_run_state :: "('s,'a) tmonad \ 's \ 'a \ 's" where "the_run_state M \ \s. THE s'. mres (M s) = {s'}" subsection "Non-Failure" text \ We can formulate non-failure separately from validity. \ definition no_fail :: "('s \ bool) \ ('s,'a) tmonad \ bool" where "no_fail P m \ \s. P s \ Failed \ snd ` (m s)" text \ It is often desired to prove non-failure and a Hoare triple simultaneously, as the reasoning is often similar. The following definitions allow such reasoning to take place. \ definition validNF ::"('s \ bool) \ ('s,'a) tmonad \ ('a \ 's \ bool) \ bool" ("\_\/ _ /\_\!") where "validNF P f Q \ valid P f Q \ no_fail P f" definition validE_NF :: "('s \ bool) \ ('s, 'a + 'b) tmonad \ ('b \ 's \ bool) \ ('a \ 's \ bool) \ bool" ("\_\/ _ /(\_\,/ \_\!)") where "validE_NF P f Q E \ validE P f Q E \ no_fail P f" lemma validE_NF_alt_def: "\ P \ B \ Q \,\ E \! = \ P \ B \ \v s. case v of Inl e \ E e s | Inr r \ Q r s \!" by (clarsimp simp: validE_NF_def validE_def validNF_def) (* text \ Usually, well-formed monads constructed from the primitives above will have the following property: if they return an empty set of results, they will have the failure flag set. \ definition empty_fail :: "('s,'a) tmonad \ bool" where "empty_fail m \ \s. fst (m s) = {} \ snd (m s)" text \ Useful in forcing otherwise unknown executions to have the @{const empty_fail} property. \ definition mk_ef :: "'a set \ bool \ 'a set \ bool" where "mk_ef S \ (fst S, fst S = {} \ snd S)" *) section "Basic exception reasoning" text \ The following predicates @{text no_throw} and @{text no_return} allow reasoning that functions in the exception monad either do no throw an exception or never return normally. \ definition "no_throw P A \ \ P \ A \ \_ _. True \,\ \_ _. False \" definition "no_return P A \ \ P \ A \\_ _. False\,\\_ _. True \" section "Trace monad Parallel" definition parallel :: "('s,'a) tmonad \ ('s,'a) tmonad \ ('s,'a) tmonad" where "parallel f g = (\s. {(xs, rv). \f_steps. length f_steps = length xs \ (map (\(f_step, (id, s)). (if f_step then id else Env, s)) (zip f_steps xs), rv) \ f s \ (map (\(f_step, (id, s)). (if f_step then Env else id, s)) (zip f_steps xs), rv) \ g s})" abbreviation(input) "parallel_mrg \ ((\((idn, s), (idn', _)). (if idn = Env then idn' else idn, s)))" lemma parallel_def2: "parallel f g = (\s. {(xs, rv). \ys zs. (ys, rv) \ f s \ (zs, rv) \ g s \ list_all2 (\y z. (fst y = Env \ fst z = Env) \ snd y = snd z) ys zs \ xs = map parallel_mrg (zip ys zs)})" apply (simp add: parallel_def fun_eq_iff set_eq_iff) apply safe apply (rule exI, rule conjI, assumption)+ apply (simp add: list_all2_conv_all_nth list_eq_iff_nth_eq split_def prod_eq_iff) apply clarsimp apply (rule_tac x="map (((\) Env) o fst) ys" in exI) apply (simp add: zip_map1 o_def split_def) apply (strengthen subst[where P="\xs. (xs, v) \ S" for v S, mk_strg I _ E]) apply (clarsimp simp: list_all2_conv_all_nth list_eq_iff_nth_eq split_def prod_eq_iff split del: if_split cong: if_cong) apply auto done lemma parallel_def3: "parallel f g = (\s. (\(ys, zs, rv). (map parallel_mrg (zip ys zs), rv)) ` {(ys, zs, rv). (ys, rv) \ f s \ (zs, rv) \ g s \ list_all2 (\y z. (fst y = Env \ fst z = Env) \ snd y = snd z) ys zs})" by (simp add: parallel_def2, rule ext, auto simp: image_def) primrec trace_steps :: "(tmid \ 's) list \ 's \ (tmid \ 's \ 's) set" where "trace_steps (elem # trace) s0 = {(fst elem, s0, snd elem)} \ trace_steps trace (snd elem)" | "trace_steps [] s0 = {}" lemma trace_steps_nth: "trace_steps xs s0 = (\i. (fst (xs ! i), (if i = 0 then s0 else snd (xs ! (i - 1))), snd (xs ! i))) ` {..< length xs}" proof (induct xs arbitrary: s0) case Nil show ?case by simp next case (Cons a xs) show ?case apply (simp add: lessThan_Suc_eq_insert_0 Cons image_image nth_Cons') apply (intro arg_cong2[where f=insert] refl image_cong) apply simp done qed definition rely_cond :: "'s rg_pred \ 's \ (tmid \ 's) list \ bool" where "rely_cond R s0s tr = (\(ident, s0, s) \ trace_steps (rev tr) s0s. ident = Env \ R s0 s)" definition guar_cond :: "'s rg_pred \ 's \ (tmid \ 's) list \ bool" where "guar_cond G s0s tr = (\(ident, s0, s) \ trace_steps (rev tr) s0s. ident = Me \ G s0 s)" lemma rg_empty_conds[simp]: "rely_cond R s0s []" "guar_cond G s0s []" by (simp_all add: rely_cond_def guar_cond_def) definition rely :: "('s, 'a) tmonad \ 's rg_pred \ 's \ ('s, 'a) tmonad" where "rely f R s0s \ (\s. f s \ ({tr. rely_cond R s0s tr} \ UNIV))" definition prefix_closed :: "('s, 'a) tmonad \ bool" where "prefix_closed f = (\s. \x xs. (x # xs) \ fst ` f s \ (xs, Incomplete) \ f s)" definition validI :: "('s \ 's \ bool) \ 's rg_pred \ ('s,'a) tmonad \ 's rg_pred \ ('a \ 's \ 's \ bool) \ bool" ("(\_\,/ \_\)/ _ /(\_\,/ \_\)") where "\P\,\R\ f \G\,\Q\ \ prefix_closed f \ (\s0 s. P s0 s \ (\tr res. (tr, res) \ (rely f R s0 s) \ guar_cond G s0 tr \ (\rv s'. res = Result (rv, s') \ Q rv (last_st_tr tr s0) s')))" lemma in_rely: "\ (tr, res) \ f s; rely_cond R s0s tr \ \ (tr, res) \ rely f R s0s s" by (simp add: rely_def) lemmas validI_D = validI_def[THEN meta_eq_to_obj_eq, THEN iffD1, THEN conjunct2, rule_format, OF _ _ in_rely] lemmas validI_GD = validI_D[THEN conjunct1] lemmas validI_rvD = validI_D[THEN conjunct2, rule_format, rotated -1, OF refl] lemmas validI_prefix_closed = validI_def[THEN meta_eq_to_obj_eq, THEN iffD1, THEN conjunct1] lemmas validI_prefix_closed_T = validI_prefix_closed[where P="\_ _. False" and R="\_ _. False" and G="\_ _. True" and Q="\_ _ _. True"] lemmas prefix_closedD1 = prefix_closed_def[THEN iffD1, rule_format] lemma in_fst_snd_image_eq: "x \ fst ` S = (\y. (x, y) \ S)" "y \ snd ` S = (\x. (x, y) \ S)" by (auto elim: image_eqI[rotated]) lemma in_fst_snd_image: "(x, y) \ S \ x \ fst ` S" "(x, y) \ S \ y \ snd ` S" by (auto simp: in_fst_snd_image_eq) lemmas prefix_closedD = prefix_closedD1[OF _ in_fst_snd_image(1)] end