(* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) theory SIMPL_Lemmas imports "Lib.HaskellLemmaBucket" "CTranslationNICTA" begin lemma hoarep_false_pre: "\\\<^bsub>/F\<^esub> {} f Q,A" by (clarsimp intro!: hoare_complete simp: HoarePartialDef.valid_def) lemma hoarep_false_pre_gen: "\s. \\\<^bsub>/F s\<^esub> {} f (Q s),(A s)" by (rule allI, rule hoarep_false_pre) lemma Cond_true: "\ \\<^bsub>/F\<^esub> P t Q, A \ \ \\<^bsub>/F\<^esub> (P \ b) Cond b t f Q, A" apply (rule hoare_complete) apply (drule hoare_sound) apply (clarsimp simp: HoarePartialDef.cvalid_def HoarePartialDef.valid_def) apply (erule exec_Normal_elim_cases) apply fastforce apply simp done lemma Cond_false: "\ \\<^bsub>/F\<^esub> P f Q, A \ \ \\<^bsub>/F\<^esub> (P \ - b) Cond b t f Q, A" apply (rule hoare_complete) apply (drule hoare_sound) apply (clarsimp simp: HoarePartialDef.cvalid_def HoarePartialDef.valid_def) apply (erule exec_Normal_elim_cases) apply simp apply fastforce done lemma update_desc_def_ext: "update_desc x y = (\d. \field_access = field_access d \ x, field_update = \bs v. y (field_update d bs (x v)) v\)" apply (rule ext) apply (simp add: update_desc_def) done lemma adjust_ti_adjust_ti: "adjust_ti (adjust_ti t xf xfu) xf' xfu' = adjust_ti t (xf \ xf') (\c a. xfu' (xfu c (xf' a)) a)" unfolding adjust_ti_def by (simp add: map_td_map comp_def update_desc_def_ext) lemma fg_cons_comp: assumes fg: "fg_cons xf xfu" and fg': "fg_cons xf' xfu'" shows "fg_cons (xf \ xf') (\c a. xfu' (xfu c (xf' a)) a)" using fg fg' unfolding fg_cons_def by clarsimp (* Generalise *) lemma hoarep_Int: fixes P' :: "'a set" and P :: "'a set" assumes s1: "\ \ P c Q" and s2: "\ \ P' c Q'" shows "\ \ (P \ P') c (Q \ Q')" using s1 s2 apply - apply (rule hoare_complete) apply (drule hoare_sound)+ apply (clarsimp simp add: cvalid_def HoarePartialDef.valid_def) apply (drule spec, drule spec, drule (1) mp) apply (drule spec, drule spec, drule (1) mp) apply fastforce done lemmas hoarep_Int_pre_fix = hoarep_Int[where P=P and P'=P for P, simplified] lemma Normal_result: "\ \ \c, s\ \ Normal t' \ \t. s = Normal t" proof (induct c arbitrary: s) case While thus ?case by - (erule exec_elim_cases, simp_all) qed (fastforce elim: exec_elim_cases)+ lemma Normal_resultE: "\ \ \ \c, s\ \ Normal t'; \t. \ \ \ \c, Normal t\ \ Normal t'; s = Normal t\ \ P \ \ P" apply (frule Normal_result) apply auto done lemma Abrupt_result: "\ \ \c, s\ \ Abrupt t' \ (\t. s = Normal t) \ s = Abrupt t'" proof (induct c arbitrary: s) case While thus ?case by - (erule exec_elim_cases, simp_all) qed (fastforce elim: exec_elim_cases)+ lemma Abrupt_resultE [consumes 1, case_names normal abrupt]: "\\ \ \c, s\ \ Abrupt t'; \t. \\ \ \c, Normal t\ \ Abrupt t'; s = Normal t\ \ P; \t. \\ \ \c, Abrupt t\ \ Abrupt t'; s = Abrupt t\ \ P\ \ P" apply (frule Abrupt_result) by auto lemma Fault_result: assumes ex: "\ \ \a, s\ \ t" and t: "t = Fault f" shows "s = Fault f \ (\s'. s = Normal s')" using ex t by induct auto lemma Fault_resultE: assumes ex: "\ \ \a, s\ \ Fault f" and r1: "\ \ \ \a, Fault f\ \ Fault f; s = Fault f \ \ P" and r2: "\s'. \ \ \ \a, Normal s'\ \ Fault f; s = Normal s'\ \ P" shows P using ex apply - apply (frule Fault_result [OF _ refl]) apply (erule disjE) apply (rule r1, simp_all) apply (erule exE) apply (rule r2, simp_all) done lemma Stuck_result: assumes ex: "\ \ \a, s\ \ t" and t: "t = Stuck" shows "s = Stuck \ (\s'. s = Normal s')" using ex t by induct auto lemma Stuck_resultE: assumes ex: "\ \ \a, s\ \ Stuck" and r1: "\ \ \ \a, Stuck\ \ Stuck; s = Stuck \ \ P" and r2: "\s'. \ \ \ \a, Normal s'\ \ Stuck; s = Normal s'\ \ P" shows P using ex apply - apply (frule Stuck_result [OF _ refl]) apply (erule disjE) apply (rule r1, simp_all) apply (erule exE) apply (rule r2, simp_all) done (* This is essentially semantic equivalence, assuming equality of xf and v at s *) definition "ceqv \ xf v s s' a b \ xf s = v \ (\ \ \a, Normal s\ \ s') = (\ \ \b, Normal s\ \ s')" lemma ceqvI: assumes rl: "xf s = v \ (\ \ \a, Normal s\ \ s') = (\ \ \b, Normal s\ \ s')" shows "ceqv \ xf v s s' a b" using rl unfolding ceqv_def by auto lemma ceqvD1: assumes lhs: "\ \ \a, Normal s\ \ s'" and xf: "xf s = v" and ceq: "ceqv \ xf v s s' a b" shows "\ \ \b, Normal s\ \ s'" using ceq xf lhs unfolding ceqv_def by auto lemma ceqvD2: assumes lhs: "\ \ \b, Normal s\ \ s'" and xf: "xf s = v" and ceq: "ceqv \ xf v s s' a b" shows "\ \ \a, Normal s\ \ s'" using ceq xf lhs unfolding ceqv_def by auto lemma ceqv_sym [sym]: "ceqv \ xf' rv' t t' c c' \ ceqv \ xf' rv' t t' c' c" unfolding ceqv_def by auto lemma exec_eq_is_valid_eq0: fixes P :: "'a set" assumes eq: "\t t'. (\ \ \a, Normal t\ \ t') = (\ \ \a', Normal t\ \ t')" and vl: "\\\<^bsub>/F\<^esub> P a Q,A" shows "\\\<^bsub>/F\<^esub> P a' Q,A" using vl apply - apply (drule hoare_sound) apply (rule hoare_complete) apply (simp add: HoarePartialDef.valid_def cvalid_def) apply (intro allI impI) apply clarsimp apply (subst (asm) eq [symmetric]) apply (drule spec, drule spec, drule (1) mp) apply simp done lemma exec_eq_is_valid_eq: fixes P :: "'a set" assumes eq: "\t t'. (\ \ \a, Normal t\ \ t') = (\ \ \a', Normal t\ \ t')" shows vl: "hoarep \ {} F P a Q A = hoarep \ {} F P a' Q A" apply rule apply (erule exec_eq_is_valid_eq0 [OF eq]) apply (erule exec_eq_is_valid_eq0 [OF eq [symmetric]]) done lemma Int_eq_symmetric: "A \ {s. x s = y s} = A \ {s. y s = x s}" by auto lemma ceqv_refl: "ceqv \ xf' rv' t t' c c" unfolding ceqv_def by auto lemma ceqv_trans: "\ ceqv \ xf' rv' t t' c c'; ceqv \ xf' rv' t t' c' c'' \ \ ceqv \ xf' rv' t t' c c''" unfolding ceqv_def by auto (* A bit yuck, might be better to define the other way around *) definition "semantic_equiv \ \ ceqv \ (\_. ()) ()" lemma semantic_equiv_sym: "semantic_equiv \ s s' a b = semantic_equiv \ s s' b a" unfolding semantic_equiv_def by (auto intro: ceqv_sym) lemma semantic_equivI: "(\\ \a,Normal s\ \ s' = \\ \b,Normal s\ \ s') \ semantic_equiv \ s s' a b" unfolding semantic_equiv_def by (auto intro: ceqvI) lemmas semantic_equivD1 = ceqvD1 [where xf = "\_. ()" and v = "()", folded semantic_equiv_def] lemmas semantic_equivD2 = ceqvD2 [where xf = "\_. ()" and v = "()", folded semantic_equiv_def] lemma Guard_Seq_semantic_equiv: "semantic_equiv \ s s' (Guard F S c ;; d) (Guard F S (c ;; d))" by (auto elim!: exec_Normal_elim_cases intro: semantic_equivI exec.intros) lemma exec_Seq_cong: "\ \s''. \ \ \a, Normal s\ \ s'' = \ \ \c, Normal s\ \ s''; \s''. \ \ \c, Normal s\ \ Normal s'' \ \ \ \b, Normal s''\ \ s' = \ \ \d, Normal s''\ \ s' \ \ \ \ \a ;; b, Normal s\ \ s' = \ \ \c ;; d, Normal s\ \ s'" apply (rule iffI) apply (erule exec_Normal_elim_cases) apply (case_tac "s'a", auto elim!: exec_elim_cases intro: exec.intros)[1] apply (erule exec_Normal_elim_cases) apply (case_tac "s'a", auto elim!: exec_elim_cases intro: exec.intros)[1] done lemma exec_While_cong': assumes c: "\s s'. \ \ \c, Normal s\ \ s' = \ \ \c', Normal s\ \ s'" assumes w: "\ \ \v, Normal s\ \ s'" assumes eq: "v = While S c" "v' = While S c'" shows "\ \ \v', Normal s\ \ s'" using w eq apply (induct, simp_all) apply (rule exec.intros, assumption) apply (simp add: c) apply simp apply (rule exec.intros, assumption) done lemma exec_While_cong: "\ \s s'. \ \ \c, Normal s\ \ s' = \ \ \c', Normal s\ \ s' \ \ \ \ \While S c, Normal s\ \ s' = \ \ \While S c', Normal s\ \ s'" apply (rule iffI) apply (erule(1) exec_While_cong'[OF _ _ refl refl]) apply (erule(1) exec_While_cong'[OF sym _ refl refl]) done lemma exec_Guard_UNIV_simp: "\ \ \Guard F UNIV c, Normal s\ \ s' = \ \ \c, Normal s\ \ s'" by (auto elim!: exec_Normal_elim_cases intro: exec.intros) lemma exec_Seq_Skip_simps: "\ \ \Skip ;; c, Normal s\ \ s' = \ \ \c, Normal s\ \ s'" "\ \ \c ;; Skip, Normal s\ \ s' = \ \ \c, Normal s\ \ s'" apply (rule iffI) apply (clarsimp elim!: exec_Normal_elim_cases exec_elim_cases) apply (fastforce intro: exec.intros) apply (rule iffI) apply (clarsimp elim!: exec_Normal_elim_cases exec_elim_cases) apply (case_tac s'a, auto elim!: exec_elim_cases)[1] apply (case_tac s', auto intro: exec.intros) done lemma hoarep_exec: assumes pre: "s \ P" assumes exec: "\\ \c, Normal s\ \ t" assumes hoare: "\\\<^bsub>/F\<^esub> P c Q,A" shows "(\f \ F. t = Fault f) \ (\t' \ Q. t = Normal t') \ (\t' \ A. t = Abrupt t')" using pre hoare_sound[OF hoare] exec apply (clarsimp simp: cvalid_def HoarePartialDef.valid_def image_def) apply (drule spec, drule spec, drule (1) mp) by auto lemmas hoarep_exec' = hoarep_exec[where s=s' and P=P' and A=A' and Q=Q' for s' P' Q' A'] lemmas exec_normal = hoarep_exec'[where t="Normal t'" for t', simplified] lemmas exec_abrupt = hoarep_exec'[where t="Abrupt t'" for t', simplified] lemmas exec_fault = hoarep_exec'[where t="Fault f" for f, simplified] lemmas exec_stuck = hoarep_exec[where t=Stuck, simplified] lemma hoarep_exec_gen: assumes h: "\s. \\\<^bsub>/F s\<^esub> (P s) c (Q s),(A s)" assumes e: "\\ \c, Normal s\ \ t" shows "s \ P x \ (\f \ F x. t = Fault f) \ (\t' \ Q x. t = Normal t') \ (\t' \ A x. t = Abrupt t')" by (rule impI, erule hoarep_exec[OF _ e spec[OF h]]) lemmas hoarep_exec_call_body = hoarep_exec_gen[OF _ exec_Call_body_aux[THEN iffD2]] (* Used so we don't simp it in ctac *) definition "guard_is_UNIV r xf Q \ \s rv. r rv (xf s) \ s \ Q rv (xf s)" lemma guard_is_UNIVI: "(\s rv. r rv (xf s) \ s \ Q rv (xf s)) \ guard_is_UNIV r xf Q" unfolding guard_is_UNIV_def by simp lemma guard_is_UNIVD: "\ guard_is_UNIV r xf Q; r rv (xf s) \ \ s \ Q rv (xf s)" unfolding guard_is_UNIV_def by auto definition isNormal::"('s,'f) xstate \ bool" where "isNormal S \ \s. S=Normal s" lemma isNormal_simps: "isNormal (Normal s) = True" "isNormal (Abrupt s) = False" "isNormal (Fault f) = False" "isNormal Stuck = False" by (auto simp add: isNormal_def) lemma hoarep_revert: "Gamm \ P (call (\s. s) fnname (\a b. b) (\i t. Basic (\s. s))) Q,R \ Gamm \ P (Call fnname) Q,R" apply (rule hoare_complete) apply (drule hoare_sound) apply (simp add: cvalid_def HoarePartialDef.valid_def) apply (erule allEI)+ apply (rule impI, drule mp) apply (erule exec.cases, simp_all)[1] apply (case_tac ta, simp_all)[1] apply (erule exec_call) apply simp apply (rule exec.intros) apply (erule exec_callAbrupt, simp_all)[1] apply (erule exec_callFault, simp_all)[1] apply (erule exec_callStuck, simp_all)[1] apply (erule exec_callUndefined) apply simp done lemma intermediate_Normal_state: "\\ \ \Seq c\<^sub>1 c\<^sub>2, Normal t\ \ t''; t \ P; \ \ P c\<^sub>1 Q\ \ \t'. \ \ \c\<^sub>1, Normal t\ \ Normal t' \ \ \ \c\<^sub>2, Normal t'\ \ t''" apply (erule exec_Normal_elim_cases(8)) apply (insert hoarep_exec) apply fastforce done lemma hoarep_ex_pre: "(\x. \ \ {s. P x s} c Q) \ \ \ {s. \x. P x s} c Q" apply (rule hoare_complete) apply (clarsimp simp: cvalid_def HoarePartialDef.valid_def) apply (fastforce dest: hoarep_exec'[rotated]) done lemma hoarep_ex_lift: "(\x. \ \ {s. P x s} c {s. Q x s}) \ \ \ {s. \x. P x s} c {s. \x. Q x s}" apply (rule hoare_complete) apply (clarsimp simp: cvalid_def HoarePartialDef.valid_def) apply (rename_tac s x) apply (drule_tac x=x in meta_spec) apply (prop_tac "s \ Collect (P x)") apply fastforce apply (frule (2) hoarep_exec) apply fastforce done lemma hoarep_conj_lift_pre_fix: "\\ \ P c {s. Q s}; \ \ P c {s. Q' s}\ \ \ \ P c {s. Q s \ Q' s}" apply (rule hoare_complete) apply (clarsimp simp: cvalid_def HoarePartialDef.valid_def) apply (frule (2) hoarep_exec[where Q="Collect Q"]) apply (frule (2) hoarep_exec[where Q="Collect Q'"]) apply fastforce done lemma exec_While_final_inv'': "\ \ \ \b, x\ \ s'; b = While C B; x = Normal s; \s. s \ C \ I s (Normal s); \t t' t''. \ t \ C; \\ \B, Normal t\ \ Normal t'; \\ \While C B, Normal t'\ \ t''; I t' t'' \ \ I t t''; \t t'. \ t \ C; \\ \B, Normal t\ \ Abrupt t' \ \ I t (Abrupt t'); \t. \ t \ C; \ \ \B, Normal t\ \ Stuck \ \ I t Stuck; \t f. \ t \ C; \\ \B, Normal t\ \ Fault f \ \ I t (Fault f) \ \ I s s'" apply (induct arbitrary: s rule: exec.induct; simp) apply (erule exec_elim_cases; fastforce simp: exec.WhileTrue exec.WhileFalse) done lemma While_inv_from_body: "\ \ (G \ C) B G \ \ \ G While C B G" apply (drule hoare_sound)+ apply (rule hoare_complete) apply (clarsimp simp: cvalid_def HoarePartialDef.valid_def) by (erule exec_While_final_inv''[where I="\s s'. s \ G \ s' \ Normal ` G", THEN impE], fastforce+) end