(* * Copyright 2023, Proofcraft Pty Ltd * * SPDX-License-Identifier: BSD-2-Clause *) (* Definition of injection_handler and supporting lemmas. *) theory Injection_Handler imports Monads.Nondet_VCG begin definition injection_handler :: "('a \ 'b) \ ('s, 'a + 'c) nondet_monad \ ('s, 'b + 'c) nondet_monad" where "injection_handler f m \ m (\ft. throwError (f ft))" lemma injection_wp: "\ t = injection_handler f; \P\ m \Q\,\\ft. E (f ft)\ \ \ \P\ t m \Q\,\E\" unfolding injection_handler_def by wpsimp lemma injection_wp_E: "\ t = injection_handler f; \P\ m \Q\,- \ \ \P\ t m \Q\,-" by (simp add: validE_R_def injection_wp) lemma injection_bindE: "\ t = injection_handler f; t2 = injection_handler f \ \ t (x >>=E y) = (t2 x) >>=E (\rv. t (y rv))" apply (simp add: injection_handler_def bindE_def handleE'_def bind_assoc) apply (rule arg_cong [where f="bind x"]) apply (fastforce simp: lift_def throwError_def split: sum.splits) done lemma injection_liftE: "t = injection_handler f \ t (liftE x) = liftE x" by (simp add: injection_handler_def handleE'_def liftE_def) lemma id_injection: "id = injection_handler id" proof - have P: "case_sum throwError (\v. return (Inr v)) = return" by (auto simp: throwError_def split: sum.splits) show ?thesis by (auto simp: injection_handler_def handleE'_def P) qed lemma injection_handler_assertE: "injection_handler inject (assertE f) = assertE f" by (simp add: assertE_liftE injection_liftE) end