(* * Copyright 2023, Proofcraft Pty Ltd * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) theory NonDetMonadLemmaBucket imports Lib Monads.More_NonDetMonadVCG Monad_Lists Monads.Monad_Equations Monad_Commute Monads.No_Fail Monads.No_Throw CutMon Oblivious Injection_Handler Monads.WhileLoopRulesCompleteness "Word_Lib.Distinct_Prop" (* for distinct_tuple_helper *) Monads.OptionMonadWP begin lemma distinct_tuple_helper: "\P\ f \\rv s. distinct (x # xs rv s)\ \ \P\ f \\rv s. distinct (x # (map fst (zip (xs rv s) (ys rv s))))\" apply (erule hoare_strengthen_post) apply (erule distinct_prefix) apply (simp add: map_fst_zip_prefix) done lemma gets_the_validE_R_wp: "\\s. f s \ None \ isRight (the (f s)) \ Q (theRight (the (f s))) s\ gets_the f \Q\,-" apply (simp add: gets_the_def validE_R_def validE_def) apply (wp | wpc | simp add: assert_opt_def)+ apply (clarsimp split: split: sum.splits) done lemma return_bindE: "isRight v \ return v >>=E f = f (theRight v)" by (cases v; clarsimp simp: return_returnOk) lemma list_case_return: (* not in Lib, because "return" is not in scope there *) "(case xs of [] \ return v | y # ys \ return (f y ys)) = return (case xs of [] \ v | y # ys \ f y ys)" by (simp split: list.split) (* We use isLeft, because isLeft=isl is the primitive concept; isRight=\isl matches on isl. *) lemma valid_isLeft_theRight_split: "\P\ f \\rv s. Q False rv s\,\\e s. \v. Q True v s\ \ \P\ f \\rv s. Q (isLeft rv) (theRight rv) s\" apply (simp add: validE_def) apply (erule hoare_strengthen_post) apply (simp split: sum.split_asm) done (* depends on Lib.list_induct_suffix *) lemma mapM_x_inv_wp2: assumes post: "\s. \ I s; V [] s \ \ Q s" and hr: "\a as. suffix (a # as) xs \ \\s. I s \ V (a # as) s\ m a \\r s. I s \ V as s\" shows "\I and V xs\ mapM_x m xs \\rv. Q\" proof (induct xs rule: list_induct_suffix) case Nil thus ?case apply (simp add: mapM_x_Nil) apply wp apply (clarsimp intro!: post) done next case (Cons x xs) thus ?case apply (simp add: mapM_x_Cons) apply wp apply (wp hr) apply assumption done qed lemma gets_the_exs_valid_no_ofail: "\no_ofail P h; ovalid P h Q\ \ \P\ gets_the h \\Q\" apply (rule exs_valid_weaken_pre) apply (rule gets_the_exs_valid) apply (fastforce simp: ovalid_def no_ofail_def) done end