diff --git a/metadata/entries/Simpl.toml b/metadata/entries/Simpl.toml --- a/metadata/entries/Simpl.toml +++ b/metadata/entries/Simpl.toml @@ -1,25 +1,26 @@ title = "A Sequential Imperative Programming Language Syntax, Semantics, Hoare Logics and Verification Environment" date = 2008-02-29 topics = [ "Computer science/Programming languages/Language definitions", "Computer science/Programming languages/Logics", ] abstract = "We present the theory of Simpl, a sequential imperative programming language. We introduce its syntax, its semantics (big and small-step operational semantics) and Hoare logics for both partial as well as total correctness. We prove soundness and completeness of the Hoare logic. We integrate and automate the Hoare logic in Isabelle/HOL to obtain a practically usable verification environment for imperative programs. Simpl is independent of a concrete programming language but expressive enough to cover all common language features: mutually recursive procedures, abrupt termination and exceptions, runtime faults, local and global variables, pointers and heap, expressions with side effects, pointers to procedures, partial application and closures, dynamic method invocation and also unbounded nondeterminism." -license = "lgpl" +license = "bsd" note = "" [authors] [authors.schirmer] [contributors] [notify] klein = "klein_email" schirmer = "schirmer_email" [history] +2024-02-18 = "License change to BSD; more canonical and versatile ML interfaces; some more derived language constructs." [extra] [related] diff --git a/thys/Simpl/AlternativeSmallStep.thy b/thys/Simpl/AlternativeSmallStep.thy --- a/thys/Simpl/AlternativeSmallStep.thy +++ b/thys/Simpl/AlternativeSmallStep.thy @@ -1,3884 +1,3863 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: AlternativeSmallStep.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2006-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA *) section \Alternative Small Step Semantics\ theory AlternativeSmallStep imports HoareTotalDef begin text \ This is the small-step semantics, which is described and used in my PhD-thesis \<^cite>\"Schirmer-PhD"\. It decomposes the statement into a list of statements and finally executes the head. So the redex is always the head of the list. The equivalence between termination (based on the big-step semantics) and the absence of infinite computations in this small-step semantics follows the same lines of reasoning as for the new small-step semantics. However, it is technically more involved since the configurations are more complicated. Thats why I switched to the new small-step semantics in the "main trunk". I keep this alternative version and the important proofs in this theory, so that one can compare both approaches. \ subsection \Small-Step Computation: \\\(cs, css, s) \ (cs', css', s')\\ type_synonym ('s,'p,'f) continuation = "('s,'p,'f) com list \ ('s,'p,'f) com list" type_synonym ('s,'p,'f) config = "('s,'p,'f)com list \ ('s,'p,'f)continuation list \ ('s,'f) xstate" inductive "step"::"[('s,'p,'f) body,('s,'p,'f) config,('s,'p,'f) config] \ bool" ("_\ (_ \/ _)" [81,81,81] 100) for \::"('s,'p,'f) body" where Skip: "\\(Skip#cs,css,Normal s) \ (cs,css,Normal s)" | Guard: "s\g \ \\(Guard f g c#cs,css,Normal s) \ (c#cs,css,Normal s)" | GuardFault: "s\g \ \\(Guard f g c#cs,css,Normal s) \ (cs,css,Fault f)" | FaultProp: "\\(c#cs,css,Fault f) \ (cs,css,Fault f)" | FaultPropBlock: "\\([],(nrms,abrs)#css,Fault f) \ (nrms,css,Fault f)" (* FaultPropBlock: "\\([],cs#css,Fault) \ ([],css,Fault)"*) | AbruptProp: "\\(c#cs,css,Abrupt s) \ (cs,css,Abrupt s)" | ExitBlockNormal: "\\([],(nrms,abrs)#css,Normal s) \ (nrms,css,Normal s)" | ExitBlockAbrupt: "\\([],(nrms,abrs)#css,Abrupt s) \ (abrs,css,Normal s)" | Basic: "\\(Basic f#cs,css,Normal s) \ (cs,css,Normal (f s))" | Spec: "(s,t) \ r \ \\(Spec r#cs,css,Normal s) \ (cs,css,Normal t)" | SpecStuck: "\t. (s,t) \ r \ \\(Spec r#cs,css,Normal s) \ (cs,css,Stuck)" | Seq: "\\(Seq c\<^sub>1 c\<^sub>2#cs,css,Normal s) \ (c\<^sub>1#c\<^sub>2#cs,css,Normal s)" | CondTrue: "s\b \ \\(Cond b c\<^sub>1 c\<^sub>2#cs,css,Normal s) \ (c\<^sub>1#cs,css,Normal s)" | CondFalse: "s\b \ \\(Cond b c\<^sub>1 c\<^sub>2#cs,css,Normal s) \ (c\<^sub>2#cs,css,Normal s)" | WhileTrue: "\s\b\ \ \\(While b c#cs,css,Normal s) \ (c#While b c#cs,css,Normal s)" | WhileFalse: "\s\b\ \ \\(While b c#cs,css,Normal s) \ (cs,css,Normal s)" | Call: "\ p=Some bdy \ \\(Call p#cs,css,Normal s) \ ([bdy],(cs,Throw#cs)#css,Normal s)" | CallUndefined: "\ p=None \ \\(Call p#cs,css,Normal s) \ (cs,css,Stuck)" | StuckProp: "\\(c#cs,css,Stuck) \ (cs,css,Stuck)" | StuckPropBlock: "\\([],(nrms,abrs)#css,Stuck) \ (nrms,css,Stuck)" | DynCom: "\\(DynCom c#cs,css,Normal s) \ (c s#cs,css,Normal s)" | Throw: "\\(Throw#cs,css,Normal s) \ (cs,css,Abrupt s)" | Catch: "\\(Catch c\<^sub>1 c\<^sub>2#cs,css,Normal s) \ ([c\<^sub>1],(cs,c\<^sub>2#cs)#css,Normal s)" lemmas step_induct = step.induct [of _ "(c,css,s)" "(c',css',s')", split_format (complete), case_names Skip Guard GuardFault FaultProp FaultPropBlock AbruptProp ExitBlockNormal ExitBlockAbrupt Basic Spec SpecStuck Seq CondTrue CondFalse WhileTrue WhileFalse Call CallUndefined StuckProp StuckPropBlock DynCom Throw Catch, induct set] inductive_cases step_elim_cases [cases set]: "\\(c#cs,css,Fault f) \ u" "\\([],css,Fault f) \ u" "\\(c#cs,css,Stuck) \ u" "\\([],css,Stuck) \ u" "\\(c#cs,css,Abrupt s) \ u" "\\([],css,Abrupt s) \ u" "\\([],css,Normal s) \ u" "\\(Skip#cs,css,s) \ u" "\\(Guard f g c#cs,css,s) \ u" "\\(Basic f#cs,css,s) \ u" "\\(Spec r#cs,css,s) \ u" "\\(Seq c1 c2#cs,css,s) \ u" "\\(Cond b c1 c2#cs,css,s) \ u" "\\(While b c#cs,css,s) \ u" "\\(Call p#cs,css,s) \ u" "\\(DynCom c#cs,css,s) \ u" "\\(Throw#cs,css,s) \ u" "\\(Catch c1 c2#cs,css,s) \ u" inductive_cases step_Normal_elim_cases [cases set]: "\\(c#cs,css,Fault f) \ u" "\\([],css,Fault f) \ u" "\\(c#cs,css,Stuck) \ u" "\\([],css,Stuck) \ u" "\\([],(nrms,abrs)#css,Normal s) \ u" "\\([],(nrms,abrs)#css,Abrupt s) \ u" "\\(Skip#cs,css,Normal s) \ u" "\\(Guard f g c#cs,css,Normal s) \ u" "\\(Basic f#cs,css,Normal s) \ u" "\\(Spec r#cs,css,Normal s) \ u" "\\(Seq c1 c2#cs,css,Normal s) \ u" "\\(Cond b c1 c2#cs,css,Normal s) \ u" "\\(While b c#cs,css,Normal s) \ u" "\\(Call p#cs,css,Normal s) \ u" "\\(DynCom c#cs,css,Normal s) \ u" "\\(Throw#cs,css,Normal s) \ u" "\\(Catch c1 c2#cs,css,Normal s) \ u" abbreviation "step_rtrancl" :: "[('s,'p,'f) body,('s,'p,'f) config,('s,'p,'f) config] \ bool" ("_\ (_ \\<^sup>*/ _)" [81,81,81] 100) where "\\cs0 \\<^sup>* cs1 == (step \)\<^sup>*\<^sup>* cs0 cs1" abbreviation "step_trancl" :: "[('s,'p,'f) body,('s,'p,'f) config,('s,'p,'f) config] \ bool" ("_\ (_ \\<^sup>+/ _)" [81,81,81] 100) where "\\cs0 \\<^sup>+ cs1 == (step \)\<^sup>+\<^sup>+ cs0 cs1" subsubsection \Structural Properties of Small Step Computations\ lemma Fault_app_steps: "\\(cs@xs,css,Fault f) \\<^sup>* (xs,css,Fault f)" proof (induct cs) case Nil thus ?case by simp next case (Cons c cs) have "\\(c#cs@xs, css, Fault f) \\<^sup>* (xs, css, Fault f)" proof - have "\\(c#cs@xs, css, Fault f) \ (cs@xs, css, Fault f)" by (rule step.FaultProp) also have "\\(cs@xs, css, Fault f) \\<^sup>* (xs, css, Fault f)" by (rule Cons.hyps) finally show ?thesis . qed thus ?case by simp qed lemma Stuck_app_steps: "\\(cs@xs,css,Stuck) \\<^sup>* (xs,css,Stuck)" proof (induct cs) case Nil thus ?case by simp next case (Cons c cs) have "\\(c#cs@xs, css, Stuck) \\<^sup>* (xs, css, Stuck)" proof - have "\\(c#cs@xs, css, Stuck) \ (cs@xs, css, Stuck)" by (rule step.StuckProp) also have "\\(cs@xs, css, Stuck) \\<^sup>* (xs, css, Stuck)" by (rule Cons.hyps) finally show ?thesis . qed thus ?case by simp qed text \We can only append commands inside a block, if execution does not enter or exit a block. \ lemma app_step: assumes step: "\\(cs,css,s) \ (cs',css',t)" shows "css=css' \ \\(cs@xs,css,s) \ (cs'@xs,css',t)" using step apply induct apply (simp_all del: fun_upd_apply,(blast intro: step.intros)+) done text \We can append whole blocks, without interfering with the actual block. Outer blocks do not influence execution of inner blocks.\ lemma app_css_step: assumes step: "\\(cs,css,s) \ (cs',css',t)" shows "\\(cs,css@xs,s) \ (cs',css'@xs,t)" using step apply induct apply (simp_all del: fun_upd_apply,(blast intro: step.intros)+) done ML \ ML_Thms.bind_thm ("trancl_induct3", Split_Rule.split_rule @{context} (Rule_Insts.read_instantiate @{context} [((("a", 0), Position.none), "(ax, ay, az)"), ((("b", 0), Position.none), "(bx, by, bz)")] [] @{thm tranclp_induct})); \ lemma app_css_steps: assumes step: "\\(cs,css,s) \\<^sup>+ (cs',css',t)" shows "\\(cs,css@xs,s) \\<^sup>+ (cs',css'@xs,t)" apply(rule trancl_induct3 [OF step]) apply (rule app_css_step [THEN tranclp.r_into_trancl [of "step \"]],assumption) apply(blast intro:app_css_step tranclp_trans) done lemma step_Cons': assumes step: "\\(ccs,css,s) \ (cs',css',t)" shows "\c cs. ccs=c#cs \ \css''. css'=css''@css \ (if css''=[] then \p. cs'=p@cs else (\pnorm pabr. css''=[(pnorm@cs,pabr@cs)]))" using step by induct force+ lemma step_Cons: assumes step: "\\(c#cs,css,s) \ (cs',css',t)" shows "\pcss. css'=pcss@css \ (if pcss=[] then \ps. cs'=ps@cs else (\pcs_normal pcs_abrupt. pcss=[(pcs_normal@cs,pcs_abrupt@cs)]))" using step_Cons' [OF step] by blast lemma step_Nil': assumes step: "\\(cs,asscss,s) \ (cs',css',t)" shows "\ass. \cs=[]; asscss=ass@css; ass\Nil\ \ css'=tl ass@css \ (case s of Abrupt s' \ cs'=snd (hd ass) \ t=Normal s' | _ \ cs'=fst (hd ass) \ t=s)" using step by (induct) (fastforce simp add: neq_Nil_conv)+ lemma step_Nil: assumes step: "\\([],ass@css,s) \ (cs',css',t)" assumes ass_not_Nil: "ass\[]" shows "css'=tl ass@css \ (case s of Abrupt s' \ cs'=snd (hd ass) \ t=Normal s' | _ \ cs'=fst (hd ass) \ t=s)" using step_Nil' [OF step _ _ ass_not_Nil] by simp lemma step_Nil'': assumes step: "\\([],(pcs_normal,pcs_abrupt)#pcss@css,s) \ (cs',pcss@css,t)" shows "(case s of Abrupt s' \ cs'=pcs_abrupt \ t=Normal s' | _ \ cs'=pcs_normal \ t=s)" using step_Nil' [OF step, where ass ="(pcs_normal,pcs_abrupt)#pcss" and css="css"] by (auto split: xstate.splits) lemma drop_suffix_css_step': assumes step: "\\(cs,cssxs,s) \ (cs',css'xs,t)" shows "\css css' xs. \cssxs = css@xs; css'xs=css'@xs\ \ \\(cs,css,s) \ (cs',css',t)" using step apply induct apply (fastforce intro: step.intros)+ done lemma drop_suffix_css_step: assumes step: "\\(cs,pcss@css,s) \ (cs',pcss'@css,t)" shows "\\(cs,pcss,s) \ (cs',pcss',t)" using step by (blast intro: drop_suffix_css_step') lemma drop_suffix_hd_css_step': assumes step: "\\ (pcs,css,s) \ (cs',css'css,t)" shows "\p ps cs pnorm pabr. \pcs=p#ps@cs; css'css=(pnorm@cs,pabr@cs)#css\ \ \\ (p#ps,css,s) \ (cs',(pnorm,pabr)#css,t)" using step by induct (force intro: step.intros)+ lemma drop_suffix_hd_css_step'': assumes step: "\\ (p#ps@cs,css,s) \ (cs',(pnorm@cs,pabr@cs)#css,t)" shows "\\ (p#ps,css,s) \ (cs',(pnorm,pabr)#css,t)" using drop_suffix_hd_css_step' [OF step] by auto lemma drop_suffix_hd_css_step: assumes step: "\\ (p#ps@cs,css,s) \ (cs',[(pnorm@ps@cs,pabr@ps@cs)]@css,t)" shows "\\ (p#ps,css,s) \ (cs',[(pnorm@ps,pabr@ps)]@css,t)" proof - from step drop_suffix_hd_css_step'' [of _ p ps cs css s cs' "pnorm@ps" "pabr@ps" t] show ?thesis by auto qed lemma drop_suffix': assumes step: "\\(csxs,css,s) \ (cs'xs,css',t)" shows "\xs cs cs'. \css=css'; csxs=cs@xs; cs'xs = cs'@xs; cs\[] \ \ \\(cs,css,s) \ (cs',css,t)" using step apply induct apply (fastforce intro: step.intros simp add: neq_Nil_conv)+ done lemma drop_suffix: assumes step: "\\(c#cs@xs,css,s) \ (cs'@xs,css,t)" shows "\\(c#cs,css,s) \ (cs',css,t)" by(rule drop_suffix' [OF step _ _ _]) auto lemma drop_suffix_same_css_step: assumes step: "\\(cs@xs,css,s) \ (cs'@xs,css,t)" assumes not_Nil: "cs\[]" shows "\\(cs,xss,s) \ (cs',xss,t)" proof- from drop_suffix' [OF step _ _ _ not_Nil] have "\\(cs,css,s) \ (cs',css,t)" by auto with drop_suffix_css_step [of _ cs "[]" css s cs' "[]" t] have "\\ (cs, [], s) \ (cs', [], t)" by auto from app_css_step [OF this] show ?thesis by auto qed lemma Cons_change_css_step: assumes step: "\\ (cs,css,s) \ (cs',css'@css,t)" shows "\\ (cs,xss,s) \ (cs',css'@xss,t)" proof - from step drop_suffix_css_step [where cs=cs and pcss="[]" and css=css and s=s and cs'=cs' and pcss'=css' and t=t] have "\\ (cs, [], s) \ (cs', css', t)" by auto from app_css_step [where xs=xss, OF this] show ?thesis by auto qed lemma Nil_change_css_step: assumes step: "\\([],ass@css,s) \ (cs',ass'@css,t)" assumes ass_not_Nil: "ass\[]" shows "\\([],ass@xss,s) \ (cs',ass'@xss,t)" proof - from step drop_suffix_css_step [of _ "[]" ass css s cs' ass' t] have "\\ ([], ass, s) \ (cs', ass', t)" by auto from app_css_step [where xs=xss, OF this] show ?thesis by auto qed subsubsection \Equivalence between Big and Small-Step Semantics\ lemma exec_impl_steps: assumes exec: "\\\c,s\ \ t" shows "\cs css. \\(c#cs,css,s) \\<^sup>* (cs,css,t)" using exec proof (induct) case Skip thus ?case by (blast intro: step.Skip) next case Guard thus ?case by (blast intro: step.Guard rtranclp_trans) next case GuardFault thus ?case by (blast intro: step.GuardFault) next case FaultProp thus ?case by (blast intro: step.FaultProp) next case Basic thus ?case by (blast intro: step.Basic) next case Spec thus ?case by (blast intro: step.Spec) next case SpecStuck thus ?case by (blast intro: step.SpecStuck) next case Seq thus ?case by (blast intro: step.Seq rtranclp_trans) next case CondTrue thus ?case by (blast intro: step.CondTrue rtranclp_trans) next case CondFalse thus ?case by (blast intro: step.CondFalse rtranclp_trans) next case WhileTrue thus ?case by (blast intro: step.WhileTrue rtranclp_trans) next case WhileFalse thus ?case by (blast intro: step.WhileFalse) next case (Call p bdy s s' cs css) have bdy: "\ p = Some bdy" by fact have steps_body: "\\([bdy],(cs,Throw#cs)#css,Normal s) \\<^sup>* ([],(cs,Throw#cs)#css, s')" by fact show ?case proof (cases s') case (Normal s'') note steps_body also from Normal have "\\([],(cs,Throw#cs)#css, s') \ (cs,css,s')" by (auto intro: step.intros) finally show ?thesis using bdy by (blast intro: step.Call rtranclp_trans) next case (Abrupt s'') with steps_body have "\\([bdy],(cs,Throw#cs)#css,Normal s) \\<^sup>* ([],(cs,Throw#cs)#css, Abrupt s'')" by simp also have "\\([],(cs,Throw#cs)#css, Abrupt s'') \ (Throw#cs,css,Normal s'')" by (rule ExitBlockAbrupt) also have "\\(Throw#cs,css,Normal s'') \ (cs,css,Abrupt s'')" by (rule Throw) finally show ?thesis using bdy Abrupt by (auto intro: step.Call rtranclp_trans) next case Fault note steps_body also from Fault have "\\([],(cs,Throw#cs)#css, s') \ (cs,css,s')" by (auto intro: step.intros) finally show ?thesis using bdy by (blast intro: step.Call rtranclp_trans) next case Stuck note steps_body also from Stuck have "\\([],(cs,Throw#cs)#css, s') \ (cs,css,s')" by (auto intro: step.intros) finally show ?thesis using bdy by (blast intro: step.Call rtranclp_trans) qed next case (CallUndefined p s cs css) have undef: "\ p = None" by fact hence "\\(Call p # cs, css, Normal s) \ (cs, css, Stuck)" by (rule step.CallUndefined) thus ?case .. next case StuckProp thus ?case by (blast intro: step.StuckProp rtrancl_trans) next case DynCom thus ?case by (blast intro: step.DynCom rtranclp_trans) next case Throw thus ?case by (blast intro: step.Throw) next case AbruptProp thus ?case by (blast intro: step.AbruptProp) next case (CatchMatch c\<^sub>1 s s' c\<^sub>2 s'' cs css) have steps_c1: "\\([c\<^sub>1],(cs,c\<^sub>2#cs)#css,Normal s) \\<^sup>* ([],(cs,c\<^sub>2#cs)#css,Abrupt s')" by fact also have "\\([],(cs,c\<^sub>2#cs)#css,Abrupt s') \ (c\<^sub>2#cs,css,Normal s')" by (rule ExitBlockAbrupt) also have steps_c2: "\\(c\<^sub>2#cs,css,Normal s') \\<^sup>* (cs,css,s'')" by fact finally show "\\(Catch c\<^sub>1 c\<^sub>2 # cs, css, Normal s) \\<^sup>* (cs, css, s'')" by (blast intro: step.Catch rtranclp_trans) next case (CatchMiss c\<^sub>1 s s' c\<^sub>2 cs css) assume notAbr: "\ isAbr s'" have steps_c1: "\\([c\<^sub>1],(cs,c\<^sub>2#cs)#css,Normal s) \\<^sup>* ([],(cs,c\<^sub>2#cs)#css,s')" by fact show "\\(Catch c\<^sub>1 c\<^sub>2 # cs, css, Normal s) \\<^sup>* (cs, css, s')" proof (cases s') case (Normal w) with steps_c1 have "\\([c\<^sub>1],(cs,c\<^sub>2#cs)#css,Normal s) \\<^sup>* ([],(cs,c\<^sub>2#cs)#css,Normal w)" by simp also have "\\([],(cs,c\<^sub>2#cs)#css,Normal w) \ (cs,css,Normal w)" by (rule ExitBlockNormal) finally show ?thesis using Normal by (auto intro: step.Catch rtranclp_trans) next case Abrupt with notAbr show ?thesis by simp next case (Fault f) with steps_c1 have "\\([c\<^sub>1],(cs,c\<^sub>2#cs)#css,Normal s) \\<^sup>* ([],(cs,c\<^sub>2#cs)#css,Fault f)" by simp also have "\\([],(cs,c\<^sub>2#cs)#css,Fault f) \ (cs,css,Fault f)" by (rule FaultPropBlock) finally show ?thesis using Fault by (auto intro: step.Catch rtranclp_trans) next case Stuck with steps_c1 have "\\([c\<^sub>1],(cs,c\<^sub>2#cs)#css,Normal s) \\<^sup>* ([],(cs,c\<^sub>2#cs)#css,Stuck)" by simp also have "\\([],(cs,c\<^sub>2#cs)#css,Stuck) \ (cs,css,Stuck)" by (rule StuckPropBlock) finally show ?thesis using Stuck by (auto intro: step.Catch rtranclp_trans) qed qed inductive "execs"::"[('s,'p,'f) body,('s,'p,'f) com list, ('s,'p,'f) continuation list, ('s,'f) xstate,('s,'f) xstate] \ bool" ("_\ \_,_,_\ \ _" [50,50,50,50,50] 50) for \:: "('s,'p,'f) body" where Nil: "\\\[],[],s\ \ s" | ExitBlockNormal: "\\\nrms,css,Normal s\ \ t \ \\\[],(nrms,abrs)#css,Normal s\ \ t" | ExitBlockAbrupt: "\\\abrs,css,Normal s\ \ t \ \\\[],(nrms,abrs)#css,Abrupt s\ \ t" | ExitBlockFault: "\\\nrms,css,Fault f\ \ t \ \\\[],(nrms,abrs)#css,Fault f\ \ t" | ExitBlockStuck: "\\\nrms,css,Stuck\ \ t \ \\\[],(nrms,abrs)#css,Stuck\ \ t" | Cons: "\\\\c,s\ \ t; \\\cs,css,t\ \ u\ \ \\\c#cs,css,s\ \ u" inductive_cases execs_elim_cases [cases set]: "\\\[],css,s\ \ t" "\\\c#cs,css,s\ \ t" ML \ ML_Thms.bind_thm ("converse_rtrancl_induct3", Split_Rule.split_rule @{context} (Rule_Insts.read_instantiate @{context} [((("a", 0), Position.none), "(cs, css, s)"), ((("b", 0), Position.none), "(cs', css', t)")] [] @{thm converse_rtranclp_induct})); \ lemma execs_Fault_end: assumes execs: "\\\cs,css,s\ \ t" shows "s=Fault f\ t=Fault f" using execs by (induct) (auto dest: Fault_end) lemma execs_Stuck_end: assumes execs: "\\\cs,css,s\ \ t" shows "s=Stuck \ t=Stuck" using execs by (induct) (auto dest: Stuck_end) theorem steps_impl_execs: assumes steps: "\\(cs,css,s) \\<^sup>* ([],[],t)" shows "\\\cs,css,s\ \ t" using steps proof (induct rule: converse_rtrancl_induct3 [consumes 1]) show "\\\[],[],t\ \ t" by (rule execs.Nil) next fix cs css s cs' css' w assume step: "\\(cs,css, s) \ (cs',css', w)" assume execs: "\\\cs',css',w\ \ t" from step show "\\\cs,css,s\ \ t" proof (cases) case (Catch c1 c2 cs s) with execs obtain t' where exec_c1: "\\\c1,Normal s\ \ t'" and execs_rest: "\\\[],(cs, c2 # cs) # css,t'\ \ t" by (clarsimp elim!: execs_elim_cases) have "\\\Catch c1 c2 # cs,css,Normal s\ \ t" proof (cases t') case (Normal t'') with exec_c1 have "\\\Catch c1 c2,Normal s\ \ t'" by (auto intro: exec.CatchMiss) moreover from execs_rest Normal have "\\\cs,css,t'\ \ t" by (cases) auto ultimately show ?thesis by (rule execs.Cons) next case (Abrupt t'') from execs_rest Abrupt have "\\\c2#cs,css,Normal t''\ \ t" by (cases) auto then obtain v where exec_c2: "\\\c2,Normal t''\ \ v" and rest: "\\\cs,css,v\ \ t" by cases from exec_c1 Abrupt exec_c2 have "\\\Catch c1 c2,Normal s\ \ v" by - (rule exec.CatchMatch, auto) from this rest show ?thesis by (rule execs.Cons) next case (Fault f) with exec_c1 have "\\\Catch c1 c2,Normal s\ \ Fault f" by (auto intro: exec.intros) moreover from execs_rest Fault have "\\\cs,css,Fault f\ \ t" by (cases) auto ultimately show ?thesis by (rule execs.Cons) next case Stuck with exec_c1 have "\\\Catch c1 c2,Normal s\ \ Stuck" by (auto intro: exec.intros) moreover from execs_rest Stuck have "\\\cs,css,Stuck\ \ t" by (cases) auto ultimately show ?thesis by (rule execs.Cons) qed with Catch show ?thesis by simp next case (Call p bdy cs s) have bdy: "\ p = Some bdy" by fact from Call execs obtain t' where exec_body: "\\\bdy,Normal s\ \ t'" and execs_rest: "\\\[],(cs,Throw#cs)#css ,t'\ \ t" by (clarsimp elim!: execs_elim_cases) have "\\\Call p # cs,css,Normal s\ \ t" proof (cases t') case (Normal t'') with exec_body bdy have "\\\Call p ,Normal s\ \ Normal t''" by (auto intro: exec.intros) moreover from execs_rest Normal have "\\\cs,css ,Normal t''\ \ t" by cases auto ultimately show ?thesis by (rule execs.Cons) next case (Abrupt t'') with exec_body bdy have "\\\Call p,Normal s\ \ Abrupt t''" by (auto intro: exec.intros) moreover from execs_rest Abrupt have "\\\Throw # cs,css,Normal t''\ \ t" by (cases) auto then obtain v where v: "\\\Throw,Normal t''\ \ v" "\\\cs,css,v\ \ t" by (clarsimp elim!: execs_elim_cases) moreover from v have "v=Abrupt t''" by (auto elim: exec_Normal_elim_cases) ultimately show ?thesis by (auto intro: execs.Cons) next case (Fault f) with exec_body bdy have "\\\Call p,Normal s\ \ Fault f" by (auto intro: exec.intros) moreover from execs_rest Fault have "\\\cs,css,Fault f\ \ t" by (cases) (auto elim: execs_elim_cases dest: Fault_end) ultimately show ?thesis by (rule execs.Cons) next case Stuck with exec_body bdy have "\\\Call p,Normal s\ \ Stuck" by (auto intro: exec.intros) moreover from execs_rest Stuck have "\\\cs,css,Stuck\ \ t" by (cases) (auto elim: execs_elim_cases dest: Stuck_end) ultimately show ?thesis by (rule execs.Cons) qed with Call show ?thesis by simp qed (insert execs, (blast intro:execs.intros exec.intros elim!: execs_elim_cases)+) qed theorem steps_impl_exec: assumes steps: "\\([c],[],s) \\<^sup>* ([],[],t)" shows "\\\c,s\ \ t" using steps_impl_execs [OF steps] by (blast elim: execs_elim_cases) corollary steps_eq_exec: "\\([c],[],s) \\<^sup>* ([],[],t) = \\\c,s\ \ t" by (blast intro: steps_impl_exec exec_impl_steps) subsection \Infinite Computations: \inf \ cs css s\\ definition inf :: "[('s,'p,'f) body,('s,'p,'f) com list,('s,'p,'f) continuation list,('s,'f) xstate] \ bool" where "inf \ cs css s = (\f. f 0 = (cs,css,s) \ (\i. \\f i \ f(Suc i)))" lemma not_infI: "\\f. \f 0 = (cs,css,s); \i. \\f i \ f (Suc i)\ \ False\ \ \inf \ cs css s" by (auto simp add: inf_def) subsection \Equivalence of Termination and Absence of Infinite Computations\ inductive "terminatess":: "[('s,'p,'f) body,('s,'p,'f) com list, ('s,'p,'f) continuation list,('s,'f) xstate] \ bool" ("_\_,_ \ _" [60,20,60] 89) for \::"('s,'p,'f) body" where Nil: "\\[],[]\s" | ExitBlockNormal: "\\nrms,css\Normal s \ \\[],(nrms,abrs)#css\Normal s" | ExitBlockAbrupt: "\\abrs,css\Normal s \ \\[],(nrms,abrs)#css\Abrupt s" | ExitBlockFault: "\\nrms,css\Fault f \ \\[],(nrms,abrs)#css\Fault f" | ExitBlockStuck: "\\nrms,css\Stuck \ \\[],(nrms,abrs)#css\Stuck" | Cons: "\\\c\s; (\t. \\\c,s\ \ t \ \\cs,css\t)\ \ \\c#cs,css\s" inductive_cases terminatess_elim_cases [cases set]: "\\[],css\t" "\\c#cs,css\t" lemma terminatess_Fault: "\cs. \\cs,css\Fault f" proof (induct css) case Nil show "\\cs,[]\Fault f" proof (induct cs) case Nil show "\\[],[]\Fault f" by (rule terminatess.Nil) next case (Cons c cs) thus ?case by (auto intro: terminatess.intros terminates.intros dest: Fault_end) qed next case (Cons d css) have hyp: "\cs. \\cs,css\Fault f" by fact obtain nrms abrs where d: "d=(nrms,abrs)" by (cases d) auto have "\\cs,(nrms,abrs)#css\Fault f" proof (induct cs) case Nil show "\\[],(nrms, abrs) # css\Fault f" by (rule terminatess.ExitBlockFault) (rule hyp) next case (Cons c cs) have hyp1: "\\cs,(nrms, abrs) # css\Fault f" by fact show "\\c#cs,(nrms, abrs)#css\Fault f" by (auto intro: hyp1 terminatess.Cons terminates.intros dest: Fault_end) qed with d show ?case by simp qed lemma terminatess_Stuck: "\cs. \\cs,css\Stuck" proof (induct css) case Nil show "\\cs,[]\Stuck" proof (induct cs) case Nil show "\\[],[]\Stuck" by (rule terminatess.Nil) next case (Cons c cs) thus ?case by (auto intro: terminatess.intros terminates.intros dest: Stuck_end) qed next case (Cons d css) have hyp: "\cs. \\cs,css\Stuck" by fact obtain nrms abrs where d: "d=(nrms,abrs)" by (cases d) auto have "\\cs,(nrms,abrs)#css\Stuck" proof (induct cs) case Nil show "\\[],(nrms, abrs) # css\Stuck" by (rule terminatess.ExitBlockStuck) (rule hyp) next case (Cons c cs) have hyp1: "\\cs,(nrms, abrs) # css\Stuck" by fact show "\\c#cs,(nrms, abrs)#css\Stuck" by (auto intro: hyp1 terminatess.Cons terminates.intros dest: Stuck_end) qed with d show ?case by simp qed lemma Basic_terminates: "\\Basic f \ t" by (cases t) (auto intro: terminates.intros) lemma step_preserves_terminations: assumes step: "\\(cs,css,s) \ (cs',css',t)" shows "\\cs,css\s \ \\cs',css'\t" using step proof (induct) case Skip thus ?case by (auto elim: terminates_Normal_elim_cases terminatess_elim_cases intro: exec.intros) next case Guard thus ?case by (blast elim: terminatess_elim_cases terminates_Normal_elim_cases intro: terminatess.intros terminates.intros exec.intros) next case GuardFault thus ?case by (blast elim: terminatess_elim_cases terminates_Normal_elim_cases intro: terminatess.intros terminates.intros exec.intros) next case FaultProp show ?case by (rule terminatess_Fault) next case FaultPropBlock show ?case by (rule terminatess_Fault) next case AbruptProp thus ?case by (blast elim: terminatess_elim_cases intro: terminatess.intros) next case ExitBlockNormal thus ?case by (blast elim: terminatess_elim_cases intro: terminatess.intros ) next case ExitBlockAbrupt thus ?case by (blast elim: terminatess_elim_cases intro: terminatess.intros ) next case Basic thus ?case by (blast elim: terminatess_elim_cases terminates_Normal_elim_cases intro: terminatess.intros terminates.intros exec.intros) next case Spec thus ?case by (blast elim: terminatess_elim_cases terminates_Normal_elim_cases intro: terminatess.intros terminates.intros exec.intros) next case SpecStuck thus ?case by (blast elim: terminatess_elim_cases terminates_Normal_elim_cases intro: terminatess.intros terminates.intros exec.intros) next case Seq thus ?case by (blast elim: terminatess_elim_cases terminates_Normal_elim_cases intro: terminatess.intros terminates.intros exec.intros) next case CondTrue thus ?case by (blast elim: terminatess_elim_cases terminates_Normal_elim_cases intro: terminatess.intros terminates.intros exec.intros) next case CondFalse thus ?case by (blast elim: terminatess_elim_cases terminates_Normal_elim_cases intro: terminatess.intros terminates.intros exec.intros) next case WhileTrue thus ?case by (blast elim: terminatess_elim_cases terminates_Normal_elim_cases intro: terminatess.intros terminates.intros exec.intros) next case WhileFalse thus ?case by (blast elim: terminatess_elim_cases terminates_Normal_elim_cases intro: terminatess.intros terminates.intros exec.intros) next case (Call p bdy cs css s) have bdy: "\ p = Some bdy" by fact from Call obtain term_body: "\\bdy \ Normal s" and term_rest: "\t. \\\Call p,Normal s\ \ t \ \\cs,css\t" by (fastforce elim!: terminatess_elim_cases terminates_Normal_elim_cases) show "\\[bdy],(cs,Throw # cs)#css\Normal s" proof (rule terminatess.Cons [OF term_body],clarsimp) fix t assume exec_body: "\\\bdy,Normal s\ \ t" show "\\[],(cs,Throw # cs) # css\t" proof (cases t) case (Normal t') with exec_body bdy have "\\\Call p,Normal s\ \ Normal t'" by (auto intro: exec.intros) with term_rest have "\\cs,css\Normal t'" by iprover with Normal show ?thesis by (auto intro: terminatess.intros terminates.intros elim: exec_Normal_elim_cases) next case (Abrupt t') with exec_body bdy have "\\\Call p,Normal s\ \ Abrupt t'" by (auto intro: exec.intros) with term_rest have "\\cs,css\Abrupt t'" by iprover with Abrupt show ?thesis by (fastforce intro: terminatess.intros terminates.intros elim: exec_Normal_elim_cases) next case Fault thus ?thesis by (iprover intro: terminatess_Fault) next case Stuck thus ?thesis by (iprover intro: terminatess_Stuck) qed qed next case CallUndefined thus ?case by (iprover intro: terminatess_Stuck) next case StuckProp show ?case by (rule terminatess_Stuck) next case StuckPropBlock show ?case by (rule terminatess_Stuck) next case DynCom thus ?case by (blast elim: terminatess_elim_cases terminates_Normal_elim_cases intro: terminatess.intros terminates.intros exec.intros) next case Throw thus ?case by (blast elim: terminatess_elim_cases terminates_Normal_elim_cases intro: terminatess.intros terminates.intros exec.intros) next case (Catch c1 c2 cs css s) then obtain term_c1: "\\c1 \ Normal s" and term_c2: "\s'. \\\c1,Normal s\ \ Abrupt s' \ \\c2 \ Normal s'"and term_rest: "\t. \\\Catch c1 c2,Normal s\ \ t \ \\cs,css\t" by (clarsimp elim!: terminatess_elim_cases terminates_Normal_elim_cases) show "\\[c1],(cs, c2 # cs) # css\Normal s" proof (rule terminatess.Cons [OF term_c1],clarsimp) fix t assume exec_c1: "\\\c1,Normal s\ \ t" show "\\[],(cs, c2 # cs) # css\t" proof (cases t) case (Normal t') with exec_c1 have "\\\Catch c1 c2,Normal s\ \ t" by (auto intro: exec.intros) with term_rest have "\\cs,css\t" by iprover with Normal show ?thesis by (iprover intro: terminatess.intros) next case (Abrupt t') with exec_c1 term_c2 have "\\c2 \ Normal t'" by auto moreover { fix w assume exec_c2: "\\\c2,Normal t'\ \ w" have "\\cs,css\w" proof - from exec_c1 Abrupt exec_c2 have "\\\Catch c1 c2,Normal s\ \ w" by (auto intro: exec.intros) with term_rest show ?thesis by simp qed } ultimately show ?thesis using Abrupt by (auto intro: terminatess.intros) next case Fault thus ?thesis by (iprover intro: terminatess_Fault) next case Stuck thus ?thesis by (iprover intro: terminatess_Stuck) qed qed qed ML \ ML_Thms.bind_thm ("rtrancl_induct3", Split_Rule.split_rule @{context} (Rule_Insts.read_instantiate @{context} [((("a", 0), Position.none), "(ax, ay, az)"), ((("b", 0), Position.none), "(bx, by, bz)")] [] @{thm rtranclp_induct})); \ lemma steps_preserves_terminations: assumes steps: "\\(cs,css,s) \\<^sup>* (cs',css',t)" shows "\\cs,css\s \ \\cs',css'\t" using steps proof (induct rule: rtrancl_induct3 [consumes 1]) assume "\\cs,css\s" then show "\\cs,css\s". next fix cs'' css'' w cs' css' t assume "\\(cs'',css'', w) \ (cs',css', t)" "\\cs,css\s \ \\cs'',css''\w" "\\cs,css\s" then show "\\cs',css'\t" by (blast dest: step_preserves_terminations) qed theorem steps_preserves_termination: assumes steps: "\\([c],[],s) \\<^sup>* (c'#cs',css',t)" assumes term_c: "\\c\s" shows "\\c'\t" proof - from term_c have "\\[c],[]\s" by (auto intro: terminatess.intros) from steps this have "\\c'#cs',css'\t" by (rule steps_preserves_terminations) thus "\\c'\t" by (auto elim: terminatess_elim_cases) qed lemma renumber': assumes f: "\i. (a,f i) \ r\<^sup>* \ (f i,f(Suc i)) \ r" assumes a_b: "(a,b) \ r\<^sup>*" shows "b = f 0 \ (\f. f 0 = a \ (\i. (f i, f(Suc i)) \ r))" using a_b proof (induct rule: converse_rtrancl_induct [consumes 1]) assume "b = f 0" with f show "\f. f 0 = b \ (\i. (f i, f (Suc i)) \ r)" by blast next fix a z assume a_z: "(a, z) \ r" and "(z, b) \ r\<^sup>*" assume "b = f 0 \ \f. f 0 = z \ (\i. (f i, f (Suc i)) \ r)" "b = f 0" then obtain f where f0: "f 0 = z" and seq: "\i. (f i, f (Suc i)) \ r" by iprover { fix i have "((\i. case i of 0 \ a | Suc i \ f i) i, f i) \ r" using seq a_z f0 by (cases i) auto } then show "\f. f 0 = a \ (\i. (f i, f (Suc i)) \ r)" by - (rule exI [where x="\i. case i of 0 \ a | Suc i \ f i"],simp) qed lemma renumber: "\i. (a,f i) \ r\<^sup>* \ (f i,f(Suc i)) \ r \ \f. f 0 = a \ (\i. (f i, f(Suc i)) \ r)" by(blast dest:renumber') lemma not_inf_Fault': assumes enum_step: "\i. \\f i \ f (Suc i)" shows "\k cs. f k = (cs,css,Fault m) \ False" proof (induct css) case Nil have f_k: "f k = (cs,[],Fault m)" by fact have "\k. f k = (cs,[],Fault m) \ False" proof (induct cs) case Nil have "f k = ([], [], Fault m)" by fact moreover from enum_step have "\\f k \ f (Suc k)".. ultimately show "False" by (fastforce elim: step_elim_cases) next case (Cons c cs) have fk: "f k = (c # cs, [], Fault m)" by fact from enum_step have "\\f k \ f (Suc k)".. with fk have "f (Suc k) = (cs,[],Fault m)" by (fastforce elim: step_elim_cases) with enum_step Cons.hyps show False by blast qed from this f_k show False by blast next case (Cons ds css) then obtain nrms abrs where ds: "ds=(nrms,abrs)" by (cases ds) auto have hyp: "\k cs. f k = (cs,css,Fault m) \ False" by fact have "\k. f k = (cs,(nrms,abrs)#css,Fault m) \ False" proof (induct cs) case Nil have fk: "f k = ([], (nrms, abrs) # css, Fault m)" by fact from enum_step have "\\f k \ f (Suc k)".. with fk have "f (Suc k) = (nrms,css,Fault m)" by (fastforce elim: step_elim_cases) thus ?case by (rule hyp) next case (Cons c cs) have fk: "f k = (c#cs, (nrms, abrs) # css, Fault m)" by fact have hyp1: "\k. f k = (cs, (nrms, abrs) # css, Fault m) \ False" by fact from enum_step have "\\f k \ f (Suc k)".. with fk have "f (Suc k) = (cs,(nrms,abrs)#css,Fault m)" by (fastforce elim: step_elim_cases) thus ?case by (rule hyp1) qed with ds Cons.prems show False by auto qed lemma not_inf_Fault: "\ inf \ cs css (Fault m)" apply (rule not_infI) apply (rule_tac f=f in not_inf_Fault' ) by auto lemma not_inf_Stuck': assumes enum_step: "\i. \\f i \ f (Suc i)" shows "\k cs. f k = (cs,css,Stuck) \ False" proof (induct css) case Nil have f_k: "f k = (cs,[],Stuck)" by fact have "\k. f k = (cs,[],Stuck) \ False" proof (induct cs) case Nil have "f k = ([], [], Stuck)" by fact moreover from enum_step have "\\f k \ f (Suc k)".. ultimately show "False" by (fastforce elim: step_elim_cases) next case (Cons c cs) have fk: "f k = (c # cs, [], Stuck)" by fact from enum_step have "\\f k \ f (Suc k)".. with fk have "f (Suc k) = (cs,[],Stuck)" by (fastforce elim: step_elim_cases) with enum_step Cons.hyps show False by blast qed from this f_k show False . next case (Cons ds css) then obtain nrms abrs where ds: "ds=(nrms,abrs)" by (cases ds) auto have hyp: "\k cs. f k = (cs,css,Stuck) \ False" by fact have "\k. f k = (cs,(nrms,abrs)#css,Stuck) \ False" proof (induct cs) case Nil have fk: "f k = ([], (nrms, abrs) # css, Stuck)" by fact from enum_step have "\\f k \ f (Suc k)".. with fk have "f (Suc k) = (nrms,css,Stuck)" by (fastforce elim: step_elim_cases) thus ?case by (rule hyp) next case (Cons c cs) have fk: "f k = (c#cs, (nrms, abrs) # css, Stuck)" by fact have hyp1: "\k. f k = (cs, (nrms, abrs) # css, Stuck) \ False" by fact from enum_step have "\\f k \ f (Suc k)".. with fk have "f (Suc k) = (cs,(nrms,abrs)#css,Stuck)" by (fastforce elim: step_elim_cases) thus ?case by (rule hyp1) qed with ds Cons.prems show False by auto qed lemma not_inf_Stuck: "\ inf \ cs css Stuck" apply (rule not_infI) apply (rule_tac f=f in not_inf_Stuck') by auto lemma last_butlast_app: assumes butlast: "butlast as = xs @ butlast bs" assumes not_Nil: "bs \ []" "as \ []" assumes last: "fst (last as) = fst (last bs)" "snd (last as) = snd (last bs)" shows "as = xs @ bs" proof - from last have "last as = last bs" by (cases "last as",cases "last bs") simp moreover from not_Nil have "as = butlast as @ [last as]" "bs = butlast bs @ [last bs]" by auto ultimately show ?thesis using butlast by simp qed lemma last_butlast_tl: assumes butlast: "butlast bs = x # butlast as" assumes not_Nil: "bs \ []" "as \ []" assumes last: "fst (last as) = fst (last bs)" "snd (last as) = snd (last bs)" shows "as = tl bs" proof - from last have "last as = last bs" by (cases "last as",cases "last bs") simp moreover from not_Nil have "as = butlast as @ [last as]" "bs = butlast bs @ [last bs]" by auto ultimately show ?thesis using butlast by simp qed locale inf = fixes CS:: "('s,'p,'f) config \ ('s, 'p,'f) com list" and CSS:: "('s,'p,'f) config \ ('s, 'p,'f) continuation list" and S:: "('s,'p,'f) config \ ('s,'f) xstate" defines CS_def : "CS \ fst" defines CSS_def : "CSS \ \c. fst (snd c)" defines S_def: "S \ \c. snd (snd c)" lemma (in inf) steps_hd_drop_suffix: assumes f_0: "f 0 = (c#cs,css,s)" assumes f_step: "\i. \\ f(i) \ f(Suc i)" assumes not_finished: "\i < k. \ (CS (f i) = cs \ CSS (f i) = css)" assumes simul: "\i\k. (if pcss i = [] then CSS (f i)=css \ CS (f i)=pcs i@cs else CS (f i)=pcs i \ CSS (f i)= butlast (pcss i)@ [(fst (last (pcss i))@cs,(snd (last (pcss i)))@cs)]@ css)" defines "p\\i. (pcs i, pcss i, S (f i))" shows "\i\ p i \ p (Suc i)" using not_finished simul proof (induct k) case 0 thus ?case by simp next case (Suc k) have simul: "\i\Suc k. (if pcss i = [] then CSS (f i)=css \ CS (f i)=pcs i@cs else CS (f i)=pcs i \ CSS (f i)= butlast (pcss i)@ [(fst (last (pcss i))@cs,(snd (last (pcss i)))@cs)]@ css)" by fact have not_finished': "\i < Suc k. \ (CS (f i) = cs \ CSS (f i) = css)" by fact with simul have not_finished: "\i (pcs i = [] \ pcss i = [])" by (auto simp add: CS_def CSS_def S_def split: if_split_asm) show ?case proof (clarify) fix i assume i_le_Suc_k: "i < Suc k" show "\\ p i \ p (Suc i)" proof (cases "i < k") case True with not_finished' simul Suc.hyps show ?thesis by auto next case False with i_le_Suc_k have eq_i_k: "i=k" by simp show "\\p i \ p (Suc i)" proof - obtain cs' css' t' where f_Suc_i: "f (Suc i) = (cs', css', t')" by (cases "f (Suc i)") obtain cs'' css'' t'' where f_i: "f i = (cs'',css'',t'')" by (cases "f i") from not_finished eq_i_k have pcs_pcss_not_Nil: "\ (pcs i = [] \ pcss i = [])" by auto from simul [rule_format, of i] i_le_Suc_k f_i have pcs_pcss_i: "if pcss i = [] then css''=css \ cs''=pcs i@cs else cs''=pcs i \ css''= butlast (pcss i)@ [(fst (last (pcss i))@cs,(snd (last (pcss i)))@cs)]@ css" by (simp add: CS_def CSS_def S_def cong: if_cong) from simul [rule_format, of "Suc i"] i_le_Suc_k f_Suc_i have pcs_pcss_Suc_i: "if pcss (Suc i) = [] then css' = css \ cs' = pcs (Suc i) @ cs else cs' = pcs (Suc i) \ css' = butlast (pcss (Suc i)) @ [(fst (last (pcss (Suc i))) @ cs, snd (last (pcss (Suc i))) @ cs)] @ css" by (simp add: CS_def CSS_def S_def cong: if_cong) show ?thesis proof (cases "pcss i = []") case True note pcss_Nil = this with pcs_pcss_i pcs_pcss_not_Nil obtain p ps where pcs_i: "pcs i = p#ps" and css'': "css''=css" and cs'': "cs''=(p#ps)@cs" by (auto simp add: neq_Nil_conv) with f_i have "f i = (p#(ps@cs),css,t'')" by simp with f_Suc_i f_step [rule_format, of i] have step_css: "\\ (p#(ps@cs),css,t'') \ (cs',css',t')" by simp from step_Cons' [OF this, of p "ps@cs"] obtain css''' where css''': "css' = css''' @ css" "if css''' = [] then \p. cs' = p @ ps @ cs else (\pnorm pabr. css'''=[(pnorm @ ps @ cs,pabr @ ps @ cs)])" by auto show ?thesis proof (cases "css''' = []") case True with css''' obtain p' where css': "css' = css" and cs': "cs' = p' @ ps @ cs" by auto (*from cs' css' f_Suc_i f_i [rule_format, of "Suc k"] have p_ps_not_Nil: "p'@ps \ Nil" by auto*) from css' cs' step_css have step: "\\ (p#(ps@cs),css,t'') \ (p'@ps@cs,css,t')" by simp hence "\\ ((p#ps)@cs,css,t'') \ ((p'@ps)@cs,css,t')" by simp from drop_suffix_css_step' [OF drop_suffix_same_css_step [OF this], where xs="css" and css="[]" and css'="[]"] have "\\ (p#ps,[],t'') \ (p'@ps,[],t')" by simp moreover from css' cs' pcs_pcss_Suc_i obtain "pcs (Suc i) = p'@ps" and "pcss (Suc i) = []" by (simp split: if_split_asm) ultimately show ?thesis using pcs_i pcss_Nil f_i f_Suc_i by (simp add: CS_def CSS_def S_def p_def) next case False with css''' obtain pnorm pabr where css': "css'=css'''@css" "css'''=[(pnorm @ ps @ cs,pabr @ ps @ cs)]" by auto with css''' step_css have "\\ (p#ps@cs,css,t'') \ (cs',[(pnorm@ps@cs,pabr@ps@cs)]@css,t')" by simp then have "\\(p#ps, css, t'') \ (cs', [(pnorm@ps, pabr@ps)] @ css, t')" by (rule drop_suffix_hd_css_step) from drop_suffix_css_step' [OF this, where css="[]" and xs="css" and css'="[(pnorm@ps, pabr@ps)]"] have "\\ (p#ps,[],t'') \ (cs',[(pnorm@ps, pabr@ps)],t')" by simp moreover from css' pcs_pcss_Suc_i obtain "pcs (Suc i) = cs'" "pcss (Suc i) = [(pnorm@ps, pabr@ps)]" apply (cases "pcss (Suc i)") apply (auto split: if_split_asm) done ultimately show ?thesis using pcs_i pcss_Nil f_i f_Suc_i by (simp add: p_def CS_def CSS_def S_def) qed next case False note pcss_i_not_Nil = this with pcs_pcss_i obtain cs'': "cs''=pcs i" and css'': "css''= butlast (pcss i)@ [(fst (last (pcss i))@cs,(snd (last (pcss i)))@cs)]@ css" by auto from f_Suc_i f_i f_step [rule_format, of i] have step_i_full: "\\ (cs'',css'',t'') \ (cs',css',t')" by simp show ?thesis proof (cases cs'') case (Cons c' cs) with step_Cons' [OF step_i_full] obtain css''' where css': "css' = css'''@css''" by auto with step_i_full have "\\ (cs'',css'',t'') \ (cs',css'''@css'',t')" by simp from Cons_change_css_step [OF this, where xss="pcss i"] Cons cs'' have "\\ (pcs i, pcss i,t'') \ (cs',css'''@pcss i,t')" by simp moreover from cs'' css'' css' False pcs_pcss_Suc_i obtain "pcs (Suc i) = cs'" "pcss (Suc i) = css'''@pcss i" apply (auto split: if_split_asm) apply (drule (4) last_butlast_app) by simp ultimately show ?thesis using f_i f_Suc_i by (simp add: p_def CS_def CSS_def S_def) next case Nil note cs''_Nil = this show ?thesis proof (cases "butlast (pcss i)") case (Cons bpcs bpcss) with cs''_Nil step_i_full css'' have *: "\\ ([],[hd css'']@tl css'',t'') \ (cs',css',t')" by simp moreover from step_Nil [OF *] have css': "css'=tl css''" by simp ultimately have step_i_full: "\\ ([],[hd css'']@tl css'',t'') \ (cs',tl css'',t')" by simp from css'' Cons pcss_i_not_Nil have "hd css'' = hd (pcss i)" by (auto simp add: neq_Nil_conv split: if_split_asm) with cs'' cs''_Nil Nil_change_css_step [where ass="[hd css'']" and css="tl css''" and ass'="[]" and xss="tl (pcss i)", simplified, OF step_i_full [simplified]] have "\\ (pcs i,[hd (pcss i)]@tl (pcss i),t'') \ (cs',tl (pcss i),t')" by simp with pcss_i_not_Nil have "\\ (pcs i,pcss i,t'') \ (cs',tl (pcss i),t')" by simp moreover from css' css'' cs''_Nil Cons pcss_i_not_Nil pcs_pcss_Suc_i obtain "pcs (Suc i) = cs'" "pcss (Suc i) = tl (pcss i)" apply (clarsimp split: if_split_asm) apply (drule (4) last_butlast_tl) by simp ultimately show ?thesis using f_i f_Suc_i by (simp add: p_def CS_def CSS_def S_def) next case Nil with css'' pcss_i_not_Nil obtain pnorm pabr where css'': "css''= [(pnorm@cs,pabr@cs)]@css" and pcss_i: "pcss i = [(pnorm,pabr)]" by (force simp add: neq_Nil_conv split: if_split_asm) with cs''_Nil step_i_full have "\\([],[(pnorm@cs,pabr@cs)]@css,t'') \ (cs',css',t')" by simp from step_Nil [OF this] obtain css': "css'=css" and cs': "(case t'' of Abrupt s' \ cs' = pabr @ cs \ t' = Normal s' | _ \ cs' = pnorm @ cs \ t' = t'')" by (simp cong: xstate.case_cong) let "?pcs_Suc_i " = "(case t'' of Abrupt s' \ pabr | _ \ pnorm)" from cs' have "\\([],[(pnorm,pabr)],t'') \ (?pcs_Suc_i,[],t')" by (auto intro: step.intros split: xstate.splits) moreover from css'' css' cs' pcss_i pcs_pcss_Suc_i obtain "pcs (Suc i) = ?pcs_Suc_i" "pcss (Suc i) = []" by (simp split: if_split_asm xstate.splits) ultimately show ?thesis using pcss_i cs'' cs''_Nil f_i f_Suc_i by (simp add: p_def CS_def CSS_def S_def) qed qed qed qed qed qed qed lemma k_steps_to_rtrancl: assumes steps: "\i\ p i \ p (Suc i)" shows "\\p 0\\<^sup>* p k" using steps proof (induct k) case 0 thus ?case by auto next case (Suc k) have "\i\ p i \ p (Suc i)" by fact then obtain step_le_k: "\i\ p i \ p (Suc i)" and step_k: "\\ p k \ p (Suc k)" by auto from Suc.hyps [OF step_le_k] have "\\ p 0 \\<^sup>* p k". also note step_k finally show ?case . qed lemma (in inf) steps_hd_drop_suffix_finite: assumes f_0: "f 0 = (c#cs,css,s)" assumes f_step: "\i. \\ f(i) \ f(Suc i)" assumes not_finished: "\i < k. \ (CS (f i) = cs \ CSS (f i) = css)" assumes simul: "\i\k. (if pcss i = [] then CSS (f i)=css \ CS (f i)=pcs i@cs else CS (f i)=pcs i \ CSS (f i)= butlast (pcss i)@ [(fst (last (pcss i))@cs,(snd (last (pcss i)))@cs)]@ css)" shows "\\([c],[],s) \\<^sup>* (pcs k, pcss k, S (f k))" proof - from steps_hd_drop_suffix [OF f_0 f_step not_finished simul] have "\i\ (pcs i, pcss i, S (f i)) \ (pcs (Suc i), pcss (Suc i), S (f (Suc i)))". from k_steps_to_rtrancl [OF this] have "\\ (pcs 0, pcss 0, S (f 0)) \\<^sup>* (pcs k, pcss k, S (f k))". moreover from f_0 simul [rule_format, of 0] have "(pcs 0, pcss 0, S (f 0)) = ([c],[],s)" by (auto split: if_split_asm simp add: CS_def CSS_def S_def) ultimately show ?thesis by simp qed lemma (in inf) steps_hd_drop_suffix_infinite: assumes f_0: "f 0 = (c#cs,css,s)" assumes f_step: "\i. \\ f(i) \ f(Suc i)" assumes not_finished: "\i. \ (CS (f i) = cs \ CSS (f i) = css)" (*assumes not_finished: "\i. \ (pcs i = [] \ pcss i = [])"*) assumes simul: "\i. (if pcss i = [] then CSS (f i)=css \ CS (f i)=pcs i@cs else CS (f i)=pcs i \ CSS (f i)= butlast (pcss i)@ [(fst (last (pcss i))@cs,(snd (last (pcss i)))@cs)]@ css)" defines "p\\i. (pcs i, pcss i, S (f i))" shows "\\ p i \ p (Suc i)" proof - from steps_hd_drop_suffix [OF f_0 f_step, of "Suc i" pcss pcs] not_finished simul show ?thesis by (auto simp add: p_def) qed lemma (in inf) steps_hd_progress: assumes f_0: "f 0 = (c#cs,css,s)" assumes f_step: "\i. \\ f(i) \ f(Suc i)" assumes c_unfinished: "\i < k. \ (CS (f i) = cs \ CSS (f i) = css)" shows "\i \ k. (\pcs pcss. (if pcss = [] then CSS (f i)=css \ CS (f i)=pcs@cs else CS (f i)=pcs \ CSS (f i)= butlast pcss@ [(fst (last pcss)@cs,(snd (last pcss))@cs)]@ css))" using c_unfinished proof (induct k) case 0 with f_0 show ?case by (simp add: CSS_def CS_def) next case (Suc k) have c_unfinished: "\i (CS (f i) = cs \ CSS (f i) = css)" by fact hence c_unfinished': "\i< k. \ (CS (f i) = cs \ CSS (f i) = css)" by simp show ?case proof (clarify) fix i assume i_le_Suc_k: "i \ Suc k" show "\pcs pcss. (if pcss = [] then CSS (f i)=css \ CS (f i)=pcs@cs else CS (f i)=pcs \ CSS (f i)= butlast pcss@ [(fst (last pcss)@cs,(snd (last pcss))@cs)]@ css)" proof (cases "i < Suc k") case True with Suc.hyps [OF c_unfinished', rule_format, of i] c_unfinished show ?thesis by auto next case False with i_le_Suc_k have eq_i_Suc_k: "i=Suc k" by auto obtain cs' css' t' where f_Suc_k: "f (Suc k) = (cs', css', t')" by (cases "f (Suc k)") obtain cs'' css'' t'' where f_k: "f k = (cs'',css'',t'')" by (cases "f k") with Suc.hyps [OF c_unfinished',rule_format, of k] obtain pcs pcss where pcs_pcss_k: "if pcss = [] then css'' = css \ cs'' = pcs @ cs else cs'' = pcs \ css'' = butlast pcss @ [(fst (last pcss) @ cs, snd (last pcss) @ cs)] @ css" by (auto simp add: CSS_def CS_def cong: if_cong) from c_unfinished [rule_format, of k] f_k pcs_pcss_k have pcs_pcss_empty: "\ (pcs = [] \ pcss = [])" by (auto simp add: CS_def CSS_def S_def split: if_split_asm) show ?thesis proof (cases "pcss = []") case True note pcss_Nil = this with pcs_pcss_k pcs_pcss_empty obtain p ps where pcs_i: "pcs = p#ps" and css'': "css''=css" and cs'': "cs''=(p#ps)@cs" by (cases "pcs") auto with f_k have "f k = (p#(ps@cs),css,t'')" by simp with f_Suc_k f_step [rule_format, of k] have step_css: "\\ (p#(ps@cs),css,t'') \ (cs',css',t')" by simp from step_Cons' [OF this, of p "ps@cs"] obtain css''' where css''': "css' = css''' @ css" "if css''' = [] then \p. cs' = p @ ps @ cs else (\pnorm pabr. css'''=[(pnorm @ ps @ cs,pabr @ ps @ cs)])" by auto show ?thesis proof (cases "css''' = []") case True with css''' obtain p' where css': "css' = css" and cs': "cs' = p' @ ps @ cs" by auto from css' cs' f_Suc_k show ?thesis apply (rule_tac x="p'@ps" in exI) apply (rule_tac x="[]" in exI) apply (simp add: CSS_def CS_def eq_i_Suc_k) done next case False with css''' obtain pnorm pabr where css': "css'=css'''@css" "css'''=[(pnorm @ ps @ cs,pabr @ ps @ cs)]" by auto with f_Suc_k eq_i_Suc_k show ?thesis apply (rule_tac x="cs'" in exI) apply (rule_tac x="[(pnorm@ps, pabr@ps)]" in exI) by (simp add: CSS_def CS_def) qed next case False note pcss_k_not_Nil = this with pcs_pcss_k obtain cs'': "cs''=pcs" and css'': "css''= butlast pcss@ [(fst (last pcss)@cs,(snd (last pcss))@cs)]@ css" by auto from f_Suc_k f_k f_step [rule_format, of k] have step_i_full: "\\ (cs'',css'',t'') \ (cs',css',t')" by simp show ?thesis proof (cases cs'') case (Cons c' cs) with step_Cons' [OF step_i_full] obtain css''' where css': "css' = css'''@css''" by auto with cs'' css'' f_Suc_k eq_i_Suc_k pcss_k_not_Nil show ?thesis apply (rule_tac x="cs'" in exI) apply (rule_tac x="css'''@pcss" in exI) by (clarsimp simp add: CSS_def CS_def butlast_append) next case Nil note cs''_Nil = this show ?thesis proof (cases "butlast pcss") case (Cons bpcs bpcss) with cs''_Nil step_i_full css'' have *: "\\ ([],[hd css'']@tl css'',t'') \ (cs',css',t')" by simp moreover from step_Nil [OF *] obtain css': "css'=tl css''" and cs': "cs' = (case t'' of Abrupt s' \ snd (hd css'') | _ \ fst (hd css''))" by (auto split: xstate.splits) from css'' Cons pcss_k_not_Nil have "hd css'' = hd pcss" by (auto simp add: neq_Nil_conv split: if_split_asm) with css' cs' css'' cs''_Nil Cons pcss_k_not_Nil f_Suc_k eq_i_Suc_k show ?thesis apply (rule_tac x="cs'" in exI) apply (rule_tac x="tl pcss" in exI) apply (clarsimp split: xstate.splits simp add: CS_def CSS_def neq_Nil_conv split: if_split_asm) done next case Nil with css'' pcss_k_not_Nil obtain pnorm pabr where css'': "css''= [(pnorm@cs,pabr@cs)]@css" and pcss_k: "pcss = [(pnorm,pabr)]" by (force simp add: neq_Nil_conv split: if_split_asm) with cs''_Nil step_i_full have "\\([],[(pnorm@cs,pabr@cs)]@css,t'') \ (cs',css',t')" by simp from step_Nil [OF this] obtain css': "css'=css" and cs': "(case t'' of Abrupt s' \ cs' = pabr @ cs \ t' = Normal s' | _ \ cs' = pnorm @ cs \ t' = t'')" by (simp cong: xstate.case_cong) let "?pcs_Suc_k " = "(case t'' of Abrupt s' \ pabr | _ \ pnorm)" from css'' css' cs' pcss_k f_Suc_k eq_i_Suc_k show ?thesis apply (rule_tac x="?pcs_Suc_k" in exI) apply (rule_tac x="[]" in exI) apply (simp split: xstate.splits add: CS_def CSS_def) done qed qed qed qed qed qed lemma (in inf) inf_progress: assumes f_0: "f 0 = (c#cs,css,s)" assumes f_step: "\i. \\ f(i) \ f(Suc i)" assumes unfinished: "\i. \ ((CS (f i) = cs) \ (CSS (f i) = css))" shows "\pcs pcss. (if pcss = [] then CSS (f i)=css \ CS (f i)=pcs@cs else CS (f i)=pcs \ CSS (f i)= butlast pcss@ [(fst (last pcss)@cs,(snd (last pcss))@cs)]@ css)" proof - from steps_hd_progress [OF f_0 f_step, of "i"] unfinished show ?thesis by auto qed lemma skolemize1: "\x. P x \ (\y. Q x y) \ \f.\x. P x \ Q x (f x)" by (rule choice) blast lemma skolemize2: "\x. P x \ (\y z. Q x y z) \ \f g.\x. P x \ Q x (f x) (g x)" apply (drule skolemize1) apply (erule exE) apply (drule skolemize1) apply fast done lemma skolemize2': "\x.\y z. P x y z \ \f g.\x. P x (f x) (g x)" apply (drule choice) apply (erule exE) apply (drule choice) apply fast done theorem (in inf) inf_cases: fixes c::"('s,'p,'f) com" assumes inf: "inf \ (c#cs) css s" shows "inf \ [c] [] s \ (\t. \\\c,s\ \ t \ inf \ cs css t)" proof - from inf obtain f where f_0: "f 0 = (c#cs,css,s)" and f_step: "(\i. \\f i \ f (Suc i))" by (auto simp add: inf_def) show ?thesis proof (cases "\i. CS (f i) = cs \ CSS (f i) = css") case True define k where "k = (LEAST i. CS (f i) = cs \ CSS (f i) = css)" from True obtain CS_f_k: "CS (f k) = cs" and CSS_f_k: "CSS (f k) = css" apply - apply (erule exE) apply (drule LeastI) apply (simp add: k_def) done have less_k_prop: "\i (CS (f i) = cs \ CSS (f i) = css)" apply (intro allI impI) apply (unfold k_def) apply (drule not_less_Least) apply simp done have "\\([c], [], s) \\<^sup>* ([],[],S (f k))" proof - have "\i\k. \pcs pcss. (if pcss = [] then CSS (f i)=css \ CS (f i)=pcs@cs else CS (f i)=pcs \ CSS (f i)= butlast pcss@ [(fst (last pcss)@cs,(snd (last pcss))@cs)]@ css)" by (rule steps_hd_progress [OF f_0 f_step, where k=k, OF less_k_prop]) from skolemize2 [OF this] obtain pcs pcss where pcs_pcss: "\i\k. (if pcss i = [] then CSS (f i)=css \ CS (f i)=pcs i@cs else CS (f i)=pcs i \ CSS (f i)= butlast (pcss i)@ [(fst (last (pcss i))@cs,(snd (last (pcss i)))@cs)]@ css)" by iprover from pcs_pcss [rule_format, of k] CS_f_k CSS_f_k have finished: "pcs k = []" "pcss k = []" by (auto simp add: CS_def CSS_def S_def split: if_split_asm) from pcs_pcss have simul: "\i\k. (if pcss i = [] then CSS (f i)=css \ CS (f i)=pcs i@cs else CS (f i)=pcs i \ CSS (f i)= butlast (pcss i)@ [(fst (last (pcss i))@cs,(snd (last (pcss i)))@cs)]@ css)" by auto from steps_hd_drop_suffix_finite [OF f_0 f_step less_k_prop simul] finished show ?thesis by simp qed hence "\\\c,s\ \ S (f k)" by (rule steps_impl_exec) moreover from CS_f_k CSS_f_k f_step have "inf \ cs css (S (f k))" apply (simp add: inf_def) apply (rule_tac x="\i. f (i + k)" in exI) apply simp apply (auto simp add: CS_def CSS_def S_def) done ultimately have "(\t. \\\c,s\ \ t \ inf \ cs css t)" by blast thus ?thesis by simp next case False hence unfinished: "\i. \ ((CS (f i) = cs) \ (CSS (f i) = css))" by simp from inf_progress [OF f_0 f_step this] have "\i. \pcs pcss. (if pcss = [] then CSS (f i)=css \ CS (f i)=pcs@cs else CS (f i)=pcs \ CSS (f i)= butlast pcss@ [(fst (last pcss)@cs,(snd (last pcss))@cs)]@ css)" by auto from skolemize2' [OF this] obtain pcs pcss where pcs_pcss: "\i. (if pcss i = [] then CSS (f i)=css \ CS (f i)=pcs i@cs else CS (f i)=pcs i \ CSS (f i)= butlast (pcss i)@ [(fst (last (pcss i))@cs,(snd (last (pcss i)))@cs)]@ css)" by iprover define g where "g i = (pcs i, pcss i, S (f i))" for i from pcs_pcss [rule_format, of 0] f_0 have "g 0 = ([c],[],s)" by (auto split: if_split_asm simp add: CS_def CSS_def S_def g_def) moreover from steps_hd_drop_suffix_infinite [OF f_0 f_step unfinished pcs_pcss] have "\i. \\g i \ g (Suc i)" by (simp add: g_def) ultimately have "inf \ [c] [] s" by (auto simp add: inf_def) thus ?thesis by simp qed qed lemma infE [consumes 1]: assumes inf: "inf \ (c#cs) css s" assumes cases: "inf \ [c] [] s \ P" "\t. \\\\c,s\ \ t; inf \ cs css t\ \ P" shows P using inf cases apply - apply (drule inf.inf_cases) apply auto done lemma inf_Seq: "inf \ (Seq c1 c2#cs) css (Normal s) = inf \ (c1#c2#cs) css (Normal s)" proof assume "inf \ (Seq c1 c2 # cs) css (Normal s)" then obtain f where f_0: "f 0 = (Seq c1 c2#cs,css,Normal s)" and f_step: "\i. \\f i \ f (Suc i)" by (auto simp add: inf_def) from f_step [rule_format, of 0] f_0 have "f 1 = (c1#c2#cs,css,Normal s)" by (auto elim: step_Normal_elim_cases) with f_step show "inf \ (c1#c2#cs) css (Normal s)" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) apply simp done next assume "inf \ (c1 # c2 # cs) css (Normal s)" then obtain f where f_0: "f 0 = (c1# c2#cs,css,Normal s)" and f_step: "\i. \\f i \ f (Suc i)" by (auto simp add: inf_def) define g where "g i = (case i of 0 \ (Seq c1 c2#cs,css,Normal s) | Suc j \ f j)" for i with f_0 have "\\g 0 \ g (Suc 0)" by (auto intro: step.intros) moreover from f_step have "\i. i\0 \ \\g i \ g (Suc i)" by (auto simp add: g_def split: nat.splits) ultimately show "inf \ (Seq c1 c2 # cs) css (Normal s)" apply (simp add: inf_def) apply (rule_tac x=g in exI) apply (simp add: g_def split: nat.splits) done qed lemma inf_WhileTrue: assumes b: "s \ b" shows "inf \ (While b c#cs) css (Normal s) = inf \ (c#While b c#cs) css (Normal s)" proof assume "inf \ (While b c # cs) css (Normal s)" then obtain f where f_0: "f 0 = (While b c#cs,css,Normal s)" and f_step: "\i. \\f i \ f (Suc i)" by (auto simp add: inf_def) from b f_step [rule_format, of 0] f_0 have "f 1 = (c#While b c#cs,css,Normal s)" by (auto elim: step_Normal_elim_cases) with f_step show "inf \ (c # While b c # cs) css (Normal s)" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) apply simp done next assume "inf \ (c # While b c # cs) css (Normal s)" then obtain f where f_0: "f 0 = (c # While b c #cs,css,Normal s)" and f_step: "\i. \\f i \ f (Suc i)" by (auto simp add: inf_def) define h where "h i = (case i of 0 \ (While b c#cs,css,Normal s) | Suc j \ f j)" for i with b f_0 have "\\h 0 \ h (Suc 0)" by (auto intro: step.intros) moreover from f_step have "\i. i\0 \ \\h i \ h (Suc i)" by (auto simp add: h_def split: nat.splits) ultimately show "inf \ (While b c # cs) css (Normal s)" apply (simp add: inf_def) apply (rule_tac x=h in exI) apply (simp add: h_def split: nat.splits) done qed lemma inf_Catch: "inf \ (Catch c1 c2#cs) css (Normal s) = inf \ [c1] ((cs,c2#cs)#css) (Normal s)" proof assume "inf \ (Catch c1 c2#cs) css (Normal s)" then obtain f where f_0: "f 0 = (Catch c1 c2#cs,css,Normal s)" and f_step: "\i. \\f i \ f (Suc i)" by (auto simp add: inf_def) from f_step [rule_format, of 0] f_0 have "f 1 = ([c1],(cs,c2#cs)#css,Normal s)" by (auto elim: step_Normal_elim_cases) with f_step show "inf \ [c1] ((cs,c2#cs)#css) (Normal s)" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) apply simp done next assume "inf \ [c1] ((cs,c2#cs)#css) (Normal s)" then obtain f where f_0: "f 0 = ([c1],(cs,c2#cs)#css,Normal s)" and f_step: "\i. \\f i \ f (Suc i)" by (auto simp add: inf_def) define h where "h i = (case i of 0 \ (Catch c1 c2#cs,css,Normal s) | Suc j \ f j)" for i with f_0 have "\\h 0 \ h (Suc 0)" by (auto intro: step.intros) moreover from f_step have "\i. i\0 \ \\h i \ h (Suc i)" by (auto simp add: h_def split: nat.splits) ultimately show "inf \ (Catch c1 c2 # cs) css (Normal s)" apply (simp add: inf_def) apply (rule_tac x=h in exI) apply (simp add: h_def split: nat.splits) done qed theorem terminates_impl_not_inf: assumes termi: "\\c \ s" shows "\inf \ [c] [] s" using termi proof induct case (Skip s) thus ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = ([Skip], [], Normal s)" from f_step [of 0] f_0 have "f (Suc 0) = ([],[],Normal s)" by (auto elim: step_Normal_elim_cases) with f_step [of 1] show False by (auto elim: step_elim_cases) qed next case (Basic g s) thus ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = ([Basic g], [], Normal s)" from f_step [of 0] f_0 have "f (Suc 0) = ([],[],Normal (g s))" by (auto elim: step_Normal_elim_cases) with f_step [of 1] show False by (auto elim: step_elim_cases) qed next case (Spec r s) thus ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = ([Spec r], [], Normal s)" with f_step [of 0] have "\\([Spec r], [], Normal s) \ f (Suc 0)" by simp then show False proof (cases) fix t assume "(s, t) \ r" "f (Suc 0) = ([], [], Normal t)" with f_step [of 1] show False by (auto elim: step_elim_cases) next assume "\t. (s, t) \ r" "f (Suc 0) = ([], [], Stuck)" with f_step [of 1] show False by (auto elim: step_elim_cases) qed qed next case (Guard s g c m) have g: "s \ g" by fact have hyp: "\ inf \ [c] [] (Normal s)" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = ([Guard m g c], [], Normal s)" from g f_step [of 0] f_0 have "f (Suc 0) = ([c],[],Normal s)" by (auto elim: step_Normal_elim_cases) with f_step have "inf \ [c] [] (Normal s)" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) by simp with hyp show False .. qed next case (GuardFault s g m c) have g: "s \ g" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = ([Guard m g c], [], Normal s)" from g f_step [of 0] f_0 have "f (Suc 0) = ([],[],Fault m)" by (auto elim: step_Normal_elim_cases) with f_step [of 1] show False by (auto elim: step_elim_cases) qed next case (Fault c m) thus ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = ([c], [], Fault m)" from f_step [of 0] f_0 have "f (Suc 0) = ([],[],Fault m)" by (auto elim: step_Normal_elim_cases) with f_step [of 1] show False by (auto elim: step_elim_cases) qed next case (Seq c1 s c2) have hyp_c1: "\ inf \ [c1] [] (Normal s)" by fact have hyp_c2: "\s'. \\\c1,Normal s\ \ s' \ \\c2 \ s' \ \ inf \ [c2] [] s'" by fact have "\ inf \ ([c1,c2]) [] (Normal s)" proof assume "inf \ [c1, c2] [] (Normal s)" then show False proof (cases rule: infE) assume "inf \ [c1] [] (Normal s)" with hyp_c1 show ?thesis by simp next fix t assume "\\\c1,Normal s\ \ t" "inf \ [c2] [] t" with hyp_c2 show ?thesis by simp qed qed thus ?case by (simp add: inf_Seq) next case (CondTrue s b c1 c2) have b: "s \ b" by fact have hyp_c1: "\ inf \ [c1] [] (Normal s)" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = ([Cond b c1 c2], [], Normal s)" from b f_step [of 0] f_0 have "f 1 = ([c1],[],Normal s)" by (auto elim: step_Normal_elim_cases) with f_step have "inf \ [c1] [] (Normal s)" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) by simp with hyp_c1 show False by simp qed next case (CondFalse s b c2 c1) have b: "s \ b" by fact have hyp_c2: "\ inf \ [c2] [] (Normal s)" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = ([Cond b c1 c2], [], Normal s)" from b f_step [of 0] f_0 have "f 1 = ([c2],[],Normal s)" by (auto elim: step_Normal_elim_cases) with f_step have "inf \ [c2] [] (Normal s)" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) by simp with hyp_c2 show False by simp qed next case (WhileTrue s b c) have b: "s \ b" by fact have hyp_c: "\ inf \ [c] [] (Normal s)" by fact have hyp_w: "\s'. \\\c,Normal s\ \ s' \ \\While b c \ s' \ \ inf \ [While b c] [] s'" by fact have "\ inf \ [c,While b c] [] (Normal s)" proof assume "inf \ [c,While b c] [] (Normal s)" from this hyp_c hyp_w show False by (cases rule: infE) auto qed with b show ?case by (simp add: inf_WhileTrue) next case (WhileFalse s b c) have b: "s \ b" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = ([While b c], [], Normal s)" from b f_step [of 0] f_0 have "f (Suc 0) = ([],[],Normal s)" by (auto elim: step_Normal_elim_cases) with f_step [of 1] show False by (auto elim: step_elim_cases) qed next case (Call p bdy s) have bdy: "\ p = Some bdy" by fact have hyp: "\ inf \ [bdy] [] (Normal s)" by fact have not_inf_bdy: "\ inf \ [bdy] [([],[Throw])] (Normal s)" proof assume "inf \ [bdy] [([],[Throw])] (Normal s)" then show False proof (rule infE) assume "inf \ [bdy] [] (Normal s)" with hyp show False by simp next fix t assume "\\\bdy,Normal s\ \ t" assume inf: "inf \ [] [([], [Throw])] t" then obtain f where f_0: "f 0 = ([],[([], [Throw])],t)" and f_step: "\i. \\f i \ f (Suc i)" by (auto simp add: inf_def) show False proof (cases t) case (Normal t') with f_0 f_step [rule_format, of 0] have "f (Suc 0) = ([],[],(Normal t'))" by (auto elim: step_Normal_elim_cases) with f_step [rule_format, of "Suc 0"] show False by (auto elim: step.cases) next case (Abrupt t') with f_0 f_step [rule_format, of 0] have "f (Suc 0) = ([Throw],[],(Normal t'))" by (auto elim: step_Normal_elim_cases) with f_step [rule_format, of "Suc 0"] have "f (Suc (Suc 0)) = ([],[],(Abrupt t'))" by (auto elim: step_Normal_elim_cases) with f_step [rule_format, of "Suc(Suc 0)"] show False by (auto elim: step.cases) next case (Fault m) with f_0 f_step [rule_format, of 0] have "f (Suc 0) = ([],[],Fault m)" by (auto elim: step_Normal_elim_cases) with f_step [rule_format, of 1] have "f (Suc (Suc 0)) = ([],[],Fault m)" by (auto elim: step_Normal_elim_cases) with f_step [rule_format, of "Suc (Suc 0)"] show False by (auto elim: step.cases) next case Stuck with f_0 f_step [rule_format, of 0] have "f (Suc 0) = ([],[],Stuck)" by (auto elim: step_Normal_elim_cases) with f_step [rule_format, of 1] have "f (Suc (Suc 0)) = ([],[],Stuck)" by (auto elim: step_Normal_elim_cases) with f_step [rule_format, of "Suc (Suc 0)"] show False by (auto elim: step.cases) qed qed qed show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = ([Call p], [], Normal s)" from bdy f_step [of 0] f_0 have "f (Suc 0) = ([bdy],[([],[Throw])],Normal s)" by (auto elim: step_Normal_elim_cases) with f_step have "inf \ [bdy] [([],[Throw])] (Normal s)" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) by simp with not_inf_bdy show False by simp qed next case (CallUndefined p s) have undef: "\ p = None" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = ([Call p], [], Normal s)" from undef f_step [of 0] f_0 have "f (Suc 0) = ([],[],Stuck)" by (auto elim: step_Normal_elim_cases) with f_step [rule_format, of "Suc 0"] show False by (auto elim: step_elim_cases) qed next case (Stuck c) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = ([c], [], Stuck)" from f_step [of 0] f_0 have "f (Suc 0) = ([],[],Stuck)" by (auto elim: step_elim_cases) with f_step [rule_format, of "Suc 0"] show False by (auto elim: step_elim_cases) qed next case (DynCom c s) have hyp: "\ inf \ [(c s)] [] (Normal s)" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = ([DynCom c], [], Normal s)" from f_step [of 0] f_0 have "f (Suc 0) = ([(c s)], [], Normal s)" by (auto elim: step_elim_cases) with f_step have "inf \ [(c s)] [] (Normal s)" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) by simp with hyp show False by simp qed next case (Throw s) thus ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = ([Throw], [], Normal s)" from f_step [of 0] f_0 have "f (Suc 0) = ([],[],Abrupt s)" by (auto elim: step_Normal_elim_cases) with f_step [of 1] show False by (auto elim: step_elim_cases) qed next case (Abrupt c s) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = ([c], [], Abrupt s)" from f_step [of 0] f_0 have "f (Suc 0) = ([],[],Abrupt s)" by (auto elim: step_elim_cases) with f_step [rule_format, of "Suc 0"] show False by (auto elim: step_elim_cases) qed next case (Catch c1 s c2) have hyp_c1: "\ inf \ [c1] [] (Normal s)" by fact have hyp_c2: "\s'. \\\c1,Normal s\ \ Abrupt s' \ \\c2 \ Normal s' \ \ inf \ [c2] [] (Normal s')" by fact have "\ inf \ [c1] [([],[c2])] (Normal s)" proof assume "inf \ [c1] [([],[c2])] (Normal s)" then show False proof (rule infE) assume "inf \ [c1] [] (Normal s)" with hyp_c1 show False by simp next fix t assume eval: "\\\c1,Normal s\ \ t" assume inf: "inf \ [] [([], [c2])] t" then obtain f where f_0: "f 0 = ([],[([], [c2] )],t)" and f_step: "\i. \\f i \ f (Suc i)" by (auto simp add: inf_def) show False proof (cases t) case (Normal t') with f_0 f_step [rule_format, of 0] have "f (Suc 0) = ([],[],Normal t')" by (auto elim: step_Normal_elim_cases) with f_step [rule_format, of 1] show False by (auto elim: step_elim_cases) next case (Abrupt t') with f_0 f_step [rule_format, of 0] have "f (Suc 0) = ([c2],[],Normal t')" by (auto elim: step_Normal_elim_cases) with f_step eval Abrupt have "inf \ [c2] [] (Normal t')" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) by simp with eval hyp_c2 Abrupt show False by simp next case (Fault m) with f_0 f_step [rule_format, of 0] have "f (Suc 0) = ([],[],Fault m)" by (auto elim: step_Normal_elim_cases) with f_step [rule_format, of 1] show False by (auto elim: step_elim_cases) next case Stuck with f_0 f_step [rule_format, of 0] have "f (Suc 0) = ([],[],Stuck)" by (auto elim: step_Normal_elim_cases) with f_step [rule_format, of 1] show False by (auto elim: step_elim_cases) qed qed qed thus ?case by (simp add: inf_Catch) qed lemma terminatess_impl_not_inf: assumes termi: "\\cs,css\s" shows "\inf \ cs css s" using termi proof (induct) case (Nil s) show ?case proof (rule not_infI) fix f assume "\i. \\f i \ f (Suc i)" hence "\\f 0 \ f (Suc 0)" by simp moreover assume "f 0 = ([], [], s)" ultimately show False by (fastforce elim: step.cases) qed next case (ExitBlockNormal nrms css s abrs) have hyp: "\ inf \ nrms css (Normal s)" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f0: "f 0 = ([], (nrms, abrs) # css, Normal s)" with f_step [of 0] have "f (Suc 0) = (nrms,css,Normal s)" by (auto elim: step_Normal_elim_cases) with f_step have "inf \ nrms css (Normal s)" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) by simp with hyp show False .. qed next case (ExitBlockAbrupt abrs css s nrms) have hyp: "\ inf \ abrs css (Normal s)" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f0: "f 0 = ([], (nrms, abrs) # css, Abrupt s)" with f_step [of 0] have "f (Suc 0) = (abrs,css,Normal s)" by (auto elim: step_Normal_elim_cases) with f_step have "inf \ abrs css (Normal s)" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) by simp with hyp show False .. qed next case (ExitBlockFault nrms css f abrs) show ?case by (rule not_inf_Fault) next case (ExitBlockStuck nrms css abrs) show ?case by (rule not_inf_Stuck) next case (Cons c s cs css) have termi_c: "\\c \ s" by fact have hyp: "\t. \\\c,s\ \ t \ \\cs,css\t \ \ inf \ cs css t" by fact show "\ inf \ (c # cs) css s" proof assume "inf \ (c # cs) css s" thus False proof (rule infE) assume "inf \ [c] [] s" with terminates_impl_not_inf [OF termi_c] show False .. next fix t assume "\\\c,s\ \ t" "inf \ cs css t" with hyp show False by simp qed qed qed lemma lem: "\y. r\<^sup>+\<^sup>+ a y \ P a \ P y \ ((b,a) \ {(y,x). P x \ r x y}\<^sup>+) = ((b,a) \ {(y,x). P x \ r\<^sup>+\<^sup>+ x y})" apply(rule iffI) apply clarify apply(erule trancl_induct) apply blast apply(blast intro:tranclp_trans) apply clarify apply(erule tranclp_induct) apply blast apply(blast intro:trancl_trans) done corollary terminatess_impl_no_inf_chain: assumes terminatess: "\\cs,css\s" shows "\(\f. f 0 = (cs,css,s) \ (\i::nat. \\f i \\<^sup>+ f(Suc i)))" proof - have "wf({(y,x). \\(cs,css,s) \\<^sup>* x \ \\x \ y}\<^sup>+)" proof (rule wf_trancl) show "wf {(y, x). \\(cs, css, s) \\<^sup>* x \ \\x \ y}" proof (simp only: wf_iff_no_infinite_down_chain,clarify,simp) fix f assume "\i. \\(cs, css, s) \\<^sup>* f i \ \\f i \ f (Suc i)" hence "\f. f 0 = (cs, css, s) \ (\i. \\f i \ f (Suc i))" by (rule renumber [to_pred]) moreover from terminatess have "\ (\f. f 0 = (cs, css, s) \ (\i. \\f i \ f (Suc i)))" by (rule terminatess_impl_not_inf [unfolded inf_def]) ultimately show False by simp qed qed hence "\ (\f. \i. (f (Suc i), f i) \ {(y, x). \\(cs, css, s) \\<^sup>* x \ \\x \ y}\<^sup>+)" by (simp add: wf_iff_no_infinite_down_chain) thus ?thesis proof (rule contrapos_nn) assume "\f. f 0 = (cs, css, s) \ (\i. \\f i \\<^sup>+ f (Suc i))" then obtain f where f0: "f 0 = (cs, css, s)" and seq: "\i. \\f i \\<^sup>+ f (Suc i)" by iprover show "\f. \i. (f (Suc i), f i) \ {(y, x). \\(cs, css, s) \\<^sup>* x \ \\x \ y}\<^sup>+" proof (rule exI [where x=f],rule allI) fix i show "(f (Suc i), f i) \ {(y, x). \\(cs, css, s) \\<^sup>* x \ \\x \ y}\<^sup>+" proof - { fix i have "\\(cs,css,s) \\<^sup>* f i" proof (induct i) case 0 show "\\(cs, css, s) \\<^sup>* f 0" by (simp add: f0) next case (Suc n) have "\\(cs, css, s) \\<^sup>* f n" by fact with seq show "\\(cs, css, s) \\<^sup>* f (Suc n)" by (blast intro: tranclp_into_rtranclp rtranclp_trans) qed } hence "\\(cs,css,s) \\<^sup>* f i" by iprover with seq have "(f (Suc i), f i) \ {(y, x). \\(cs, css, s) \\<^sup>* x \ \\x \\<^sup>+ y}" by clarsimp moreover have "\y. \\f i \\<^sup>+ y\\\(cs, css, s) \\<^sup>* f i\\\(cs, css, s) \\<^sup>* y" by (blast intro: tranclp_into_rtranclp rtranclp_trans) ultimately show ?thesis by (subst lem) qed qed qed qed corollary terminates_impl_no_inf_chain: "\\c\s \ \(\f. f 0 = ([c],[],s) \ (\i::nat. \\f i \\<^sup>+ f(Suc i)))" by (rule terminatess_impl_no_inf_chain) (iprover intro: terminatess.intros) definition termi_call_steps :: "('s,'p,'f) body \ (('s \ 'p) \ ('s \ 'p))set" where "termi_call_steps \ = {((t,q),(s,p)). \\the (\ p)\Normal s \ (\css. \\([the (\ p)],[],Normal s) \\<^sup>+ ([the (\ q)],css,Normal t))}" text \Sequencing computations, or more exactly continuation stacks\ primrec seq:: "(nat \ 'a list) \ nat \ 'a list" where "seq css 0 = []" | "seq css (Suc i) = css i@seq css i" theorem wf_termi_call_steps: "wf (termi_call_steps \)" proof (simp only: termi_call_steps_def wf_iff_no_infinite_down_chain, clarify,simp) fix S assume inf: "\i. (\(t,q) (s,p). \\(the (\ p)) \ Normal s \ (\css. \\([the (\ p)],[],Normal s) \\<^sup>+ ([the (\ q)],css,Normal t))) (S (Suc i)) (S i)" obtain s p where "s = (\i. fst (S i))" and "p = (\i. snd (S i))" by auto with inf have inf': "\i. \\(the (\ (p i))) \ Normal (s i) \ (\css. \\([the (\ (p i))],[],Normal (s i)) \\<^sup>+ ([the (\ (p (Suc i)))],css,Normal (s (Suc i))))" apply - apply (rule allI) apply (erule_tac x=i in allE) apply auto done show False proof - from inf' \ \Skolemization of css with axiom of choice\ have "\css. \i. \\(the (\ (p i))) \ Normal (s i) \ \\([the (\ (p i))],[],Normal (s i)) \\<^sup>+ ([the (\ (p (Suc i)))],css i,Normal (s (Suc i)))" apply - apply (rule choice) by blast then obtain css where termi_css: "\i. \\(the (\ (p i))) \ Normal (s i)" and step_css: "\i. \\([the (\ (p i))],[],Normal (s i)) \\<^sup>+ ([the (\ (p (Suc i)))],css i,Normal (s (Suc i)))" by blast define f where "f i = ([the (\ (p i))], seq css i,Normal (s i)::('a,'c) xstate)" for i have "f 0 = ([the (\ (p 0))],[],Normal (s 0))" by (simp add: f_def) moreover have "\i. \\ (f i) \\<^sup>+ (f (i+1))" proof fix i from step_css [rule_format, of i] have "\\([the (\ (p i))], [], Normal (s i)) \\<^sup>+ ([the (\ (p (Suc i)))], css i, Normal (s (Suc i)))". from app_css_steps [OF this,simplified] have "\\([the (\ (p i))], seq css i, Normal (s i)) \\<^sup>+ ([the (\ (p (Suc i)))], css i@seq css i, Normal (s (Suc i)))". thus "\\ (f i) \\<^sup>+ (f (i+1))" by (simp add: f_def) qed moreover from termi_css [rule_format, of 0] have "\ (\f. (f 0 = ([the (\ (p 0))],[],Normal (s 0)) \ (\i. \\(f i) \\<^sup>+ f(Suc i))))" by (rule terminates_impl_no_inf_chain) ultimately show False by auto qed qed text \An alternative proof using Hilbert-choice instead of axiom of choice.\ theorem "wf (termi_call_steps \)" proof (simp only: termi_call_steps_def wf_iff_no_infinite_down_chain, clarify,simp) fix S assume inf: "\i. (\(t,q) (s,p). \\(the (\ p)) \ Normal s \ (\css. \\([the (\ p)],[],Normal s) \\<^sup>+ ([the (\ q)],css,Normal t))) (S (Suc i)) (S i)" obtain s p where "s = (\i. fst (S i))" and "p = (\i. snd (S i))" by auto with inf have inf': "\i. \\(the (\ (p i))) \ Normal (s i) \ (\css. \\([the (\ (p i))],[],Normal (s i)) \\<^sup>+ ([the (\ (p (Suc i)))],css,Normal (s (Suc i))))" apply - apply (rule allI) apply (erule_tac x=i in allE) apply auto done show "False" proof - define CSS where "CSS i = (SOME css. \\([the (\ (p i))],[], Normal (s i)) \\<^sup>+ ([the (\ (p (i+1)))],css,Normal (s (i+1))))" for i define f where "f i = ([the (\ (p i))], seq CSS i,Normal (s i)::('a,'c) xstate)" for i have "f 0 = ([the (\ (p 0))],[],Normal (s 0))" by (simp add: f_def) moreover have "\i. \\ (f i) \\<^sup>+ (f (i+1))" proof fix i from inf' [rule_format, of i] obtain css where css: "\\([the (\ (p i))],[],Normal (s i)) \\<^sup>+ ([the (\ (p (i+1)))],css,Normal (s (i+1)))" by fastforce hence "\\([the (\ (p i))], seq CSS i, Normal (s i)) \\<^sup>+ ([the (\ (p (i+1)))], CSS i @ seq CSS i, Normal (s (i+1)))" apply - apply (unfold CSS_def) apply (rule someI2 [where P="\css. \\([the (\ (p i))],[],Normal (s i))\\<^sup>+ ([the (\ (p (i+1)))],css, Normal (s (i+1)))"]) apply (rule css) apply (fastforce dest: app_css_steps) done thus "\\ (f i) \\<^sup>+ (f (i+1))" by (simp add: f_def) qed moreover from inf' [rule_format, of 0] have "\\the (\ (p 0)) \ Normal (s 0)" by iprover then have "\ (\f. (f 0 = ([the (\ (p 0))],[],Normal (s 0)) \ (\i. \\(f i) \\<^sup>+ f(Suc i))))" by (rule terminates_impl_no_inf_chain) ultimately show False by auto qed qed lemma not_inf_implies_wf: assumes not_inf: "\ inf \ cs css s" shows "wf {(c2,c1). \ \ (cs,css,s) \\<^sup>* c1 \ \ \ c1 \ c2}" proof (simp only: wf_iff_no_infinite_down_chain,clarify, simp) fix f assume "\i. \\(cs, css, s) \\<^sup>* f i \ \\f i \ f (Suc i)" hence "\f. f 0 = (cs, css, s) \ (\i. \\f i \ f (Suc i))" by (rule renumber [to_pred]) moreover from not_inf have "\ (\f. f 0 = (cs, css, s) \ (\i. \\f i \ f (Suc i)))" by (unfold inf_def) ultimately show False by simp qed lemma wf_implies_termi_reach: assumes wf: "wf {(c2,c1). \ \ (cs,css,s) \\<^sup>* c1 \ \ \ c1 \ c2}" shows "\cs1 css1 s1. \\ \ (cs,css,s) \\<^sup>* c1; c1=(cs1,css1,s1)\\ \\cs1,css1\s1" using wf proof (induct c1, simp) fix cs1 css1 s1 assume reach: "\\(cs, css, s) \\<^sup>* (cs1, css1, s1)" assume hyp_raw: "\y cs2 css2 s2. \\ \ (cs1,css1,s1) \ (cs2,css2,s2); \ \ (cs,css,s) \\<^sup>* (cs2,css2,s2); y=(cs2,css2,s2)\ \ \\cs2,css2\s2" have hyp: "\cs2 css2 s2. \\ \ (cs1,css1,s1) \ (cs2,css2,s2)\ \ \\cs2,css2\s2" apply - apply (rule hyp_raw) apply assumption using reach apply simp apply (rule refl) done show "\\cs1,css1\s1" proof (cases s1) case (Normal s1') show ?thesis proof (cases cs1) case Nil note cs1_Nil = this show ?thesis proof (cases css1) case Nil with cs1_Nil show ?thesis by (auto intro: terminatess.intros) next case (Cons nrms_abrs css1') then obtain nrms abrs where nrms_abrs: "nrms_abrs=(nrms,abrs)" by (cases "nrms_abrs") have "\ \ ([],(nrms,abrs)#css1',Normal s1') \ (nrms,css1',Normal s1')" by (rule step.intros) from hyp [simplified cs1_Nil Cons nrms_abrs Normal, OF this] have "\\nrms,css1'\Normal s1'". from ExitBlockNormal [OF this] cs1_Nil Cons nrms_abrs Normal show ?thesis by auto qed next case (Cons c1 cs1') have "\\c1#cs1',css1\Normal s1'" proof (cases c1) case Skip have "\ \ (Skip#cs1',css1,Normal s1') \ (cs1',css1,Normal s1')" by (rule step.intros) from hyp [simplified Cons Skip Normal, OF this] have "\\cs1',css1\Normal s1'". with Normal Skip show ?thesis by (auto intro: terminatess.intros terminates.intros elim: exec_Normal_elim_cases) next case (Basic f) have "\ \ (Basic f#cs1',css1,Normal s1') \ (cs1',css1,Normal (f s1'))" by (rule step.intros) from hyp [simplified Cons Basic Normal, OF this] have "\\cs1',css1\Normal (f s1')". with Normal Basic show ?thesis by (auto intro: terminatess.intros terminates.intros elim: exec_Normal_elim_cases) next case (Spec r) with Normal show ?thesis apply simp apply (rule terminatess.Cons) apply (fastforce intro: terminates.intros) apply (clarify) apply (erule exec_Normal_elim_cases) apply clarsimp apply (rule hyp) apply (fastforce intro: step.intros simp add: Cons Spec Normal ) apply (fastforce intro: terminatess_Stuck) done next case (Seq c\<^sub>1 c\<^sub>2) have "\ \ (Seq c\<^sub>1 c\<^sub>2#cs1',css1,Normal s1') \ (c\<^sub>1#c\<^sub>2#cs1',css1,Normal s1')" by (rule step.intros) from hyp [simplified Cons Seq Normal, OF this] have "\\c\<^sub>1 # c\<^sub>2 # cs1',css1\Normal s1'". with Normal Seq show ?thesis by (fastforce intro: terminatess.intros terminates.intros elim: terminatess_elim_cases exec_Normal_elim_cases) next case (Cond b c\<^sub>1 c\<^sub>2) show ?thesis proof (cases "s1' \ b") case True hence "\\(Cond b c\<^sub>1 c\<^sub>2#cs1',css1,Normal s1') \ (c\<^sub>1#cs1',css1,Normal s1')" by (rule step.intros) from hyp [simplified Cons Cond Normal, OF this] have "\\c\<^sub>1 # cs1',css1\Normal s1'". with Normal Cond True show ?thesis by (fastforce intro: terminatess.intros terminates.intros elim: terminatess_elim_cases exec_Normal_elim_cases) next case False hence "\\(Cond b c\<^sub>1 c\<^sub>2#cs1',css1,Normal s1') \ (c\<^sub>2#cs1',css1,Normal s1')" by (rule step.intros) from hyp [simplified Cons Cond Normal, OF this] have "\\c\<^sub>2 # cs1',css1\Normal s1'". with Normal Cond False show ?thesis by (fastforce intro: terminatess.intros terminates.intros elim: terminatess_elim_cases exec_Normal_elim_cases) qed next case (While b c') show ?thesis proof (cases "s1' \ b") case True then have "\\(While b c' # cs1', css1, Normal s1') \ (c' # While b c' # cs1', css1, Normal s1')" by (rule step.intros) from hyp [simplified Cons While Normal, OF this] have "\\c' # While b c' # cs1',css1\Normal s1'". with Cons While True Normal show ?thesis by (fastforce intro: terminatess.intros terminates.intros exec.intros elim: terminatess_elim_cases exec_Normal_elim_cases) next case False then have "\\(While b c' # cs1', css1, Normal s1') \ (cs1', css1, Normal s1')" by (rule step.intros) from hyp [simplified Cons While Normal, OF this] have "\\cs1',css1\Normal s1'". with Cons While False Normal show ?thesis by (fastforce intro: terminatess.intros terminates.intros exec.intros elim: terminatess_elim_cases exec_Normal_elim_cases) qed next case (Call p) show ?thesis proof (cases "\ p") case None with Call Normal show ?thesis by (fastforce intro: terminatess.intros terminates.intros terminatess_Stuck elim: exec_Normal_elim_cases) next case (Some bdy) then have "\ \ (Call p#cs1',css1,Normal s1') \ ([bdy],(cs1',Throw#cs1')#css1,Normal s1')" by (rule step.intros) from hyp [simplified Cons Call Normal Some, OF this] have "\\[bdy],(cs1', Throw # cs1') # css1\Normal s1'". with Some Call Normal show ?thesis apply simp apply (rule terminatess.intros) apply (blast elim: terminatess_elim_cases intro: terminates.intros) apply clarify apply (erule terminatess_elim_cases) apply (erule exec_Normal_elim_cases) prefer 2 apply simp apply (erule_tac x=t in allE) apply (case_tac t) apply (auto intro: terminatess_Stuck terminatess_Fault exec.intros elim: terminatess_elim_cases exec_Normal_elim_cases) done qed next case (DynCom c') have "\ \ (DynCom c'#cs1',css1,Normal s1') \ (c' s1'#cs1',css1,Normal s1')" by (rule step.intros) from hyp [simplified Cons DynCom Normal, OF this] have "\\c' s1'#cs1',css1\Normal s1'". with Normal DynCom show ?thesis by (fastforce intro: terminatess.intros terminates.intros exec.intros elim: terminatess_elim_cases exec_Normal_elim_cases) next case (Guard f g c') show ?thesis proof (cases "s1' \ g") case True then have "\ \ (Guard f g c'#cs1',css1,Normal s1') \ (c'#cs1',css1,Normal s1')" by (rule step.intros) from hyp [simplified Cons Guard Normal, OF this] have "\\c'#cs1',css1\Normal s1'". with Normal Guard True show ?thesis by (fastforce intro: terminatess.intros terminates.intros exec.intros elim: terminatess_elim_cases exec_Normal_elim_cases) next case False with Guard Normal show ?thesis by (fastforce intro: terminatess.intros terminatess_Fault terminates.intros elim: exec_Normal_elim_cases) qed next case Throw have "\ \ (Throw#cs1',css1,Normal s1') \ (cs1',css1,Abrupt s1')" by (rule step.intros) from hyp [simplified Cons Throw Normal, OF this] have "\\cs1',css1\Abrupt s1'". with Normal Throw show ?thesis by (auto intro: terminatess.intros terminates.intros elim: exec_Normal_elim_cases) next case (Catch c\<^sub>1 c\<^sub>2) have "\ \ (Catch c\<^sub>1 c\<^sub>2#cs1',css1,Normal s1') \ ([c\<^sub>1], (cs1',c\<^sub>2#cs1')# css1,Normal s1')" by (rule step.intros) from hyp [simplified Cons Catch Normal, OF this] have "\\[c\<^sub>1],(cs1', c\<^sub>2 # cs1') # css1\Normal s1'". with Normal Catch show ?thesis by (fastforce intro: terminatess.intros terminates.intros exec.intros elim: terminatess_elim_cases exec_Normal_elim_cases) qed with Cons Normal show ?thesis by simp qed next case (Abrupt s1') show ?thesis proof (cases cs1) case Nil note cs1_Nil = this show ?thesis proof (cases css1) case Nil with cs1_Nil show ?thesis by (auto intro: terminatess.intros) next case (Cons nrms_abrs css1') then obtain nrms abrs where nrms_abrs: "nrms_abrs=(nrms,abrs)" by (cases "nrms_abrs") have "\ \ ([],(nrms,abrs)#css1',Abrupt s1') \ (abrs,css1',Normal s1')" by (rule step.intros) from hyp [simplified cs1_Nil Cons nrms_abrs Abrupt, OF this] have "\\abrs,css1'\Normal s1'". from ExitBlockAbrupt [OF this] cs1_Nil Cons nrms_abrs Abrupt show ?thesis by auto qed next case (Cons c1 cs1') have "\\c1#cs1',css1\Abrupt s1'" proof - have "\ \ (c1#cs1',css1,Abrupt s1') \ (cs1',css1,Abrupt s1')" by (rule step.intros) from hyp [simplified Cons Abrupt, OF this] have "\\cs1',css1\Abrupt s1'". with Cons Abrupt show ?thesis by (fastforce intro: terminatess.intros terminates.intros exec.intros elim: terminatess_elim_cases exec_Normal_elim_cases) qed with Cons Abrupt show ?thesis by simp qed next case (Fault f) thus ?thesis by (auto intro: terminatess_Fault) next case Stuck thus ?thesis by (auto intro: terminatess_Stuck) qed qed lemma not_inf_impl_terminatess: assumes not_inf: "\ inf \ cs css s" shows "\\cs,css\s" proof - from not_inf_implies_wf [OF not_inf] have wf: "wf {(c2, c1). \\(cs, css, s) \\<^sup>* c1 \ \\c1 \ c2}". show ?thesis by (rule wf_implies_termi_reach [OF wf]) auto qed lemma not_inf_impl_terminates: assumes not_inf: "\ inf \ [c] [] s" shows "\\c\s" proof - from not_inf_impl_terminatess [OF not_inf] have "\\[c],[]\s". thus ?thesis by (auto elim: terminatess_elim_cases) qed theorem terminatess_iff_not_inf: "\\cs,css\s = (\ inf \ cs css s)" apply rule apply (erule terminatess_impl_not_inf) apply (erule not_inf_impl_terminatess) done corollary terminates_iff_not_inf: "\\c\s = (\ inf \ [c] [] s)" apply (rule) apply (erule terminates_impl_not_inf) apply (erule not_inf_impl_terminates) done subsection \Completeness of Total Correctness Hoare Logic\ lemma ConseqMGT: assumes modif: "\Z::'a. \,\ \\<^sub>t\<^bsub>/F\<^esub> (P' Z::'a assn) c (Q' Z),(A' Z)" assumes impl: "\s. s \ P \ s \ P' s \ (\t. t \ Q' s \ t \ Q) \ (\t. t \ A' s \ t \ A)" shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P c Q,A" using impl by - (rule conseq [OF modif],blast) lemma conseq_extract_state_indep_prop: assumes state_indep_prop:"\s \ P. R" assumes to_show: "R \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" apply (rule Conseq) apply (clarify) apply (rule_tac x="P" in exI) apply (rule_tac x="Q" in exI) apply (rule_tac x="A" in exI) using state_indep_prop to_show by blast lemma Call_lemma': assumes Call_hyp: "\q\dom \. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub>{s. s=Z \ \\\Call q,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call q\Normal s \ ((s,q),(\,p)) \ termi_call_steps \} (Call q) {t. \\\Call q,Normal Z\ \ Normal t}, {t. \\\Call q,Normal Z\ \ Abrupt t}" shows "\Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c#cs,css,Normal s))} c {t. \\\c,Normal Z\ \ Normal t}, {t. \\\c,Normal Z\ \ Abrupt t}" proof (induct c) case Skip show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s = Z \ \\\Skip,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p) \ Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (Skip # cs,css,Normal s))} Skip {t. \\\Skip,Normal Z\ \ Normal t}, {t. \\\Skip,Normal Z\ \ Abrupt t}" by (rule hoaret.Skip [THEN conseqPre]) (blast intro: exec.Skip) next case (Basic f) show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Basic f,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (Basic f#cs,css,Normal s))} Basic f {t. \\\Basic f,Normal Z\ \ Normal t}, {t. \\\Basic f,Normal Z\ \ Abrupt t}" by (rule hoaret.Basic [THEN conseqPre]) (blast intro: exec.Basic) next case (Spec r) show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Spec r,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (Spec r#cs,css,Normal s))} Spec r {t. \\\Spec r,Normal Z\ \ Normal t}, {t. \\\Spec r,Normal Z\ \ Abrupt t}" apply (rule hoaret.Spec [THEN conseqPre]) apply (clarsimp) apply (case_tac "\t. (Z,t) \ r") apply (auto elim: exec_elim_cases simp add: final_notin_def intro: exec.intros) done next case (Seq c1 c2) have hyp_c1: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c1,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c1#cs,css,Normal s))} c1 {t. \\\c1,Normal Z\ \ Normal t}, {t. \\\c1,Normal Z\ \ Abrupt t}" using Seq.hyps by iprover have hyp_c2: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c2#cs,css,Normal s))} c2 {t. \\\c2,Normal Z\ \ Normal t}, {t. \\\c2,Normal Z\ \ Abrupt t}" using Seq.hyps by iprover have c1: "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Seq c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (Seq c1 c2#cs,css,Normal s))} c1 {t. \\\c1,Normal Z\ \ Normal t \ \\\c2,Normal t\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c2#cs,css,Normal t))}, {t. \\\Seq c1 c2,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [OF hyp_c1],clarify,safe) assume "\\\Seq c1 c2,Normal Z\ \\({Stuck} \ Fault ` (-F))" thus "\\\c1,Normal Z\ \\({Stuck} \ Fault ` (-F))" by (blast dest: Seq_NoFaultStuckD1) next fix cs css assume "\\([the (\ p)],[],Normal \) \\<^sup>* (Seq c1 c2 # cs,css,Normal Z)" thus "\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c1 # cs,css, Normal Z)" by (blast intro: rtranclp_into_tranclp1 [THEN tranclp_into_rtranclp] step.Seq) next fix t assume "\\\Seq c1 c2,Normal Z\ \\({Stuck} \ Fault ` (-F))" "\\\c1,Normal Z\ \ Normal t" thus "\\\c2,Normal t\ \\({Stuck} \ Fault ` (-F))" by (blast dest: Seq_NoFaultStuckD2) next fix cs css t assume "\\([the (\ p)],[],Normal \) \\<^sup>* (Seq c1 c2#cs,css,Normal Z)" also have "\\(Seq c1 c2 # cs,css,Normal Z) \ (c1#c2#cs,css,Normal Z)" by (rule step.Seq) also assume "\\\c1,Normal Z\ \ Normal t" hence "\\ (c1#c2#cs,css,Normal Z) \\<^sup>* (c2#cs,css,Normal t)" by (rule exec_impl_steps) finally show "\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c2 # cs,css, Normal t)" by iprover next fix t assume "\\\c1,Normal Z\ \ Abrupt t" thus "\\\Seq c1 c2,Normal Z\ \ Abrupt t" by (blast intro: exec.intros) qed show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Seq c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (Seq c1 c2#cs,css,Normal s))} Seq c1 c2 {t. \\\Seq c1 c2,Normal Z\ \ Normal t}, {t. \\\Seq c1 c2,Normal Z\ \ Abrupt t}" by (rule hoaret.Seq [OF c1 ConseqMGT [OF hyp_c2]]) (blast intro: exec.intros) next case (Cond b c1 c2) have hyp_c1: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c1,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[], Normal \) \\<^sup>* (c1#cs,css,Normal s))} c1 {t. \\\c1,Normal Z\ \ Normal t}, {t. \\\c1,Normal Z\ \ Abrupt t}" using Cond.hyps by iprover have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. s=Z \ \\\Cond b c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (Cond b c1 c2#cs,css,Normal s))} \ b) c1 {t. \\\Cond b c1 c2,Normal Z\ \ Normal t}, {t. \\\Cond b c1 c2,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [OF hyp_c1],safe) assume "Z \ b" "\\\Cond b c1 c2,Normal Z\ \\({Stuck} \ Fault ` (-F))" thus "\\\c1,Normal Z\ \\({Stuck} \ Fault ` (-F))" by (auto simp add: final_notin_def intro: exec.CondTrue) next fix cs css assume "Z \ b" "\\([the (\ p)],[], Normal \) \\<^sup>* (Cond b c1 c2 # cs,css, Normal Z)" thus "\cs css. \\([the (\ p)],[], Normal \) \\<^sup>* (c1 # cs,css, Normal Z)" by (blast intro: rtranclp_into_tranclp1 [THEN tranclp_into_rtranclp] step.CondTrue) next fix t assume "Z \ b" "\\\c1,Normal Z\ \ Normal t" thus "\\\Cond b c1 c2,Normal Z\ \ Normal t" by (blast intro: exec.CondTrue) next fix t assume "Z \ b" "\\\c1,Normal Z\ \ Abrupt t" thus "\\\Cond b c1 c2,Normal Z\ \ Abrupt t" by (blast intro: exec.CondTrue) qed moreover have hyp_c2: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c2#cs,css,Normal s))} c2 {t. \\\c2,Normal Z\ \ Normal t}, {t. \\\c2,Normal Z\ \ Abrupt t}" using Cond.hyps by iprover have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. s=Z \ \\\Cond b c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>*(Cond b c1 c2#cs,css, Normal s))} \ -b) c2 {t. \\\Cond b c1 c2,Normal Z\ \ Normal t}, {t. \\\Cond b c1 c2,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [OF hyp_c2],safe) assume "Z \ b" "\\\Cond b c1 c2,Normal Z\ \\({Stuck} \ Fault ` (-F))" thus "\\\c2,Normal Z\ \\({Stuck} \ Fault ` (-F))" by (auto simp add: final_notin_def intro: exec.CondFalse) next fix cs css assume "Z \ b" "\\([the (\ p)],[], Normal \) \\<^sup>* (Cond b c1 c2 # cs,css, Normal Z)" thus "\cs css. \\([the (\ p)],[], Normal \) \\<^sup>* (c2 # cs,css,Normal Z)" by (blast intro: rtranclp_into_tranclp1 [THEN tranclp_into_rtranclp] step.CondFalse) next fix t assume "Z \ b" "\\\c2,Normal Z\ \ Normal t" thus "\\\Cond b c1 c2,Normal Z\ \ Normal t" by (blast intro: exec.CondFalse) next fix t assume "Z \ b" "\\\c2,Normal Z\ \ Abrupt t" thus "\\\Cond b c1 c2,Normal Z\ \ Abrupt t" by (blast intro: exec.CondFalse) qed ultimately show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Cond b c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (Cond b c1 c2#cs,css,Normal s))} (Cond b c1 c2) {t. \\\Cond b c1 c2,Normal Z\ \ Normal t}, {t. \\\Cond b c1 c2,Normal Z\ \ Abrupt t}" by (rule hoaret.Cond) next case (While b c) let ?unroll = "({(s,t). s\b \ \\\c,Normal s\ \ Normal t})\<^sup>*" let ?P' = "\Z. {t. (Z,t)\?unroll \ (\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (While b c#cs,css,Normal t))}" let ?A = "\Z. {t. \\\While b c,Normal Z\ \ Abrupt t}" let ?r = "{(t,s). \\(While b c)\Normal s \ s\b \ \\\c,Normal s\ \ Normal t}" show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\While b c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \)\\<^sup>*(While b c#cs,css,Normal s))} (While b c) {t. \\\While b c,Normal Z\ \ Normal t}, {t. \\\While b c,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [where ?P'="\ Z. ?P' Z" and ?Q'="\ Z. ?P' Z \ - b"]) have wf_r: "wf ?r" by (rule wf_terminates_while) show "\ Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (?P' Z) (While b c) (?P' Z \ - b),(?A Z)" proof (rule allI, rule hoaret.While [OF wf_r]) fix Z from While have hyp_c: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c # cs,css,Normal s))} c {t. \\\c,Normal Z\ \ Normal t}, {t. \\\c,Normal Z\ \ Abrupt t}" by iprover show "\\. \,\\\<^sub>t\<^bsub>/F\<^esub> ({\} \ ?P' Z \ b) c ({t. (t, \) \ ?r} \ ?P' Z),(?A Z)" proof (rule allI, rule ConseqMGT [OF hyp_c]) fix \ s assume asm: "s\ {\} \ {t. (Z, t) \ ?unroll \ (\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)) \ \\the (\ p)\ Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (While b c#cs,css,Normal t))} \ b" then obtain cs css where s_eq_\: "s=\" and Z_s_unroll: "(Z,s) \ ?unroll" and noabort:"\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)" and termi: "\\the (\ p) \ Normal \" and reach: "\\([the (\ p)],[],Normal \) \\<^sup>* (While b c#cs,css,Normal s)"and s_in_b: "s\b" by blast have reach_c: "\\([the (\ p)],[],Normal \) \\<^sup>* (c#While b c#cs,css,Normal s)" proof - note reach also from s_in_b have "\\(While b c#cs,css,Normal s)\ (c#While b c#cs,css,Normal s)" by (rule step.WhileTrue) finally show ?thesis . qed from reach termi have termi_while: "\\While b c \ Normal s" by (rule steps_preserves_termination) show "s \ {t. t = s \ \\\c,Normal t\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p) \ Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c#cs,css,Normal t))} \ (\t. t \ {t. \\\c,Normal s\ \ Normal t} \ t \ {t. (t,\) \ ?r} \ {t. (Z, t) \ ?unroll \ (\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)) \ \\the (\ p) \ Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (While b c # cs,css,Normal t))}) \ (\t. t \ {t. \\\c,Normal s\ \ Abrupt t} \ t \ {t. \\\While b c,Normal Z\ \ Abrupt t})" (is "?C1 \ ?C2 \ ?C3") proof (intro conjI) from Z_s_unroll noabort s_in_b termi reach_c show ?C1 by blast next { fix t assume s_t: "\\\c,Normal s\ \ Normal t" with s_eq_\ termi_while s_in_b have "(t,\) \ ?r" by blast moreover from Z_s_unroll s_t s_in_b have "(Z, t) \ ?unroll" by (blast intro: rtrancl_into_rtrancl) moreover have "\\([the (\ p)],[],Normal \) \\<^sup>* (While b c#cs,css,Normal t)" proof - note reach_c also from s_t have "\\(c#While b c#cs,css,Normal s)\\<^sup>* (While b c#cs,css, Normal t)" by (rule exec_impl_steps) finally show ?thesis . qed moreover note noabort termi ultimately have "(t,\) \ ?r \ (Z, t) \ ?unroll \ (\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)) \ \\the (\ p) \ Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (While b c # cs,css, Normal t))" by iprover } then show ?C2 by blast next { fix t assume s_t: "\\\c,Normal s\ \ Abrupt t" from Z_s_unroll noabort s_t s_in_b have "\\\While b c,Normal Z\ \ Abrupt t" by blast } thus ?C3 by simp qed qed qed next fix s assume P: "s \ {s. s=Z \ \\\While b c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (While b c#cs,css,Normal s))}" hence WhileNoFault: "\\\While b c,Normal Z\ \\({Stuck} \ Fault ` (-F))" by auto show "s \ ?P' s \ (\t. t\(?P' s \ - b)\ t\{t. \\\While b c,Normal Z\ \ Normal t})\ (\t. t\?A s \ t\?A Z)" proof (intro conjI) { fix e assume "(Z,e) \ ?unroll" "e \ b" from this WhileNoFault have "\\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)" (is "?Prop Z e") proof (induct rule: converse_rtrancl_induct [consumes 1]) assume e_in_b: "e \ b" assume WhileNoFault: "\\\While b c,Normal e\ \\({Stuck} \ Fault ` (-F))" with e_in_b WhileNoFault have cNoFault: "\\\c,Normal e\ \\({Stuck} \ Fault ` (-F))" by (auto simp add: final_notin_def intro: exec.intros) moreover { fix u assume "\\\c,Normal e\ \ Abrupt u" with e_in_b have "\\\While b c,Normal e\ \ Abrupt u" by (blast intro: exec.intros) } ultimately show "?Prop e e" by iprover next fix Z r assume e_in_b: "e\b" assume WhileNoFault: "\\\While b c,Normal Z\ \\({Stuck} \ Fault ` (-F))" assume hyp: "\e\b;\\\While b c,Normal r\ \\({Stuck} \ Fault ` (-F))\ \ ?Prop r e" assume Z_r: "(Z, r) \ {(Z, r). Z \ b \ \\\c,Normal Z\ \ Normal r}" with WhileNoFault have "\\\While b c,Normal r\ \\({Stuck} \ Fault ` (-F))" by (auto simp add: final_notin_def intro: exec.intros) from hyp [OF e_in_b this] obtain cNoFault: "\\\c,Normal e\ \\({Stuck} \ Fault ` (-F))" and Abrupt_r: "\u. \\\c,Normal e\ \ Abrupt u \ \\\While b c,Normal r\ \ Abrupt u" by simp { fix u assume "\\\c,Normal e\ \ Abrupt u" with Abrupt_r have "\\\While b c,Normal r\ \ Abrupt u" by simp moreover from Z_r obtain "Z \ b" "\\\c,Normal Z\ \ Normal r" by simp ultimately have "\\\While b c,Normal Z\ \ Abrupt u" by (blast intro: exec.intros) } with cNoFault show "?Prop Z e" by iprover qed } with P show "s \ ?P' s" by blast next { fix t assume "termination": "t \ b" assume "(Z, t) \ ?unroll" hence "\\\While b c,Normal Z\ \ Normal t" proof (induct rule: converse_rtrancl_induct [consumes 1]) from "termination" show "\\\While b c,Normal t\ \ Normal t" by (blast intro: exec.WhileFalse) next fix Z r assume first_body: "(Z, r) \ {(s, t). s \ b \ \\\c,Normal s\ \ Normal t}" assume "(r, t) \ ?unroll" assume rest_loop: "\\\While b c, Normal r\ \ Normal t" show "\\\While b c,Normal Z\ \ Normal t" proof - from first_body obtain "Z \ b" "\\\c,Normal Z\ \ Normal r" by fast moreover from rest_loop have "\\\While b c,Normal r\ \ Normal t" by fast ultimately show "\\\While b c,Normal Z\ \ Normal t" by (rule exec.WhileTrue) qed qed } with P show "\t. t\(?P' s \ - b) \t\{t. \\\While b c,Normal Z\ \ Normal t}" by blast next from P show "\t. t\?A s \ t\?A Z" by simp qed qed next case (Call q) let ?P = "{s. s=Z \ \\\Call q ,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (Call q # cs,css,Normal s))}" from noStuck_Call have "\s \ ?P. q \ dom \" by (fastforce simp add: final_notin_def) then show ?case proof (rule conseq_extract_state_indep_prop) assume q_defined: "q \ dom \" from Call_hyp have "\q\dom \. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub>{s. s=Z \ \\\the (\ q),Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\(the (\ q))\Normal s \ ((s,q),(\,p)) \ termi_call_steps \} (Call q) {t. \\\the (\ q),Normal Z\ \ Normal t}, {t. \\\the (\ q),Normal Z\ \ Abrupt t}" by (simp add: exec_Call_body' noFaultStuck_Call_body' [simplified] terminates_Normal_Call_body) from Call_hyp q_defined have Call_hyp': "\Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Call q,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call q\Normal s \ ((s,q),(\,p)) \ termi_call_steps \} (Call q) {t. \\\Call q,Normal Z\ \ Normal t}, {t. \\\Call q,Normal Z\ \ Abrupt t}" by auto show "\,\\\<^sub>t\<^bsub>/F\<^esub> ?P (Call q) {t. \\\Call q ,Normal Z\ \ Normal t}, {t. \\\Call q ,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [OF Call_hyp'],safe) fix cs css assume "\\([the (\ p)],[],Normal \)\\<^sup>* (Call q # cs,css,Normal Z)" "\\the (\ p) \ Normal \" hence "\\Call q \ Normal Z" by (rule steps_preserves_termination) with q_defined show "\\Call q \ Normal Z" by (auto elim: terminates_Normal_elim_cases) next fix cs css assume reach: "\\([the (\ p)],[],Normal \) \\<^sup>* (Call q#cs,css,Normal Z)" moreover have "\\(Call q # cs,css, Normal Z) \ ([the (\ q)],(cs,Throw#cs)#css, Normal Z)" by (rule step.Call) (insert q_defined,auto) ultimately have "\\([the (\ p)],[],Normal \) \\<^sup>+ ([the (\ q)],(cs,Throw#cs)#css,Normal Z)" by (rule rtranclp_into_tranclp1) moreover assume termi: "\\the (\ p) \ Normal \" ultimately show "((Z,q), \,p) \ termi_call_steps \" by (auto simp add: termi_call_steps_def) qed qed next case (DynCom c) have hyp: "\s'. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s = Z \ \\\c s',Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p) \ Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c s'#cs,css,Normal s))} (c s') {t. \\\c s',Normal Z\ \ Normal t},{t. \\\c s',Normal Z\ \ Abrupt t}" using DynCom by simp have hyp': "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\DynCom c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p) \ Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (DynCom c#cs,css,Normal s))} (c Z) {t. \\\DynCom c,Normal Z\ \ Normal t},{t. \\\DynCom c,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [OF hyp],safe) assume "\\\DynCom c,Normal Z\ \\({Stuck} \ Fault ` (-F))" then show "\\\c Z,Normal Z\ \\({Stuck} \ Fault ` (-F))" by (fastforce simp add: final_notin_def intro: exec.intros) next fix cs css assume "\\([the (\ p)], [], Normal \) \\<^sup>* (DynCom c # cs, css, Normal Z)" also have "\\(DynCom c # cs, css, Normal Z) \ (c Z#cs,css,Normal Z)" by (rule step.DynCom) finally show "\cs css. \\([the (\ p)], [], Normal \) \\<^sup>* (c Z # cs, css, Normal Z)" by blast next fix t assume "\\\c Z,Normal Z\ \ Normal t" thus "\\\DynCom c,Normal Z\ \ Normal t" by (auto intro: exec.intros) next fix t assume "\\\c Z,Normal Z\ \ Abrupt t" thus "\\\DynCom c,Normal Z\ \ Abrupt t" by (auto intro: exec.intros) qed show ?case apply (rule hoaret.DynCom) apply safe apply (rule hyp') done next case (Guard f g c) have hyp_c: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c#cs,css,Normal s))} c {t. \\\c,Normal Z\ \ Normal t}, {t. \\\c,Normal Z\ \ Abrupt t}" using Guard.hyps by iprover show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Guard f g c ,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (Guard f g c#cs,css,Normal s))} Guard f g c {t. \\\Guard f g c ,Normal Z\ \ Normal t}, {t. \\\Guard f g c ,Normal Z\ \ Abrupt t}" proof (cases "f \ F") case True have "\,\\\<^sub>t\<^bsub>/F\<^esub> (g \ {s. s=Z \ \\\Guard f g c ,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (Guard f g c#cs,css,Normal s))}) c {t. \\\Guard f g c ,Normal Z\ \ Normal t}, {t. \\\Guard f g c ,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [OF hyp_c], safe) assume "\\\Guard f g c ,Normal Z\ \\({Stuck} \ Fault ` (-F))" "Z\g" thus "\\\c,Normal Z\ \\({Stuck} \ Fault ` (-F))" by (auto simp add: final_notin_def intro: exec.intros) next fix cs css assume "\\([the (\ p)],[],Normal \) \\<^sup>* (Guard f g c#cs,css,Normal Z)" also assume "Z \ g" hence "\\(Guard f g c#cs,css,Normal Z) \ (c#cs,css,Normal Z)" by (rule step.Guard) finally show "\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c#cs,css,Normal Z)" by iprover next fix t assume "\\\Guard f g c ,Normal Z\ \\({Stuck} \ Fault ` (-F))" "\\\c,Normal Z\ \ Normal t" "Z \ g" thus "\\\Guard f g c ,Normal Z\ \ Normal t" by (auto simp add: final_notin_def intro: exec.intros ) next fix t assume "\\\Guard f g c ,Normal Z\ \\({Stuck} \ Fault ` (-F))" "\\\c,Normal Z\ \ Abrupt t" "Z \ g" thus "\\\Guard f g c ,Normal Z\ \ Abrupt t" by (auto simp add: final_notin_def intro: exec.intros ) qed from True this show ?thesis by (rule conseqPre [OF Guarantee]) auto next case False have "\,\\\<^sub>t\<^bsub>/F\<^esub> (g \ {s. s=Z \ \\\Guard f g c ,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (Guard f g c#cs,css,Normal s))}) c {t. \\\Guard f g c ,Normal Z\ \ Normal t}, {t. \\\Guard f g c ,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [OF hyp_c], safe) assume "\\\Guard f g c ,Normal Z\ \\({Stuck} \ Fault ` (-F))" thus "\\\c,Normal Z\ \\({Stuck} \ Fault ` (-F))" using False by (cases "Z\ g") (auto simp add: final_notin_def intro: exec.intros) next fix cs css assume "\\([the (\ p)],[],Normal \) \\<^sup>* (Guard f g c#cs,css,Normal Z)" also assume "\\\Guard f g c ,Normal Z\ \\({Stuck} \ Fault ` (-F))" hence "Z \ g" using False by (auto simp add: final_notin_def intro: exec.GuardFault) hence "\\(Guard f g c#cs,css,Normal Z) \ (c#cs,css,Normal Z)" by (rule step.Guard) finally show "\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c#cs,css,Normal Z)" by iprover next fix t assume "\\\Guard f g c ,Normal Z\ \\({Stuck} \ Fault ` (-F))" "\\\c,Normal Z\ \ Normal t" thus "\\\Guard f g c ,Normal Z\ \ Normal t" using False by (cases "Z\ g") (auto simp add: final_notin_def intro: exec.intros ) next fix t assume "\\\Guard f g c ,Normal Z\ \\({Stuck} \ Fault ` (-F))" "\\\c,Normal Z\ \ Abrupt t" thus "\\\Guard f g c ,Normal Z\ \ Abrupt t" using False by (cases "Z\ g") (auto simp add: final_notin_def intro: exec.intros ) qed then show ?thesis apply (rule conseqPre [OF hoaret.Guard]) apply clarify apply (frule Guard_noFaultStuckD [OF _ False]) apply auto done qed next case Throw show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Throw,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal \ \ (\cs css. \\([the (\ p)],[], Normal \) \\<^sup>* (Throw#cs,css,Normal s))} Throw {t. \\\Throw,Normal Z\ \ Normal t}, {t. \\\Throw,Normal Z\ \ Abrupt t}" by (rule conseqPre [OF hoaret.Throw]) (blast intro: exec.intros terminates.intros) next case (Catch c\<^sub>1 c\<^sub>2) have hyp_c1: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s= Z \ \\\c\<^sub>1,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p) \ Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c\<^sub>1# cs, css,Normal s))} c\<^sub>1 {t. \\\c\<^sub>1,Normal Z\ \ Normal t},{t. \\\c\<^sub>1,Normal Z\ \ Abrupt t}" using Catch.hyps by iprover have hyp_c2: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s= Z \ \\\c\<^sub>2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p) \ Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c\<^sub>2# cs, css,Normal s))} c\<^sub>2 {t. \\\c\<^sub>2,Normal Z\ \ Normal t},{t. \\\c\<^sub>2,Normal Z\ \ Abrupt t}" using Catch.hyps by iprover have "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s = Z \ \\\Catch c\<^sub>1 c\<^sub>2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p) \ Normal \ \ (\cs css. \\([the (\ p)],[],Normal \)\\<^sup>*(Catch c\<^sub>1 c\<^sub>2 #cs,css,Normal s))} c\<^sub>1 {t. \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Normal t}, {t. \\\c\<^sub>1,Normal Z\ \ Abrupt t \ \\\c\<^sub>2,Normal t\ \\({Stuck} \ Fault`(-F)) \ \\the (\ p) \ Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c\<^sub>2# cs, css,Normal t))}" proof (rule ConseqMGT [OF hyp_c1],clarify,safe) assume "\\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \\({Stuck} \ Fault ` (-F))" thus "\\\c\<^sub>1,Normal Z\ \\({Stuck} \ Fault ` (-F))" by (fastforce simp add: final_notin_def intro: exec.intros) next fix cs css assume "\\([the (\ p)], [], Normal \) \\<^sup>* (Catch c\<^sub>1 c\<^sub>2 # cs, css, Normal Z)" also have "\\(Catch c\<^sub>1 c\<^sub>2 # cs, css, Normal Z) \ ([c\<^sub>1],(cs,c\<^sub>2#cs)#css,Normal Z)" by (rule step.Catch) finally show "\cs css. \\([the (\ p)], [], Normal \) \\<^sup>* (c\<^sub>1 # cs, css, Normal Z)" by iprover next fix t assume "\\\c\<^sub>1,Normal Z\ \ Normal t" thus "\\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Normal t" by (auto intro: exec.intros) next fix t assume "\\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \\({Stuck} \ Fault ` (-F))" "\\\c\<^sub>1,Normal Z\ \ Abrupt t" thus "\\\c\<^sub>2,Normal t\ \\({Stuck} \ Fault ` (-F))" by (auto simp add: final_notin_def intro: exec.intros) next fix cs css t assume "\\([the (\ p)], [], Normal \) \\<^sup>* (Catch c\<^sub>1 c\<^sub>2 # cs, css, Normal Z)" also have "\\(Catch c\<^sub>1 c\<^sub>2 # cs, css, Normal Z) \ ([c\<^sub>1],(cs,c\<^sub>2#cs)#css,Normal Z)" by (rule step.Catch) also assume "\\\c\<^sub>1,Normal Z\ \ Abrupt t" hence "\\([c\<^sub>1],(cs,c\<^sub>2#cs)#css,Normal Z) \\<^sup>* ([],(cs,c\<^sub>2#cs)#css,Abrupt t)" by (rule exec_impl_steps) also have "\\([],(cs,c\<^sub>2#cs)#css,Abrupt t) \ (c\<^sub>2#cs,css,Normal t)" by (rule step.intros) finally show "\cs css. \\([the (\ p)], [], Normal \) \\<^sup>* (c\<^sub>2 # cs, css, Normal t)" by iprover qed moreover have "\,\\\<^sub>t\<^bsub>/F\<^esub> {t. \\\c\<^sub>1,Normal Z\ \ Abrupt t \ \\\c\<^sub>2,Normal t\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p) \ Normal \ \ (\cs css. \\([the (\ p)],[],Normal \) \\<^sup>* (c\<^sub>2# cs, css,Normal t))} c\<^sub>2 {t. \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Normal t}, {t. \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Abrupt t}" by (rule ConseqMGT [OF hyp_c2]) (fastforce intro: exec.intros) ultimately show ?case by (rule hoaret.Catch) qed text \To prove a procedure implementation correct it suffices to assume only the procedure specifications of procedures that actually occur during evaluation of the body. \ lemma Call_lemma: assumes Call: "\q \ dom \. \Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Call q,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call q\Normal s \ ((s,q),(\,p)) \ termi_call_steps \} (Call q) {t. \\\Call q,Normal Z\ \ Normal t}, {t. \\\Call q,Normal Z\ \ Abrupt t}" shows "\Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> ({\} \ {s. s=Z \ \\\the (\ p),Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal s}) the (\ p) {t. \\\the (\ p),Normal Z\ \ Normal t}, {t. \\\the (\ p),Normal Z\ \ Abrupt t}" apply (rule conseqPre) apply (rule Call_lemma') apply (rule Call) apply blast done lemma Call_lemma_switch_Call_body: assumes call: "\q \ dom \. \Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Call q,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call q\Normal s \ ((s,q),(\,p)) \ termi_call_steps \} (Call q) {t. \\\Call q,Normal Z\ \ Normal t}, {t. \\\Call q,Normal Z\ \ Abrupt t}" assumes p_defined: "p \ dom \" shows "\Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> ({\} \ {s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal s}) the (\ p) {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t}" apply (simp only: exec_Call_body' [OF p_defined] noFaultStuck_Call_body' [OF p_defined] terminates_Normal_Call_body [OF p_defined]) apply (rule conseqPre) apply (rule Call_lemma') apply (rule call) apply blast done lemma MGT_Call: "\p \ dom \. \Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\(Call p)\Normal s} (Call p) {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t}" apply (intro ballI allI) apply (rule CallRec' [where Procs="dom \" and P="\p Z. {s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal s}" and Q="\p Z. {t. \\\Call p,Normal Z\ \ Normal t}" and A="\p Z. {t. \\\Call p,Normal Z\ \ Abrupt t}" and r="termi_call_steps \" ]) apply simp apply simp apply (rule wf_termi_call_steps) apply (intro ballI allI) apply simp apply (rule Call_lemma_switch_Call_body [rule_format, simplified]) apply (rule hoaret.Asm) apply fastforce apply assumption done lemma CollInt_iff: "{s. P s} \ {s. Q s} = {s. P s \ Q s}" by auto lemma image_Un_conv: "f ` (\p\dom \. \Z. {x p Z}) = (\p\dom \. \Z. {f (x p Z)})" by (auto iff: not_None_eq) text \Another proof of \MGT_Call\, maybe a little more readable\ lemma "\p \ dom \. \Z. \,{} \\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\(Call p)\Normal s} (Call p) {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t}" proof - { fix p Z \ assume defined: "p \ dom \" define Specs where "Specs = (\p\dom \. \Z. {({s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal s}, p, {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t})})" define Specs_wf where "Specs_wf p \ = (\(P,q,Q,A). (P \ {s. ((s,q),\,p) \ termi_call_steps \}, q, Q, A)) ` Specs" for p \ have "\,Specs_wf p \ \\<^sub>t\<^bsub>/F\<^esub>({\} \ {s. s = Z \ \\\the (\ p),Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal s}) (the (\ p)) {t. \\\the (\ p),Normal Z\ \ Normal t}, {t. \\\the (\ p),Normal Z\ \ Abrupt t}" apply (rule Call_lemma [rule_format]) apply (rule hoaret.Asm) apply (clarsimp simp add: Specs_wf_def Specs_def image_Un_conv) apply (rule_tac x=q in bexI) apply (rule_tac x=Z in exI) apply (clarsimp simp add: CollInt_iff) apply auto done hence "\,Specs_wf p \ \\<^sub>t\<^bsub>/F\<^esub>({\} \ {s. s = Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal s}) (the (\ p)) {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t}" by (simp only: exec_Call_body' [OF defined] noFaultStuck_Call_body' [OF defined] terminates_Normal_Call_body [OF defined]) } note bdy=this show ?thesis apply (intro ballI allI) apply (rule hoaret.CallRec [where Specs="(\p\dom \. \Z. {({s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal s}, p, {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t})})", OF _ wf_termi_call_steps [of \] refl]) apply fastforce apply clarify apply (rule conjI) apply fastforce apply (rule allI) apply (simp (no_asm_use) only : Un_empty_left) apply (rule bdy) apply auto done qed end diff --git a/thys/Simpl/COPYRIGHT b/thys/Simpl/COPYRIGHT deleted file mode 100644 --- a/thys/Simpl/COPYRIGHT +++ /dev/null @@ -1,423 +0,0 @@ -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - - - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - - - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA - - -TERMS AND CONDITIONS FOR COPYING, DISTRIBUTION AND MODIFICATION - - - -0. This License Agreement applies to any software library or other -program which contains a notice placed by the copyright holder or -other authorized party saying it may be distributed under the terms of -this Lesser General Public License (also called "this License"). Each -licensee is addressed as "you". - - - -A "library" means a collection of software functions and/or data -prepared so as to be conveniently linked with application programs -(which use some of those functions and data) to form executables. - - -The "Library", below, refers to any such software library or work -which has been distributed under these terms. A "work based on the -Library" means either the Library or any derivative work under -copyright law: that is to say, a work containing the Library or a -portion of it, either verbatim or with modifications and/or translated -straightforwardly into another language. (Hereinafter, translation is -included without limitation in the term "modification".) - - -"Source code" for a work means the preferred form of the work for -making modifications to it. For a library, complete source code means -all the source code for all modules it contains, plus any associated -interface definition files, plus the scripts used to control -compilation and installation of the library. - - - -Activities other than copying, distribution and modification are not -covered by this License; they are outside its scope. The act of -running a program using the Library is not restricted, and output from -such a program is covered only if its contents constitute a work based -on the Library (independent of the use of the Library in a tool for -writing it). Whether that is true depends on what the Library does and -what the program that uses the Library does. - - -1. You may copy and distribute verbatim copies of the Library's -complete source code as you receive it, in any medium, provided that -you conspicuously and appropriately publish on each copy an -appropriate copyright notice and disclaimer of warranty; keep intact -all the notices that refer to this License and to the absence of any -warranty; and distribute a copy of this License along with the -Library. - - -You may charge a fee for the physical act of transferring a copy, and -you may at your option offer warranty protection in exchange for a -fee. - - -2. You may modify your copy or copies of the Library or any portion of -it, thus forming a work based on the Library, and copy and distribute -such modifications or work under the terms of Section 1 above, -provided that you also meet all of these conditions: - - -a) The modified work must itself be a software library. - - -b) You must cause the files modified to carry prominent notices -stating that you changed the files and the date of any change. - - - -c) You must cause the whole of the work to be licensed at no charge to -all third parties under the terms of this License. - - -d) If a facility in the modified Library refers to a function or a -table of data to be supplied by an application program that uses the -facility, other than as an argument passed when the facility is -invoked, then you must make a good faith effort to ensure that, in the -event an application does not supply such function or table, the -facility still operates, and performs whatever part of its purpose -remains meaningful. - - -(For example, a function in a library to compute square roots has a -purpose that is entirely well-defined independent of the -application. Therefore, Subsection 2d requires that any -application-supplied function or table used by this function must be -optional: if the application does not supply it, the square root -function must still compute square roots.) - - -These requirements apply to the modified work as a whole. If -identifiable sections of that work are not derived from the Library, -and can be reasonably considered independent and separate works in -themselves, then this License, and its terms, do not apply to those -sections when you distribute them as separate works. But when you -distribute the same sections as part of a whole which is a work based -on the Library, the distribution of the whole must be on the terms of -this License, whose permissions for other licensees extend to the -entire whole, and thus to each and every part regardless of who wrote -it. - - -Thus, it is not the intent of this section to claim rights or contest -your rights to work written entirely by you; rather, the intent is to -exercise the right to control the distribution of derivative or -collective works based on the Library. - - -In addition, mere aggregation of another work not based on the Library -with the Library (or with a work based on the Library) on a volume of -a storage or distribution medium does not bring the other work under -the scope of this License. - - -3. You may opt to apply the terms of the ordinary GNU General Public -License instead of this License to a given copy of the Library. To do -this, you must alter all the notices that refer to this License, so -that they refer to the ordinary GNU General Public License, version 2, -instead of to this License. (If a newer version than version 2 of the -ordinary GNU General Public License has appeared, then you can specify -that version instead if you wish.) Do not make any other change in -these notices. - - -Once this change is made in a given copy, it is irreversible for that -copy, so the ordinary GNU General Public License applies to all -subsequent copies and derivative works made from that copy. - - -This option is useful when you wish to copy part of the code of the -Library into a program that is not a library. - - -4. You may copy and distribute the Library (or a portion or derivative -of it, under Section 2) in object code or executable form under the -terms of Sections 1 and 2 above provided that you accompany it with -the complete corresponding machine-readable source code, which must be -distributed under the terms of Sections 1 and 2 above on a medium -customarily used for software interchange. - - -If distribution of object code is made by offering access to copy from -a designated place, then offering equivalent access to copy the source -code from the same place satisfies the requirement to distribute the -source code, even though third parties are not compelled to copy the -source along with the object code. - - - -5. A program that contains no derivative of any portion of the -Library, but is designed to work with the Library by being compiled or -linked with it, is called a "work that uses the Library". Such a work, -in isolation, is not a derivative work of the Library, and therefore -falls outside the scope of this License. - - -However, linking a "work that uses the Library" with the Library -creates an executable that is a derivative of the Library (because it -contains portions of the Library), rather than a "work that uses the -library". The executable is therefore covered by this License. Section -6 states terms for distribution of such executables. - - -When a "work that uses the Library" uses material from a header file -that is part of the Library, the object code for the work may be a -derivative work of the Library even though the source code is -not. Whether this is true is especially significant if the work can be -linked without the Library, or if the work is itself a library. The -threshold for this to be true is not precisely defined by law. - - -If such an object file uses only numerical parameters, data structure -layouts and accessors, and small macros and small inline functions -(ten lines or less in length), then the use of the object file is -unrestricted, regardless of whether it is legally a derivative -work. (Executables containing this object code plus portions of the -Library will still fall under Section 6.) - - -Otherwise, if the work is a derivative of the Library, you may -distribute the object code for the work under the terms of Section -6. Any executables containing that work also fall under Section 6, -whether or not they are linked directly with the Library itself. - - -6. As an exception to the Sections above, you may also combine or link -a "work that uses the Library" with the Library to produce a work -containing portions of the Library, and distribute that work under -terms of your choice, provided that the terms permit modification of -the work for the customer's own use and reverse engineering for -debugging such modifications. - - -You must give prominent notice with each copy of the work that the -Library is used in it and that the Library and its use are covered by -this License. You must supply a copy of this License. If the work -during execution displays copyright notices, you must include the -copyright notice for the Library among them, as well as a reference -directing the user to the copy of this License. Also, you must do one -of these things: - - -a) Accompany the work with the complete corresponding machine-readable -source code for the Library including whatever changes were used in -the work (which must be distributed under Sections 1 and 2 above); -and, if the work is an executable linked with the Library, with the -complete machine-readable "work that uses the Library", as object code -and/or source code, so that the user can modify the Library and then -relink to produce a modified executable containing the modified -Library. (It is understood that the user who changes the contents of -definitions files in the Library will not necessarily be able to -recompile the application to use the modified definitions.) - - -b) Use a suitable shared library mechanism for linking with the -Library. A suitable mechanism is one that (1) uses at run time a copy -of the library already present on the user's computer system, rather -than copying library functions into the executable, and (2) will -operate properly with a modified version of the library, if the user -installs one, as long as the modified version is interface-compatible -with the version that the work was made with. - - -c) Accompany the work with a written offer, valid for at least three -years, to give the same user the materials specified in Subsection 6a, -above, for a charge no more than the cost of performing this -distribution. - - -d) If distribution of the work is made by offering access to copy from -a designated place, offer equivalent access to copy the above -specified materials from the same place. - - -e) Verify that the user has already received a copy of these materials -or that you have already sent this user a copy. - - - -For an executable, the required form of the "work that uses the -Library" must include any data and utility programs needed for -reproducing the executable from it. However, as a special exception, -the materials to be distributed need not include anything that is -normally distributed (in either source or binary form) with the major -components (compiler, kernel, and so on) of the operating system on -which the executable runs, unless that component itself accompanies -the executable. - - -It may happen that this requirement contradicts the license -restrictions of other proprietary libraries that do not normally -accompany the operating system. Such a contradiction means you cannot -use both them and the Library together in an executable that you -distribute. - - -7. You may place library facilities that are a work based on the -Library side-by-side in a single library together with other library -facilities not covered by this License, and distribute such a combined -library, provided that the separate distribution of the work based on -the Library and of the other library facilities is otherwise -permitted, and provided that you do these two things: - - -a) Accompany the combined library with a copy of the same work based -on the Library, uncombined with any other library facilities. This -must be distributed under the terms of the Sections above. - - -b) Give prominent notice with the combined library of the fact that -part of it is a work based on the Library, and explaining where to -find the accompanying uncombined form of the same work. - - - -8. You may not copy, modify, sublicense, link with, or distribute the -Library except as expressly provided under this License. Any attempt -otherwise to copy, modify, sublicense, link with, or distribute the -Library is void, and will automatically terminate your rights under -this License. However, parties who have received copies, or rights, -from you under this License will not have their licenses terminated so -long as such parties remain in full compliance. - - -9. You are not required to accept this License, since you have not -signed it. However, nothing else grants you permission to modify or -distribute the Library or its derivative works. These actions are -prohibited by law if you do not accept this License. Therefore, by -modifying or distributing the Library (or any work based on the -Library), you indicate your acceptance of this License to do so, and -all its terms and conditions for copying, distributing or modifying -the Library or works based on it. - - -10. Each time you redistribute the Library (or any work based on the -Library), the recipient automatically receives a license from the -original licensor to copy, distribute, link with or modify the Library -subject to these terms and conditions. You may not impose any further -restrictions on the recipients' exercise of the rights granted -herein. You are not responsible for enforcing compliance by third -parties with this License. - - -11. If, as a consequence of a court judgment or allegation of patent -infringement or for any other reason (not limited to patent issues), -conditions are imposed on you (whether by court order, agreement or -otherwise) that contradict the conditions of this License, they do not -excuse you from the conditions of this License. If you cannot -distribute so as to satisfy simultaneously your obligations under this -License and any other pertinent obligations, then as a consequence you -may not distribute the Library at all. For example, if a patent -license would not permit royalty-free redistribution of the Library by -all those who receive copies directly or indirectly through you, then -the only way you could satisfy both it and this License would be to -refrain entirely from distribution of the Library. - - -If any portion of this section is held invalid or unenforceable under -any particular circumstance, the balance of the section is intended to -apply, and the section as a whole is intended to apply in other -circumstances. - - -It is not the purpose of this section to induce you to infringe any -patents or other property right claims or to contest validity of any -such claims; this section has the sole purpose of protecting the -integrity of the free software distribution system which is -implemented by public license practices. Many people have made -generous contributions to the wide range of software distributed -through that system in reliance on consistent application of that -system; it is up to the author/donor to decide if he or she is willing -to distribute software through any other system and a licensee cannot -impose that choice. - - -This section is intended to make thoroughly clear what is believed to -be a consequence of the rest of this License. - - -12. If the distribution and/or use of the Library is restricted in -certain countries either by patents or by copyrighted interfaces, the -original copyright holder who places the Library under this License -may add an explicit geographical distribution limitation excluding -those countries, so that distribution is permitted only in or among -countries not thus excluded. In such case, this License incorporates -the limitation as if written in the body of this License. - - -13. The Free Software Foundation may publish revised and/or new -versions of the Lesser General Public License from time to time. Such -new versions will be similar in spirit to the present version, but may -differ in detail to address new problems or concerns. - - - -Each version is given a distinguishing version number. If the Library -specifies a version number of this License which applies to it and -"any later version", you have the option of following the terms and -conditions either of that version or of any later version published by -the Free Software Foundation. If the Library does not specify a -license version number, you may choose any version ever published by -the Free Software Foundation. - - -14. If you wish to incorporate parts of the Library into other free -programs whose distribution conditions are incompatible with these, -write to the author to ask for permission. For software which is -copyrighted by the Free Software Foundation, write to the Free -Software Foundation; we sometimes make exceptions for this. Our -decision will be guided by the two goals of preserving the free status -of all derivatives of our free software and of promoting the sharing -and reuse of software generally. - - -NO WARRANTY - - -15. BECAUSE THE LIBRARY IS LICENSED FREE OF CHARGE, THERE IS NO -WARRANTY FOR THE LIBRARY, TO THE EXTENT PERMITTED BY APPLICABLE -LAW. EXCEPT WHEN OTHERWISE STATED IN WRITING THE COPYRIGHT HOLDERS -AND/OR OTHER PARTIES PROVIDE THE LIBRARY "AS IS" WITHOUT WARRANTY OF -ANY KIND, EITHER EXPRESSED OR IMPLIED, INCLUDING, BUT NOT LIMITED TO, -THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR A PARTICULAR -PURPOSE. THE ENTIRE RISK AS TO THE QUALITY AND PERFORMANCE OF THE -LIBRARY IS WITH YOU. SHOULD THE LIBRARY PROVE DEFECTIVE, YOU ASSUME -THE COST OF ALL NECESSARY SERVICING, REPAIR OR CORRECTION. - - -16. IN NO EVENT UNLESS REQUIRED BY APPLICABLE LAW OR AGREED TO IN -WRITING WILL ANY COPYRIGHT HOLDER, OR ANY OTHER PARTY WHO MAY MODIFY -AND/OR REDISTRIBUTE THE LIBRARY AS PERMITTED ABOVE, BE LIABLE TO YOU -FOR DAMAGES, INCLUDING ANY GENERAL, SPECIAL, INCIDENTAL OR -CONSEQUENTIAL DAMAGES ARISING OUT OF THE USE OR INABILITY TO USE THE -LIBRARY (INCLUDING BUT NOT LIMITED TO LOSS OF DATA OR DATA BEING -RENDERED INACCURATE OR LOSSES SUSTAINED BY YOU OR THIRD PARTIES OR A -FAILURE OF THE LIBRARY TO OPERATE WITH ANY OTHER SOFTWARE), EVEN IF -SUCH HOLDER OR OTHER PARTY HAS BEEN ADVISED OF THE POSSIBILITY OF SUCH -DAMAGES. - - -END OF TERMS AND CONDITIONS \ No newline at end of file diff --git a/thys/Simpl/DPC0Expressions.thy b/thys/Simpl/DPC0Expressions.thy --- a/thys/Simpl/DPC0Expressions.thy +++ b/thys/Simpl/DPC0Expressions.thy @@ -1,23 +1,24 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL + +Copyright (C) 2006-2008 Norbert Schirmer *) section \SHORTENED! Parallel expressions in DPC/Hoare.\ theory DPC0Expressions imports Main begin definition p_not :: "bool list => bool list" ("\\<^sub>p") where "p_not = map Not" definition elem_wise :: "('a \ 'b \ 'c) \ 'a list \ 'b list \ 'c list" where "elem_wise f xs ys = map (\ (x, y). f x y) (zip xs ys)" definition p_and :: "bool list => bool list => bool list" (infixl "\\<^sub>p" 35) where "p_and = elem_wise (&)" notation (ASCII) p_and (infixl "pand" 35) end diff --git a/thys/Simpl/DPC0Library.thy b/thys/Simpl/DPC0Library.thy --- a/thys/Simpl/DPC0Library.thy +++ b/thys/Simpl/DPC0Library.thy @@ -1,104 +1,105 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL + +Copyright (C) 2006-2008 Norbert Schirmer *) section \DPC0 standard library\ theory DPC0Library imports DPC0Expressions Vcg begin text \This theory constitutes a standard library for DPC0 programs. At first, it includes (indirectly) the C0 library and (directly) its extensions for DPC0 specific expressions. Secondly, the Vcg of the verification environment is included and its syntax extended by the DPC0 specific statement constructs for contextualization. \ (* =================================================== *) section \Auxiliary functions for the concrete syntax\ (* =================================================== *) primrec pfilter:: "bool list \ 'a list \ 'a list" where "pfilter c [] = []" | "pfilter c (v#vs) = (if hd c then v#pfilter (tl c) vs else pfilter (tl c) vs)" primrec pmask:: "nat \ bool list \ nat list" where "pmask i [] = []" | "pmask i (b#bs) = (if b then i#(pmask (Suc i) bs) else pmask (Suc i) bs)" (* ============================================= *) section \Concrete syntax for Contextualization\ (* ============================================= *) syntax "_In":: "[ident,'a,'a] \ ('s,'p,'f) com" ("(2 IN \_:/ _ :== _)" [1000,30,30] 21) "_Where":: "['a,ident,('s,'p,'f) com] \ ('s,'p,'f) com" ("(0 WHERE (_)/ FOR \_ DO/ _ EREHW)" [0,0,0] 71) "_WhereElse":: "['a,ident,('s,'p,'f) com,('s,'p,'f) com] \ ('s,'p,'f) com" ("(0 WHERE (_)/ FOR \_ DO/ _ ELSE _ EREHW)" [0,0,0,0] 71) translations "_In c (x!!i) y" => "x!!(pfilter \c i) :== pfilter \c y" "_In c x y" => "x!!(pmask 0 \c) :== pfilter \c y" "_Where m c s" => "_Loc (_locinit c (p_and \c m)) s" "_WhereElse m c s1 s2" => "(_Loc (_locinit c (p_and \c m)) s1);; (_Loc (_locinit c (p_and \c (p_not m))) s2)" print_translation \ let fun in_tr' [Const (@{const_syntax list_multsel}, _) $ x $ (Const (@{const_syntax pfilter}, _) $ (Const (@{syntax_const "_antiquoteCur"}, _) $ c) $ i), Const (@{const_syntax pfilter}, _) $ (Const (@{syntax_const "_antiquoteCur"}, _) $ c') $ y] = if c = c' then Syntax.const @{syntax_const "_In"} $ c $ (Syntax.const @{const_syntax list_multsel} $ x $ i) $ y else raise Match | in_tr' [Const (@{const_syntax list_multsel}, _) $ x $ (Const (@{const_syntax pmask}, _) $ z $ (Const (@{syntax_const "_antiquoteCur"}, _) $ c)), Const (@{const_syntax pfilter}, _) $ (Const (@{syntax_const "_antiquoteCur"}, _) $ c') $ y] = if c = c' then Syntax.const @{syntax_const "_In"} $ c $ x $ y else raise Match fun where_tr' [Const (@{syntax_const "_locinit"}, _) $ Const (c, _) $ (Const (@{const_syntax p_and}, _) $ (Const (@{syntax_const "_antiquoteCur"}, _) $ Const (c', _)) $ m), s] = if c = c' then Syntax.const @{syntax_const "_Where"} $ m $ Syntax.const c $ s else raise Match | where_tr' ts = raise Match in [(@{syntax_const "_Assign"}, K in_tr'), (@{syntax_const "_Loc"}, K where_tr')] end \ print_ast_translation \ let fun where_else_tr' [Appl [Constant @{syntax_const "_Where"}, m, c, s1], Appl [Constant @{syntax_const "_Where"}, Appl [Constant @{const_syntax p_not}, m'], c', s2]] = if c = c' andalso m = m' then Appl [Constant @{syntax_const "_WhereElse"}, m, c, s1, s2] else raise Match in [(@{syntax_const "_seq"}, K where_else_tr')] end \ end diff --git a/thys/Simpl/Generalise.thy b/thys/Simpl/Generalise.thy --- a/thys/Simpl/Generalise.thy +++ b/thys/Simpl/Generalise.thy @@ -1,154 +1,133 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: Generalise.thy - Author: Norbert Schirmer, TU Muenchen - -Copyright (C) 2005-2008 Norbert Schirmer -Some rights reserved, TU Muenchen -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (C) 2006-2008 Norbert Schirmer *) theory Generalise imports "HOL-Statespace.DistinctTreeProver" begin lemma protectRefl: "PROP Pure.prop (PROP C) \ PROP Pure.prop (PROP C)" by (simp add: prop_def) lemma protectImp: assumes i: "PROP Pure.prop (PROP P \ PROP Q)" shows "PROP Pure.prop (PROP Pure.prop P \ PROP Pure.prop Q)" proof - { assume P: "PROP Pure.prop P" from i [unfolded prop_def, OF P [unfolded prop_def]] have "PROP Pure.prop Q" by (simp add: prop_def) } note i' = this show "PROP ?thesis" apply (rule protectI) apply (rule i') apply assumption done qed lemma generaliseConj: assumes i1: "PROP Pure.prop (PROP Pure.prop (Trueprop P) \ PROP Pure.prop (Trueprop Q))" assumes i2: "PROP Pure.prop (PROP Pure.prop (Trueprop P') \ PROP Pure.prop (Trueprop Q'))" shows "PROP Pure.prop (PROP Pure.prop (Trueprop (P \ P')) \ (PROP Pure.prop (Trueprop (Q \ Q'))))" using i1 i2 by (auto simp add: prop_def) lemma generaliseAll: assumes i: "PROP Pure.prop (\s. PROP Pure.prop (Trueprop (P s)) \ PROP Pure.prop (Trueprop (Q s)))" shows "PROP Pure.prop (PROP Pure.prop (Trueprop (\s. P s)) \ PROP Pure.prop (Trueprop (\s. Q s)))" using i by (auto simp add: prop_def) lemma generalise_all: assumes i: "PROP Pure.prop (\s. PROP Pure.prop (PROP P s) \ PROP Pure.prop (PROP Q s))" shows "PROP Pure.prop ((PROP Pure.prop (\s. PROP P s)) \ (PROP Pure.prop (\s. PROP Q s)))" using i proof (unfold prop_def) assume i1: "\s. (PROP P s) \ (PROP Q s)" assume i2: "\s. PROP P s" show "\s. PROP Q s" by (rule i1) (rule i2) qed lemma generaliseTrans: assumes i1: "PROP Pure.prop (PROP P \ PROP Q)" assumes i2: "PROP Pure.prop (PROP Q \ PROP R)" shows "PROP Pure.prop (PROP P \ PROP R)" using i1 i2 proof (unfold prop_def) assume P_Q: "PROP P \ PROP Q" assume Q_R: "PROP Q \ PROP R" assume P: "PROP P" show "PROP R" by (rule Q_R [OF P_Q [OF P]]) qed lemma meta_spec: assumes "\x. PROP P x" shows "PROP P x" by fact lemma meta_spec_protect: assumes g: "\x. PROP P x" shows "PROP Pure.prop (PROP P x)" using g by (auto simp add: prop_def) lemma generaliseImp: assumes i: "PROP Pure.prop (PROP Pure.prop (Trueprop P) \ PROP Pure.prop (Trueprop Q))" shows "PROP Pure.prop (PROP Pure.prop (Trueprop (X \ P)) \ PROP Pure.prop (Trueprop (X \ Q)))" using i by (auto simp add: prop_def) lemma generaliseEx: assumes i: "PROP Pure.prop (\s. PROP Pure.prop (Trueprop (P s)) \ PROP Pure.prop (Trueprop (Q s)))" shows "PROP Pure.prop (PROP Pure.prop (Trueprop (\s. P s)) \ PROP Pure.prop (Trueprop (\s. Q s)))" using i by (auto simp add: prop_def) lemma generaliseRefl: "PROP Pure.prop (PROP Pure.prop (Trueprop P) \ PROP Pure.prop (Trueprop P))" by (auto simp add: prop_def) lemma generaliseRefl': "PROP Pure.prop (PROP P \ PROP P)" by (auto simp add: prop_def) lemma generaliseAllShift: assumes i: "PROP Pure.prop (\s. P \ Q s)" shows "PROP Pure.prop (PROP Pure.prop (Trueprop P) \ PROP Pure.prop (Trueprop (\s. Q s)))" using i by (auto simp add: prop_def) lemma generalise_allShift: assumes i: "PROP Pure.prop (\s. PROP P \ PROP Q s)" shows "PROP Pure.prop (PROP Pure.prop (PROP P) \ PROP Pure.prop (\s. PROP Q s))" using i proof (unfold prop_def) assume P_Q: "\s. PROP P \ PROP Q s" assume P: "PROP P" show "\s. PROP Q s" by (rule P_Q [OF P]) qed lemma generaliseImpl: assumes i: "PROP Pure.prop (PROP Pure.prop P \ PROP Pure.prop Q)" shows "PROP Pure.prop ((PROP Pure.prop (PROP X \ PROP P)) \ (PROP Pure.prop (PROP X \ PROP Q)))" using i proof (unfold prop_def) assume i1: "PROP P \ PROP Q" assume i2: "PROP X \ PROP P" assume X: "PROP X" show "PROP Q" by (rule i1 [OF i2 [OF X]]) qed ML_file \generalise_state.ML\ end diff --git a/thys/Simpl/HeapList.thy b/thys/Simpl/HeapList.thy --- a/thys/Simpl/HeapList.thy +++ b/thys/Simpl/HeapList.thy @@ -1,297 +1,276 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: HeapList.thy - Author: Norbert Schirmer, TU Muenchen - -Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (C) 2006-2008 Norbert Schirmer *) section \Paths and Lists in the Heap\ theory HeapList imports Simpl_Heap begin text \Adapted from 'HOL/Hoare/Heap.thy'.\ subsection "Paths in The Heap" primrec Path :: "ref \ (ref \ ref) \ ref \ ref list \ bool" where "Path x h y [] = (x = y)" | "Path x h y (p#ps) = (x = p \ x \ Null \ Path (h x) h y ps)" lemma Path_Null_iff [iff]: "Path Null h y xs = (xs = [] \ y = Null)" apply(case_tac xs) apply fastforce apply fastforce done lemma Path_not_Null_iff [simp]: "p\Null \ Path p h q as = (as = [] \ q = p \ (\ps. as = p#ps \ Path (h p) h q ps ))" apply(case_tac as) apply fastforce apply fastforce done lemma Path_append [simp]: "\p. Path p f q (as@bs) = (\y. Path p f y as \ Path y f q bs)" by(induct as, simp+) lemma notin_Path_update[simp]: "\p. u \ set ps \ Path p (f(u := v)) q ps = Path p f q ps " by(induct ps, simp, simp add:eq_sym_conv) lemma Path_upd_same [simp]: "Path p (f(p:=p)) q qs = ((p=Null \ q=Null \ qs = []) \ (p\Null \ q=p \ (\x\set qs. x=p)))" by (induct qs) auto text \@{thm[source] Path_upd_same} prevents @{term "p\Null \ Path p (f(p:=p)) q qs = X"} from looping, because of @{thm[source] Path_not_Null_iff} and @{thm[source]fun_upd_apply}. \ lemma notin_Path_updateI [intro]: "\Path p h q ps ; r \ set ps\ \ Path p (h(r := y)) q ps " by simp lemma Path_update_new [simp]: "\set ps \ set alloc\ \ Path p (f(new (set alloc) := x)) q ps = Path p f q ps " by (rule notin_Path_update) fastforce lemma Null_notin_Path [simp,intro]: "\p. Path p f q ps \ Null \ set ps" by(induct ps) auto lemma Path_snoc: "\Path p (f(a := q)) a as ; a\Null\ \ Path p (f(a := q)) q (as @ [a])" by simp subsection "Lists on The Heap" subsubsection "Relational Abstraction" definition List :: "ref \ (ref \ ref) \ ref list \ bool" where "List p h ps = Path p h Null ps " lemma List_empty [simp]: "List p h [] = (p = Null)" by(simp add:List_def) lemma List_cons [simp]: "List p h (a#ps) = (p = a \ p\Null \ List (h p) h ps)" by(simp add:List_def) lemma List_Null [simp]: "List Null h ps = (ps = [])" by(case_tac ps, simp_all) lemma List_not_Null [simp]: "p\Null \ List p h as = (\ps. as = p#ps \ List (h p) h ps)" by(case_tac as, simp_all, fast) lemma Null_notin_List [simp,intro]: "\p. List p h ps \ Null \ set ps" by (simp add : List_def) theorem notin_List_update[simp]: "\p. q \ set ps \ List p (h(q := y)) ps = List p h ps" apply(induct ps) apply simp apply clarsimp done lemma List_upd_same_lemma: "\p. p \ Null \ \ List p (h(p := p)) ps" apply (induct ps) apply simp apply (simp (no_asm_simp) del: fun_upd_apply) apply (simp (no_asm_simp) only: fun_upd_apply refl if_True) apply blast done lemma List_upd_same [simp]: "List p (h(p:=p)) ps = (p = Null \ ps = [])" apply (cases "p=Null") apply simp apply (fast dest: List_upd_same_lemma) done text \@{thm[source] List_upd_same} prevents @{term "p\Null \ List p (h(p:=p)) as = X"} from looping, because of @{thm[source] List_not_Null} and @{thm[source] fun_upd_apply}. \ lemma List_update_new [simp]: "\set ps \ set alloc\ \ List p (h(new (set alloc) := x)) ps = List p h ps" by (rule notin_List_update) fastforce lemma List_updateI [intro]: "\List p h ps; q \ set ps\ \ List p (h(q := y)) ps" by simp lemma List_unique: "\p bs. List p h as \ List p h bs \ as = bs" by(induct as, simp, clarsimp) lemma List_unique1: "List p h as \ \!as. List p h as" by(blast intro:List_unique) lemma List_app: "\p. List p h (as@bs) = (\y. Path p h y as \ List y h bs)" by(induct as, simp, clarsimp) lemma List_hd_not_in_tl[simp]: "List (h p) h ps \ p \ set ps" apply (clarsimp simp add:in_set_conv_decomp) apply(frule List_app[THEN iffD1]) apply(fastforce dest: List_unique) done lemma List_distinct[simp]: "\p. List p h ps \ distinct ps" apply(induct ps, simp) apply(fastforce dest:List_hd_not_in_tl) done lemma heap_eq_List_eq: "\p. \x \ set ps. h x = g x \ List p h ps = List p g ps" by (induct ps) auto lemma heap_eq_ListI: assumes list: "List p h ps" assumes hp_eq: "\x \ set ps. h x = g x" shows "List p g ps" using list by (simp add: heap_eq_List_eq [OF hp_eq]) lemma heap_eq_ListI1: assumes list: "List p h ps" assumes hp_eq: "\x \ set ps. g x = h x" shows "List p g ps" using list by (simp add: heap_eq_List_eq [OF hp_eq]) text \The following lemmata are usefull for the simplifier to instantiate bound variables in the assumptions resp. conclusion, using the uniqueness of the List predicate\ lemma conj_impl_simp: "(P \ Q \ K) = (P \ Q \ K)" by auto lemma List_unique_all_impl_simp [simp]: "List p h ps \ (\ps. List p h ps \ P ps) = P ps" by (auto dest: List_unique) (* lemma List_unique_all_impl_simp1 [simp]: "List p h ps \ (\ps. Q ps \ List p h ps \ P ps) = Q ps \ P ps" by (auto dest: List_unique) *) lemma List_unique_ex_conj_simp [simp]: "List p h ps \ (\ps. List p h ps \ P ps) = P ps" by (auto dest: List_unique) subsection "Functional abstraction" definition islist :: "ref \ (ref \ ref) \ bool" where "islist p h = (\ps. List p h ps)" definition list :: "ref \ (ref \ ref) \ ref list" where "list p h = (THE ps. List p h ps)" lemma List_conv_islist_list: "List p h ps = (islist p h \ ps = list p h)" apply(simp add:islist_def list_def) apply(rule iffI) apply(rule conjI) apply blast apply(subst the1_equality) apply(erule List_unique1) apply assumption apply(rule refl) apply simp apply (clarify) apply (rule theI) apply assumption by (rule List_unique) lemma List_islist [intro]: "List p h ps \ islist p h" apply (simp add: List_conv_islist_list) done lemma List_list: "List p h ps \ list p h = ps" apply (simp only: List_conv_islist_list) done lemma [simp]: "islist Null h" by(simp add:islist_def) lemma [simp]: "p\Null \ islist (h p) h = islist p h" by(simp add:islist_def) lemma [simp]: "list Null h = []" by(simp add:list_def) lemma list_Ref_conv[simp]: "\islist (h p) h; p\Null \ \ list p h = p # list (h p) h" apply(insert List_not_Null[of _ h]) apply(fastforce simp:List_conv_islist_list) done lemma [simp]: "islist (h p) h \ p \ set(list (h p) h)" apply(insert List_hd_not_in_tl[of h]) apply(simp add:List_conv_islist_list) done lemma list_upd_conv[simp]: "islist p h \ y \ set(list p h) \ list p (h(y := q)) = list p h" apply(drule notin_List_update[of _ _ p h q]) apply(simp add:List_conv_islist_list) done lemma islist_upd[simp]: "islist p h \ y \ set(list p h) \ islist p (h(y := q))" apply(frule notin_List_update[of _ _ p h q]) apply(simp add:List_conv_islist_list) done lemma list_distinct[simp]: "islist p h \ distinct (list p h)" apply (clarsimp simp add: list_def islist_def) apply (frule List_unique1) apply (drule (1) the1_equality) apply simp done lemma Null_notin_list [simp,intro]: "islist p h \ Null \ set (list p h)" apply (clarsimp simp add: list_def islist_def) apply (frule List_unique1) apply (drule (1) the1_equality) apply simp done end diff --git a/thys/Simpl/Hoare.thy b/thys/Simpl/Hoare.thy --- a/thys/Simpl/Hoare.thy +++ b/thys/Simpl/Hoare.thy @@ -1,424 +1,403 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: Hoare.thy - Author: Norbert Schirmer, TU Muenchen - -Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (C) 2006-2008 Norbert Schirmer *) section \Auxiliary Definitions/Lemmas to Facilitate Hoare Logic\ theory Hoare imports HoarePartial HoareTotal begin syntax "_hoarep_emptyFaults":: "[('s,'p,'f) body,('s,'p) quadruple set, 'f set,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_,_/\ (_/ (_)/ _,/_))" [61,60,1000,20,1000,1000]60) "_hoarep_emptyCtx":: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_/\\<^bsub>'/_\<^esub> (_/ (_)/ _,/_))" [61,60,1000,20,1000,1000]60) "_hoarep_emptyCtx_emptyFaults":: "[('s,'p,'f) body,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_/\ (_/ (_)/ _,/_))" [61,1000,20,1000,1000]60) "_hoarep_noAbr":: "[('s,'p,'f) body,('s,'p) quadruple set,'f set, 's assn,('s,'p,'f) com, 's assn] => bool" ("(3_,_/\\<^bsub>'/_\<^esub> (_/ (_)/ _))" [61,60,60,1000,20,1000]60) "_hoarep_noAbr_emptyFaults":: "[('s,'p,'f) body,('s,'p) quadruple set,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_,_/\ (_/ (_)/ _))" [61,60,1000,20,1000]60) "_hoarep_emptyCtx_noAbr":: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_/\\<^bsub>'/_\<^esub> (_/ (_)/ _))" [61,60,1000,20,1000]60) "_hoarep_emptyCtx_noAbr_emptyFaults":: "[('s,'p,'f) body,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_/\ (_/ (_)/ _))" [61,1000,20,1000]60) "_hoaret_emptyFaults":: "[('s,'p,'f) body,('s,'p) quadruple set, 's assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_,_/\\<^sub>t (_/ (_)/ _,/_))" [61,60,1000,20,1000,1000]60) "_hoaret_emptyCtx":: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_/\\<^sub>t\<^bsub>'/_\<^esub> (_/ (_)/ _,/_))" [61,60,1000,20,1000,1000]60) "_hoaret_emptyCtx_emptyFaults":: "[('s,'p,'f) body,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_/\\<^sub>t (_/ (_)/ _,/_))" [61,1000,20,1000,1000]60) "_hoaret_noAbr":: "[('s,'p,'f) body,'f set, ('s,'p) quadruple set, 's assn,('s,'p,'f) com, 's assn] => bool" ("(3_,_/\\<^sub>t\<^bsub>'/_\<^esub> (_/ (_)/ _))" [61,60,60,1000,20,1000]60) "_hoaret_noAbr_emptyFaults":: "[('s,'p,'f) body,('s,'p) quadruple set,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_,_/\\<^sub>t (_/ (_)/ _))" [61,60,1000,20,1000]60) "_hoaret_emptyCtx_noAbr":: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_/\\<^sub>t\<^bsub>'/_\<^esub> (_/ (_)/ _))" [61,60,1000,20,1000]60) "_hoaret_emptyCtx_noAbr_emptyFaults":: "[('s,'p,'f) body,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_/\\<^sub>t (_/ (_)/ _))" [61,1000,20,1000]60) syntax (ASCII) "_hoarep_emptyFaults":: "[('s,'p,'f) body,('s,'p) quadruple set, 's assn,('s,'p,'f) com, 's assn,'s assn] \ bool" ("(3_,_/|- (_/ (_)/ _,/_))" [61,60,1000,20,1000,1000]60) "_hoarep_emptyCtx":: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_/|-'/_ (_/ (_)/ _,/_))" [61,60,1000,20,1000,1000]60) "_hoarep_emptyCtx_emptyFaults":: "[('s,'p,'f) body,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_/|-(_/ (_)/ _,/_))" [61,1000,20,1000,1000]60) "_hoarep_noAbr":: "[('s,'p,'f) body,('s,'p) quadruple set,'f set, 's assn,('s,'p,'f) com, 's assn] => bool" ("(3_,_/|-'/_ (_/ (_)/ _))" [61,60,60,1000,20,1000]60) "_hoarep_noAbr_emptyFaults":: "[('s,'p,'f) body,('s,'p) quadruple set,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_,_/|-(_/ (_)/ _))" [61,60,1000,20,1000]60) "_hoarep_emptyCtx_noAbr":: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_/|-'/_ (_/ (_)/ _))" [61,60,1000,20,1000]60) "_hoarep_emptyCtx_noAbr_emptyFaults":: "[('s,'p,'f) body,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_/|-(_/ (_)/ _))" [61,1000,20,1000]60) "_hoaret_emptyFault":: "[('s,'p,'f) body,('s,'p) quadruple set, 's assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_,_/|-t (_/ (_)/ _,/_))" [61,60,1000,20,1000,1000]60) "_hoaret_emptyCtx":: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_/|-t'/_ (_/ (_)/ _,/_))" [61,60,1000,20,1000,1000]60) "_hoaret_emptyCtx_emptyFaults":: "[('s,'p,'f) body,'s assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_/|-t(_/ (_)/ _,/_))" [61,1000,20,1000,1000]60) "_hoaret_noAbr":: "[('s,'p,'f) body,('s,'p) quadruple set,'f set, 's assn,('s,'p,'f) com, 's assn] => bool" ("(3_,_/|-t'/_ (_/ (_)/ _))" [61,60,60,1000,20,1000]60) "_hoaret_noAbr_emptyFaults":: "[('s,'p,'f) body,('s,'p) quadruple set,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_,_/|-t(_/ (_)/ _))" [61,60,1000,20,1000]60) "_hoaret_emptyCtx_noAbr":: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_/|-t'/_ (_/ (_)/ _))" [61,60,1000,20,1000]60) "_hoaret_emptyCtx_noAbr_emptyFaults":: "[('s,'p,'f) body,'s assn,('s,'p,'f) com, 's assn] => bool" ("(3_/|-t(_/ (_)/ _))" [61,1000,20,1000]60) translations "\\ P c Q,A" == "\\\<^bsub>/{}\<^esub> P c Q,A" "\\\<^bsub>/F\<^esub> P c Q,A" == "\,{}\\<^bsub>/F\<^esub> P c Q,A" "\,\\ P c Q" == "\,\\\<^bsub>/{}\<^esub> P c Q" "\,\\\<^bsub>/F\<^esub> P c Q" == "\,\\\<^bsub>/F\<^esub> P c Q,{}" "\,\\ P c Q,A" == "\,\\\<^bsub>/{}\<^esub> P c Q,A" "\\ P c Q" == "\\\<^bsub>/{}\<^esub> P c Q" "\\\<^bsub>/F\<^esub> P c Q" == "\,{}\\<^bsub>/F\<^esub> P c Q" "\\\<^bsub>/F\<^esub> P c Q" <= "\\\<^bsub>/F\<^esub> P c Q,{}" "\\ P c Q" <= "\\ P c Q,{}" "\\\<^sub>t P c Q,A" == "\\\<^sub>t\<^bsub>/{}\<^esub> P c Q,A" "\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" == "\,{}\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" "\,\\\<^sub>t P c Q" == "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c Q" "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q" == "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,{}" "\,\\\<^sub>t P c Q,A" == "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c Q,A" "\\\<^sub>t P c Q" == "\\\<^sub>t\<^bsub>/{}\<^esub> P c Q" "\\\<^sub>t\<^bsub>/F\<^esub> P c Q" == "\,{}\\<^sub>t\<^bsub>/F\<^esub> P c Q" "\\\<^sub>t\<^bsub>/F\<^esub> P c Q" <= "\\\<^sub>t\<^bsub>/F\<^esub> P c Q,{}" "\\\<^sub>t P c Q" <= "\\\<^sub>t P c Q,{}" term "\\ P c Q" term "\\ P c Q,A" term "\\\<^bsub>/F\<^esub> P c Q" term "\\\<^bsub>/F\<^esub> P c Q,A" term "\,\\ P c Q" term "\,\\\<^bsub>/F\<^esub> P c Q" term "\,\\ P c Q,A" term "\,\\\<^bsub>/F\<^esub> P c Q,A" term "\\\<^sub>t P c Q" term "\\\<^sub>t P c Q,A" term "\\\<^sub>t\<^bsub>/F\<^esub> P c Q" term "\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" term "\,\\ P c Q" term "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q" term "\,\\ P c Q,A" term "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" locale hoare = fixes \::"('s,'p,'f) body" primrec assoc:: "('a \'b) list \ 'a \ 'b " where "assoc [] x = undefined" | "assoc (p#ps) x = (if fst p = x then (snd p) else assoc ps x)" lemma conjE_simp: "(P \ Q \ PROP R) \ (P \ Q \ PROP R)" by rule simp_all lemma CollectInt_iff: "{s. P s} \ {s. Q s} = {s. P s \ Q s}" by auto lemma Compl_Collect:"-(Collect b) = {x. \(b x)}" by fastforce lemma Collect_False: "{s. False} = {}" by simp lemma Collect_True: "{s. True} = UNIV" by simp lemma triv_All_eq: "\x. P \ P" by simp lemma triv_Ex_eq: "\x. P \ P" by simp lemma Ex_True: "\b. b" by blast lemma Ex_False: "\b. \b" by blast definition mex::"('a \ bool) \ bool" where "mex P = Ex P" definition meq::"'a \ 'a \ bool" where "meq s Z = (s = Z)" lemma subset_unI1: "A \ B \ A \ B \ C" by blast lemma subset_unI2: "A \ C \ A \ B \ C" by blast lemma split_paired_UN: "(\p. (P p)) = (\a b. (P (a,b)))" by auto lemma in_insert_hd: "f \ insert f X" by simp lemma lookup_Some_in_dom: "\ p = Some bdy \ p \ dom \" by auto lemma unit_object: "(\u::unit. P u) = P ()" by auto lemma unit_ex: "(\u::unit. P u) = P ()" by auto lemma unit_meta: "(\(u::unit). PROP P u) \ PROP P ()" by auto lemma unit_UN: "(\z::unit. P z) = P ()" by auto lemma subset_singleton_insert1: "y = x \ {y} \ insert x A" by auto lemma subset_singleton_insert2: "{y} \ A \ {y} \ insert x A" by auto lemma in_Specs_simp: "(\x\\Z. {(P Z, p, Q Z, A Z)}. Prop x) = (\Z. Prop (P Z,p,Q Z,A Z))" by auto lemma in_set_Un_simp: "(\x\A \ B. P x) = ((\x \ A. P x) \ (\x \ B. P x))" by auto lemma split_all_conj: "(\x. P x \ Q x) = ((\x. P x) \ (\x. Q x))" by blast lemma image_Un_single_simp: "f ` (\Z. {P Z}) = (\Z. {f (P Z)}) " by auto lemma measure_lex_prod_def': "f <*mlex*> r \ ({(x,y). (x,y) \ measure f \ f x=f y \ (x,y) \ r})" by (auto simp add: mlex_prod_def inv_image_def) lemma in_measure_iff: "(x,y) \ measure f = (f x < f y)" by (simp add: measure_def inv_image_def) lemma in_lex_iff: "((a,b),(x,y)) \ r <*lex*> s = ((a,x) \ r \ (a=x \ (b,y)\s))" by (simp add: lex_prod_def) lemma in_mlex_iff: "(x,y) \ f <*mlex*> r = (f x < f y \ (f x=f y \ (x,y) \ r))" by (simp add: measure_lex_prod_def' in_measure_iff) lemma in_inv_image_iff: "(x,y) \ inv_image r f = ((f x, f y) \ r)" by (simp add: inv_image_def) text \This is actually the same as @{thm [source] wf_mlex}. However, this basic proof took me so long that I'm not willing to delete it. \ lemma wf_measure_lex_prod [simp,intro]: assumes wf_r: "wf r" shows "wf (f <*mlex*> r)" proof (rule ccontr) assume " \ wf (f <*mlex*> r)" then obtain g where "\i. (g (Suc i), g i) \ f <*mlex*> r" by (auto simp add: wf_iff_no_infinite_down_chain) hence g: "\i. (g (Suc i), g i) \ measure f \ f (g (Suc i)) = f (g i) \ (g (Suc i), g i) \ r" by (simp add: measure_lex_prod_def') hence le_g: "\i. f (g (Suc i)) \ f (g i)" by (auto simp add: in_measure_iff order_le_less) have "wf (measure f)" by simp hence " \Q. (\x. x \ Q) \ (\z\Q. \y. (y, z) \ measure f \ y \ Q)" by (simp add: wf_eq_minimal) from this [rule_format, of "g ` UNIV"] have "\z. z \ range g \ (\y. (y, z) \ measure f \ y \ range g)" by auto then obtain z where z: "z \ range g" and min_z: "\y. f y < f z \ y \ range g" by (auto simp add: in_measure_iff) from z obtain k where k: "z = g k" by auto have "\i. k \ i \ f (g i) = f (g k)" proof (intro allI impI) fix i assume "k \ i" then show "f (g i) = f (g k)" proof (induct i) case 0 have "k \ 0" by fact hence "k = 0" by simp thus "f (g 0) = f (g k)" by simp next case (Suc n) have k_Suc_n: "k \ Suc n" by fact then show "f (g (Suc n)) = f (g k)" proof (cases "k = Suc n") case True thus ?thesis by simp next case False with k_Suc_n have "k \ n" by simp with Suc.hyps have n_k: "f (g n) = f (g k)" by simp from le_g have le: "f (g (Suc n)) \ f (g n)" by simp show ?thesis proof (cases "f (g (Suc n)) = f (g n)") case True with n_k show ?thesis by simp next case False with le have "f (g (Suc n)) < f (g n)" by simp with n_k k have "f (g (Suc n)) < f z" by simp with min_z have "g (Suc n) \ range g" by blast hence False by simp thus ?thesis by simp qed qed qed qed with k [symmetric] have "\i. k \ i \ f (g i) = f z" by simp hence "\i. k \ i \ f (g (Suc i)) = f (g i)" by simp with g have "\i. k \ i \ (g (Suc i),(g i)) \ r" by (auto simp add: in_measure_iff order_less_le ) hence "\i. (g (Suc (i+k)),(g (i+k))) \ r" by simp then have "\f. \i. (f (Suc i), f i) \ r" by - (rule exI [where x="\i. g (i+k)"],simp) with wf_r show False by (simp add: wf_iff_no_infinite_down_chain) qed lemmas all_imp_to_ex = all_simps (5) (*"\P Q. (\x. P x \ Q) = ((\x. P x) \ Q)" Avoid introduction of existential quantification of states on negative position. *) lemma all_imp_eq_triv: "(\x. x = k \ Q) = Q" "(\x. k = x \ Q) = Q" by auto end diff --git a/thys/Simpl/HoarePartial.thy b/thys/Simpl/HoarePartial.thy --- a/thys/Simpl/HoarePartial.thy +++ b/thys/Simpl/HoarePartial.thy @@ -1,1212 +1,1420 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: HoarePartial.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. +Copyright (c) 2022 Apple Inc. All rights reserved. -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA *) section \Derived Hoare Rules for Partial Correctness\ theory HoarePartial imports HoarePartialProps begin lemma conseq_no_aux: "\\,\\\<^bsub>/F\<^esub> P' c Q',A'; \s. s \ P \ (s\P' \ (Q' \ Q) \ (A' \ A))\ \ \,\\\<^bsub>/F\<^esub> P c Q,A" by (rule conseq [where P'="\Z. P'" and Q'="\Z. Q'" and A'="\Z. A'"]) auto lemma conseq_exploit_pre: "\\s \ P. \,\\\<^bsub>/F\<^esub> ({s} \ P) c Q,A\ \ \,\\\<^bsub>/F\<^esub> P c Q,A" apply (rule Conseq) apply clarify apply (rule_tac x="{s} \ P" in exI) apply (rule_tac x="Q" in exI) apply (rule_tac x="A" in exI) by simp lemma conseq:"\\Z. \,\\\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z); \s. s \ P \ (\ Z. s\P' Z \ (Q' Z \ Q) \ (A' Z \ A))\ \ \,\\\<^bsub>/F\<^esub> P c Q,A" by (rule Conseq') blast lemma Lem: "\\Z. \,\\\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z); P \ {s. \ Z. s\P' Z \ (Q' Z \ Q) \ (A' Z \ A)}\ \ \,\\\<^bsub>/F\<^esub> P (lem x c) Q,A" apply (unfold lem_def) apply (erule conseq) apply blast done lemma LemAnno: assumes conseq: "P \ {s. \Z. s\P' Z \ (\t. t \ Q' Z \ t \ Q) \ (\t. t \ A' Z \ t \ A)}" assumes lem: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z)" shows "\,\\\<^bsub>/F\<^esub> P (lem x c) Q,A" apply (rule Lem [OF lem]) using conseq by blast lemma LemAnnoNoAbrupt: assumes conseq: "P \ {s. \Z. s\P' Z \ (\t. t \ Q' Z \ t \ Q)}" assumes lem: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) c (Q' Z),{}" shows "\,\\\<^bsub>/F\<^esub> P (lem x c) Q,{}" apply (rule Lem [OF lem]) using conseq by blast lemma TrivPost: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z) \ \Z. \,\\\<^bsub>/F\<^esub> (P' Z) c UNIV,UNIV" apply (rule allI) apply (erule conseq) apply auto done lemma TrivPostNoAbr: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) c (Q' Z),{} \ \Z. \,\\\<^bsub>/F\<^esub> (P' Z) c UNIV,{}" apply (rule allI) apply (erule conseq) apply auto done lemma conseq_under_new_pre:"\\,\\\<^bsub>/F \<^esub>P' c Q',A'; \s \ P. s \ P' \ Q' \ Q \ A' \ A\ \ \,\\\<^bsub>/F \<^esub>P c Q,A" apply (rule conseq) apply (rule allI) apply assumption apply auto done lemma conseq_Kleymann:"\\Z. \,\\\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z); \s \ P. (\Z. s\P' Z \ (Q' Z \ Q) \ (A' Z \ A))\ \ \,\\\<^bsub>/F\<^esub> P c Q,A" by (rule Conseq') blast lemma DynComConseq: assumes "P \ {s. \P' Q' A'. \,\\\<^bsub>/F \<^esub>P' (c s) Q',A' \ P \ P' \ Q' \ Q \ A' \ A}" shows "\,\\\<^bsub>/F \<^esub>P DynCom c Q,A" using assms apply - apply (rule DynCom) apply clarsimp apply (rule Conseq) apply clarsimp apply blast done lemma SpecAnno: assumes consequence: "P \ {s. (\ Z. s\P' Z \ (Q' Z \ Q) \ (A' Z \ A))}" assumes spec: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) (c Z) (Q' Z),(A' Z)" assumes bdy_constant: "\Z. c Z = c undefined" shows "\,\\\<^bsub>/F\<^esub> P (specAnno P' c Q' A') Q,A" proof - from spec bdy_constant have "\Z. \,\\\<^bsub>/F\<^esub> ((P' Z)) (c undefined) (Q' Z),(A' Z)" apply - apply (rule allI) apply (erule_tac x=Z in allE) apply (erule_tac x=Z in allE) apply simp done with consequence show ?thesis apply (simp add: specAnno_def) apply (erule conseq) apply blast done qed lemma SpecAnno': "\P \ {s. \ Z. s\P' Z \ (\t. t \ Q' Z \ t \ Q) \ (\t. t \ A' Z \ t \ A)}; \Z. \,\\\<^bsub>/F\<^esub> (P' Z) (c Z) (Q' Z),(A' Z); \Z. c Z = c undefined \ \ \,\\\<^bsub>/F\<^esub> P (specAnno P' c Q' A') Q,A" apply (simp only: subset_iff [THEN sym]) apply (erule (1) SpecAnno) apply assumption done lemma SpecAnnoNoAbrupt: "\P \ {s. \ Z. s\P' Z \ (\t. t \ Q' Z \ t \ Q)}; \Z. \,\\\<^bsub>/F\<^esub> (P' Z) (c Z) (Q' Z),{}; \Z. c Z = c undefined \ \ \,\\\<^bsub>/F\<^esub> P (specAnno P' c Q' (\s. {})) Q,A" apply (rule SpecAnno') apply auto done lemma Skip: "P \ Q \ \,\\\<^bsub>/F\<^esub> P Skip Q,A" by (rule hoarep.Skip [THEN conseqPre],simp) lemma Basic: "P \ {s. (f s) \ Q} \ \,\\\<^bsub>/F\<^esub> P (Basic f) Q,A" by (rule hoarep.Basic [THEN conseqPre]) lemma BasicCond: "\P \ {s. (b s \ f s\Q) \ (\ b s \ g s\Q)}\ \ \,\\\<^bsub>/F\<^esub> P Basic (\s. if b s then f s else g s) Q,A" apply (rule Basic) apply auto done lemma Spec: "P \ {s. (\t. (s,t) \ r \ t \ Q) \ (\t. (s,t) \ r)} \ \,\\\<^bsub>/F\<^esub> P (Spec r) Q,A" by (rule hoarep.Spec [THEN conseqPre]) lemma SpecIf: "\P \ {s. (b s \ f s \ Q) \ (\ b s \ g s \ Q \ h s \ Q)}\ \ \,\\\<^bsub>/F\<^esub> P Spec (if_rel b f g h) Q,A" apply (rule Spec) apply (auto simp add: if_rel_def) done lemma Seq [trans, intro?]: "\\,\\\<^bsub>/F\<^esub> P c\<^sub>1 R,A; \,\\\<^bsub>/F\<^esub> R c\<^sub>2 Q,A\ \ \,\\\<^bsub>/F\<^esub> P (Seq c\<^sub>1 c\<^sub>2) Q,A" by (rule hoarep.Seq) +lemma SeqSame: + "\\,\\\<^bsub>/F\<^esub> P c\<^sub>1 Q,A; \,\\\<^bsub>/F\<^esub> Q c\<^sub>2 Q,A\ \ \,\\\<^bsub>/F\<^esub> P (Seq c\<^sub>1 c\<^sub>2) Q,A" + by (rule hoarep.Seq) + + lemma SeqSwap: "\\,\\\<^bsub>/F\<^esub> R c2 Q,A; \,\\\<^bsub>/F\<^esub> P c1 R,A\ \ \,\\\<^bsub>/F\<^esub> P (Seq c1 c2) Q,A" by (rule Seq) lemma BSeq: "\\,\\\<^bsub>/F\<^esub> P c\<^sub>1 R,A; \,\\\<^bsub>/F\<^esub> R c\<^sub>2 Q,A\ \ \,\\\<^bsub>/F\<^esub> P (bseq c\<^sub>1 c\<^sub>2) Q,A" by (unfold bseq_def) (rule Seq) +lemma BSeqSame: + "\\,\\\<^bsub>/F\<^esub> P c\<^sub>1 Q,A; \,\\\<^bsub>/F\<^esub> Q c\<^sub>2 Q,A\ \ \,\\\<^bsub>/F\<^esub> P (bseq c\<^sub>1 c\<^sub>2) Q,A" + by (rule BSeq) lemma Cond: assumes wp: "P \ {s. (s\b \ s\P\<^sub>1) \ (s\b \ s\P\<^sub>2)}" assumes deriv_c1: "\,\\\<^bsub>/F\<^esub> P\<^sub>1 c\<^sub>1 Q,A" assumes deriv_c2: "\,\\\<^bsub>/F\<^esub> P\<^sub>2 c\<^sub>2 Q,A" shows "\,\\\<^bsub>/F\<^esub> P (Cond b c\<^sub>1 c\<^sub>2) Q,A" proof (rule hoarep.Cond [THEN conseqPre]) from deriv_c1 show "\,\\\<^bsub>/F\<^esub> ({s. (s \ b \ s \ P\<^sub>1) \ (s \ b \ s \ P\<^sub>2)} \ b) c\<^sub>1 Q,A" by (rule conseqPre) blast next from deriv_c2 show "\,\\\<^bsub>/F\<^esub> ({s. (s \ b \ s \ P\<^sub>1) \ (s \ b \ s \ P\<^sub>2)} \ - b) c\<^sub>2 Q,A" by (rule conseqPre) blast next show "P \ {s. (s\b \ s\P\<^sub>1) \ (s\b \ s\P\<^sub>2)}" by (rule wp) qed lemma CondSwap: "\\,\\\<^bsub>/F\<^esub> P1 c1 Q,A; \,\\\<^bsub>/F\<^esub> P2 c2 Q,A; P \ {s. (s\b \ s\P1) \ (s\b \ s\P2)}\ \ \,\\\<^bsub>/F\<^esub> P (Cond b c1 c2) Q,A" by (rule Cond) lemma Cond': "\P \ {s. (b \ P1) \ (- b \ P2)};\,\\\<^bsub>/F\<^esub> P1 c1 Q,A; \,\\\<^bsub>/F\<^esub> P2 c2 Q,A\ \ \,\\\<^bsub>/F\<^esub> P (Cond b c1 c2) Q,A" by (rule CondSwap) blast+ lemma CondInv: assumes wp: "P \ Q" assumes inv: "Q \ {s. (s\b \ s\P\<^sub>1) \ (s\b \ s\P\<^sub>2)}" assumes deriv_c1: "\,\\\<^bsub>/F\<^esub> P\<^sub>1 c\<^sub>1 Q,A" assumes deriv_c2: "\,\\\<^bsub>/F\<^esub> P\<^sub>2 c\<^sub>2 Q,A" shows "\,\\\<^bsub>/F\<^esub> P (Cond b c\<^sub>1 c\<^sub>2) Q,A" proof - from wp inv have "P \ {s. (s\b \ s\P\<^sub>1) \ (s\b \ s\P\<^sub>2)}" by blast from Cond [OF this deriv_c1 deriv_c2] show ?thesis . qed lemma CondInv': assumes wp: "P \ I" assumes inv: "I \ {s. (s\b \ s\P\<^sub>1) \ (s\b \ s\P\<^sub>2)}" assumes wp': "I \ Q" assumes deriv_c1: "\,\\\<^bsub>/F\<^esub> P\<^sub>1 c\<^sub>1 I,A" assumes deriv_c2: "\,\\\<^bsub>/F\<^esub> P\<^sub>2 c\<^sub>2 I,A" shows "\,\\\<^bsub>/F\<^esub> P (Cond b c\<^sub>1 c\<^sub>2) Q,A" proof - from CondInv [OF wp inv deriv_c1 deriv_c2] have "\,\\\<^bsub>/F\<^esub> P (Cond b c\<^sub>1 c\<^sub>2) I,A". from conseqPost [OF this wp' subset_refl] show ?thesis . qed lemma switchNil: "P \ Q \ \,\\\<^bsub>/F \<^esub>P (switch v []) Q,A" by (simp add: Skip) lemma switchCons: "\P \ {s. (v s \ V \ s \ P\<^sub>1) \ (v s \ V \ s \ P\<^sub>2)}; \,\\\<^bsub>/F \<^esub>P\<^sub>1 c Q,A; \,\\\<^bsub>/F \<^esub>P\<^sub>2 (switch v vs) Q,A\ \ \,\\\<^bsub>/F \<^esub>P (switch v ((V,c)#vs)) Q,A" by (simp add: Cond) lemma Guard: "\P \ g \ R; \,\\\<^bsub>/F\<^esub> R c Q,A\ \ \,\\\<^bsub>/F\<^esub> P (Guard f g c) Q,A" apply (rule Guard [THEN conseqPre, of _ _ _ _ R]) apply (erule conseqPre) apply auto done lemma GuardSwap: "\ \,\\\<^bsub>/F\<^esub> R c Q,A; P \ g \ R\ \ \,\\\<^bsub>/F\<^esub> P (Guard f g c) Q,A" by (rule Guard) lemma Guarantee: "\P \ {s. s \ g \ s \ R}; \,\\\<^bsub>/F\<^esub> R c Q,A; f \ F\ \ \,\\\<^bsub>/F\<^esub> P (Guard f g c) Q,A" apply (rule Guarantee [THEN conseqPre, of _ _ _ _ _ "{s. s \ g \ s \ R}"]) apply assumption apply (erule conseqPre) apply auto done lemma GuaranteeSwap: "\ \,\\\<^bsub>/F\<^esub> R c Q,A; P \ {s. s \ g \ s \ R}; f \ F\ \ \,\\\<^bsub>/F\<^esub> P (Guard f g c) Q,A" by (rule Guarantee) lemma GuardStrip: "\P \ R; \,\\\<^bsub>/F\<^esub> R c Q,A; f \ F\ \ \,\\\<^bsub>/F\<^esub> P (Guard f g c) Q,A" apply (rule Guarantee [THEN conseqPre]) apply auto done +lemma GuardStripSame: + "\\,\\\<^bsub>/F\<^esub> P c Q,A; f \ F\ + \ \,\\\<^bsub>/F\<^esub> P (Guard f g c) Q,A" + by (rule GuardStrip [OF subset_refl]) + lemma GuardStripSwap: "\\,\\\<^bsub>/F\<^esub> R c Q,A; P \ R; f \ F\ \ \,\\\<^bsub>/F\<^esub> P (Guard f g c) Q,A" by (rule GuardStrip) lemma GuaranteeStrip: "\P \ R; \,\\\<^bsub>/F\<^esub> R c Q,A; f \ F\ \ \,\\\<^bsub>/F\<^esub> P (guaranteeStrip f g c) Q,A" by (unfold guaranteeStrip_def) (rule GuardStrip) lemma GuaranteeStripSwap: "\\,\\\<^bsub>/F\<^esub> R c Q,A; P \ R; f \ F\ \ \,\\\<^bsub>/F\<^esub> P (guaranteeStrip f g c) Q,A" by (unfold guaranteeStrip_def) (rule GuardStrip) lemma GuaranteeAsGuard: "\P \ g \ R; \,\\\<^bsub>/F\<^esub> R c Q,A\ \ \,\\\<^bsub>/F\<^esub> P (guaranteeStrip f g c) Q,A" by (unfold guaranteeStrip_def) (rule Guard) lemma GuaranteeAsGuardSwap: "\ \,\\\<^bsub>/F\<^esub> R c Q,A; P \ g \ R\ \ \,\\\<^bsub>/F\<^esub> P (guaranteeStrip f g c) Q,A" by (rule GuaranteeAsGuard) lemma GuardsNil: "\,\\\<^bsub>/F\<^esub> P c Q,A \ \,\\\<^bsub>/F\<^esub> P (guards [] c) Q,A" by simp lemma GuardsCons: "\,\\\<^bsub>/F\<^esub> P Guard f g (guards gs c) Q,A \ \,\\\<^bsub>/F\<^esub> P (guards ((f,g)#gs) c) Q,A" by simp lemma GuardsConsGuaranteeStrip: "\,\\\<^bsub>/F\<^esub> P guaranteeStrip f g (guards gs c) Q,A \ \,\\\<^bsub>/F\<^esub> P (guards (guaranteeStripPair f g#gs) c) Q,A" by (simp add: guaranteeStripPair_def guaranteeStrip_def) lemma While: assumes P_I: "P \ I" assumes deriv_body: "\,\\\<^bsub>/F\<^esub> (I \ b) c I,A" assumes I_Q: "I \ -b \ Q" shows "\,\\\<^bsub>/F\<^esub> P (whileAnno b I V c) Q,A" proof - from deriv_body P_I I_Q show ?thesis apply (simp add: whileAnno_def) apply (erule conseqPrePost [OF HoarePartialDef.While]) apply simp_all done qed text \@{term "J"} will be instantiated by tactic with @{term "gs' \ I"} for those guards that are not stripped.\ lemma WhileAnnoG: "\,\\\<^bsub>/F\<^esub> P (guards gs (whileAnno b J V (Seq c (guards gs Skip)))) Q,A \ \,\\\<^bsub>/F\<^esub> P (whileAnnoG gs b I V c) Q,A" by (simp add: whileAnnoG_def whileAnno_def while_def) text \This form stems from @{term "strip_guards F (whileAnnoG gs b I V c)"}\ lemma WhileNoGuard': assumes P_I: "P \ I" assumes deriv_body: "\,\\\<^bsub>/F\<^esub> (I \ b) c I,A" assumes I_Q: "I \ -b \ Q" shows "\,\\\<^bsub>/F\<^esub> P (whileAnno b I V (Seq c Skip)) Q,A" apply (rule While [OF P_I _ I_Q]) apply (rule Seq) apply (rule deriv_body) apply (rule hoarep.Skip) done lemma WhileAnnoFix: assumes consequence: "P \ {s. (\ Z. s\I Z \ (I Z \ -b \ Q)) }" assumes bdy: "\Z. \,\\\<^bsub>/F\<^esub> (I Z \ b) (c Z) (I Z),A" assumes bdy_constant: "\Z. c Z = c undefined" shows "\,\\\<^bsub>/F\<^esub> P (whileAnnoFix b I V c) Q,A" proof - from bdy bdy_constant have bdy': "\Z. \,\\\<^bsub>/F\<^esub> (I Z \ b) (c undefined) (I Z),A" apply - apply (rule allI) apply (erule_tac x=Z in allE) apply (erule_tac x=Z in allE) apply simp done have "\Z. \,\\\<^bsub>/F\<^esub> (I Z) (whileAnnoFix b I V c) (I Z \ -b),A" apply rule apply (unfold whileAnnoFix_def) apply (rule hoarep.While) apply (rule bdy' [rule_format]) done then show ?thesis apply (rule conseq) using consequence by blast qed lemma WhileAnnoFix': assumes consequence: "P \ {s. (\ Z. s\I Z \ (\t. t \ I Z \ -b \ t \ Q)) }" assumes bdy: "\Z. \,\\\<^bsub>/F\<^esub> (I Z \ b) (c Z) (I Z),A" assumes bdy_constant: "\Z. c Z = c undefined" shows "\,\\\<^bsub>/F\<^esub> P (whileAnnoFix b I V c) Q,A" apply (rule WhileAnnoFix [OF _ bdy bdy_constant]) using consequence by blast lemma WhileAnnoGFix: assumes whileAnnoFix: "\,\\\<^bsub>/F\<^esub> P (guards gs (whileAnnoFix b J V (\Z. (Seq (c Z) (guards gs Skip))))) Q,A" shows "\,\\\<^bsub>/F\<^esub> P (whileAnnoGFix gs b I V c) Q,A" using whileAnnoFix by (simp add: whileAnnoGFix_def whileAnnoFix_def while_def) lemma Bind: assumes adapt: "P \ {s. s \ P' s}" assumes c: "\s. \,\\\<^bsub>/F\<^esub> (P' s) (c (e s)) Q,A" shows "\,\\\<^bsub>/F\<^esub> P (bind e c) Q,A" apply (rule conseq [where P'="\Z. {s. s=Z \ s \ P' Z}" and Q'="\Z. Q" and A'="\Z. A"]) apply (rule allI) apply (unfold bind_def) apply (rule DynCom) apply (rule ballI) apply simp apply (rule conseqPre) apply (rule c [rule_format]) apply blast using adapt apply blast done -lemma Block: +lemma Block_exn: assumes adapt: "P \ {s. init s \ P' s}" -assumes bdy: "\s. \,\\\<^bsub>/F\<^esub> (P' s) bdy {t. return s t \ R s t},{t. return s t \ A}" +assumes bdy: "\s. \,\\\<^bsub>/F\<^esub> (P' s) bdy {t. return s t \ R s t},{t. result_exn (return s t) t \ A}" assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" -shows "\,\\\<^bsub>/F\<^esub> P (block init bdy return c) Q,A" +shows "\,\\\<^bsub>/F\<^esub> P (block_exn init bdy return result_exn c) Q,A" apply (rule conseq [where P'="\Z. {s. s=Z \ init s \ P' Z}" and Q'="\Z. Q" and A'="\Z. A"]) prefer 2 using adapt apply blast apply (rule allI) -apply (unfold block_def) +apply (unfold block_exn_def) apply (rule DynCom) apply (rule ballI) apply clarsimp apply (rule_tac R="{t. return Z t \ R Z t}" in SeqSwap ) apply (rule_tac P'="\Z'. {t. t=Z' \ return Z t \ R Z t}" and Q'="\Z'. Q" and A'="\Z'. A" in conseq) prefer 2 apply simp apply (rule allI) apply (rule DynCom) apply (clarsimp) apply (rule SeqSwap) apply (rule c [rule_format]) apply (rule Basic) apply clarsimp -apply (rule_tac R="{t. return Z t \ A}" in Catch) +apply (rule_tac R="{t. result_exn (return Z t) t \ A}" in Catch) apply (rule_tac R="{i. i \ P' Z}" in Seq) apply (rule Basic) apply clarsimp apply simp apply (rule bdy [rule_format]) apply (rule SeqSwap) apply (rule Throw) apply (rule Basic) apply simp done +lemma Block: +assumes adapt: "P \ {s. init s \ P' s}" +assumes bdy: "\s. \,\\\<^bsub>/F\<^esub> (P' s) bdy {t. return s t \ R s t},{t. return s t \ A}" +assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" +shows "\,\\\<^bsub>/F\<^esub> P (block init bdy return c) Q,A" + unfolding block_def + by (rule Block_exn [OF adapt bdy c]) + lemma BlockSwap: assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes bdy: "\s. \,\\\<^bsub>/F\<^esub> (P' s) bdy {t. return s t \ R s t},{t. return s t \ A}" assumes adapt: "P \ {s. init s \ P' s}" shows "\,\\\<^bsub>/F\<^esub> P (block init bdy return c) Q,A" using adapt bdy c by (rule Block) -lemma BlockSpec: +lemma Block_exnSpec: assumes adapt: "P \ {s. \Z. init s \ P' Z \ (\t. t \ Q' Z \ return s t \ R s t) \ - (\t. t \ A' Z \ return s t \ A)}" + (\t. t \ A' Z \ (result_exn (return s t) t) \ A)}" assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes bdy: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) bdy (Q' Z),(A' Z)" - shows "\,\\\<^bsub>/F\<^esub> P (block init bdy return c) Q,A" + shows "\,\\\<^bsub>/F\<^esub> P (block_exn init bdy return result_exn c) Q,A" apply (rule conseq [where P'="\Z. {s. init s \ P' Z \ (\t. t \ Q' Z \ return s t \ R s t) \ - (\t. t \ A' Z \ return s t \ A)}" and Q'="\Z. Q" and + (\t. t \ A' Z \ (result_exn (return s t) t) \ A)}" and Q'="\Z. Q" and A'="\Z. A"]) prefer 2 using adapt apply blast apply (rule allI) -apply (unfold block_def) +apply (unfold block_exn_def) apply (rule DynCom) apply (rule ballI) apply clarsimp apply (rule_tac R="{t. return s t \ R s t}" in SeqSwap ) apply (rule_tac P'="\Z'. {t. t=Z' \ return s t \ R s t}" and Q'="\Z'. Q" and A'="\Z'. A" in conseq) prefer 2 apply simp apply (rule allI) apply (rule DynCom) apply (clarsimp) apply (rule SeqSwap) apply (rule c [rule_format]) apply (rule Basic) apply clarsimp -apply (rule_tac R="{t. return s t \ A}" in Catch) +apply (rule_tac R="{t. (result_exn (return s t) t) \ A}" in Catch) apply (rule_tac R="{i. i \ P' Z}" in Seq) apply (rule Basic) apply clarsimp apply simp apply (rule conseq [OF bdy]) apply clarsimp apply blast apply (rule SeqSwap) apply (rule Throw) apply (rule Basic) apply simp -done + done + +lemma BlockSpec: + assumes adapt: "P \ {s. \Z. init s \ P' Z \ + (\t. t \ Q' Z \ return s t \ R s t) \ + (\t. t \ A' Z \ return s t \ A)}" + assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" + assumes bdy: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) bdy (Q' Z),(A' Z)" + shows "\,\\\<^bsub>/F\<^esub> P (block init bdy return c) Q,A" + unfolding block_def + by (rule Block_exnSpec [OF adapt c bdy]) + lemma Throw: "P \ A \ \,\\\<^bsub>/F\<^esub> P Throw Q,A" by (rule hoarep.Throw [THEN conseqPre]) lemmas Catch = hoarep.Catch lemma CatchSwap: "\\,\\\<^bsub>/F\<^esub> R c\<^sub>2 Q,A; \,\\\<^bsub>/F\<^esub> P c\<^sub>1 Q,R\ \ \,\\\<^bsub>/F\<^esub> P Catch c\<^sub>1 c\<^sub>2 Q,A" by (rule hoarep.Catch) +lemma CatchSame: "\\,\\\<^bsub>/F\<^esub> P c\<^sub>1 Q,A; \,\\\<^bsub>/F\<^esub> A c\<^sub>2 Q,A\ \ \,\\\<^bsub>/F\<^esub> P Catch c\<^sub>1 c\<^sub>2 Q,A" + by (rule hoarep.Catch) + lemma raise: "P \ {s. f s \ A} \ \,\\\<^bsub>/F\<^esub> P raise f Q,A" apply (simp add: raise_def) apply (rule Seq) apply (rule Basic) apply (assumption) apply (rule Throw) apply (rule subset_refl) done lemma condCatch: "\\,\\\<^bsub>/F\<^esub> P c\<^sub>1 Q,((b \ R) \ (-b \ A));\,\\\<^bsub>/F\<^esub> R c\<^sub>2 Q,A\ \ \,\\\<^bsub>/F\<^esub>P condCatch c\<^sub>1 b c\<^sub>2 Q,A" apply (simp add: condCatch_def) apply (rule Catch) apply assumption apply (rule CondSwap) apply (assumption) apply (rule hoarep.Throw) apply blast done lemma condCatchSwap: "\\,\\\<^bsub>/F\<^esub> R c\<^sub>2 Q,A;\,\\\<^bsub>/F\<^esub> P c\<^sub>1 Q,((b \ R) \ (-b \ A))\ \ \,\\\<^bsub>/F\<^esub>P condCatch c\<^sub>1 b c\<^sub>2 Q,A" by (rule condCatch) +lemma condCatchSame: + assumes c1: "\,\\\<^bsub>/F\<^esub> P c\<^sub>1 Q,A" + assumes c2: "\,\\\<^bsub>/F\<^esub> A c\<^sub>2 Q,A" + shows "\,\\\<^bsub>/F\<^esub>P condCatch c\<^sub>1 b c\<^sub>2 Q,A" +proof - + have eq: "((b \ A) \ (-b \ A)) = A" by blast + show ?thesis + apply (rule condCatch [OF _ c2]) + apply (simp add: eq) + apply (rule c1) + done +qed lemma ProcSpec: assumes adapt: "P \ {s. \Z. init s \ P' Z \ (\t. t \ Q' Z \ return s t \ R s t) \ (\t. t \ A' Z \ return s t \ A)}" assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes p: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) Call p (Q' Z),(A' Z)" shows "\,\\\<^bsub>/F\<^esub> P (call init p return c) Q,A" using adapt c p apply (unfold call_def) by (rule BlockSpec) +lemma Proc_exnSpec: + assumes adapt: "P \ {s. \Z. init s \ P' Z \ + (\t. t \ Q' Z \ return s t \ R s t) \ + (\t. t \ A' Z \ result_exn (return s t) t \ A)}" + assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" + assumes p: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) Call p (Q' Z),(A' Z)" + shows "\,\\\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" +using adapt c p +apply (unfold call_exn_def) + by (rule Block_exnSpec) + lemma ProcSpec': assumes adapt: "P \ {s. \Z. init s \ P' Z \ (\t \ Q' Z. return s t \ R s t) \ (\t \ A' Z. return s t \ A)}" assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes p: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) Call p (Q' Z),(A' Z)" shows "\,\\\<^bsub>/F\<^esub> P (call init p return c) Q,A" apply (rule ProcSpec [OF _ c p]) apply (insert adapt) apply clarsimp apply (drule (1) subsetD) apply (clarsimp) apply (rule_tac x=Z in exI) apply blast done +lemma Proc_exnSpecNoAbrupt: + assumes adapt: "P \ {s. \Z. init s \ P' Z \ + (\t. t \ Q' Z \ return s t \ R s t)}" + assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" + assumes p: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) Call p (Q' Z),{}" + shows "\,\\\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" +apply (rule Proc_exnSpec [OF _ c p]) +using adapt +apply simp + done + lemma ProcSpecNoAbrupt: assumes adapt: "P \ {s. \Z. init s \ P' Z \ (\t. t \ Q' Z \ return s t \ R s t)}" assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes p: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) Call p (Q' Z),{}" shows "\,\\\<^bsub>/F\<^esub> P (call init p return c) Q,A" apply (rule ProcSpec [OF _ c p]) using adapt apply simp done lemma FCall: "\,\\\<^bsub>/F\<^esub> P (call init p return (\s t. c (result t))) Q,A \ \,\\\<^bsub>/F\<^esub> P (fcall init p return result c) Q,A" by (simp add: fcall_def) lemma ProcRec: assumes deriv_bodies: "\p\Procs. \Z. \,\\(\p\Procs. \Z. {(P p Z,p,Q p Z,A p Z)}) \\<^bsub>/F\<^esub> (P p Z) (the (\ p)) (Q p Z),(A p Z)" assumes Procs_defined: "Procs \ dom \" shows "\p\Procs. \Z. \,\\\<^bsub>/F\<^esub>(P p Z) Call p (Q p Z),(A p Z)" by (intro strip) (rule CallRec' [OF _ Procs_defined deriv_bodies], simp_all) lemma ProcRec': assumes ctxt: "\' = \\(\p\Procs. \Z. {(P p Z,p,Q p Z,A p Z)})" assumes deriv_bodies: "\p\Procs. \Z. \,\'\\<^bsub>/F\<^esub> (P p Z) (the (\ p)) (Q p Z),(A p Z)" assumes Procs_defined: "Procs \ dom \" shows "\p\Procs. \Z. \,\\\<^bsub>/F\<^esub>(P p Z) Call p (Q p Z),(A p Z)" using ctxt deriv_bodies apply simp apply (erule ProcRec [OF _ Procs_defined]) done lemma ProcRecList: assumes deriv_bodies: "\p\set Procs. \Z. \,\\(\p\set Procs. \Z. {(P p Z,p,Q p Z,A p Z)}) \\<^bsub>/F\<^esub> (P p Z) (the (\ p)) (Q p Z),(A p Z)" assumes dist: "distinct Procs" assumes Procs_defined: "set Procs \ dom \" shows "\p\set Procs. \Z. \,\\\<^bsub>/F\<^esub>(P p Z) Call p (Q p Z),(A p Z)" using deriv_bodies Procs_defined by (rule ProcRec) lemma ProcRecSpecs: "\\(P,p,Q,A) \ Specs. \,\\Specs\\<^bsub>/F\<^esub> P (the (\ p)) Q,A; \(P,p,Q,A) \ Specs. p \ dom \\ \ \(P,p,Q,A) \ Specs. \,\\\<^bsub>/F\<^esub> P (Call p) Q,A" apply (auto intro: CallRec) done lemma ProcRec1: assumes deriv_body: "\Z. \,\\(\Z. {(P Z,p,Q Z,A Z)})\\<^bsub>/F\<^esub> (P Z) (the (\ p)) (Q Z),(A Z)" assumes p_defined: "p \ dom \" shows "\Z. \,\\\<^bsub>/F\<^esub> (P Z) Call p (Q Z),(A Z)" proof - from deriv_body p_defined have "\p\{p}. \Z. \,\\\<^bsub>/F\<^esub> (P Z) Call p (Q Z),(A Z)" by - (rule ProcRec [where A="\p. A" and P="\p. P" and Q="\p. Q"], simp_all) thus ?thesis by simp qed lemma ProcNoRec1: assumes deriv_body: "\Z. \,\\\<^bsub>/F\<^esub> (P Z) (the (\ p)) (Q Z),(A Z)" assumes p_def: "p \ dom \" shows "\Z. \,\\\<^bsub>/F\<^esub> (P Z) Call p (Q Z),(A Z)" proof - from deriv_body have "\Z. \,\\(\Z. {(P Z,p,Q Z,A Z)}) \\<^bsub>/F\<^esub> (P Z) (the (\ p)) (Q Z),(A Z)" by (blast intro: hoare_augment_context) from this p_def show ?thesis by (rule ProcRec1) qed lemma ProcBody: assumes WP: "P \ P'" assumes deriv_body: "\,\\\<^bsub>/F\<^esub> P' body Q,A" assumes body: "\ p = Some body" shows "\,\\\<^bsub>/F\<^esub> P Call p Q,A" apply (rule conseqPre [OF _ WP]) apply (rule ProcNoRec1 [rule_format, where P="\Z. P'" and Q="\Z. Q" and A="\Z. A"]) apply (insert body) apply simp apply (rule hoare_augment_context [OF deriv_body]) apply blast apply fastforce done lemma CallBody: assumes adapt: "P \ {s. init s \ P' s}" assumes bdy: "\s. \,\\\<^bsub>/F\<^esub> (P' s) body {t. return s t \ R s t},{t. return s t \ A}" assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes body: "\ p = Some body" shows "\,\\\<^bsub>/F\<^esub> P (call init p return c) Q,A" apply (unfold call_def) apply (rule Block [OF adapt _ c]) apply (rule allI) apply (rule ProcBody [where \=\, OF _ bdy [rule_format] body]) apply simp done +lemma Call_exnBody: +assumes adapt: "P \ {s. init s \ P' s}" +assumes bdy: "\s. \,\\\<^bsub>/F\<^esub> (P' s) body {t. return s t \ R s t},{t. result_exn (return s t) t \ A}" +assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" +assumes body: "\ p = Some body" +shows "\,\\\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" +apply (unfold call_exn_def) +apply (rule Block_exn [OF adapt _ c]) +apply (rule allI) +apply (rule ProcBody [where \=\, OF _ bdy [rule_format] body]) +apply simp +done + lemmas ProcModifyReturn = HoarePartialProps.ProcModifyReturn lemmas ProcModifyReturnSameFaults = HoarePartialProps.ProcModifyReturnSameFaults +lemmas Proc_exnModifyReturn = HoarePartialProps.Proc_exnModifyReturn +lemmas Proc_exnModifyReturnSameFaults = HoarePartialProps.Proc_exnModifyReturnSameFaults + +lemma Proc_exnModifyReturnNoAbr: + assumes spec: "\,\\\<^bsub>/F\<^esub> P (call_exn init p return' result_exn c) Q,A" + assumes result_conform: + "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" + assumes modifies_spec: + "\\. \,\\\<^bsub>/UNIV\<^esub> {\} Call p (Modif \),{}" + shows "\,\\\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" + by (rule Proc_exnModifyReturn [OF spec result_conform _ modifies_spec]) simp lemma ProcModifyReturnNoAbr: assumes spec: "\,\\\<^bsub>/F\<^esub> P (call init p return' c) Q,A" assumes result_conform: "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" assumes modifies_spec: "\\. \,\\\<^bsub>/UNIV\<^esub> {\} Call p (Modif \),{}" shows "\,\\\<^bsub>/F\<^esub> P (call init p return c) Q,A" by (rule ProcModifyReturn [OF spec result_conform _ modifies_spec]) simp +lemma Proc_exnModifyReturnNoAbrSameFaults: + assumes spec: "\,\\\<^bsub>/F\<^esub> P (call_exn init p return' result_exn c) Q,A" + assumes result_conform: + "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" + assumes modifies_spec: + "\\. \,\\\<^bsub>/F\<^esub> {\} Call p (Modif \),{}" + shows "\,\\\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" + by (rule Proc_exnModifyReturnSameFaults [OF spec result_conform _ modifies_spec]) simp + lemma ProcModifyReturnNoAbrSameFaults: assumes spec: "\,\\\<^bsub>/F\<^esub> P (call init p return' c) Q,A" assumes result_conform: "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" assumes modifies_spec: "\\. \,\\\<^bsub>/F\<^esub> {\} Call p (Modif \),{}" shows "\,\\\<^bsub>/F\<^esub> P (call init p return c) Q,A" by (rule ProcModifyReturnSameFaults [OF spec result_conform _ modifies_spec]) simp - -lemma DynProc: +lemma DynProc_exn: assumes adapt: "P \ {s. \Z. init s \ P' s Z \ (\t. t \ Q' s Z \ return s t \ R s t) \ - (\t. t \ A' s Z \ return s t \ A)}" + (\t. t \ A' s Z \ result_exn (return s t) t \ A)}" assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes p: "\s\ P. \Z. \,\\\<^bsub>/F\<^esub> (P' s Z) Call (p s) (Q' s Z),(A' s Z)" - shows "\,\\\<^bsub>/F\<^esub> P dynCall init p return c Q,A" + shows "\,\\\<^bsub>/F\<^esub> P dynCall_exn f UNIV init p return result_exn c Q,A" apply (rule conseq [where P'="\Z. {s. s=Z \ s \ P}" and Q'="\Z. Q" and A'="\Z. A"]) prefer 2 using adapt apply blast apply (rule allI) -apply (unfold dynCall_def call_def block_def) +apply (unfold dynCall_exn_def call_exn_def maybe_guard_UNIV block_exn_def guards.simps) apply (rule DynCom) apply clarsimp apply (rule DynCom) apply clarsimp apply (frule in_mono [rule_format, OF adapt]) apply clarsimp apply (rename_tac Z') apply (rule_tac R="Q' Z Z'" in Seq) apply (rule CatchSwap) apply (rule SeqSwap) apply (rule Throw) apply (rule subset_refl) apply (rule Basic) apply (rule subset_refl) apply (rule_tac R="{i. i \ P' Z Z'}" in Seq) apply (rule Basic) apply clarsimp apply simp apply (rule_tac Q'="Q' Z Z'" and A'="A' Z Z'" in conseqPost) using p apply clarsimp apply simp apply clarsimp apply (rule_tac P'="\Z''. {t. t=Z'' \ return Z t \ R Z t}" and Q'="\Z''. Q" and A'="\Z''. A" in conseq) prefer 2 apply simp apply (rule allI) apply (rule DynCom) apply clarsimp apply (rule SeqSwap) apply (rule c [rule_format]) apply (rule Basic) apply clarsimp -done + done + +lemma DynProc_exn_guards_cons: + assumes p: "\,\\\<^bsub>/F\<^esub> P dynCall_exn f UNIV init p return result_exn c Q,A" + shows "\,\\\<^bsub>/F\<^esub> (g \ P) dynCall_exn f g init p return result_exn c Q,A" + using p apply (clarsimp simp add: dynCall_exn_def maybe_guard_def) + apply (rule Guard) + apply (rule subset_refl) + apply assumption + done + +lemma DynProc: + assumes adapt: "P \ {s. \Z. init s \ P' s Z \ + (\t. t \ Q' s Z \ return s t \ R s t) \ + (\t. t \ A' s Z \ return s t \ A)}" + assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" + assumes p: "\s\ P. \Z. \,\\\<^bsub>/F\<^esub> (P' s Z) Call (p s) (Q' s Z),(A' s Z)" + shows "\,\\\<^bsub>/F\<^esub> P dynCall init p return c Q,A" + using adapt c p unfolding dynCall_dynCall_exn by (rule DynProc_exn) + +lemma DynProc_exn': + assumes adapt: "P \ {s. \Z. init s \ P' s Z \ + (\t \ Q' s Z. return s t \ R s t) \ + (\t \ A' s Z. result_exn (return s t) t \ A)}" + assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" + assumes p: "\s\ P. \Z. \,\\\<^bsub>/F\<^esub> (P' s Z) Call (p s) (Q' s Z),(A' s Z)" + shows "\,\\\<^bsub>/F\<^esub> P dynCall_exn f UNIV init p return result_exn c Q,A" +proof - + from adapt have "P \ {s. \Z. init s \ P' s Z \ + (\t. t \ Q' s Z \ return s t \ R s t) \ + (\t. t \ A' s Z \ result_exn (return s t) t \ A)}" + by blast + from this c p show ?thesis + by (rule DynProc_exn) +qed lemma DynProc': assumes adapt: "P \ {s. \Z. init s \ P' s Z \ (\t \ Q' s Z. return s t \ R s t) \ (\t \ A' s Z. return s t \ A)}" assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes p: "\s\ P. \Z. \,\\\<^bsub>/F\<^esub> (P' s Z) Call (p s) (Q' s Z),(A' s Z)" shows "\,\\\<^bsub>/F\<^esub> P dynCall init p return c Q,A" -proof - - from adapt have "P \ {s. \Z. init s \ P' s Z \ - (\t. t \ Q' s Z \ return s t \ R s t) \ - (\t. t \ A' s Z \ return s t \ A)}" - by blast - from this c p show ?thesis - by (rule DynProc) -qed + using adapt c p unfolding dynCall_dynCall_exn by (rule DynProc_exn') - -lemma DynProcStaticSpec: +lemma DynProc_exnStaticSpec: assumes adapt: "P \ {s. s \ S \ (\Z. init s \ P' Z \ (\\. \ \ Q' Z \ return s \ \ R s \) \ - (\\. \ \ A' Z \ return s \ \ A))}" + (\\. \ \ A' Z \ result_exn (return s \) \ \ A))}" assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes spec: "\s\S. \Z. \,\\\<^bsub>/F\<^esub> (P' Z) Call (p s) (Q' Z),(A' Z)" -shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" +shows "\,\\\<^bsub>/F\<^esub> P (dynCall_exn f UNIV init p return result_exn c) Q,A" proof - from adapt have P_S: "P \ S" by blast - have "\,\\\<^bsub>/F\<^esub> (P \ S) (dynCall init p return c) Q,A" - apply (rule DynProc [where P'="\s Z. P' Z" and Q'="\s Z. Q' Z" + have "\,\\\<^bsub>/F\<^esub> (P \ S) (dynCall_exn f UNIV init p return result_exn c) Q,A" + apply (rule DynProc_exn [where P'="\s Z. P' Z" and Q'="\s Z. Q' Z" and A'="\s Z. A' Z", OF _ c]) apply clarsimp apply (frule in_mono [rule_format, OF adapt]) apply clarsimp using spec apply clarsimp done thus ?thesis by (rule conseqPre) (insert P_S,blast) qed +lemma DynProcStaticSpec: +assumes adapt: "P \ {s. s \ S \ (\Z. init s \ P' Z \ + (\\. \ \ Q' Z \ return s \ \ R s \) \ + (\\. \ \ A' Z \ return s \ \ A))}" +assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" +assumes spec: "\s\S. \Z. \,\\\<^bsub>/F\<^esub> (P' Z) Call (p s) (Q' Z),(A' Z)" +shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using adapt c spec unfolding dynCall_dynCall_exn by (rule DynProc_exnStaticSpec) + +lemma DynProc_exnProcPar: +assumes adapt: "P \ {s. p s = q \ (\Z. init s \ P' Z \ + (\\. \ \ Q' Z \ return s \ \ R s \) \ + (\\. \ \ A' Z \ result_exn (return s \) \ \ A))}" +assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" +assumes spec: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) Call q (Q' Z),(A' Z)" +shows "\,\\\<^bsub>/F\<^esub> P (dynCall_exn f UNIV init p return result_exn c) Q,A" + apply (rule DynProc_exnStaticSpec [where S="{s. p s = q}",simplified, OF adapt c]) + using spec + apply simp + done lemma DynProcProcPar: assumes adapt: "P \ {s. p s = q \ (\Z. init s \ P' Z \ (\\. \ \ Q' Z \ return s \ \ R s \) \ (\\. \ \ A' Z \ return s \ \ A))}" assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes spec: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) Call q (Q' Z),(A' Z)" shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" apply (rule DynProcStaticSpec [where S="{s. p s = q}",simplified, OF adapt c]) using spec apply simp done - -lemma DynProcProcParNoAbrupt: +lemma DynProc_exnProcParNoAbrupt: assumes adapt: "P \ {s. p s = q \ (\Z. init s \ P' Z \ (\\. \ \ Q' Z \ return s \ \ R s \))}" assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes spec: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) Call q (Q' Z),{}" -shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" +shows "\,\\\<^bsub>/F\<^esub> P (dynCall_exn f UNIV init p return result_exn c) Q,A" proof - have "P \ {s. p s = q \ (\ Z. init s \ P' Z \ (\t. t \ Q' Z \ return s t \ R s t) \ - (\t. t \ {} \ return s t \ A))}" + (\t. t \ {} \ result_exn (return s t) t \ A))}" (is "P \ ?P'") proof fix s assume P: "s\P" with adapt obtain Z where Pre: "p s = q \ init s \ P' Z" and adapt_Norm: "\\. \ \ Q' Z \ return s \ \ R s \" by blast from adapt_Norm have "\t. t \ Q' Z \ return s t \ R s t" by auto then show "s\?P'" using Pre by blast qed note P = this show ?thesis apply - - apply (rule DynProcStaticSpec [where S="{s. p s = q}",simplified, OF P c]) + apply (rule DynProc_exnStaticSpec [where S="{s. p s = q}",simplified, OF P c]) apply (insert spec) apply auto done qed +lemma DynProcProcParNoAbrupt: +assumes adapt: "P \ {s. p s = q \ (\Z. init s \ P' Z \ + (\\. \ \ Q' Z \ return s \ \ R s \))}" +assumes c: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" +assumes spec: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) Call q (Q' Z),{}" +shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using adapt c spec unfolding dynCall_dynCall_exn by (rule DynProc_exnProcParNoAbrupt) -lemma DynProcModifyReturnNoAbr: - assumes to_prove: "\,\\\<^bsub>/F\<^esub> P (dynCall init p return' c) Q,A" +lemma DynProc_exnModifyReturnNoAbr: + assumes to_prove: "\,\\\<^bsub>/F\<^esub> P (dynCall_exn f g init p return' result_exn c) Q,A" assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) \ return' s t = return s t" assumes modif_clause: "\s \ P. \\. \,\\\<^bsub>/UNIV\<^esub> {\} Call (p s) (Modif \),{}" - shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + shows "\,\\\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" proof - from ret_nrm_modif have "\s t. t \ (Modif (init s)) \ return' s t = return s t" by iprover then have ret_nrm_modif': "\s t. t \ (Modif (init s)) \ return' s t = return s t" by simp have ret_abr_modif': "\s t. t \ {} - \ return' s t = return s t" + \ result_exn (return' s t) t = result_exn (return s t) t" by simp from to_prove ret_nrm_modif' ret_abr_modif' modif_clause show ?thesis - by (rule dynProcModifyReturn) + by (rule dynProc_exnModifyReturn) qed - -lemma ProcDynModifyReturnNoAbrSameFaults: +lemma DynProcModifyReturnNoAbr: assumes to_prove: "\,\\\<^bsub>/F\<^esub> P (dynCall init p return' c) Q,A" assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) \ return' s t = return s t" assumes modif_clause: + "\s \ P. \\. \,\\\<^bsub>/UNIV\<^esub> {\} Call (p s) (Modif \),{}" + shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using to_prove ret_nrm_modif modif_clause unfolding dynCall_dynCall_exn + by (rule DynProc_exnModifyReturnNoAbr) + +lemma ProcDyn_exnModifyReturnNoAbrSameFaults: + assumes to_prove: "\,\\\<^bsub>/F\<^esub> P (dynCall_exn f g init p return' result_exn c) Q,A" + assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) + \ return' s t = return s t" + assumes modif_clause: "\s \ P. \\. \,\\\<^bsub>/F\<^esub> {\} (Call (p s)) (Modif \),{}" - shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + shows "\,\\\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" proof - from ret_nrm_modif have "\s t. t \ (Modif (init s)) \ return' s t = return s t" by iprover then have ret_nrm_modif': "\s t. t \ (Modif (init s)) \ return' s t = return s t" by simp have ret_abr_modif': "\s t. t \ {} - \ return' s t = return s t" + \ result_exn (return' s t) t = result_exn (return s t) t" by simp - from to_prove ret_nrm_modif' ret_abr_modif' modif_clause show ?thesis - by (rule dynProcModifyReturnSameFaults) + from to_prove ret_nrm_modif' ret_abr_modif' modif_clause show ?thesis + by (rule dynProc_exnModifyReturnSameFaults) +qed + +lemma ProcDynModifyReturnNoAbrSameFaults: + assumes to_prove: "\,\\\<^bsub>/F\<^esub> P (dynCall init p return' c) Q,A" + assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) + \ return' s t = return s t" + assumes modif_clause: + "\s \ P. \\. \,\\\<^bsub>/F\<^esub> {\} (Call (p s)) (Modif \),{}" + shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using to_prove ret_nrm_modif modif_clause unfolding dynCall_dynCall_exn + by (rule ProcDyn_exnModifyReturnNoAbrSameFaults) + +lemma Proc_exnProcParModifyReturn: + assumes q: "P \ {s. p s = q} \ P'" + \ \@{thm[source] DynProcProcPar} introduces the same constraint as first conjunction in + @{term P'}, so the vcg can simplify it.\ + assumes to_prove: "\,\\\<^bsub>/F\<^esub> P' (dynCall_exn f g init p return' result_exn c) Q,A" + assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) + \ return' s t = return s t" + assumes ret_abr_modif: "\s t. t \ (ModifAbr (init s)) + \ result_exn (return' s t) t = result_exn (return s t) t" + assumes modif_clause: + "\\. \,\\\<^bsub>/UNIV\<^esub> {\} (Call q) (Modif \),(ModifAbr \)" + shows "\,\\\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" +proof - + from to_prove have "\,\\\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall_exn f g init p return' result_exn c) Q,A" + by (rule conseqPre) blast + from this ret_nrm_modif + ret_abr_modif + have "\,\\\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall_exn f g init p return result_exn c) Q,A" + by (rule dynProc_exnModifyReturn) (insert modif_clause,auto) + from this q show ?thesis + by (rule conseqPre) qed lemma ProcProcParModifyReturn: assumes q: "P \ {s. p s = q} \ P'" \ \@{thm[source] DynProcProcPar} introduces the same constraint as first conjunction in @{term P'}, so the vcg can simplify it.\ assumes to_prove: "\,\\\<^bsub>/F\<^esub> P' (dynCall init p return' c) Q,A" assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) \ return' s t = return s t" assumes ret_abr_modif: "\s t. t \ (ModifAbr (init s)) \ return' s t = return s t" assumes modif_clause: "\\. \,\\\<^bsub>/UNIV\<^esub> {\} (Call q) (Modif \),(ModifAbr \)" - shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using q to_prove ret_nrm_modif ret_abr_modif modif_clause unfolding dynCall_dynCall_exn + by (rule Proc_exnProcParModifyReturn) + +lemma Proc_exnProcParModifyReturnSameFaults: + assumes q: "P \ {s. p s = q} \ P'" + \ \@{thm[source] DynProcProcPar} introduces the same constraint as first conjunction in + @{term P'}, so the vcg can simplify it.\ + assumes to_prove: "\,\\\<^bsub>/F\<^esub> P' (dynCall_exn f g init p return' result_exn c) Q,A" + assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) + \ return' s t = return s t" + assumes ret_abr_modif: "\s t. t \ (ModifAbr (init s)) + \ result_exn (return' s t) t = result_exn (return s t) t" + assumes modif_clause: + "\\. \,\\\<^bsub>/F\<^esub> {\} Call q (Modif \),(ModifAbr \)" + shows "\,\\\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" proof - - from to_prove have "\,\\\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall init p return' c) Q,A" + from to_prove + have "\,\\\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall_exn f g init p return' result_exn c) Q,A" by (rule conseqPre) blast from this ret_nrm_modif ret_abr_modif - have "\,\\\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall init p return c) Q,A" - by (rule dynProcModifyReturn) (insert modif_clause,auto) + have "\,\\\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall_exn f g init p return result_exn c) Q,A" + by (rule dynProc_exnModifyReturnSameFaults) (insert modif_clause,auto) from this q show ?thesis by (rule conseqPre) qed - lemma ProcProcParModifyReturnSameFaults: assumes q: "P \ {s. p s = q} \ P'" \ \@{thm[source] DynProcProcPar} introduces the same constraint as first conjunction in @{term P'}, so the vcg can simplify it.\ assumes to_prove: "\,\\\<^bsub>/F\<^esub> P' (dynCall init p return' c) Q,A" assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) \ return' s t = return s t" assumes ret_abr_modif: "\s t. t \ (ModifAbr (init s)) \ return' s t = return s t" assumes modif_clause: "\\. \,\\\<^bsub>/F\<^esub> {\} Call q (Modif \),(ModifAbr \)" - shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using q to_prove ret_nrm_modif ret_abr_modif modif_clause unfolding dynCall_dynCall_exn + by (rule Proc_exnProcParModifyReturnSameFaults) + +lemma Proc_exnProcParModifyReturnNoAbr: + assumes q: "P \ {s. p s = q} \ P'" + \ \@{thm[source] DynProcProcParNoAbrupt} introduces the same constraint as + first conjunction in @{term P'}, so the vcg can simplify it.\ + assumes to_prove: "\,\\\<^bsub>/F\<^esub> P' (dynCall_exn f g init p return' result_exn c) Q,A" + assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) + \ return' s t = return s t" + assumes modif_clause: + "\\. \,\\\<^bsub>/UNIV\<^esub> {\} (Call q) (Modif \),{}" + shows "\,\\\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" proof - - from to_prove - have "\,\\\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall init p return' c) Q,A" + from to_prove have "\,\\\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall_exn f g init p return' result_exn c) Q,A" by (rule conseqPre) blast from this ret_nrm_modif - ret_abr_modif - have "\,\\\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall init p return c) Q,A" - by (rule dynProcModifyReturnSameFaults) (insert modif_clause,auto) + have "\,\\\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall_exn f g init p return result_exn c) Q,A" + by (rule DynProc_exnModifyReturnNoAbr) (insert modif_clause,auto) from this q show ?thesis by (rule conseqPre) qed - lemma ProcProcParModifyReturnNoAbr: assumes q: "P \ {s. p s = q} \ P'" \ \@{thm[source] DynProcProcParNoAbrupt} introduces the same constraint as first conjunction in @{term P'}, so the vcg can simplify it.\ assumes to_prove: "\,\\\<^bsub>/F\<^esub> P' (dynCall init p return' c) Q,A" assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) \ return' s t = return s t" assumes modif_clause: "\\. \,\\\<^bsub>/UNIV\<^esub> {\} (Call q) (Modif \),{}" - shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using q to_prove ret_nrm_modif modif_clause unfolding dynCall_dynCall_exn + by (rule Proc_exnProcParModifyReturnNoAbr) + +lemma Proc_exnProcParModifyReturnNoAbrSameFaults: + assumes q: "P \ {s. p s = q} \ P'" + \ \@{thm[source] DynProcProcParNoAbrupt} introduces the same constraint as + first conjunction in @{term P'}, so the vcg can simplify it.\ + assumes to_prove: "\,\\\<^bsub>/F\<^esub> P' (dynCall_exn f g init p return' result_exn c) Q,A" + assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) + \ return' s t = return s t" + assumes modif_clause: + "\\. \,\\\<^bsub>/F\<^esub> {\} (Call q) (Modif \),{}" + shows "\,\\\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" proof - - from to_prove have "\,\\\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall init p return' c) Q,A" + from to_prove have + "\,\\\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall_exn f g init p return' result_exn c) Q,A" by (rule conseqPre) blast from this ret_nrm_modif - have "\,\\\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall init p return c) Q,A" - by (rule DynProcModifyReturnNoAbr) (insert modif_clause,auto) + have "\,\\\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall_exn f g init p return result_exn c) Q,A" + by (rule ProcDyn_exnModifyReturnNoAbrSameFaults) (insert modif_clause,auto) from this q show ?thesis by (rule conseqPre) qed lemma ProcProcParModifyReturnNoAbrSameFaults: assumes q: "P \ {s. p s = q} \ P'" \ \@{thm[source] DynProcProcParNoAbrupt} introduces the same constraint as first conjunction in @{term P'}, so the vcg can simplify it.\ assumes to_prove: "\,\\\<^bsub>/F\<^esub> P' (dynCall init p return' c) Q,A" assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) \ return' s t = return s t" assumes modif_clause: "\\. \,\\\<^bsub>/F\<^esub> {\} (Call q) (Modif \),{}" - shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" -proof - - from to_prove have - "\,\\\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall init p return' c) Q,A" - by (rule conseqPre) blast - from this ret_nrm_modif - have "\,\\\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall init p return c) Q,A" - by (rule ProcDynModifyReturnNoAbrSameFaults) (insert modif_clause,auto) - from this q show ?thesis - by (rule conseqPre) -qed + shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using q to_prove ret_nrm_modif modif_clause unfolding dynCall_dynCall_exn + by (rule Proc_exnProcParModifyReturnNoAbrSameFaults) + lemma MergeGuards_iff: "\,\\\<^bsub>/F\<^esub> P merge_guards c Q,A = \,\\\<^bsub>/F\<^esub> P c Q,A" by (auto intro: MergeGuardsI MergeGuardsD) lemma CombineStrip': assumes deriv: "\,\\\<^bsub>/F\<^esub> P c' Q,A" assumes deriv_strip_triv: "\,{}\\<^bsub>/{}\<^esub> P c'' UNIV,UNIV" assumes c'': "c''= mark_guards False (strip_guards (-F) c')" assumes c: "merge_guards c = merge_guards (mark_guards False c')" shows "\,\\\<^bsub>/{}\<^esub> P c Q,A" proof - from deriv_strip_triv have deriv_strip: "\,\\\<^bsub>/{}\<^esub> P c'' UNIV,UNIV" by (auto intro: hoare_augment_context) from deriv_strip [simplified c''] have "\,\\\<^bsub>/{}\<^esub> P (strip_guards (- F) c') UNIV,UNIV" by (rule MarkGuardsD) with deriv have "\,\\\<^bsub>/{}\<^esub> P c' Q,A" by (rule CombineStrip) hence "\,\\\<^bsub>/{}\<^esub> P mark_guards False c' Q,A" by (rule MarkGuardsI) hence "\,\\\<^bsub>/{}\<^esub> P merge_guards (mark_guards False c') Q,A" by (rule MergeGuardsI) hence "\,\\\<^bsub>/{}\<^esub> P merge_guards c Q,A" by (simp add: c) thus ?thesis by (rule MergeGuardsD) qed lemma CombineStrip'': assumes deriv: "\,\\\<^bsub>/{True}\<^esub> P c' Q,A" assumes deriv_strip_triv: "\,{}\\<^bsub>/{}\<^esub> P c'' UNIV,UNIV" assumes c'': "c''= mark_guards False (strip_guards ({False}) c')" assumes c: "merge_guards c = merge_guards (mark_guards False c')" shows "\,\\\<^bsub>/{}\<^esub> P c Q,A" apply (rule CombineStrip' [OF deriv deriv_strip_triv _ c]) apply (insert c'') apply (subgoal_tac "- {True} = {False}") apply auto done lemma AsmUN: "(\Z. {(P Z, p, Q Z,A Z)}) \ \ \ \Z. \,\\\<^bsub>/F\<^esub> (P Z) (Call p) (Q Z),(A Z)" by (blast intro: hoarep.Asm) lemma augment_context': "\\ \ \'; \Z. \,\\\<^bsub>/F\<^esub> (P Z) p (Q Z),(A Z)\ \ \Z. \,\'\\<^bsub>/F\<^esub> (P Z) p (Q Z),(A Z)" by (iprover intro: hoare_augment_context) lemma hoarep_strip: "\\Z. \,{}\\<^bsub>/F\<^esub> (P Z) p (Q Z),(A Z); F' \ -F\ \ \Z. strip F' \,{}\\<^bsub>/F\<^esub> (P Z) p (Q Z),(A Z)" by (iprover intro: hoare_strip_\) lemma augment_emptyFaults: "\\Z. \,{}\\<^bsub>/{}\<^esub> (P Z) p (Q Z),(A Z)\ \ \Z. \,{}\\<^bsub>/F\<^esub> (P Z) p (Q Z),(A Z)" by (blast intro: augment_Faults) lemma augment_FaultsUNIV: "\\Z. \,{}\\<^bsub>/F\<^esub> (P Z) p (Q Z),(A Z)\ \ \Z. \,{}\\<^bsub>/UNIV\<^esub> (P Z) p (Q Z),(A Z)" by (blast intro: augment_Faults) lemma PostConjI [trans]: "\\,\\\<^bsub>/F\<^esub> P c Q,A; \,\\\<^bsub>/F\<^esub> P c R,B\ \ \,\\\<^bsub>/F\<^esub> P c (Q \ R),(A \ B)" by (rule PostConjI) lemma PostConjI' : "\\,\\\<^bsub>/F\<^esub> P c Q,A; \,\\\<^bsub>/F\<^esub> P c Q,A \ \,\\\<^bsub>/F\<^esub> P c R,B\ \ \,\\\<^bsub>/F\<^esub> P c (Q \ R),(A \ B)" by (rule PostConjI) iprover+ lemma PostConjE [consumes 1]: assumes conj: "\,\\\<^bsub>/F\<^esub> P c (Q \ R),(A \ B)" assumes E: "\\,\\\<^bsub>/F\<^esub> P c Q,A; \,\\\<^bsub>/F\<^esub> P c R,B\ \ S" shows "S" proof - from conj have "\,\\\<^bsub>/F\<^esub> P c Q,A" by (rule conseqPost) blast+ moreover from conj have "\,\\\<^bsub>/F\<^esub> P c R,B" by (rule conseqPost) blast+ ultimately show "S" by (rule E) qed subsection \Rules for Single-Step Proof \label{sec:hoare-isar}\ text \ We are now ready to introduce a set of Hoare rules to be used in single-step structured proofs in Isabelle/Isar. \medskip Assertions of Hoare Logic may be manipulated in calculational proofs, with the inclusion expressed in terms of sets or predicates. Reversed order is supported as well. \ lemma annotateI [trans]: "\\,\\\<^bsub>/F\<^esub>P anno Q,A; c = anno\ \ \,\\\<^bsub>/F\<^esub>P c Q,A" by simp lemma annotate_normI: assumes deriv_anno: "\,\\\<^bsub>/F\<^esub>P anno Q,A" assumes norm_eq: "normalize c = normalize anno" shows "\,\\\<^bsub>/F\<^esub>P c Q,A" proof - from NormalizeI [OF deriv_anno] norm_eq have "\,\\\<^bsub>/F \<^esub>P normalize c Q,A" by simp from NormalizeD [OF this] show ?thesis . qed lemma annotateWhile: "\\,\\\<^bsub>/F\<^esub> P (whileAnnoG gs b I V c) Q,A\ \ \,\\\<^bsub>/F\<^esub> P (while gs b c) Q,A" by (simp add: whileAnnoG_def) lemma reannotateWhile: "\\,\\\<^bsub>/F\<^esub> P (whileAnnoG gs b I V c) Q,A\ \ \,\\\<^bsub>/F\<^esub> P (whileAnnoG gs b J V c) Q,A" by (simp add: whileAnnoG_def) lemma reannotateWhileNoGuard: "\\,\\\<^bsub>/F\<^esub> P (whileAnno b I V c) Q,A\ \ \,\\\<^bsub>/F\<^esub> P (whileAnno b J V c) Q,A" by (simp add: whileAnno_def) lemma [trans] : "P' \ P \ \,\\\<^bsub>/F\<^esub> P c Q,A \ \,\\\<^bsub>/F\<^esub> P' c Q,A" by (rule conseqPre) lemma [trans]: "Q \ Q' \ \,\\\<^bsub>/F\<^esub> P c Q,A \ \,\\\<^bsub>/F\<^esub> P c Q',A" by (rule conseqPost) blast+ lemma [trans]: "\,\\\<^bsub>/F\<^esub> {s. P s} c Q,A \ (\s. P' s \ P s) \ \,\\\<^bsub>/F\<^esub> {s. P' s} c Q,A" by (rule conseqPre) auto lemma [trans]: "(\s. P' s \ P s) \ \,\\\<^bsub>/F\<^esub> {s. P s} c Q,A \ \,\\\<^bsub>/F\<^esub> {s. P' s} c Q,A" by (rule conseqPre) auto lemma [trans]: "\,\\\<^bsub>/F\<^esub>P c {s. Q s},A \ (\s. Q s \ Q' s) \ \,\\\<^bsub>/F\<^esub>P c {s. Q' s},A" by (rule conseqPost) auto lemma [trans]: "(\s. Q s \ Q' s) \ \,\\\<^bsub>/F\<^esub>P c {s. Q s},A \ \,\\\<^bsub>/F\<^esub>P c {s. Q' s},A" by (rule conseqPost) auto lemma [intro?]: "\,\\\<^bsub>/F\<^esub> P Skip P,A" by (rule Skip) auto lemma CondInt [trans,intro?]: "\\,\\\<^bsub>/F\<^esub> (P \ b) c1 Q,A; \,\\\<^bsub>/F\<^esub> (P \ - b) c2 Q,A\ \ \,\\\<^bsub>/F\<^esub> P (Cond b c1 c2) Q,A" by (rule Cond) auto lemma CondConj [trans, intro?]: "\\,\\\<^bsub>/F\<^esub> {s. P s \ b s} c1 Q,A; \,\\\<^bsub>/F\<^esub> {s. P s \ \ b s} c2 Q,A\ \ \,\\\<^bsub>/F\<^esub> {s. P s} (Cond {s. b s} c1 c2) Q,A" by (rule Cond) auto lemma WhileInvInt [intro?]: "\,\\\<^bsub>/F\<^esub> (P \ b) c P,A \ \,\\\<^bsub>/F\<^esub> P (whileAnno b P V c) (P \ -b),A" by (rule While) auto lemma WhileInt [intro?]: "\,\\\<^bsub>/F\<^esub> (P \ b) c P,A \ \,\\\<^bsub>/F\<^esub> P (whileAnno b {s. undefined} V c) (P \ -b),A" by (unfold whileAnno_def) (rule HoarePartialDef.While [THEN conseqPrePost],auto) lemma WhileInvConj [intro?]: "\,\\\<^bsub>/F\<^esub> {s. P s \ b s} c {s. P s},A \ \,\\\<^bsub>/F\<^esub> {s. P s} (whileAnno {s. b s} {s. P s} V c) {s. P s \ \ b s},A" by (simp add: While Collect_conj_eq Collect_neg_eq) lemma WhileConj [intro?]: "\,\\\<^bsub>/F\<^esub> {s. P s \ b s} c {s. P s},A \ \,\\\<^bsub>/F\<^esub> {s. P s} (whileAnno {s. b s} {s. undefined} V c) {s. P s \ \ b s},A" by (unfold whileAnno_def) (simp add: HoarePartialDef.While [THEN conseqPrePost] Collect_conj_eq Collect_neg_eq) -(* FIXME: Add rules for guarded while *) +(* fixme: Add rules for guarded while *) end diff --git a/thys/Simpl/HoarePartialDef.thy b/thys/Simpl/HoarePartialDef.thy --- a/thys/Simpl/HoarePartialDef.thy +++ b/thys/Simpl/HoarePartialDef.thy @@ -1,434 +1,413 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: HoarePartialDef.thy - Author: Norbert Schirmer, TU Muenchen - -Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (C) 2006-2008 Norbert Schirmer *) section \Hoare Logic for Partial Correctness\ theory HoarePartialDef imports Semantic begin type_synonym ('s,'p) quadruple = "('s assn \ 'p \ 's assn \ 's assn)" subsection \Validity of Hoare Tuples: \\,\\\<^bsub>/F\<^esub> P c Q,A\\ definition valid :: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com,'s assn,'s assn] => bool" ("_\\<^bsub>'/_\<^esub>/ _ _ _,_" [61,60,1000, 20, 1000,1000] 60) where "\\\<^bsub>/F\<^esub> P c Q,A \ \s t. \\\c,s\ \ t \ s \ Normal ` P \ t \ Fault ` F \ t \ Normal ` Q \ Abrupt ` A" definition cvalid:: "[('s,'p,'f) body,('s,'p) quadruple set,'f set, 's assn,('s,'p,'f) com,'s assn,'s assn] =>bool" ("_,_\\<^bsub>'/_\<^esub>/ _ _ _,_" [61,60,60,1000, 20, 1000,1000] 60) where "\,\\\<^bsub>/F\<^esub> P c Q,A \ (\(P,p,Q,A)\\. \\\<^bsub>/F\<^esub> P (Call p) Q,A) \ \ \\<^bsub>/F\<^esub> P c Q,A" definition nvalid :: "[('s,'p,'f) body,nat,'f set, 's assn,('s,'p,'f) com,'s assn,'s assn] => bool" ("_\_:\<^bsub>'/_\<^esub>/ _ _ _,_" [61,60,60,1000, 20, 1000,1000] 60) where "\\n:\<^bsub>/F\<^esub> P c Q,A \ \s t. \\\c,s \ =n\ t \ s \ Normal ` P \ t \ Fault ` F \ t \ Normal ` Q \ Abrupt ` A" definition cnvalid:: "[('s,'p,'f) body,('s,'p) quadruple set,nat,'f set, 's assn,('s,'p,'f) com,'s assn,'s assn] \ bool" ("_,_\_:\<^bsub>'/_\<^esub>/ _ _ _,_" [61,60,60,60,1000, 20, 1000,1000] 60) where "\,\\n:\<^bsub>/F\<^esub> P c Q,A \ (\(P,p,Q,A)\\. \\n:\<^bsub>/F\<^esub> P (Call p) Q,A) \ \ \n:\<^bsub>/F\<^esub> P c Q,A" notation (ASCII) valid ("_|='/_/ _ _ _,_" [61,60,1000, 20, 1000,1000] 60) and cvalid ("_,_|='/_/ _ _ _,_" [61,60,60,1000, 20, 1000,1000] 60) and nvalid ("_|=_:'/_/ _ _ _,_" [61,60,60,1000, 20, 1000,1000] 60) and cnvalid ("_,_|=_:'/_/ _ _ _,_" [61,60,60,60,1000, 20, 1000,1000] 60) subsection \Properties of Validity\ lemma valid_iff_nvalid: "\\\<^bsub>/F\<^esub> P c Q,A = (\n. \\n:\<^bsub>/F\<^esub> P c Q,A)" apply (simp only: valid_def nvalid_def exec_iff_execn ) apply (blast dest: exec_final_notin_to_execn) done lemma cnvalid_to_cvalid: "(\n. \,\\n:\<^bsub>/F\<^esub> P c Q,A) \ \,\\\<^bsub>/F\<^esub> P c Q,A" apply (unfold cvalid_def cnvalid_def valid_iff_nvalid [THEN eq_reflection]) apply fast done lemma nvalidI: "\\s t. \\\\c,Normal s \ =n\ t;s \ P; t\ Fault ` F\ \ t \ Normal ` Q \ Abrupt ` A\ \ \\n:\<^bsub>/F\<^esub> P c Q,A" by (auto simp add: nvalid_def) lemma validI: "\\s t. \\\\c,Normal s \ \ t;s \ P; t\Fault ` F\ \ t \ Normal ` Q \ Abrupt ` A\ \ \\\<^bsub>/F\<^esub> P c Q,A" by (auto simp add: valid_def) lemma cvalidI: "\\s t. \\(P,p,Q,A)\\. \\\<^bsub>/F\<^esub> P (Call p) Q,A;\\\c,Normal s\ \ t;s \ P;t\Fault ` F\ \ t \ Normal ` Q \ Abrupt ` A\ \ \,\\\<^bsub>/F\<^esub> P c Q,A" by (auto simp add: cvalid_def valid_def) lemma cvalidD: "\\,\\\<^bsub>/F\<^esub> P c Q,A;\(P,p,Q,A)\\. \\\<^bsub>/F\<^esub> P (Call p) Q,A;\\\c,Normal s\ \ t;s \ P;t\Fault ` F\ \ t \ Normal ` Q \ Abrupt ` A" by (auto simp add: cvalid_def valid_def) lemma cnvalidI: "\\s t. \\(P,p,Q,A)\\. \\n:\<^bsub>/F\<^esub> P (Call p) Q,A; \\\c,Normal s \ =n\ t;s \ P;t\Fault ` F\ \ t \ Normal ` Q \ Abrupt ` A\ \ \,\\n:\<^bsub>/F\<^esub> P c Q,A" by (auto simp add: cnvalid_def nvalid_def) lemma cnvalidD: "\\,\\n:\<^bsub>/F\<^esub> P c Q,A;\(P,p,Q,A)\\. \\n:\<^bsub>/F\<^esub> P (Call p) Q,A; \\\c,Normal s \ =n\ t;s \ P; t\Fault ` F\ \ t \ Normal ` Q \ Abrupt ` A" by (auto simp add: cnvalid_def nvalid_def) lemma nvalid_augment_Faults: assumes validn:"\\n:\<^bsub>/F\<^esub> P c Q,A" assumes F': "F \ F'" shows "\\n:\<^bsub>/F'\<^esub> P c Q,A" proof (rule nvalidI) fix s t assume exec: "\\\c,Normal s \ =n\ t" assume P: "s \ P" assume F: "t \ Fault ` F'" with F' have "t \ Fault ` F" by blast with exec P validn show "t \ Normal ` Q \ Abrupt ` A" by (auto simp add: nvalid_def) qed lemma valid_augment_Faults: assumes validn:"\\\<^bsub>/F\<^esub> P c Q,A" assumes F': "F \ F'" shows "\\\<^bsub>/F'\<^esub> P c Q,A" proof (rule validI) fix s t assume exec: "\\\c,Normal s \ \ t" assume P: "s \ P" assume F: "t \ Fault ` F'" with F' have "t \ Fault ` F" by blast with exec P validn show "t \ Normal ` Q \ Abrupt ` A" by (auto simp add: valid_def) qed lemma nvalid_to_nvalid_strip: assumes validn:"\\n:\<^bsub>/F\<^esub> P c Q,A" assumes F': "F' \ -F" shows "strip F' \\n:\<^bsub>/F\<^esub> P c Q,A" proof (rule nvalidI) fix s t assume exec_strip: "strip F' \\\c,Normal s \ =n\ t" assume P: "s \ P" assume F: "t \ Fault ` F" from exec_strip obtain t' where exec: "\\\c,Normal s \ =n\ t'" and t': "t' \ Fault ` (-F') \ t'=t" "\ isFault t' \ t'=t" by (blast dest: execn_strip_to_execn) show "t \ Normal ` Q \ Abrupt ` A" proof (cases "t' \ Fault ` F") case True with t' F F' have False by blast thus ?thesis .. next case False with exec P validn have *: "t' \ Normal ` Q \ Abrupt ` A" by (auto simp add: nvalid_def) with t' have "t'=t" by auto with * show ?thesis by simp qed qed lemma valid_to_valid_strip: assumes valid:"\\\<^bsub>/F\<^esub> P c Q,A" assumes F': "F' \ -F" shows "strip F' \\\<^bsub>/F\<^esub> P c Q,A" proof (rule validI) fix s t assume exec_strip: "strip F' \\\c,Normal s \ \ t" assume P: "s \ P" assume F: "t \ Fault ` F" from exec_strip obtain t' where exec: "\\\c,Normal s \ \ t'" and t': "t' \ Fault ` (-F') \ t'=t" "\ isFault t' \ t'=t" by (blast dest: exec_strip_to_exec) show "t \ Normal ` Q \ Abrupt ` A" proof (cases "t' \ Fault ` F") case True with t' F F' have False by blast thus ?thesis .. next case False with exec P valid have *: "t' \ Normal ` Q \ Abrupt ` A" by (auto simp add: valid_def) with t' have "t'=t" by auto with * show ?thesis by simp qed qed subsection \The Hoare Rules: \\,\\\<^bsub>/F\<^esub> P c Q,A\\ lemma mono_WeakenContext: "A \ B \ (\(P, c, Q, A'). (\, \, F, P, c, Q, A') \ A) x \ (\(P, c, Q, A'). (\, \, F, P, c, Q, A') \ B) x" apply blast done inductive "hoarep"::"[('s,'p,'f) body,('s,'p) quadruple set,'f set, 's assn,('s,'p,'f) com, 's assn,'s assn] => bool" ("(3_,_/\\<^bsub>'/_ \<^esub>(_/ (_)/ _,/_))" [60,60,60,1000,20,1000,1000]60) for \::"('s,'p,'f) body" where Skip: "\,\\\<^bsub>/F\<^esub> Q Skip Q,A" | Basic: "\,\\\<^bsub>/F\<^esub> {s. f s \ Q} (Basic f) Q,A" | Spec: "\,\\\<^bsub>/F\<^esub> {s. (\t. (s,t) \ r \ t \ Q) \ (\t. (s,t) \ r)} (Spec r) Q,A" | Seq: "\\,\\\<^bsub>/F\<^esub> P c\<^sub>1 R,A; \,\\\<^bsub>/F\<^esub> R c\<^sub>2 Q,A\ \ \,\\\<^bsub>/F\<^esub> P (Seq c\<^sub>1 c\<^sub>2) Q,A" | Cond: "\\,\\\<^bsub>/F\<^esub> (P \ b) c\<^sub>1 Q,A; \,\\\<^bsub>/F\<^esub> (P \ - b) c\<^sub>2 Q,A\ \ \,\\\<^bsub>/F\<^esub> P (Cond b c\<^sub>1 c\<^sub>2) Q,A" | While: "\,\\\<^bsub>/F\<^esub> (P \ b) c P,A \ \,\\\<^bsub>/F\<^esub> P (While b c) (P \ - b),A" | Guard: "\,\\\<^bsub>/F\<^esub> (g \ P) c Q,A \ \,\\\<^bsub>/F\<^esub> (g \ P) (Guard f g c) Q,A" | Guarantee: "\f \ F; \,\\\<^bsub>/F\<^esub> (g \ P) c Q,A\ \ \,\\\<^bsub>/F\<^esub> P (Guard f g c) Q,A" | CallRec: "\(P,p,Q,A) \ Specs; \(P,p,Q,A) \ Specs. p \ dom \ \ \,\\Specs\\<^bsub>/F\<^esub> P (the (\ p)) Q,A \ \ \,\\\<^bsub>/F\<^esub> P (Call p) Q,A" | DynCom: "\s \ P. \,\\\<^bsub>/F\<^esub> P (c s) Q,A \ \,\\\<^bsub>/F\<^esub> P (DynCom c) Q,A" | Throw: "\,\\\<^bsub>/F\<^esub> A Throw Q,A" | Catch: "\\,\\\<^bsub>/F\<^esub> P c\<^sub>1 Q,R; \,\\\<^bsub>/F\<^esub> R c\<^sub>2 Q,A\ \ \,\\\<^bsub>/F\<^esub> P Catch c\<^sub>1 c\<^sub>2 Q,A" | Conseq: "\s \ P. \P' Q' A'. \,\\\<^bsub>/F\<^esub> P' c Q',A' \ s \ P' \ Q' \ Q \ A' \ A \ \,\\\<^bsub>/F\<^esub> P c Q,A" | Asm: "\(P,p,Q,A) \ \\ \ \,\\\<^bsub>/F\<^esub> P (Call p) Q,A" | ExFalso: "\\n. \,\\n:\<^bsub>/F\<^esub> P c Q,A; \ \\\<^bsub>/F\<^esub> P c Q,A\ \ \,\\\<^bsub>/F\<^esub> P c Q,A" \ \This is a hack rule that enables us to derive completeness for an arbitrary context \\\, from completeness for an empty context.\ text \Does not work, because of rule ExFalso, the context \\\ is to blame. A weaker version with empty context can be derived from soundness and completeness later on.\ lemma hoare_strip_\: assumes deriv: "\,\\\<^bsub>/F\<^esub> P p Q,A" shows "strip (-F) \,\\\<^bsub>/F\<^esub> P p Q,A" using deriv proof induct case Skip thus ?case by (iprover intro: hoarep.Skip) next case Basic thus ?case by (iprover intro: hoarep.Basic) next case Spec thus ?case by (iprover intro: hoarep.Spec) next case Seq thus ?case by (iprover intro: hoarep.Seq) next case Cond thus ?case by (iprover intro: hoarep.Cond) next case While thus ?case by (iprover intro: hoarep.While) next case Guard thus ?case by (iprover intro: hoarep.Guard) (*next case CallSpec thus ?case by (iprover intro: hoarep.CallSpec) next case (CallRec A Abr Abr' Init P Post Pre Procs Q R Result Return Z \ \ init p result return ) from CallRec.hyps have "\p\Procs. \Z. (strip \),\ \ (\\<^bsub>p\Procs\<^esub> \\<^bsub>Z\<^esub> {(Pre p Z, Call (Init p) p (Return p) (Result p), Post p Z, Abr p Z)})\ (Pre p Z) (the (\ p)) (R p Z),(Abr' p Z)" by blast hence "\p\Procs. \Z. (strip \),\ \ (\\<^bsub>p\Procs\<^esub> \\<^bsub>Z\<^esub> {(Pre p Z, Call (Init p) p (Return p) (Result p), Post p Z, Abr p Z)})\ (Pre p Z) (the ((strip \) p)) (R p Z),(Abr' p Z)" by (auto intro: hoarep.StripI) then show ?case apply - apply (rule hoarep.CallRec) apply (assumption | simp only:dom_strip)+ done*) next case DynCom thus ?case by - (rule hoarep.DynCom,best elim!: ballE exE) next case Throw thus ?case by (iprover intro: hoarep.Throw) next case Catch thus ?case by (iprover intro: hoarep.Catch) (*next case CONSEQ thus ?case apply (auto intro: hoarep.CONSEQ)*) next case Asm thus ?case by (iprover intro: hoarep.Asm) next case ExFalso thus ?case oops lemma hoare_augment_context: assumes deriv: "\,\\\<^bsub>/F\<^esub> P p Q,A" shows "\\'. \ \ \' \ \,\'\\<^bsub>/F\<^esub> P p Q,A" using deriv proof (induct) case CallRec case (CallRec P p Q A Specs \ F \') from CallRec.prems have "\\Specs \ \'\Specs" by blast with CallRec.hyps (2) have "\(P,p,Q,A)\Specs. p \ dom \ \ \,\'\Specs \\<^bsub>/F\<^esub> P (the (\ p)) Q,A" by fastforce with CallRec show ?case by - (rule hoarep.CallRec) next case DynCom thus ?case by (blast intro: hoarep.DynCom) next case (Conseq P \ F c Q A \') from Conseq have "\s \ P. (\P' Q' A'. \,\' \\<^bsub>/F\<^esub> P' c Q',A' \ s \ P' \ Q' \ Q \ A' \ A)" by blast with Conseq show ?case by - (rule hoarep.Conseq) next case (ExFalso \ F P c Q A \') have valid_ctxt: "\n. \,\\n:\<^bsub>/F\<^esub> P c Q,A" "\ \ \'" by fact+ hence "\n. \,\'\n:\<^bsub>/F\<^esub> P c Q,A" by (simp add: cnvalid_def) blast moreover have invalid: "\ \\\<^bsub>/F\<^esub> P c Q,A" by fact ultimately show ?case by (rule hoarep.ExFalso) qed (blast intro: hoarep.intros)+ subsection \Some Derived Rules\ lemma Conseq': "\s. s \ P \ (\P' Q' A'. (\ Z. \,\\\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z)) \ (\Z. s \ P' Z \ (Q' Z \ Q) \ (A' Z \ A))) \ \,\\\<^bsub>/F\<^esub> P c Q,A" apply (rule Conseq) apply (rule ballI) apply (erule_tac x=s in allE) apply (clarify) apply (rule_tac x="P' Z" in exI) apply (rule_tac x="Q' Z" in exI) apply (rule_tac x="A' Z" in exI) apply blast done lemma conseq:"\\Z. \,\ \\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z); \s. s \ P \ (\ Z. s\P' Z \ (Q' Z \ Q) \ (A' Z \ A))\ \ \,\\\<^bsub>/F\<^esub> P c Q,A" by (rule Conseq) blast theorem conseqPrePost [trans]: "\,\\\<^bsub>/F\<^esub> P' c Q',A' \ P \ P' \ Q' \ Q \ A' \ A \ \,\\\<^bsub>/F\<^esub> P c Q,A" by (rule conseq [where ?P'="\Z. P'" and ?Q'="\Z. Q'"]) auto lemma conseqPre [trans]: "\,\\\<^bsub>/F\<^esub> P' c Q,A \ P \ P' \ \,\\\<^bsub>/F\<^esub> P c Q,A" by (rule conseq) auto lemma conseqPost [trans]: "\,\\\<^bsub>/F\<^esub> P c Q',A' \ Q' \ Q \ A' \ A \ \,\\\<^bsub>/F\<^esub> P c Q,A" by (rule conseq) auto lemma CallRec': "\p\Procs; Procs \ dom \; \p\Procs. \Z. \,\ \ (\p\Procs. \Z. {((P p Z),p,Q p Z,A p Z)}) \\<^bsub>/F\<^esub> (P p Z) (the (\ p)) (Q p Z),(A p Z)\ \ \,\\\<^bsub>/F\<^esub> (P p Z) (Call p) (Q p Z),(A p Z)" apply (rule CallRec [where Specs="\p\Procs. \Z. {((P p Z),p,Q p Z,A p Z)}"]) apply blast apply blast done end diff --git a/thys/Simpl/HoarePartialProps.thy b/thys/Simpl/HoarePartialProps.thy --- a/thys/Simpl/HoarePartialProps.thy +++ b/thys/Simpl/HoarePartialProps.thy @@ -1,2399 +1,2538 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: HoarePartialProps.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. +Copyright (c) 2022 Apple Inc. All rights reserved. -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA *) section \Properties of Partial Correctness Hoare Logic\ theory HoarePartialProps imports HoarePartialDef begin subsection \Soundness\ lemma hoare_cnvalid: assumes hoare: "\,\\\<^bsub>/F\<^esub> P c Q,A" shows "\n. \,\\n:\<^bsub>/F\<^esub> P c Q,A" using hoare proof (induct) case (Skip \ F P A) show "\,\ \n:\<^bsub>/F\<^esub> P Skip P,A" proof (rule cnvalidI) fix s t assume "\\\Skip,Normal s\ =n\ t" "s \ P" thus "t \ Normal ` P \ Abrupt ` A" by cases auto qed next case (Basic \ F f P A) show "\,\ \n:\<^bsub>/F\<^esub> {s. f s \ P} (Basic f) P,A" proof (rule cnvalidI) fix s t assume "\\\Basic f,Normal s\ =n\ t" "s \ {s. f s \ P}" thus "t \ Normal ` P \ Abrupt ` A" by cases auto qed next case (Spec \ F r Q A) show "\,\\n:\<^bsub>/F\<^esub> {s. (\t. (s, t) \ r \ t \ Q) \ (\t. (s, t) \ r)} Spec r Q,A" proof (rule cnvalidI) fix s t assume exec: "\\\Spec r,Normal s\ =n\ t" assume P: "s \ {s. (\t. (s, t) \ r \ t \ Q) \ (\t. (s, t) \ r)}" from exec P show "t \ Normal ` Q \ Abrupt ` A" by cases auto qed next case (Seq \ F P c1 R A c2 Q) have valid_c1: "\n. \,\ \n:\<^bsub>/F\<^esub> P c1 R,A" by fact have valid_c2: "\n. \,\ \n:\<^bsub>/F\<^esub> R c2 Q,A" by fact show "\,\ \n:\<^bsub>/F\<^esub> P Seq c1 c2 Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\Seq c1 c2,Normal s\ =n\ t" assume t_notin_F: "t \ Fault ` F" assume P: "s \ P" from exec P obtain r where exec_c1: "\\\c1,Normal s\ =n\ r" and exec_c2: "\\\c2,r\ =n\ t" by cases auto with t_notin_F have "r \ Fault ` F" by (auto dest: execn_Fault_end) with valid_c1 ctxt exec_c1 P have r: "r\Normal ` R \ Abrupt ` A" by (rule cnvalidD) show "t\Normal ` Q \ Abrupt ` A" proof (cases r) case (Normal r') with exec_c2 r show "t\Normal ` Q \ Abrupt ` A" apply - apply (rule cnvalidD [OF valid_c2 ctxt _ _ t_notin_F]) apply auto done next case (Abrupt r') with exec_c2 have "t=Abrupt r'" by (auto elim: execn_elim_cases) with Abrupt r show ?thesis by auto next case Fault with r show ?thesis by blast next case Stuck with r show ?thesis by blast qed qed next case (Cond \ F P b c1 Q A c2) have valid_c1: "\n. \,\ \n:\<^bsub>/F\<^esub> (P \ b) c1 Q,A" by fact have valid_c2: "\n. \,\ \n:\<^bsub>/F\<^esub> (P \ - b) c2 Q,A" by fact show "\,\ \n:\<^bsub>/F\<^esub> P Cond b c1 c2 Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\Cond b c1 c2,Normal s\ =n\ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" show "t \ Normal ` Q \ Abrupt ` A" proof (cases "s\b") case True with exec have "\\\c1,Normal s\ =n\ t" by cases auto with P True show ?thesis by - (rule cnvalidD [OF valid_c1 ctxt _ _ t_notin_F],auto) next case False with exec P have "\\\c2,Normal s\ =n\ t" by cases auto with P False show ?thesis by - (rule cnvalidD [OF valid_c2 ctxt _ _ t_notin_F],auto) qed qed next case (While \ F P b c A n) have valid_c: "\n. \,\ \n:\<^bsub>/F\<^esub> (P \ b) c P,A" by fact show "\,\ \n:\<^bsub>/F\<^esub> P While b c (P \ - b),A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\While b c,Normal s\ =n\ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" show "t \ Normal ` (P \ - b) \ Abrupt ` A" proof (cases "s \ b") case True { fix d::"('b,'a,'c) com" fix s t assume exec: "\\\d,s\ =n\ t" assume d: "d=While b c" assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" from exec d ctxt have "\s \ Normal ` P; t \ Fault ` F\ \ t \ Normal ` (P \ - b) \ Abrupt`A" proof (induct) case (WhileTrue s b' c' n r t) have t_notin_F: "t \ Fault ` F" by fact have eqs: "While b' c' = While b c" by fact note valid_c moreover have ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" by fact moreover from WhileTrue obtain "\\\c,Normal s\ =n\ r" and "\\\While b c,r\ =n\ t" and "Normal s \ Normal `(P \ b)" by auto moreover with t_notin_F have "r \ Fault ` F" by (auto dest: execn_Fault_end) ultimately have r: "r \ Normal ` P \ Abrupt ` A" by - (rule cnvalidD,auto) from this _ ctxt show "t \ Normal ` (P \ - b) \ Abrupt ` A " proof (cases r) case (Normal r') with r ctxt eqs t_notin_F show ?thesis by - (rule WhileTrue.hyps,auto) next case (Abrupt r') have "\\\While b' c',r\ =n\ t" by fact with Abrupt have "t=r" by (auto dest: execn_Abrupt_end) with r Abrupt show ?thesis by blast next case Fault with r show ?thesis by blast next case Stuck with r show ?thesis by blast qed qed auto } with exec ctxt P t_notin_F show ?thesis by auto next case False with exec P have "t=Normal s" by cases auto with P False show ?thesis by auto qed qed next case (Guard \ F g P c Q A f) have valid_c: "\n. \,\ \n:\<^bsub>/F\<^esub> (g \ P) c Q,A" by fact show "\,\ \n:\<^bsub>/F\<^esub> (g \ P) Guard f g c Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\Guard f g c,Normal s\ =n\ t" assume t_notin_F: "t \ Fault ` F" assume P:"s \ (g \ P)" from exec P have "\\\c,Normal s\ =n\ t" by cases auto from valid_c ctxt this P t_notin_F show "t \ Normal ` Q \ Abrupt ` A" by (rule cnvalidD) qed next case (Guarantee f F \ g P c Q A) have valid_c: "\n. \,\ \n:\<^bsub>/F\<^esub> (g \ P) c Q,A" by fact have f_F: "f \ F" by fact show "\,\ \n:\<^bsub>/F\<^esub> P Guard f g c Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\Guard f g c,Normal s\ =n\ t" assume t_notin_F: "t \ Fault ` F" assume P:"s \ P" from exec f_F t_notin_F have g: "s \ g" by cases auto with P have P': "s \ g \ P" by blast from exec P g have "\\\c,Normal s\ =n\ t" by cases auto from valid_c ctxt this P' t_notin_F show "t \ Normal ` Q \ Abrupt ` A" by (rule cnvalidD) qed next case (CallRec P p Q A Specs \ F) have p: "(P,p,Q,A) \ Specs" by fact have valid_body: "\(P,p,Q,A) \ Specs. p \ dom \ \ (\n. \,\ \ Specs \n:\<^bsub>/F\<^esub> P (the (\ p)) Q,A)" using CallRec.hyps by blast show "\,\\n:\<^bsub>/F\<^esub> P Call p Q,A" proof - { fix n have "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A \ \(P,p,Q,A) \Specs. \\n:\<^bsub>/F\<^esub> P (Call p) Q,A" proof (induct n) case 0 show "\(P,p,Q,A) \Specs. \\0:\<^bsub>/F\<^esub> P (Call p) Q,A" by (fastforce elim!: execn_elim_cases simp add: nvalid_def) next case (Suc m) have hyp: "\(P, p, Q, A)\\. \ \m:\<^bsub>/F\<^esub> P (Call p) Q,A \ \(P,p,Q,A) \Specs. \\m:\<^bsub>/F\<^esub> P (Call p) Q,A" by fact have "\(P, p, Q, A)\\. \ \Suc m:\<^bsub>/F\<^esub> P (Call p) Q,A" by fact hence ctxt_m: "\(P, p, Q, A)\\. \ \m:\<^bsub>/F\<^esub> P (Call p) Q,A" by (fastforce simp add: nvalid_def intro: execn_Suc) hence valid_Proc: "\(P,p,Q,A) \Specs. \\m:\<^bsub>/F\<^esub> P (Call p) Q,A" by (rule hyp) let ?\'= "\ \ Specs" from valid_Proc ctxt_m have "\(P, p, Q, A)\?\'. \ \m:\<^bsub>/F\<^esub> P (Call p) Q,A" by fastforce with valid_body have valid_body_m: "\(P,p,Q,A) \Specs. \n. \ \m:\<^bsub>/F\<^esub> P (the (\ p)) Q,A" by (fastforce simp add: cnvalid_def) show "\(P,p,Q,A) \Specs. \ \Suc m:\<^bsub>/F\<^esub> P (Call p) Q,A" proof (clarify) fix P p Q A assume p: "(P,p,Q,A) \ Specs" show "\ \Suc m:\<^bsub>/F\<^esub> P (Call p) Q,A" proof (rule nvalidI) fix s t assume exec_call: "\\\Call p,Normal s\ =Suc m\ t" assume Pre: "s \ P" assume t_notin_F: "t \ Fault ` F" from exec_call show "t \ Normal ` Q \ Abrupt ` A" proof (cases) fix bdy m' assume m: "Suc m = Suc m'" assume bdy: "\ p = Some bdy" assume exec_body: "\\\bdy,Normal s\ =m'\ t" from Pre valid_body_m exec_body bdy m p t_notin_F show ?thesis by (fastforce simp add: nvalid_def) next assume "\ p = None" with valid_body p have False by auto thus ?thesis .. qed qed qed qed } with p show ?thesis by (fastforce simp add: cnvalid_def) qed next case (DynCom P \ F c Q A) hence valid_c: "\s\P. (\n. \,\\n:\<^bsub>/F\<^esub> P (c s) Q,A)" by auto show "\,\\n:\<^bsub>/F\<^esub> P DynCom c Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\DynCom c,Normal s\ =n\ t" assume P: "s \ P" assume t_notin_Fault: "t \ Fault ` F" from exec show "t \ Normal ` Q \ Abrupt ` A" proof (cases) assume "\\\c s,Normal s\ =n\ t" from cnvalidD [OF valid_c [rule_format, OF P] ctxt this P t_notin_Fault] show ?thesis . qed qed next case (Throw \ F A Q) show "\,\ \n:\<^bsub>/F\<^esub> A Throw Q,A" proof (rule cnvalidI) fix s t assume "\\\Throw,Normal s\ =n\ t" "s \ A" then show "t \ Normal ` Q \ Abrupt ` A" by cases simp qed next case (Catch \ F P c\<^sub>1 Q R c\<^sub>2 A) have valid_c1: "\n. \,\ \n:\<^bsub>/F\<^esub> P c\<^sub>1 Q,R" by fact have valid_c2: "\n. \,\ \n:\<^bsub>/F\<^esub> R c\<^sub>2 Q,A" by fact show "\,\ \n:\<^bsub>/F\<^esub> P Catch c\<^sub>1 c\<^sub>2 Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\Catch c\<^sub>1 c\<^sub>2,Normal s\ =n\ t" assume P: "s \ P" assume t_notin_Fault: "t \ Fault ` F" from exec show "t \ Normal ` Q \ Abrupt ` A" proof (cases) fix s' assume exec_c1: "\\\c\<^sub>1,Normal s\ =n\ Abrupt s'" assume exec_c2: "\\\c\<^sub>2,Normal s'\ =n\ t" from cnvalidD [OF valid_c1 ctxt exec_c1 P ] have "Abrupt s' \ Abrupt ` R" by auto with cnvalidD [OF valid_c2 ctxt _ _ t_notin_Fault] exec_c2 show ?thesis by fastforce next assume exec_c1: "\\\c\<^sub>1,Normal s\ =n\ t" assume notAbr: "\ isAbr t" from cnvalidD [OF valid_c1 ctxt exec_c1 P t_notin_Fault] have "t \ Normal ` Q \ Abrupt ` R" . with notAbr show ?thesis by auto qed qed next case (Conseq P \ F c Q A) hence adapt: "\s \ P. (\P' Q' A'. \,\ \n:\<^bsub>/F\<^esub> P' c Q',A' \ s \ P' \ Q' \ Q \ A' \ A)" by blast show "\,\ \n:\<^bsub>/F\<^esub> P c Q,A" proof (rule cnvalidI) fix s t assume ctxt:"\(P, p, Q, A)\\. \\n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\c,Normal s\ =n\ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" show "t \ Normal ` Q \ Abrupt ` A" proof - from P adapt obtain P' Q' A' Z where spec: "\,\\n:\<^bsub>/F\<^esub> P' c Q',A'" and P': "s \ P'" and strengthen: "Q' \ Q \ A' \ A" by auto from spec [rule_format] ctxt exec P' t_notin_F have "t \ Normal ` Q' \ Abrupt ` A'" by (rule cnvalidD) with strengthen show ?thesis by blast qed qed next case (Asm P p Q A \ F) have asm: "(P, p, Q, A) \ \" by fact show "\,\ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\Call p,Normal s\ =n\ t" from asm ctxt have "\ \n:\<^bsub>/F\<^esub> P Call p Q,A" by auto moreover assume "s \ P" "t \ Fault ` F" ultimately show "t \ Normal ` Q \ Abrupt ` A" using exec by (auto simp add: nvalid_def) qed next case ExFalso thus ?case by iprover qed theorem hoare_sound: "\,\\\<^bsub>/F\<^esub> P c Q,A \ \,\\\<^bsub>/F\<^esub> P c Q,A" by (iprover intro: cnvalid_to_cvalid hoare_cnvalid) subsection \Completeness\ lemma MGT_valid: "\\\<^bsub>/F\<^esub>{s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F))} c {t. \\\c,Normal Z\ \ Normal t}, {t. \\\c,Normal Z\ \ Abrupt t}" proof (rule validI) fix s t assume "\\\c,Normal s\ \ t" "s \ {s. s = Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F))}" "t \ Fault ` F" thus "t \ Normal ` {t. \\\c,Normal Z\ \ Normal t} \ Abrupt ` {t. \\\c,Normal Z\ \ Abrupt t}" by (cases t) (auto simp add: final_notin_def) qed text \The consequence rule where the existential @{term Z} is instantiated to @{term s}. Usefull in proof of \MGT_lemma\.\ lemma ConseqMGT: assumes modif: "\Z. \,\ \\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z)" assumes impl: "\s. s \ P \ s \ P' s \ (\t. t \ Q' s \ t \ Q) \ (\t. t \ A' s \ t \ A)" shows "\,\ \\<^bsub>/F\<^esub> P c Q,A" using impl by - (rule conseq [OF modif],blast) lemma Seq_NoFaultStuckD1: assumes noabort: "\\\Seq c1 c2,s\ \\({Stuck} \ Fault ` F)" shows "\\\c1,s\ \\({Stuck} \ Fault ` F)" proof (rule final_notinI) fix t assume exec_c1: "\\\c1,s\ \ t" show "t \ {Stuck} \ Fault ` F" proof assume "t \ {Stuck} \ Fault ` F" moreover { assume "t = Stuck" with exec_c1 have "\\\Seq c1 c2,s\ \ Stuck" by (auto intro: exec_Seq') with noabort have False by (auto simp add: final_notin_def) hence False .. } moreover { assume "t \ Fault ` F" then obtain f where t: "t=Fault f" and f: "f \ F" by auto from t exec_c1 have "\\\Seq c1 c2,s\ \ Fault f" by (auto intro: exec_Seq') with noabort f have False by (auto simp add: final_notin_def) hence False .. } ultimately show False by auto qed qed lemma Seq_NoFaultStuckD2: assumes noabort: "\\\Seq c1 c2,s\ \\({Stuck} \ Fault ` F)" shows "\t. \\\c1,s\ \ t \ t\ ({Stuck} \ Fault ` F) \ \\\c2,t\ \\({Stuck} \ Fault ` F)" using noabort by (auto simp add: final_notin_def intro: exec_Seq') lemma MGT_implies_complete: assumes MGT: "\Z. \,{}\\<^bsub>/F\<^esub> {s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F))} c {t. \\\c,Normal Z\ \ Normal t}, {t. \\\c,Normal Z\ \ Abrupt t}" assumes valid: "\ \\<^bsub>/F\<^esub> P c Q,A" shows "\,{} \\<^bsub>/F\<^esub> P c Q,A" using MGT apply (rule ConseqMGT) apply (insert valid) apply (auto simp add: valid_def intro!: final_notinI) done text \Equipped only with the classic consequence rule @{thm "conseqPrePost"} we can only derive this syntactically more involved version of completeness. But semantically it is equivalent to the "real" one (see below)\ lemma MGT_implies_complete': assumes MGT: "\Z. \,{}\\<^bsub>/F\<^esub> {s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F))} c {t. \\\c,Normal Z\ \ Normal t}, {t. \\\c,Normal Z\ \ Abrupt t}" assumes valid: "\ \\<^bsub>/F\<^esub> P c Q,A" shows "\,{} \\<^bsub>/F\<^esub> {s. s=Z \ s \ P} c {t. Z \ P \ t \ Q},{t. Z \ P \ t \ A}" using MGT [rule_format, of Z] apply (rule conseqPrePost) apply (insert valid) apply (fastforce simp add: valid_def final_notin_def) apply (fastforce simp add: valid_def) apply (fastforce simp add: valid_def) done text \Semantic equivalence of both kind of formulations\ lemma valid_involved_to_valid: assumes valid: "\Z. \\\<^bsub>/F\<^esub> {s. s=Z \ s \ P} c {t. Z \ P \ t \ Q},{t. Z \ P \ t \ A}" shows "\ \\<^bsub>/F\<^esub> P c Q,A" using valid apply (simp add: valid_def) apply clarsimp apply (erule_tac x="x" in allE) apply (erule_tac x="Normal x" in allE) apply (erule_tac x=t in allE) apply fastforce done text \The sophisticated consequence rule allow us to do this semantical transformation on the hoare-level, too. The magic is, that it allow us to choose the instance of @{term Z} under the assumption of an state @{term "s \ P"}\ lemma assumes deriv: "\Z. \,{}\\<^bsub>/F\<^esub> {s. s=Z \ s \ P} c {t. Z \ P \ t \ Q},{t. Z \ P \ t \ A}" shows "\,{} \\<^bsub>/F\<^esub> P c Q,A" apply (rule ConseqMGT [OF deriv]) apply auto done lemma valid_to_valid_involved: "\ \\<^bsub>/F\<^esub> P c Q,A \ \\\<^bsub>/F\<^esub> {s. s=Z \ s \ P} c {t. Z \ P \ t \ Q},{t. Z \ P \ t \ A}" by (simp add: valid_def Collect_conv_if) lemma assumes deriv: "\,{} \\<^bsub>/F\<^esub> P c Q,A" shows "\,{}\\<^bsub>/F\<^esub> {s. s=Z \ s \ P} c {t. Z \ P \ t \ Q},{t. Z \ P \ t \ A}" apply (rule conseqPrePost [OF deriv]) apply auto done lemma conseq_extract_state_indep_prop: assumes state_indep_prop:"\s \ P. R" assumes to_show: "R \ \,\\\<^bsub>/F\<^esub> P c Q,A" shows "\,\\\<^bsub>/F\<^esub> P c Q,A" apply (rule Conseq) apply (clarify) apply (rule_tac x="P" in exI) apply (rule_tac x="Q" in exI) apply (rule_tac x="A" in exI) using state_indep_prop to_show by blast lemma MGT_lemma: assumes MGT_Calls: "\p\dom \. \Z. \,\ \\<^bsub>/F\<^esub> {s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F))} (Call p) {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t}" shows "\Z. \,\\\<^bsub>/F\<^esub> {s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F))} c {t. \\\c,Normal Z\ \ Normal t},{t. \\\c,Normal Z\ \ Abrupt t}" proof (induct c) case Skip show "\,\\\<^bsub>/F\<^esub> {s. s = Z \ \\\Skip,Normal s\ \\({Stuck} \ Fault ` (-F))} Skip {t. \\\Skip,Normal Z\ \ Normal t},{t. \\\Skip,Normal Z\ \ Abrupt t}" by (rule hoarep.Skip [THEN conseqPre]) (auto elim: exec_elim_cases simp add: final_notin_def intro: exec.intros) next case (Basic f) show "\,\\\<^bsub>/F\<^esub> {s. s = Z \ \\\Basic f,Normal s\ \\({Stuck} \ Fault ` (-F))} Basic f {t. \\\Basic f,Normal Z\ \ Normal t}, {t. \\\Basic f,Normal Z\ \ Abrupt t}" by (rule hoarep.Basic [THEN conseqPre]) (auto elim: exec_elim_cases simp add: final_notin_def intro: exec.intros) next case (Spec r) show "\,\\\<^bsub>/F\<^esub> {s. s = Z \ \\\Spec r,Normal s\ \\({Stuck} \ Fault ` (-F))} Spec r {t. \\\Spec r,Normal Z\ \ Normal t}, {t. \\\Spec r,Normal Z\ \ Abrupt t}" apply (rule hoarep.Spec [THEN conseqPre]) apply (clarsimp simp add: final_notin_def) apply (case_tac "\t. (Z,t) \ r") apply (auto elim: exec_elim_cases simp add: final_notin_def intro: exec.intros) done next case (Seq c1 c2) have hyp_c1: "\Z. \,\\\<^bsub>/F\<^esub> {s. s=Z \ \\\c1,Normal s\ \\({Stuck} \ Fault ` (-F))} c1 {t. \\\c1,Normal Z\ \ Normal t}, {t. \\\c1,Normal Z\ \ Abrupt t}" using Seq.hyps by iprover have hyp_c2: "\Z. \,\\\<^bsub>/F\<^esub> {s. s=Z \ \\\c2,Normal s\ \\({Stuck} \ Fault ` (-F))} c2 {t. \\\c2,Normal Z\ \ Normal t}, {t. \\\c2,Normal Z\ \ Abrupt t}" using Seq.hyps by iprover from hyp_c1 have "\,\\\<^bsub>/F\<^esub> {s. s=Z \ \\\Seq c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F))} c1 {t. \\\c1,Normal Z\ \ Normal t \ \\\c2,Normal t\ \\({Stuck} \ Fault ` (-F))}, {t. \\\Seq c1 c2,Normal Z\ \ Abrupt t}" by (rule ConseqMGT) (auto dest: Seq_NoFaultStuckD1 [simplified] Seq_NoFaultStuckD2 [simplified] intro: exec.Seq) thus "\,\\\<^bsub>/F\<^esub> {s. s=Z \ \\\Seq c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F))} Seq c1 c2 {t. \\\Seq c1 c2,Normal Z\ \ Normal t}, {t. \\\Seq c1 c2,Normal Z\ \ Abrupt t}" proof (rule hoarep.Seq ) show "\,\\\<^bsub>/F\<^esub> {t. \\\c1,Normal Z\ \ Normal t \ \\\c2,Normal t\ \\({Stuck} \ Fault ` (-F))} c2 {t. \\\Seq c1 c2,Normal Z\ \ Normal t}, {t. \\\Seq c1 c2,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [OF hyp_c2],safe) fix r t assume "\\\c1,Normal Z\ \ Normal r" "\\\c2,Normal r\ \ Normal t" then show "\\\Seq c1 c2,Normal Z\ \ Normal t" by (iprover intro: exec.intros) next fix r t assume "\\\c1,Normal Z\ \ Normal r" "\\\c2,Normal r\ \ Abrupt t" then show "\\\Seq c1 c2,Normal Z\ \ Abrupt t" by (iprover intro: exec.intros) qed qed next case (Cond b c1 c2) have "\Z. \,\\\<^bsub>/F\<^esub>{s. s=Z \ \\\c1,Normal s\ \\({Stuck} \ Fault ` (-F))} c1 {t. \\\c1,Normal Z\ \ Normal t}, {t. \\\c1,Normal Z\ \ Abrupt t}" using Cond.hyps by iprover hence "\,\\\<^bsub>/F\<^esub> ({s. s=Z \ \\\Cond b c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F))}\b) c1 {t. \\\Cond b c1 c2,Normal Z\ \ Normal t}, {t. \\\Cond b c1 c2,Normal Z\ \ Abrupt t}" by (rule ConseqMGT) (fastforce intro: exec.CondTrue simp add: final_notin_def) moreover have "\Z. \,\\\<^bsub>/F\<^esub> {s. s=Z \ \\\c2,Normal s\ \\({Stuck} \ Fault ` (-F))} c2 {t. \\\c2,Normal Z\ \ Normal t}, {t. \\\c2,Normal Z\ \ Abrupt t}" using Cond.hyps by iprover hence "\,\\\<^bsub>/F\<^esub>({s. s=Z \ \\\Cond b c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F))}\-b) c2 {t. \\\Cond b c1 c2,Normal Z\ \ Normal t}, {t. \\\Cond b c1 c2,Normal Z\ \ Abrupt t}" by (rule ConseqMGT) (fastforce intro: exec.CondFalse simp add: final_notin_def) ultimately show "\,\\\<^bsub>/F\<^esub> {s. s=Z \ \\\Cond b c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F))} Cond b c1 c2 {t. \\\Cond b c1 c2,Normal Z\ \ Normal t}, {t. \\\Cond b c1 c2,Normal Z\ \ Abrupt t}" by (rule hoarep.Cond) next case (While b c) let ?unroll = "({(s,t). s\b \ \\\c,Normal s\ \ Normal t})\<^sup>*" let ?P' = "\Z. {t. (Z,t)\?unroll \ (\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u))}" let ?A' = "\Z. {t. \\\While b c,Normal Z\ \ Abrupt t}" show "\,\\\<^bsub>/F\<^esub> {s. s=Z \ \\\While b c,Normal s\ \\({Stuck} \ Fault ` (-F))} While b c {t. \\\While b c,Normal Z\ \ Normal t}, {t. \\\While b c,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [where ?P'="?P'" and ?Q'="\Z. ?P' Z \ - b" and ?A'="?A'"]) show "\Z. \,\\\<^bsub>/F\<^esub> (?P' Z) (While b c) (?P' Z \ - b),(?A' Z)" proof (rule allI, rule hoarep.While) fix Z from While have "\Z. \,\\\<^bsub>/F\<^esub> {s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F))} c {t. \\\c,Normal Z\ \ Normal t}, {t. \\\c,Normal Z\ \ Abrupt t}" by iprover then show "\,\\\<^bsub>/F\<^esub> (?P' Z \ b) c (?P' Z),(?A' Z)" proof (rule ConseqMGT) fix s assume "s\ {t. (Z, t) \ ?unroll \ (\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u))} \ b" then obtain Z_s_unroll: "(Z,s) \ ?unroll" and noabort:"\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)" and s_in_b: "s\b" by blast show "s \ {t. t = s \ \\\c,Normal t\ \\({Stuck} \ Fault ` (-F))} \ (\t. t \ {t. \\\c,Normal s\ \ Normal t} \ t \ {t. (Z, t) \ ?unroll \ (\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u))}) \ (\t. t \ {t. \\\c,Normal s\ \ Abrupt t} \ t \ {t. \\\While b c,Normal Z\ \ Abrupt t})" (is "?C1 \ ?C2 \ ?C3") proof (intro conjI) from Z_s_unroll noabort s_in_b show ?C1 by blast next { fix t assume s_t: "\\\c,Normal s\ \ Normal t" moreover from Z_s_unroll s_t s_in_b have "(Z, t) \ ?unroll" by (blast intro: rtrancl_into_rtrancl) moreover note noabort ultimately have "(Z, t) \ ?unroll \ (\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u))" by iprover } then show ?C2 by blast next { fix t assume s_t: "\\\c,Normal s\ \ Abrupt t" from Z_s_unroll noabort s_t s_in_b have "\\\While b c,Normal Z\ \ Abrupt t" by blast } thus ?C3 by simp qed qed qed next fix s assume P: "s \ {s. s=Z \ \\\While b c,Normal s\ \\({Stuck} \ Fault ` (-F))}" hence WhileNoFault: "\\\While b c,Normal Z\ \\({Stuck} \ Fault ` (-F))" by auto show "s \ ?P' s \ (\t. t\(?P' s \ - b)\ t\{t. \\\While b c,Normal Z\ \ Normal t})\ (\t. t\?A' s \ t\?A' Z)" proof (intro conjI) { fix e assume "(Z,e) \ ?unroll" "e \ b" from this WhileNoFault have "\\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)" (is "?Prop Z e") proof (induct rule: converse_rtrancl_induct [consumes 1]) assume e_in_b: "e \ b" assume WhileNoFault: "\\\While b c,Normal e\ \\({Stuck} \ Fault ` (-F))" with e_in_b WhileNoFault have cNoFault: "\\\c,Normal e\ \\({Stuck} \ Fault ` (-F))" by (auto simp add: final_notin_def intro: exec.intros) moreover { fix u assume "\\\c,Normal e\ \ Abrupt u" with e_in_b have "\\\While b c,Normal e\ \ Abrupt u" by (blast intro: exec.intros) } ultimately show "?Prop e e" by iprover next fix Z r assume e_in_b: "e\b" assume WhileNoFault: "\\\While b c,Normal Z\ \\({Stuck} \ Fault ` (-F))" assume hyp: "\e\b;\\\While b c,Normal r\ \\({Stuck} \ Fault ` (-F))\ \ ?Prop r e" assume Z_r: "(Z, r) \ {(Z, r). Z \ b \ \\\c,Normal Z\ \ Normal r}" with WhileNoFault have "\\\While b c,Normal r\ \\({Stuck} \ Fault ` (-F))" by (auto simp add: final_notin_def intro: exec.intros) from hyp [OF e_in_b this] obtain cNoFault: "\\\c,Normal e\ \\({Stuck} \ Fault ` (-F))" and Abrupt_r: "\u. \\\c,Normal e\ \ Abrupt u \ \\\While b c,Normal r\ \ Abrupt u" by simp { fix u assume "\\\c,Normal e\ \ Abrupt u" with Abrupt_r have "\\\While b c,Normal r\ \ Abrupt u" by simp moreover from Z_r obtain "Z \ b" "\\\c,Normal Z\ \ Normal r" by simp ultimately have "\\\While b c,Normal Z\ \ Abrupt u" by (blast intro: exec.intros) } with cNoFault show "?Prop Z e" by iprover qed } with P show "s \ ?P' s" by blast next { fix t assume "termination": "t \ b" assume "(Z, t) \ ?unroll" hence "\\\While b c,Normal Z\ \ Normal t" proof (induct rule: converse_rtrancl_induct [consumes 1]) from "termination" show "\\\While b c,Normal t\ \ Normal t" by (blast intro: exec.WhileFalse) next fix Z r assume first_body: "(Z, r) \ {(s, t). s \ b \ \\\c,Normal s\ \ Normal t}" assume "(r, t) \ ?unroll" assume rest_loop: "\\\While b c, Normal r\ \ Normal t" show "\\\While b c,Normal Z\ \ Normal t" proof - from first_body obtain "Z \ b" "\\\c,Normal Z\ \ Normal r" by fast moreover from rest_loop have "\\\While b c,Normal r\ \ Normal t" by fast ultimately show "\\\While b c,Normal Z\ \ Normal t" by (rule exec.WhileTrue) qed qed } with P show "(\t. t\(?P' s \ - b) \t\{t. \\\While b c,Normal Z\ \ Normal t})" by blast next from P show "\t. t\?A' s \ t\?A' Z" by simp qed qed next case (Call p) let ?P = "{s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F))}" from noStuck_Call have "\s \ ?P. p \ dom \" by (fastforce simp add: final_notin_def ) then show "\,\\\<^bsub>/F\<^esub> ?P (Call p) {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t}" proof (rule conseq_extract_state_indep_prop) assume p_definied: "p \ dom \" with MGT_Calls show "\,\\\<^bsub>/F\<^esub>{s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F))} (Call p) {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t}" by (auto) qed next case (DynCom c) have hyp: "\s'. \Z. \,\\\<^bsub>/F\<^esub>{s. s = Z \ \\\c s',Normal s\ \\({Stuck} \ Fault ` (-F))} c s' {t. \\\c s',Normal Z\ \ Normal t},{t. \\\c s',Normal Z\ \ Abrupt t}" using DynCom by simp have hyp': "\,\\\<^bsub>/F\<^esub>{s. s = Z \ \\\DynCom c,Normal s\ \\({Stuck} \ Fault ` (-F))} c Z {t. \\\DynCom c,Normal Z\ \ Normal t},{t. \\\DynCom c,Normal Z\ \ Abrupt t}" by (rule ConseqMGT [OF hyp]) (fastforce simp add: final_notin_def intro: exec.intros) show "\,\\\<^bsub>/F\<^esub>{s. s = Z \ \\\DynCom c,Normal s\ \\({Stuck} \ Fault ` (-F))} DynCom c {t. \\\DynCom c,Normal Z\ \ Normal t}, {t. \\\DynCom c,Normal Z\ \ Abrupt t}" apply (rule hoarep.DynCom) apply (clarsimp) apply (rule hyp' [simplified]) done next case (Guard f g c) have hyp_c: "\Z. \,\\\<^bsub>/F\<^esub> {s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F))} c {t. \\\c,Normal Z\ \ Normal t}, {t. \\\c,Normal Z\ \ Abrupt t}" using Guard by iprover show ?case proof (cases "f \ F") case True from hyp_c have "\,\\\<^bsub>/F \<^esub>(g \ {s. s = Z \ \\\Guard f g c,Normal s\ \\({Stuck} \ Fault ` (- F))}) c {t. \\\Guard f g c,Normal Z\ \ Normal t}, {t. \\\Guard f g c,Normal Z\ \ Abrupt t}" apply (rule ConseqMGT) apply (insert True) apply (auto simp add: final_notin_def intro: exec.intros) done from True this show ?thesis by (rule conseqPre [OF Guarantee]) auto next case False from hyp_c have "\,\\\<^bsub>/F\<^esub> (g \ {s. s=Z \ \\\Guard f g c,Normal s\ \\({Stuck} \ Fault ` (-F))}) c {t. \\\Guard f g c,Normal Z\ \ Normal t}, {t. \\\Guard f g c,Normal Z\ \ Abrupt t}" apply (rule ConseqMGT) apply clarify apply (frule Guard_noFaultStuckD [OF _ False]) apply (auto simp add: final_notin_def intro: exec.intros) done then show ?thesis apply (rule conseqPre [OF hoarep.Guard]) apply clarify apply (frule Guard_noFaultStuckD [OF _ False]) apply auto done qed next case Throw show "\,\\\<^bsub>/F\<^esub> {s. s = Z \ \\\Throw,Normal s\ \\({Stuck} \ Fault ` (-F))} Throw {t. \\\Throw,Normal Z\ \ Normal t}, {t. \\\Throw,Normal Z\ \ Abrupt t}" by (rule conseqPre [OF hoarep.Throw]) (blast intro: exec.intros) next case (Catch c\<^sub>1 c\<^sub>2) have "\Z. \,\\\<^bsub>/F\<^esub> {s. s = Z \ \\\c\<^sub>1,Normal s\ \\({Stuck} \ Fault ` (-F))} c\<^sub>1 {t. \\\c\<^sub>1,Normal Z\ \ Normal t}, {t. \\\c\<^sub>1,Normal Z\ \ Abrupt t}" using Catch.hyps by iprover hence "\,\\\<^bsub>/F\<^esub> {s. s = Z \ \\\Catch c\<^sub>1 c\<^sub>2,Normal s\ \\({Stuck} \ Fault ` (-F))} c\<^sub>1 {t. \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Normal t}, {t. \\\c\<^sub>1,Normal Z\ \ Abrupt t \ \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \\({Stuck} \ Fault ` (-F))}" by (rule ConseqMGT) (fastforce intro: exec.intros simp add: final_notin_def) moreover have "\Z. \,\\\<^bsub>/F\<^esub> {s. s=Z \ \\\c\<^sub>2,Normal s\ \\({Stuck} \ Fault ` (-F))} c\<^sub>2 {t. \\\c\<^sub>2,Normal Z\ \ Normal t}, {t. \\\c\<^sub>2,Normal Z\ \ Abrupt t}" using Catch.hyps by iprover hence "\,\\\<^bsub>/F\<^esub>{s. \\\c\<^sub>1,Normal Z\ \Abrupt s \ \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \\({Stuck} \ Fault ` (-F))} c\<^sub>2 {t. \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Normal t}, {t. \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Abrupt t}" by (rule ConseqMGT) (fastforce intro: exec.intros simp add: final_notin_def) ultimately show "\,\\\<^bsub>/F\<^esub> {s. s = Z \ \\\Catch c\<^sub>1 c\<^sub>2,Normal s\ \\({Stuck} \ Fault ` (-F))} Catch c\<^sub>1 c\<^sub>2 {t. \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Normal t}, {t. \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Abrupt t}" by (rule hoarep.Catch) qed lemma MGT_Calls: "\p\dom \. \Z. \,{}\\<^bsub>/F\<^esub>{s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F))} (Call p) {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t}" proof - { fix p Z assume defined: "p \ dom \" have "\,(\p\dom \. \Z. {({s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F))}, p, {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t})}) \\<^bsub>/F\<^esub>{s. s = Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F))} (the (\ p)) {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t}" (is "\,?\ \\<^bsub>/F\<^esub> (?Pre p Z) (the (\ p)) (?Post p Z),(?Abr p Z)") proof - have MGT_Calls: "\p\dom \. \Z. \,?\ \\<^bsub>/F\<^esub> {s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F))} (Call p) {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t}" by (intro ballI allI, rule HoarePartialDef.Asm,auto) have "\Z. \,?\\\<^bsub>/F\<^esub> {s. s=Z \ \\\the (\ p) ,Normal s\ \\({Stuck} \ Fault`(-F))} (the (\ p)) {t. \\\the (\ p),Normal Z\ \ Normal t}, {t. \\\the (\ p),Normal Z\ \ Abrupt t}" by (iprover intro: MGT_lemma [OF MGT_Calls]) thus "\,?\\\<^bsub>/F\<^esub> (?Pre p Z) (the (\ p)) (?Post p Z),(?Abr p Z)" apply (rule ConseqMGT) apply (clarify,safe) proof - assume "\\\Call p,Normal Z\ \\({Stuck} \ Fault ` (-F))" with defined show "\\\the (\ p),Normal Z\ \\({Stuck} \ Fault ` (-F))" by (fastforce simp add: final_notin_def intro: exec.intros) next fix t assume "\\\the (\ p),Normal Z\ \ Normal t" with defined show "\\\Call p,Normal Z\ \Normal t" by (auto intro: exec.Call) next fix t assume "\\\the (\ p),Normal Z\ \ Abrupt t" with defined show "\\\Call p,Normal Z\ \Abrupt t" by (auto intro: exec.Call) qed qed } then show ?thesis apply - apply (intro ballI allI) apply (rule CallRec' [where Procs="dom \" and P="\p Z. {s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F))}"and Q="\p Z. {t. \\\Call p,Normal Z\ \ Normal t}" and A="\p Z. {t. \\\Call p,Normal Z\ \ Abrupt t}"] ) apply simp+ done qed theorem hoare_complete: "\\\<^bsub>/F\<^esub> P c Q,A \ \,{}\\<^bsub>/F\<^esub> P c Q,A" by (iprover intro: MGT_implies_complete MGT_lemma [OF MGT_Calls]) lemma hoare_complete': assumes cvalid: "\n. \,\\n:\<^bsub>/F\<^esub> P c Q,A" shows "\,\\\<^bsub>/F\<^esub> P c Q,A" proof (cases "\\\<^bsub>/F\<^esub> P c Q,A") case True hence "\,{}\\<^bsub>/F\<^esub> P c Q,A" by (rule hoare_complete) thus "\,\\\<^bsub>/F \<^esub>P c Q,A" by (rule hoare_augment_context) simp next case False with cvalid show ?thesis by (rule ExFalso) qed lemma hoare_strip_\: assumes deriv: "\,{}\\<^bsub>/F\<^esub> P p Q,A" assumes F': "F' \ -F" shows "strip F' \,{}\\<^bsub>/F\<^esub> P p Q,A" proof (rule hoare_complete) from hoare_sound [OF deriv] have "\\\<^bsub>/F\<^esub> P p Q,A" by (simp add: cvalid_def) from this F' show "strip F' \\\<^bsub>/F\<^esub> P p Q,A" by (rule valid_to_valid_strip) qed subsection \And Now: Some Useful Rules\ subsubsection \Consequence\ lemma LiberalConseq_sound: fixes F::"'f set" assumes cons: "\s \ P. \(t::('s,'f) xstate). \P' Q' A'. (\n. \,\\n:\<^bsub>/F\<^esub> P' c Q',A') \ ((s \ P' \ t \ Normal ` Q' \ Abrupt ` A') \ t \ Normal ` Q \ Abrupt ` A)" shows "\,\\n:\<^bsub>/F\<^esub> P c Q,A " proof (rule cnvalidI) fix s t assume ctxt:"\(P, p, Q, A)\\. \\n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\c,Normal s\ =n\ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" show "t \ Normal ` Q \ Abrupt ` A" proof - from P cons obtain P' Q' A' where spec: "\n. \,\\n:\<^bsub>/F\<^esub> P' c Q',A'" and adapt: "(s \ P' \ t \ Normal ` Q' \ Abrupt ` A') \ t \ Normal ` Q \ Abrupt ` A" apply - apply (drule (1) bspec) apply (erule_tac x=t in allE) apply (elim exE conjE) apply iprover done from exec spec ctxt t_notin_F have "s \ P' \ t \ Normal ` Q' \ Abrupt ` A'" by (simp add: cnvalid_def nvalid_def) with adapt show ?thesis by simp qed qed lemma LiberalConseq: fixes F:: "'f set" assumes cons: "\s \ P. \(t::('s,'f) xstate). \P' Q' A'. \,\\\<^bsub>/F\<^esub> P' c Q',A' \ ((s \ P' \ t \ Normal ` Q' \ Abrupt ` A') \ t \ Normal ` Q \ Abrupt ` A)" shows "\,\\\<^bsub>/F\<^esub> P c Q,A " apply (rule hoare_complete') apply (rule allI) apply (rule LiberalConseq_sound) using cons apply (clarify) apply (drule (1) bspec) apply (erule_tac x=t in allE) apply clarify apply (rule_tac x=P' in exI) apply (rule_tac x=Q' in exI) apply (rule_tac x=A' in exI) apply (rule conjI) apply (blast intro: hoare_cnvalid) apply assumption done lemma "\s \ P. \P' Q' A'. \,\\\<^bsub>/F\<^esub> P' c Q',A' \ s \ P' \ Q' \ Q \ A' \ A \ \,\\\<^bsub>/F\<^esub> P c Q,A" apply (rule LiberalConseq) apply (rule ballI) apply (drule (1) bspec) apply clarify apply (rule_tac x=P' in exI) apply (rule_tac x=Q' in exI) apply (rule_tac x=A' in exI) apply auto done lemma fixes F:: "'f set" assumes cons: "\s \ P. \P' Q' A'. \,\\\<^bsub>/F\<^esub> P' c Q',A' \ (\(t::('s,'f) xstate). (s \ P' \ t \ Normal ` Q' \ Abrupt ` A') \ t \ Normal ` Q \ Abrupt ` A)" shows "\,\\\<^bsub>/F\<^esub> P c Q,A " apply (rule Conseq) apply (rule ballI) apply (insert cons) apply (drule (1) bspec) apply clarify apply (rule_tac x=P' in exI) apply (rule_tac x=Q' in exI) apply (rule_tac x=A' in exI) apply (rule conjI) apply assumption (* no way to get s \ P' *) oops lemma LiberalConseq': fixes F:: "'f set" assumes cons: "\s \ P. \P' Q' A'. \,\\\<^bsub>/F\<^esub> P' c Q',A' \ (\(t::('s,'f) xstate). (s \ P' \ t \ Normal ` Q' \ Abrupt ` A') \ t \ Normal ` Q \ Abrupt ` A)" shows "\,\\\<^bsub>/F\<^esub> P c Q,A " apply (rule LiberalConseq) apply (rule ballI) apply (rule allI) apply (insert cons) apply (drule (1) bspec) apply clarify apply (rule_tac x=P' in exI) apply (rule_tac x=Q' in exI) apply (rule_tac x=A' in exI) apply iprover done lemma LiberalConseq'': fixes F:: "'f set" assumes spec: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z)" assumes cons: "\s (t::('s,'f) xstate). (\Z. s \ P' Z \ t \ Normal ` Q' Z \ Abrupt ` A' Z) \ (s \ P \ t \ Normal ` Q \ Abrupt ` A)" shows "\,\\\<^bsub>/F\<^esub> P c Q,A " apply (rule LiberalConseq) apply (rule ballI) apply (rule allI) apply (insert cons) apply (erule_tac x=s in allE) apply (erule_tac x=t in allE) apply (case_tac "t \ Normal ` Q \ Abrupt ` A") apply (insert spec) apply iprover apply auto done primrec procs:: "('s,'p,'f) com \ 'p set" where "procs Skip = {}" | "procs (Basic f) = {}" | "procs (Seq c\<^sub>1 c\<^sub>2) = (procs c\<^sub>1 \ procs c\<^sub>2)" | "procs (Cond b c\<^sub>1 c\<^sub>2) = (procs c\<^sub>1 \ procs c\<^sub>2)" | "procs (While b c) = procs c" | "procs (Call p) = {p}" | "procs (DynCom c) = (\s. procs (c s))" | "procs (Guard f g c) = procs c" | "procs Throw = {}" | "procs (Catch c\<^sub>1 c\<^sub>2) = (procs c\<^sub>1 \ procs c\<^sub>2)" primrec noSpec:: "('s,'p,'f) com \ bool" where "noSpec Skip = True" | "noSpec (Basic f) = True" | "noSpec (Spec r) = False" | "noSpec (Seq c\<^sub>1 c\<^sub>2) = (noSpec c\<^sub>1 \ noSpec c\<^sub>2)" | "noSpec (Cond b c\<^sub>1 c\<^sub>2) = (noSpec c\<^sub>1 \ noSpec c\<^sub>2)" | "noSpec (While b c) = noSpec c" | "noSpec (Call p) = True" | "noSpec (DynCom c) = (\s. noSpec (c s))" | "noSpec (Guard f g c) = noSpec c" | "noSpec Throw = True" | "noSpec (Catch c\<^sub>1 c\<^sub>2) = (noSpec c\<^sub>1 \ noSpec c\<^sub>2)" lemma exec_noSpec_no_Stuck: assumes exec: "\\\c,s\ \ t" assumes noSpec_c: "noSpec c" assumes noSpec_\: "\p \ dom \. noSpec (the (\ p))" assumes procs_subset: "procs c \ dom \" assumes procs_subset_\: "\p \ dom \. procs (the (\ p)) \ dom \" assumes s_no_Stuck: "s\Stuck" shows "t\Stuck" using exec noSpec_c procs_subset s_no_Stuck proof induct case (Call p bdy s t) with noSpec_\ procs_subset_\ show ?case by (auto dest!: bspec [of _ _ p]) next case (DynCom c s t) then show ?case by auto blast qed auto lemma execn_noSpec_no_Stuck: assumes exec: "\\\c,s\ =n\ t" assumes noSpec_c: "noSpec c" assumes noSpec_\: "\p \ dom \. noSpec (the (\ p))" assumes procs_subset: "procs c \ dom \" assumes procs_subset_\: "\p \ dom \. procs (the (\ p)) \ dom \" assumes s_no_Stuck: "s\Stuck" shows "t\Stuck" using exec noSpec_c procs_subset s_no_Stuck proof induct case (Call p bdy n s t) with noSpec_\ procs_subset_\ show ?case by (auto dest!: bspec [of _ _ p]) next case (DynCom c s t) then show ?case by auto blast qed auto lemma LiberalConseq_noguards_nothrows_sound: assumes spec: "\Z. \n. \,\\n:\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z)" assumes cons: "\s t. (\Z. s \ P' Z \ t \ Q' Z ) \ (s \ P \ t \ Q )" assumes noguards_c: "noguards c" assumes noguards_\: "\p \ dom \. noguards (the (\ p))" assumes nothrows_c: "nothrows c" assumes nothrows_\: "\p \ dom \. nothrows (the (\ p))" assumes noSpec_c: "noSpec c" assumes noSpec_\: "\p \ dom \. noSpec (the (\ p))" assumes procs_subset: "procs c \ dom \" assumes procs_subset_\: "\p \ dom \. procs (the (\ p)) \ dom \" shows "\,\\n:\<^bsub>/F\<^esub> P c Q,A " proof (rule cnvalidI) fix s t assume ctxt:"\(P, p, Q, A)\\. \\n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\c,Normal s\ =n\ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" show "t \ Normal ` Q \ Abrupt ` A" proof - from execn_noguards_no_Fault [OF exec noguards_c noguards_\] execn_nothrows_no_Abrupt [OF exec nothrows_c nothrows_\ ] execn_noSpec_no_Stuck [OF exec noSpec_c noSpec_\ procs_subset procs_subset_\] obtain t' where t: "t=Normal t'" by (cases t) auto with exec spec ctxt have "(\Z. s \ P' Z \ t' \ Q' Z)" by (unfold cnvalid_def nvalid_def) blast with cons P t show ?thesis by simp qed qed lemma LiberalConseq_noguards_nothrows: assumes spec: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z)" assumes cons: "\s t. (\Z. s \ P' Z \ t \ Q' Z ) \ (s \ P \ t \ Q )" assumes noguards_c: "noguards c" assumes noguards_\: "\p \ dom \. noguards (the (\ p))" assumes nothrows_c: "nothrows c" assumes nothrows_\: "\p \ dom \. nothrows (the (\ p))" assumes noSpec_c: "noSpec c" assumes noSpec_\: "\p \ dom \. noSpec (the (\ p))" assumes procs_subset: "procs c \ dom \" assumes procs_subset_\: "\p \ dom \. procs (the (\ p)) \ dom \" shows "\,\\\<^bsub>/F\<^esub> P c Q,A " apply (rule hoare_complete') apply (rule allI) apply (rule LiberalConseq_noguards_nothrows_sound [OF _ cons noguards_c noguards_\ nothrows_c nothrows_\ noSpec_c noSpec_\ procs_subset procs_subset_\]) apply (insert spec) apply (intro allI) apply (erule_tac x=Z in allE) by (rule hoare_cnvalid) lemma assumes spec: "\Z. \,\\\<^bsub>/F\<^esub>{s. s=fst Z \ P s (snd Z)} c {t. Q (fst Z) (snd Z) t},{}" assumes noguards_c: "noguards c" assumes noguards_\: "\p \ dom \. noguards (the (\ p))" assumes nothrows_c: "nothrows c" assumes nothrows_\: "\p \ dom \. nothrows (the (\ p))" assumes noSpec_c: "noSpec c" assumes noSpec_\: "\p \ dom \. noSpec (the (\ p))" assumes procs_subset: "procs c \ dom \" assumes procs_subset_\: "\p \ dom \. procs (the (\ p)) \ dom \" shows "\\. \,\\\<^bsub>/F\<^esub>{s. s=\} c {t. \l. P \ l \ Q \ l t},{}" apply (rule allI) apply (rule LiberalConseq_noguards_nothrows [OF spec _ noguards_c noguards_\ nothrows_c nothrows_\ noSpec_c noSpec_\ procs_subset procs_subset_\]) apply auto done subsubsection \Modify Return\ -lemma ProcModifyReturn_sound: - assumes valid_call: "\n. \,\ \n:\<^bsub>/F\<^esub> P call init p return' c Q,A" +lemma Proc_exnModifyReturn_sound: + assumes valid_call: "\n. \,\ \n:\<^bsub>/F\<^esub> P call_exn init p return' result_exn c Q,A" assumes valid_modif: "\\. \n. \,\\n:\<^bsub>/UNIV\<^esub> {\} Call p (Modif \),(ModifAbr \)" assumes ret_modif: "\s t. t \ Modif (init s) \ return' s t = return s t" assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) - \ return' s t = return s t" - shows "\,\ \n:\<^bsub>/F\<^esub> P (call init p return c) Q,A" + \ result_exn (return' s t) t = result_exn (return s t) t" + shows "\,\ \n:\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" then have ctxt': "\(P, p, Q, A)\\. \ \n:\<^bsub>/UNIV\<^esub> P (Call p) Q,A" by (auto intro: nvalid_augment_Faults) - assume exec: "\\\call init p return c,Normal s\ =n\ t" + assume exec: "\\\call_exn init p return result_exn c,Normal s\ =n\ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" from exec show "t \ Normal ` Q \ Abrupt ` A" - proof (cases rule: execn_call_Normal_elim) + proof (cases rule: execn_call_exn_Normal_elim) fix bdy m t' assume bdy: "\ p = Some bdy" assume exec_body: "\\\bdy,Normal (init s)\ =m\ Normal t'" assume exec_c: "\\\c s t',Normal (return s t')\ =Suc m\ t" assume n: "n = Suc m" from exec_body n bdy have "\\\Call p,Normal (init s)\ =n\ Normal t'" by (auto simp add: intro: execn.Call) from cnvalidD [OF valid_modif [rule_format, of n "init s"] ctxt' this] P have "t' \ Modif (init s)" by auto with ret_modif have "Normal (return' s t') = Normal (return s t')" by simp with exec_body exec_c bdy n - have "\\\call init p return' c,Normal s\ =n\ t" - by (auto intro: execn_call) + have "\\\call_exn init p return' result_exn c,Normal s\ =n\ t" + by (auto intro: execn_call_exn) from cnvalidD [OF valid_call [rule_format] ctxt this] P t_notin_F show ?thesis by simp next fix bdy m t' assume bdy: "\ p = Some bdy" assume exec_body: "\\\bdy,Normal (init s)\ =m\ Abrupt t'" assume n: "n = Suc m" - assume t: "t = Abrupt (return s t')" + assume t: "t = Abrupt (result_exn (return s t') t')" also from exec_body n bdy have "\\\Call p,Normal (init s)\ =n\ Abrupt t'" by (auto simp add: intro: execn.intros) from cnvalidD [OF valid_modif [rule_format, of n "init s"] ctxt' this] P have "t' \ ModifAbr (init s)" by auto - with ret_modifAbr have "Abrupt (return s t') = Abrupt (return' s t')" + with ret_modifAbr have "Abrupt (result_exn (return s t') t') = Abrupt (result_exn (return' s t') t')" by simp - finally have "t = Abrupt (return' s t')" . + finally have "t = Abrupt (result_exn (return' s t') t')" . with exec_body bdy n - have "\\\call init p return' c,Normal s\ =n\ t" - by (auto intro: execn_callAbrupt) + have "\\\call_exn init p return' result_exn c,Normal s\ =n\ t" + by (auto intro: execn_call_exnAbrupt) from cnvalidD [OF valid_call [rule_format] ctxt this] P t_notin_F show ?thesis by simp next fix bdy m f assume bdy: "\ p = Some bdy" assume "\\\bdy,Normal (init s)\ =m\ Fault f" "n = Suc m" "t = Fault f" - with bdy have "\\\call init p return' c ,Normal s\ =n\ t" - by (auto intro: execn_callFault) + with bdy have "\\\call_exn init p return' result_exn c ,Normal s\ =n\ t" + by (auto intro: execn_call_exnFault) from valid_call [rule_format] ctxt this P t_notin_F show ?thesis by (rule cnvalidD) next fix bdy m assume bdy: "\ p = Some bdy" assume "\\\bdy,Normal (init s)\ =m\ Stuck" "n = Suc m" "t = Stuck" - with bdy have "\\\call init p return' c ,Normal s\ =n\ t" - by (auto intro: execn_callStuck) + with bdy have "\\\call_exn init p return' result_exn c ,Normal s\ =n\ t" + by (auto intro: execn_call_exnStuck) from valid_call [rule_format] ctxt this P t_notin_F show ?thesis by (rule cnvalidD) next fix m assume "\ p = None" and "n = Suc m" "t = Stuck" - then have "\\\call init p return' c ,Normal s\ =n\ t" - by (auto intro: execn_callUndefined) + then have "\\\call_exn init p return' result_exn c ,Normal s\ =n\ t" + by (auto intro: execn_call_exnUndefined) from valid_call [rule_format] ctxt this P t_notin_F show ?thesis by (rule cnvalidD) qed qed +lemma ProcModifyReturn_sound: + assumes valid_call: "\n. \,\ \n:\<^bsub>/F\<^esub> P call init p return' c Q,A" + assumes valid_modif: + "\\. \n. \,\\n:\<^bsub>/UNIV\<^esub> {\} Call p (Modif \),(ModifAbr \)" + assumes ret_modif: + "\s t. t \ Modif (init s) + \ return' s t = return s t" + assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) + \ return' s t = return s t" + shows "\,\ \n:\<^bsub>/F\<^esub> P (call init p return c) Q,A" + using valid_call valid_modif ret_modif ret_modifAbr + unfolding call_call_exn + by (rule Proc_exnModifyReturn_sound) + +lemma Proc_exnModifyReturn: + assumes spec: "\,\\\<^bsub>/F\<^esub> P (call_exn init p return' result_exn c) Q,A" + assumes result_conform: + "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" + assumes return_conform: + "\s t. t \ ModifAbr (init s) + \ (result_exn (return' s t) t) = (result_exn (return s t) t)" + assumes modifies_spec: + "\\. \,\\\<^bsub>/UNIV\<^esub> {\} Call p (Modif \),(ModifAbr \)" + shows "\,\\\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" +apply (rule hoare_complete') +apply (rule allI) +apply (rule Proc_exnModifyReturn_sound + [where Modif=Modif and ModifAbr=ModifAbr, + OF _ _ result_conform return_conform] ) +using spec +apply (blast intro: hoare_cnvalid) +using modifies_spec +apply (blast intro: hoare_cnvalid) +done lemma ProcModifyReturn: assumes spec: "\,\\\<^bsub>/F\<^esub> P (call init p return' c) Q,A" assumes result_conform: "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" assumes return_conform: "\s t. t \ ModifAbr (init s) \ (return' s t) = (return s t)" assumes modifies_spec: "\\. \,\\\<^bsub>/UNIV\<^esub> {\} Call p (Modif \),(ModifAbr \)" - shows "\,\\\<^bsub>/F\<^esub> P (call init p return c) Q,A" -apply (rule hoare_complete') -apply (rule allI) -apply (rule ProcModifyReturn_sound - [where Modif=Modif and ModifAbr=ModifAbr, - OF _ _ result_conform return_conform] ) -using spec -apply (blast intro: hoare_cnvalid) -using modifies_spec -apply (blast intro: hoare_cnvalid) -done +shows "\,\\\<^bsub>/F\<^esub> P (call init p return c) Q,A" + using spec result_conform return_conform modifies_spec + unfolding call_call_exn + by (rule Proc_exnModifyReturn) -lemma ProcModifyReturnSameFaults_sound: - assumes valid_call: "\n. \,\ \n:\<^bsub>/F\<^esub> P call init p return' c Q,A" +lemma Proc_exnModifyReturnSameFaults_sound: + assumes valid_call: "\n. \,\ \n:\<^bsub>/F\<^esub> P call_exn init p return' result_exn c Q,A" assumes valid_modif: "\\. \n. \,\\n:\<^bsub>/F\<^esub> {\} Call p (Modif \),(ModifAbr \)" assumes ret_modif: "\s t. t \ Modif (init s) \ return' s t = return s t" assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) - \ return' s t = return s t" - shows "\,\ \n:\<^bsub>/F\<^esub> P (call init p return c) Q,A" + \ result_exn (return' s t) t = result_exn (return s t) t" + shows "\,\ \n:\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" - assume exec: "\\\call init p return c,Normal s\ =n\ t" + assume exec: "\\\call_exn init p return result_exn c,Normal s\ =n\ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" from exec show "t \ Normal ` Q \ Abrupt ` A" - proof (cases rule: execn_call_Normal_elim) + proof (cases rule: execn_call_exn_Normal_elim) fix bdy m t' assume bdy: "\ p = Some bdy" assume exec_body: "\\\bdy,Normal (init s)\ =m\ Normal t'" assume exec_c: "\\\c s t',Normal (return s t')\ =Suc m\ t" assume n: "n = Suc m" from exec_body n bdy have "\\\Call p,Normal (init s)\ =n\ Normal t'" by (auto simp add: intro: execn.intros) from cnvalidD [OF valid_modif [rule_format, of n "init s"] ctxt this] P have "t' \ Modif (init s)" by auto with ret_modif have "Normal (return' s t') = Normal (return s t')" by simp with exec_body exec_c bdy n - have "\\\call init p return' c,Normal s\ =n\ t" - by (auto intro: execn_call) + have "\\\call_exn init p return' result_exn c,Normal s\ =n\ t" + by (auto intro: execn_call_exn) from cnvalidD [OF valid_call [rule_format] ctxt this] P t_notin_F show ?thesis by simp next fix bdy m t' assume bdy: "\ p = Some bdy" assume exec_body: "\\\bdy,Normal (init s)\ =m\ Abrupt t'" assume n: "n = Suc m" - assume t: "t = Abrupt (return s t')" + assume t: "t = Abrupt (result_exn (return s t') t')" also from exec_body n bdy have "\\\Call p,Normal (init s)\ =n \ Abrupt t'" by (auto simp add: intro: execn.intros) from cnvalidD [OF valid_modif [rule_format, of n "init s"] ctxt this] P have "t' \ ModifAbr (init s)" by auto - with ret_modifAbr have "Abrupt (return s t') = Abrupt (return' s t')" + with ret_modifAbr have "Abrupt (result_exn (return s t') t') = Abrupt (result_exn (return' s t') t')" by simp - finally have "t = Abrupt (return' s t')" . + finally have "t = Abrupt (result_exn (return' s t') t')" . with exec_body bdy n - have "\\\call init p return' c,Normal s\ =n\ t" - by (auto intro: execn_callAbrupt) + have "\\\call_exn init p return' result_exn c,Normal s\ =n\ t" + by (auto intro: execn_call_exnAbrupt) from cnvalidD [OF valid_call [rule_format] ctxt this] P t_notin_F show ?thesis by simp next fix bdy m f assume bdy: "\ p = Some bdy" assume "\\\bdy,Normal (init s)\ =m\ Fault f" "n = Suc m" and t: "t = Fault f" - with bdy have "\\\call init p return' c ,Normal s\ =n\ t" - by (auto intro: execn_callFault) + with bdy have "\\\call_exn init p return' result_exn c ,Normal s\ =n\ t" + by (auto intro: execn_call_exnFault) from cnvalidD [OF valid_call [rule_format] ctxt this P] t t_notin_F show ?thesis by simp next fix bdy m assume bdy: "\ p = Some bdy" assume "\\\bdy,Normal (init s)\ =m\ Stuck" "n = Suc m" "t = Stuck" - with bdy have "\\\call init p return' c ,Normal s\ =n\ t" - by (auto intro: execn_callStuck) + with bdy have "\\\call_exn init p return' result_exn c ,Normal s\ =n\ t" + by (auto intro: execn_call_exnStuck) from valid_call [rule_format] ctxt this P t_notin_F show ?thesis by (rule cnvalidD) next fix m assume "\ p = None" and "n = Suc m" "t = Stuck" - then have "\\\call init p return' c ,Normal s\ =n\ t" - by (auto intro: execn_callUndefined) + then have "\\\call_exn init p return' result_exn c ,Normal s\ =n\ t" + by (auto intro: execn_call_exnUndefined) from valid_call [rule_format] ctxt this P t_notin_F show ?thesis by (rule cnvalidD) qed qed +lemma ProcModifyReturnSameFaults_sound: + assumes valid_call: "\n. \,\ \n:\<^bsub>/F\<^esub> P call init p return' c Q,A" + assumes valid_modif: + "\\. \n. \,\\n:\<^bsub>/F\<^esub> {\} Call p (Modif \),(ModifAbr \)" + assumes ret_modif: + "\s t. t \ Modif (init s) + \ return' s t = return s t" + assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) + \ return' s t = return s t" + shows "\,\ \n:\<^bsub>/F\<^esub> P (call init p return c) Q,A" + using valid_call valid_modif ret_modif ret_modifAbr + unfolding call_call_exn + by (rule Proc_exnModifyReturnSameFaults_sound) + + +lemma Proc_exnModifyReturnSameFaults: + assumes spec: "\,\\\<^bsub>/F\<^esub> P (call_exn init p return' result_exn c) Q,A" + assumes result_conform: + "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" + assumes return_conform: + "\s t. t \ ModifAbr (init s) \ (result_exn (return' s t) t) = (result_exn (return s t) t)" + assumes modifies_spec: + "\\. \,\\\<^bsub>/F\<^esub> {\} Call p (Modif \),(ModifAbr \)" + shows "\,\\\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" +apply (rule hoare_complete') +apply (rule allI) +apply (rule Proc_exnModifyReturnSameFaults_sound + [where Modif=Modif and ModifAbr=ModifAbr, + OF _ _ result_conform return_conform]) +using spec +apply (blast intro: hoare_cnvalid) +using modifies_spec +apply (blast intro: hoare_cnvalid) +done + lemma ProcModifyReturnSameFaults: assumes spec: "\,\\\<^bsub>/F\<^esub> P (call init p return' c) Q,A" assumes result_conform: "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" assumes return_conform: "\s t. t \ ModifAbr (init s) \ (return' s t) = (return s t)" assumes modifies_spec: "\\. \,\\\<^bsub>/F\<^esub> {\} Call p (Modif \),(ModifAbr \)" - shows "\,\\\<^bsub>/F\<^esub> P (call init p return c) Q,A" -apply (rule hoare_complete') -apply (rule allI) -apply (rule ProcModifyReturnSameFaults_sound - [where Modif=Modif and ModifAbr=ModifAbr, - OF _ _ result_conform return_conform]) -using spec -apply (blast intro: hoare_cnvalid) -using modifies_spec -apply (blast intro: hoare_cnvalid) -done +shows "\,\\\<^bsub>/F\<^esub> P (call init p return c) Q,A" + using spec result_conform return_conform modifies_spec + unfolding call_call_exn + by (rule Proc_exnModifyReturnSameFaults) + subsubsection \DynCall\ + + +lemma dynProc_exnModifyReturn_sound: +assumes valid_call: "\n. \,\ \n:\<^bsub>/F\<^esub> P dynCall_exn f g init p return' result_exn c Q,A" +assumes valid_modif: + "\s \ P. \\. \n. + \,\\n:\<^bsub>/UNIV\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" +assumes ret_modif: + "\s t. t \ Modif (init s) + \ return' s t = return s t" +assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) + \ result_exn (return' s t) t = result_exn (return s t) t" +shows "\,\ \n:\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" +proof (rule cnvalidI) + fix s t + assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" + then have ctxt': "\(P, p, Q, A)\\. \ \n:\<^bsub>/UNIV\<^esub> P (Call p) Q,A" + by (auto intro: nvalid_augment_Faults) + assume exec: "\\\dynCall_exn f g init p return result_exn c,Normal s\ =n\ t" + assume t_notin_F: "t \ Fault ` F" + assume P: "s \ P" + with valid_modif + have valid_modif': "\\. \n. + \,\\n:\<^bsub>/UNIV\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" + by blast + from exec thm execn_Normal_elim_cases + have exec_call: "\\\maybe_guard f g (call_exn init (p s) return result_exn c),Normal s\ =n\ t" + by (cases rule: execn_dynCall_exn_Normal_elim) + then show "t \ Normal ` Q \ Abrupt ` A" + proof (cases rule: execn_maybe_guard_Normal_elim_cases) + case noFault + from noFault have guards_ok: "s \ g" by simp + from noFault have "\\ \call_exn init (p s) return result_exn c,Normal s\ =n\ t" by simp + then show "t \ Normal ` Q \ Abrupt ` A" + proof (cases rule: execn_call_exn_Normal_elim) + fix bdy m t' + assume bdy: "\ (p s) = Some bdy" + assume exec_body: "\\\bdy,Normal (init s)\ =m\ Normal t'" + assume exec_c: "\\\c s t',Normal (return s t')\ =Suc m\ t" + assume n: "n = Suc m" + from exec_body n bdy + have "\\\Call (p s) ,Normal (init s)\ =n\ Normal t'" + by (auto simp add: intro: execn.intros) + from cnvalidD [OF valid_modif' [rule_format, of n "init s"] ctxt' this] P + have "t' \ Modif (init s)" + by auto + with ret_modif have "Normal (return' s t') = Normal (return s t')" + by simp + with exec_body exec_c bdy n + have "\\\call_exn init (p s) return' result_exn c,Normal s\ =n\ t" + by (auto intro: execn_call_exn) + from execn_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ =n\ t" + by (simp add: dynCall_exn_def execn_maybe_guard_DynCom) + from cnvalidD [OF valid_call ctxt this] P t_notin_F + show ?thesis + by simp + next + fix bdy m t' + assume bdy: "\ (p s) = Some bdy" + assume exec_body: "\\\bdy,Normal (init s)\ =m\ Abrupt t'" + assume n: "n = Suc m" + assume t: "t = Abrupt (result_exn (return s t') t')" + also from exec_body n bdy + have "\\\Call (p s) ,Normal (init s)\ =n\ Abrupt t'" + by (auto simp add: intro: execn.intros) + from cnvalidD [OF valid_modif' [rule_format, of n "init s"] ctxt' this] P + have "t' \ ModifAbr (init s)" + by auto + with ret_modifAbr have "Abrupt (result_exn (return s t') t') = Abrupt (result_exn (return' s t') t')" + by simp + finally have "t = Abrupt (result_exn (return' s t') t')" . + with exec_body bdy n + have "\\\call_exn init (p s) return' result_exn c,Normal s\ =n\ t" + by (auto intro: execn_call_exnAbrupt) + from execn_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ =n\ t" + by (simp add: dynCall_exn_def execn_maybe_guard_DynCom) + from cnvalidD [OF valid_call ctxt this] P t_notin_F + show ?thesis + by simp + next + fix bdy m f' + assume bdy: "\ (p s) = Some bdy" + assume "\\\bdy,Normal (init s)\ =m\ Fault f'" "n = Suc m" + "t = Fault f'" + with bdy have "\\\call_exn init (p s) return' result_exn c ,Normal s\ =n\ t" + by (auto intro: execn_call_exnFault) + from execn_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ =n\ t" + by (simp add: dynCall_exn_def execn_maybe_guard_DynCom) + from valid_call ctxt this P t_notin_F + show ?thesis + by (rule cnvalidD) + next + fix bdy m + assume bdy: "\ (p s) = Some bdy" + assume "\\\bdy,Normal (init s)\ =m\ Stuck" "n = Suc m" + "t = Stuck" + with bdy have "\\\call_exn init (p s) return' result_exn c ,Normal s\ =n\ t" + by (auto intro: execn_call_exnStuck) + from execn_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ =n\ t" + by (simp add: dynCall_exn_def execn_maybe_guard_DynCom) + from valid_call ctxt this P t_notin_F + show ?thesis + by (rule cnvalidD) + next + fix m + assume "\ (p s) = None" + and "n = Suc m" "t = Stuck" + hence "\\\call_exn init (p s) return' result_exn c ,Normal s\ =n\ t" + by (auto intro: execn_call_exnUndefined) + from execn_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ =n\ t" + by (simp add: dynCall_exn_def execn_maybe_guard_DynCom) + from valid_call ctxt this P t_notin_F + show ?thesis + by (rule cnvalidD) + qed + next + case (someFault) + then obtain guards_fail:"s \ g" + and t: "t = Fault f" by simp + from execn_maybe_guard_Fault [OF guards_fail] t + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ =n\ t" + by (simp add: dynCall_exn_def execn_maybe_guard_DynCom) + from cnvalidD [OF valid_call ctxt this] P t_notin_F + show ?thesis by simp + qed +qed + lemma dynProcModifyReturn_sound: assumes valid_call: "\n. \,\ \n:\<^bsub>/F\<^esub> P dynCall init p return' c Q,A" assumes valid_modif: "\s \ P. \\. \n. \,\\n:\<^bsub>/UNIV\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" assumes ret_modif: "\s t. t \ Modif (init s) \ return' s t = return s t" assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) \ return' s t = return s t" shows "\,\ \n:\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" -proof (rule cnvalidI) - fix s t - assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" - then have ctxt': "\(P, p, Q, A)\\. \ \n:\<^bsub>/UNIV\<^esub> P (Call p) Q,A" - by (auto intro: nvalid_augment_Faults) - assume exec: "\\\dynCall init p return c,Normal s\ =n\ t" - assume t_notin_F: "t \ Fault ` F" - assume P: "s \ P" - with valid_modif - have valid_modif': "\\. \n. - \,\\n:\<^bsub>/UNIV\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" - by blast - from exec - have "\\\call init (p s) return c,Normal s\ =n\ t" - by (cases rule: execn_dynCall_Normal_elim) - then show "t \ Normal ` Q \ Abrupt ` A" - proof (cases rule: execn_call_Normal_elim) - fix bdy m t' - assume bdy: "\ (p s) = Some bdy" - assume exec_body: "\\\bdy,Normal (init s)\ =m\ Normal t'" - assume exec_c: "\\\c s t',Normal (return s t')\ =Suc m\ t" - assume n: "n = Suc m" - from exec_body n bdy - have "\\\Call (p s) ,Normal (init s)\ =n\ Normal t'" - by (auto simp add: intro: execn.intros) - from cnvalidD [OF valid_modif' [rule_format, of n "init s"] ctxt' this] P - have "t' \ Modif (init s)" - by auto - with ret_modif have "Normal (return' s t') = Normal (return s t')" - by simp - with exec_body exec_c bdy n - have "\\\call init (p s) return' c,Normal s\ =n\ t" - by (auto intro: execn_call) - hence "\\\dynCall init p return' c,Normal s\ =n\ t" - by (rule execn_dynCall) - from cnvalidD [OF valid_call ctxt this] P t_notin_F - show ?thesis - by simp - next - fix bdy m t' - assume bdy: "\ (p s) = Some bdy" - assume exec_body: "\\\bdy,Normal (init s)\ =m\ Abrupt t'" - assume n: "n = Suc m" - assume t: "t = Abrupt (return s t')" - also from exec_body n bdy - have "\\\Call (p s) ,Normal (init s)\ =n\ Abrupt t'" - by (auto simp add: intro: execn.intros) - from cnvalidD [OF valid_modif' [rule_format, of n "init s"] ctxt' this] P - have "t' \ ModifAbr (init s)" - by auto - with ret_modifAbr have "Abrupt (return s t') = Abrupt (return' s t')" - by simp - finally have "t = Abrupt (return' s t')" . - with exec_body bdy n - have "\\\call init (p s) return' c,Normal s\ =n\ t" - by (auto intro: execn_callAbrupt) - hence "\\\dynCall init p return' c,Normal s\ =n\ t" - by (rule execn_dynCall) - from cnvalidD [OF valid_call ctxt this] P t_notin_F - show ?thesis - by simp - next - fix bdy m f - assume bdy: "\ (p s) = Some bdy" - assume "\\\bdy,Normal (init s)\ =m\ Fault f" "n = Suc m" - "t = Fault f" - with bdy have "\\\call init (p s) return' c ,Normal s\ =n\ t" - by (auto intro: execn_callFault) - hence "\\\dynCall init p return' c,Normal s\ =n\ t" - by (rule execn_dynCall) - from valid_call ctxt this P t_notin_F - show ?thesis - by (rule cnvalidD) - next - fix bdy m - assume bdy: "\ (p s) = Some bdy" - assume "\\\bdy,Normal (init s)\ =m\ Stuck" "n = Suc m" - "t = Stuck" - with bdy have "\\\call init (p s) return' c ,Normal s\ =n\ t" - by (auto intro: execn_callStuck) - hence "\\\dynCall init p return' c,Normal s\ =n\ t" - by (rule execn_dynCall) - from valid_call ctxt this P t_notin_F - show ?thesis - by (rule cnvalidD) - next - fix m - assume "\ (p s) = None" - and "n = Suc m" "t = Stuck" - hence "\\\call init (p s) return' c ,Normal s\ =n\ t" - by (auto intro: execn_callUndefined) - hence "\\\dynCall init p return' c,Normal s\ =n\ t" - by (rule execn_dynCall) - from valid_call ctxt this P t_notin_F - show ?thesis - by (rule cnvalidD) - qed -qed + using valid_call valid_modif ret_modif ret_modifAbr + unfolding dynCall_dynCall_exn + by (rule dynProc_exnModifyReturn_sound) + + +lemma dynProc_exnModifyReturn: +assumes dyn_call: "\,\\\<^bsub>/F\<^esub> P dynCall_exn f g init p return' result_exn c Q,A" +assumes ret_modif: + "\s t. t \ Modif (init s) + \ return' s t = return s t" +assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) + \ result_exn (return' s t) t = result_exn (return s t) t" +assumes modif: + "\s \ P. \\. + \,\\\<^bsub>/UNIV\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" +shows "\,\\\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" +apply (rule hoare_complete') +apply (rule allI) +apply (rule dynProc_exnModifyReturn_sound [where Modif=Modif and ModifAbr=ModifAbr, + OF hoare_cnvalid [OF dyn_call] _ ret_modif ret_modifAbr]) +apply (intro ballI allI) +apply (rule hoare_cnvalid [OF modif [rule_format]]) +apply assumption +done lemma dynProcModifyReturn: assumes dyn_call: "\,\\\<^bsub>/F\<^esub> P dynCall init p return' c Q,A" assumes ret_modif: "\s t. t \ Modif (init s) \ return' s t = return s t" assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) \ return' s t = return s t" assumes modif: "\s \ P. \\. \,\\\<^bsub>/UNIV\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" -shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" -apply (rule hoare_complete') -apply (rule allI) -apply (rule dynProcModifyReturn_sound [where Modif=Modif and ModifAbr=ModifAbr, - OF hoare_cnvalid [OF dyn_call] _ ret_modif ret_modifAbr]) -apply (intro ballI allI) -apply (rule hoare_cnvalid [OF modif [rule_format]]) -apply assumption -done + shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using dyn_call ret_modif ret_modifAbr modif + unfolding dynCall_dynCall_exn + by (rule dynProc_exnModifyReturn) + +lemma dynProc_exnModifyReturnSameFaults_sound: +assumes valid_call: "\n. \,\ \n:\<^bsub>/F\<^esub> P dynCall_exn f g init p return' result_exn c Q,A" +assumes valid_modif: + "\s \ P. \\. \n. + \,\\n:\<^bsub>/F\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" +assumes ret_modif: + "\s t. t \ Modif (init s) \ return' s t = return s t" +assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) \ result_exn (return' s t) t = result_exn (return s t) t" +shows "\,\ \n:\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" +proof (rule cnvalidI) + fix s t + assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" + assume exec: "\\\dynCall_exn f g init p return result_exn c,Normal s\ =n\ t" + assume t_notin_F: "t \ Fault ` F" + assume P: "s \ P" + with valid_modif + have valid_modif': "\\. \n. + \,\\n:\<^bsub>/F\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" + by blast + from exec + have exec_call: "\\\maybe_guard f g (call_exn init (p s) return result_exn c),Normal s\ =n\ t" + by (cases rule: execn_dynCall_exn_Normal_elim) + then show "t \ Normal ` Q \ Abrupt ` A" + proof (cases rule: execn_maybe_guard_Normal_elim_cases) + case noFault + from noFault have guards_ok: "s \ g" by simp + from noFault have "\\ \call_exn init (p s) return result_exn c,Normal s\ =n\ t" by simp + then show "t \ Normal ` Q \ Abrupt ` A" + + proof (cases rule: execn_call_exn_Normal_elim) + fix bdy m t' + assume bdy: "\ (p s) = Some bdy" + assume exec_body: "\\\bdy,Normal (init s)\ =m\ Normal t'" + assume exec_c: "\\\c s t',Normal (return s t')\ =Suc m\ t" + assume n: "n = Suc m" + from exec_body n bdy + have "\\\Call (p s) ,Normal (init s)\ =n \ Normal t'" + by (auto simp add: intro: execn.Call) + from cnvalidD [OF valid_modif' [rule_format, of n "init s"] ctxt this] P + have "t' \ Modif (init s)" + by auto + with ret_modif have "Normal (return' s t') = Normal (return s t')" + by simp + with exec_body exec_c bdy n + have "\\\call_exn init (p s) return' result_exn c,Normal s\ =n\ t" + by (auto intro: execn_call_exn) + from execn_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ =n\ t" + by (simp add: dynCall_exn_def execn_maybe_guard_DynCom) + from cnvalidD [OF valid_call ctxt this] P t_notin_F + show ?thesis + by simp + next + fix bdy m t' + assume bdy: "\ (p s) = Some bdy" + assume exec_body: "\\\bdy,Normal (init s)\ =m\ Abrupt t'" + assume n: "n = Suc m" + assume t: "t = Abrupt (result_exn (return s t') t')" + also from exec_body n bdy + have "\\\Call (p s) ,Normal (init s)\ =n \ Abrupt t'" + by (auto simp add: intro: execn.intros) + from cnvalidD [OF valid_modif' [rule_format, of n "init s"] ctxt this] P + have "t' \ ModifAbr (init s)" + by auto + with ret_modifAbr have "Abrupt (result_exn (return s t') t') = Abrupt (result_exn (return' s t') t')" + by simp + finally have "t = Abrupt (result_exn (return' s t') t')" . + with exec_body bdy n + have "\\\call_exn init (p s) return' result_exn c,Normal s\ =n\ t" + by (auto intro: execn_call_exnAbrupt) + from execn_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ =n\ t" + by (simp add: dynCall_exn_def execn_maybe_guard_DynCom) + from cnvalidD [OF valid_call ctxt this] P t_notin_F + show ?thesis + by simp + next + fix bdy m f' + assume bdy: "\ (p s) = Some bdy" + assume "\\\bdy,Normal (init s)\ =m\ Fault f'" "n = Suc m" and + t: "t = Fault f'" + with bdy have "\\\call_exn init (p s) return' result_exn c ,Normal s\ =n\ t" + by (auto intro: execn_call_exnFault) + from execn_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ =n\ t" + by (simp add: dynCall_exn_def execn_maybe_guard_DynCom) + from cnvalidD [OF valid_call ctxt this P] t t_notin_F + show ?thesis + by simp + next + fix bdy m + assume bdy: "\ (p s) = Some bdy" + assume "\\\bdy,Normal (init s)\ =m\ Stuck" "n = Suc m" + "t = Stuck" + with bdy have "\\\call_exn init (p s) return' result_exn c ,Normal s\ =n\ t" + by (auto intro: execn_call_exnStuck) + from execn_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ =n\ t" + by (simp add: dynCall_exn_def execn_maybe_guard_DynCom) + from valid_call ctxt this P t_notin_F + show ?thesis + by (rule cnvalidD) + next + fix m + assume "\ (p s) = None" + and "n = Suc m" "t = Stuck" + hence "\\\call_exn init (p s) return' result_exn c ,Normal s\ =n\ t" + by (auto intro: execn_call_exnUndefined) + from execn_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ =n\ t" + by (simp add: dynCall_exn_def execn_maybe_guard_DynCom) + from valid_call ctxt this P t_notin_F + show ?thesis + by (rule cnvalidD) + qed + next + case (someFault) + then obtain guards_fail:"s \ g" + and t: "t = Fault f" by simp + from execn_maybe_guard_Fault [OF guards_fail] t + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ =n\ t" + by (simp add: dynCall_exn_def execn_guards_DynCom) + from cnvalidD [OF valid_call ctxt this] P t_notin_F + show ?thesis by simp + qed +qed + lemma dynProcModifyReturnSameFaults_sound: assumes valid_call: "\n. \,\ \n:\<^bsub>/F\<^esub> P dynCall init p return' c Q,A" assumes valid_modif: "\s \ P. \\. \n. \,\\n:\<^bsub>/F\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" assumes ret_modif: "\s t. t \ Modif (init s) \ return' s t = return s t" assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) \ return' s t = return s t" shows "\,\ \n:\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" -proof (rule cnvalidI) - fix s t - assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" - assume exec: "\\\dynCall init p return c,Normal s\ =n\ t" - assume t_notin_F: "t \ Fault ` F" - assume P: "s \ P" - with valid_modif - have valid_modif': "\\. \n. - \,\\n:\<^bsub>/F\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" - by blast - from exec - have "\\\call init (p s) return c,Normal s\ =n\ t" - by (cases rule: execn_dynCall_Normal_elim) - then show "t \ Normal ` Q \ Abrupt ` A" - proof (cases rule: execn_call_Normal_elim) - fix bdy m t' - assume bdy: "\ (p s) = Some bdy" - assume exec_body: "\\\bdy,Normal (init s)\ =m\ Normal t'" - assume exec_c: "\\\c s t',Normal (return s t')\ =Suc m\ t" - assume n: "n = Suc m" - from exec_body n bdy - have "\\\Call (p s) ,Normal (init s)\ =n \ Normal t'" - by (auto simp add: intro: execn.Call) - from cnvalidD [OF valid_modif' [rule_format, of n "init s"] ctxt this] P - have "t' \ Modif (init s)" - by auto - with ret_modif have "Normal (return' s t') = Normal (return s t')" - by simp - with exec_body exec_c bdy n - have "\\\call init (p s) return' c,Normal s\ =n\ t" - by (auto intro: execn_call) - hence "\\\dynCall init p return' c,Normal s\ =n\ t" - by (rule execn_dynCall) - from cnvalidD [OF valid_call ctxt this] P t_notin_F - show ?thesis - by simp - next - fix bdy m t' - assume bdy: "\ (p s) = Some bdy" - assume exec_body: "\\\bdy,Normal (init s)\ =m\ Abrupt t'" - assume n: "n = Suc m" - assume t: "t = Abrupt (return s t')" - also from exec_body n bdy - have "\\\Call (p s) ,Normal (init s)\ =n \ Abrupt t'" - by (auto simp add: intro: execn.intros) - from cnvalidD [OF valid_modif' [rule_format, of n "init s"] ctxt this] P - have "t' \ ModifAbr (init s)" - by auto - with ret_modifAbr have "Abrupt (return s t') = Abrupt (return' s t')" - by simp - finally have "t = Abrupt (return' s t')" . - with exec_body bdy n - have "\\\call init (p s) return' c,Normal s\ =n\ t" - by (auto intro: execn_callAbrupt) - hence "\\\dynCall init p return' c,Normal s\ =n\ t" - by (rule execn_dynCall) - from cnvalidD [OF valid_call ctxt this] P t_notin_F - show ?thesis - by simp - next - fix bdy m f - assume bdy: "\ (p s) = Some bdy" - assume "\\\bdy,Normal (init s)\ =m\ Fault f" "n = Suc m" and - t: "t = Fault f" - with bdy have "\\\call init (p s) return' c ,Normal s\ =n\ t" - by (auto intro: execn_callFault) - hence "\\\dynCall init p return' c,Normal s\ =n\ t" - by (rule execn_dynCall) - from cnvalidD [OF valid_call ctxt this P] t t_notin_F - show ?thesis - by simp - next - fix bdy m - assume bdy: "\ (p s) = Some bdy" - assume "\\\bdy,Normal (init s)\ =m\ Stuck" "n = Suc m" - "t = Stuck" - with bdy have "\\\call init (p s) return' c ,Normal s\ =n\ t" - by (auto intro: execn_callStuck) - hence "\\\dynCall init p return' c,Normal s\ =n\ t" - by (rule execn_dynCall) - from valid_call ctxt this P t_notin_F - show ?thesis - by (rule cnvalidD) - next - fix m - assume "\ (p s) = None" - and "n = Suc m" "t = Stuck" - hence "\\\call init (p s) return' c ,Normal s\ =n\ t" - by (auto intro: execn_callUndefined) - hence "\\\dynCall init p return' c,Normal s\ =n\ t" - by (rule execn_dynCall) - from valid_call ctxt this P t_notin_F - show ?thesis - by (rule cnvalidD) - qed -qed + using valid_call valid_modif ret_modif ret_modifAbr + unfolding dynCall_dynCall_exn + by (rule dynProc_exnModifyReturnSameFaults_sound) + +lemma dynProc_exnModifyReturnSameFaults: +assumes dyn_call: "\,\\\<^bsub>/F\<^esub> P dynCall_exn f g init p return' result_exn c Q,A" +assumes ret_modif: + "\s t. t \ Modif (init s) + \ return' s t = return s t" +assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) + \ result_exn (return' s t) t = result_exn (return s t) t" +assumes modif: + "\s \ P. \\. \,\\\<^bsub>/F\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" +shows "\,\\\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" +apply (rule hoare_complete') +apply (rule allI) +apply (rule dynProc_exnModifyReturnSameFaults_sound + [where Modif=Modif and ModifAbr=ModifAbr, + OF hoare_cnvalid [OF dyn_call] _ ret_modif ret_modifAbr]) +apply (intro ballI allI) +apply (rule hoare_cnvalid [OF modif [rule_format]]) +apply assumption + done lemma dynProcModifyReturnSameFaults: assumes dyn_call: "\,\\\<^bsub>/F\<^esub> P dynCall init p return' c Q,A" assumes ret_modif: "\s t. t \ Modif (init s) \ return' s t = return s t" assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) \ return' s t = return s t" assumes modif: "\s \ P. \\. \,\\\<^bsub>/F\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" -shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" -apply (rule hoare_complete') -apply (rule allI) -apply (rule dynProcModifyReturnSameFaults_sound - [where Modif=Modif and ModifAbr=ModifAbr, - OF hoare_cnvalid [OF dyn_call] _ ret_modif ret_modifAbr]) -apply (intro ballI allI) -apply (rule hoare_cnvalid [OF modif [rule_format]]) -apply assumption -done + shows "\,\\\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using dyn_call ret_modif ret_modifAbr modif + unfolding dynCall_dynCall_exn + by (rule dynProc_exnModifyReturnSameFaults) subsubsection \Conjunction of Postcondition\ lemma PostConjI_sound: assumes valid_Q: "\n. \,\ \n:\<^bsub>/F\<^esub> P c Q,A" assumes valid_R: "\n. \,\ \n:\<^bsub>/F\<^esub> P c R,B" shows "\,\ \n:\<^bsub>/F\<^esub> P c (Q \ R),(A \ B)" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\c,Normal s\ =n\ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" from valid_Q [rule_format] ctxt exec P t_notin_F have "t \ Normal ` Q \ Abrupt ` A" by (rule cnvalidD) moreover from valid_R [rule_format] ctxt exec P t_notin_F have "t \ Normal ` R \ Abrupt ` B" by (rule cnvalidD) ultimately show "t \ Normal ` (Q \ R) \ Abrupt ` (A \ B)" by blast qed lemma PostConjI: assumes deriv_Q: "\,\\\<^bsub>/F\<^esub> P c Q,A" assumes deriv_R: "\,\\\<^bsub>/F\<^esub> P c R,B" shows "\,\\\<^bsub>/F\<^esub> P c (Q \ R),(A \ B)" apply (rule hoare_complete') apply (rule allI) apply (rule PostConjI_sound) using deriv_Q apply (blast intro: hoare_cnvalid) using deriv_R apply (blast intro: hoare_cnvalid) done lemma Merge_PostConj_sound: assumes validF: "\n. \,\\n:\<^bsub>/F\<^esub> P c Q,A" assumes validG: "\n. \,\\n:\<^bsub>/G\<^esub> P' c R,X" assumes F_G: "F \ G" assumes P_P': "P \ P'" shows "\,\\n:\<^bsub>/F\<^esub> P c (Q \ R),(A \ X)" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\n:\<^bsub>/F\<^esub> P (Call p) Q,A" with F_G have ctxt': "\(P, p, Q, A)\\. \\n:\<^bsub>/G\<^esub> P (Call p) Q,A" by (auto intro: nvalid_augment_Faults) assume exec: "\\\c,Normal s\ =n\ t" assume P: "s \ P" with P_P' have P': "s \ P'" by auto assume t_noFault: "t \ Fault ` F" show "t \ Normal ` (Q \ R) \ Abrupt ` (A \ X)" proof - from cnvalidD [OF validF [rule_format] ctxt exec P t_noFault] have *: "t \ Normal ` Q \ Abrupt ` A". then have "t \ Fault ` G" by auto from cnvalidD [OF validG [rule_format] ctxt' exec P' this] have "t \ Normal ` R \ Abrupt ` X" . with * show ?thesis by auto qed qed lemma Merge_PostConj: assumes validF: "\,\\\<^bsub>/F\<^esub> P c Q,A" assumes validG: "\,\\\<^bsub>/G\<^esub> P' c R,X" assumes F_G: "F \ G" assumes P_P': "P \ P'" shows "\,\\\<^bsub>/F\<^esub> P c (Q \ R),(A \ X)" apply (rule hoare_complete') apply (rule allI) apply (rule Merge_PostConj_sound [OF _ _ F_G P_P']) using validF apply (blast intro:hoare_cnvalid) using validG apply (blast intro:hoare_cnvalid) done subsubsection \Weaken Context\ lemma WeakenContext_sound: assumes valid_c: "\n. \,\'\n:\<^bsub>/F\<^esub> P c Q,A" assumes valid_ctxt: "\(P, p, Q, A)\\'. \,\\n:\<^bsub>/F\<^esub> P (Call p) Q,A" shows "\,\\n:\<^bsub>/F\<^esub> P c Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" with valid_ctxt have ctxt': "\(P, p, Q, A)\\'. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" by (simp add: cnvalid_def) assume exec: "\\\c,Normal s\ =n\ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" from valid_c [rule_format] ctxt' exec P t_notin_F show "t \ Normal ` Q \ Abrupt ` A" by (rule cnvalidD) qed lemma WeakenContext: assumes deriv_c: "\,\'\\<^bsub>/F\<^esub> P c Q,A" assumes deriv_ctxt: "\(P,p,Q,A)\\'. \,\\\<^bsub>/F\<^esub> P (Call p) Q,A" shows "\,\\\<^bsub>/F\<^esub> P c Q,A" apply (rule hoare_complete') apply (rule allI) apply (rule WeakenContext_sound) using deriv_c apply (blast intro: hoare_cnvalid) using deriv_ctxt apply (blast intro: hoare_cnvalid) done subsubsection \Guards and Guarantees\ lemma SplitGuards_sound: assumes valid_c1: "\n. \,\\n:\<^bsub>/F\<^esub> P c\<^sub>1 Q,A" assumes valid_c2: "\n. \,\\n:\<^bsub>/F\<^esub> P c\<^sub>2 UNIV,UNIV" assumes c: "(c\<^sub>1 \\<^sub>g c\<^sub>2) = Some c" shows "\,\\n:\<^bsub>/F\<^esub> P c Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\c,Normal s\ =n\ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" show "t \ Normal ` Q \ Abrupt ` A" proof (cases t) case Normal with inter_guards_execn_noFault [OF c exec] have "\\\c\<^sub>1,Normal s\ =n\ t" by simp from valid_c1 [rule_format] ctxt this P t_notin_F show ?thesis by (rule cnvalidD) next case Abrupt with inter_guards_execn_noFault [OF c exec] have "\\\c\<^sub>1,Normal s\ =n\ t" by simp from valid_c1 [rule_format] ctxt this P t_notin_F show ?thesis by (rule cnvalidD) next case (Fault f) with exec inter_guards_execn_Fault [OF c] have "\\\c\<^sub>1,Normal s\ =n\ Fault f \ \\\c\<^sub>2,Normal s\ =n\ Fault f" by auto then show ?thesis proof (cases rule: disjE [consumes 1]) assume "\\\c\<^sub>1,Normal s\ =n\ Fault f" from Fault cnvalidD [OF valid_c1 [rule_format] ctxt this P] t_notin_F show ?thesis by blast next assume "\\\c\<^sub>2,Normal s\ =n\ Fault f" from Fault cnvalidD [OF valid_c2 [rule_format] ctxt this P] t_notin_F show ?thesis by blast qed next case Stuck with inter_guards_execn_noFault [OF c exec] have "\\\c\<^sub>1,Normal s\ =n\ t" by simp from valid_c1 [rule_format] ctxt this P t_notin_F show ?thesis by (rule cnvalidD) qed qed lemma SplitGuards: assumes c: "(c\<^sub>1 \\<^sub>g c\<^sub>2) = Some c" assumes deriv_c1: "\,\\\<^bsub>/F\<^esub> P c\<^sub>1 Q,A" assumes deriv_c2: "\,\\\<^bsub>/F\<^esub> P c\<^sub>2 UNIV,UNIV" shows "\,\\\<^bsub>/F\<^esub> P c Q,A" apply (rule hoare_complete') apply (rule allI) apply (rule SplitGuards_sound [OF _ _ c]) using deriv_c1 apply (blast intro: hoare_cnvalid) using deriv_c2 apply (blast intro: hoare_cnvalid) done lemma CombineStrip_sound: assumes valid: "\n. \,\\n:\<^bsub>/F\<^esub> P c Q,A" assumes valid_strip: "\n. \,\\n:\<^bsub>/{}\<^esub> P (strip_guards (-F) c) UNIV,UNIV" shows "\,\\n:\<^bsub>/{}\<^esub> P c Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\n:\<^bsub>/{}\<^esub> P (Call p) Q,A" hence ctxt': "\(P, p, Q, A)\\. \\n:\<^bsub>/F\<^esub> P (Call p) Q,A" by (auto intro: nvalid_augment_Faults) assume exec: "\\\c,Normal s\ =n\ t" assume P: "s \ P" assume t_noFault: "t \ Fault ` {}" show "t \ Normal ` Q \ Abrupt ` A" proof (cases t) case (Normal t') from cnvalidD [OF valid [rule_format] ctxt' exec P] Normal show ?thesis by auto next case (Abrupt t') from cnvalidD [OF valid [rule_format] ctxt' exec P] Abrupt show ?thesis by auto next case (Fault f) show ?thesis proof (cases "f \ F") case True hence "f \ -F" by simp with exec Fault have "\\\strip_guards (-F) c,Normal s\ =n\ Fault f" by (auto intro: execn_to_execn_strip_guards_Fault) from cnvalidD [OF valid_strip [rule_format] ctxt this P] Fault have False by auto thus ?thesis .. next case False with cnvalidD [OF valid [rule_format] ctxt' exec P] Fault show ?thesis by auto qed next case Stuck from cnvalidD [OF valid [rule_format] ctxt' exec P] Stuck show ?thesis by auto qed qed lemma CombineStrip: assumes deriv: "\,\\\<^bsub>/F\<^esub> P c Q,A" assumes deriv_strip: "\,\\\<^bsub>/{}\<^esub> P (strip_guards (-F) c) UNIV,UNIV" shows "\,\\\<^bsub>/{}\<^esub> P c Q,A" apply (rule hoare_complete') apply (rule allI) apply (rule CombineStrip_sound) apply (iprover intro: hoare_cnvalid [OF deriv]) apply (iprover intro: hoare_cnvalid [OF deriv_strip]) done lemma GuardsFlip_sound: assumes valid: "\n. \,\\n:\<^bsub>/F\<^esub> P c Q,A" assumes validFlip: "\n. \,\\n:\<^bsub>/-F\<^esub> P c UNIV,UNIV" shows "\,\\n:\<^bsub>/{}\<^esub> P c Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\n:\<^bsub>/{}\<^esub> P (Call p) Q,A" hence ctxt': "\(P, p, Q, A)\\. \\n:\<^bsub>/F\<^esub> P (Call p) Q,A" by (auto intro: nvalid_augment_Faults) from ctxt have ctxtFlip: "\(P, p, Q, A)\\. \\n:\<^bsub>/-F\<^esub> P (Call p) Q,A" by (auto intro: nvalid_augment_Faults) assume exec: "\\\c,Normal s\ =n\ t" assume P: "s \ P" assume t_noFault: "t \ Fault ` {}" show "t \ Normal ` Q \ Abrupt ` A" proof (cases t) case (Normal t') from cnvalidD [OF valid [rule_format] ctxt' exec P] Normal show ?thesis by auto next case (Abrupt t') from cnvalidD [OF valid [rule_format] ctxt' exec P] Abrupt show ?thesis by auto next case (Fault f) show ?thesis proof (cases "f \ F") case True hence "f \ -F" by simp with cnvalidD [OF validFlip [rule_format] ctxtFlip exec P] Fault have False by auto thus ?thesis .. next case False with cnvalidD [OF valid [rule_format] ctxt' exec P] Fault show ?thesis by auto qed next case Stuck from cnvalidD [OF valid [rule_format] ctxt' exec P] Stuck show ?thesis by auto qed qed lemma GuardsFlip: assumes deriv: "\,\\\<^bsub>/F\<^esub> P c Q,A" assumes derivFlip: "\,\\\<^bsub>/-F\<^esub> P c UNIV,UNIV" shows "\,\\\<^bsub>/{}\<^esub> P c Q,A" apply (rule hoare_complete') apply (rule allI) apply (rule GuardsFlip_sound) apply (iprover intro: hoare_cnvalid [OF deriv]) apply (iprover intro: hoare_cnvalid [OF derivFlip]) done lemma MarkGuardsI_sound: assumes valid: "\n. \,\\n:\<^bsub>/{}\<^esub> P c Q,A" shows "\,\\n:\<^bsub>/{}\<^esub> P mark_guards f c Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\n:\<^bsub>/{}\<^esub> P (Call p) Q,A" assume exec: "\\\mark_guards f c,Normal s\ =n\ t" from execn_mark_guards_to_execn [OF exec] obtain t' where exec_c: "\\\c,Normal s\ =n\ t'" and t'_noFault: "\ isFault t' \ t' = t" by blast assume P: "s \ P" assume t_noFault: "t \ Fault ` {}" show "t \ Normal ` Q \ Abrupt ` A" proof - from cnvalidD [OF valid [rule_format] ctxt exec_c P] have "t' \ Normal ` Q \ Abrupt ` A" by blast with t'_noFault show ?thesis by auto qed qed lemma MarkGuardsI: assumes deriv: "\,\\\<^bsub>/{}\<^esub> P c Q,A" shows "\,\\\<^bsub>/{}\<^esub> P mark_guards f c Q,A" apply (rule hoare_complete') apply (rule allI) apply (rule MarkGuardsI_sound) apply (iprover intro: hoare_cnvalid [OF deriv]) done lemma MarkGuardsD_sound: assumes valid: "\n. \,\\n:\<^bsub>/{}\<^esub> P mark_guards f c Q,A" shows "\,\\n:\<^bsub>/{}\<^esub> P c Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\n:\<^bsub>/{}\<^esub> P (Call p) Q,A" assume exec: "\\\c,Normal s\ =n\ t" assume P: "s \ P" assume t_noFault: "t \ Fault ` {}" show "t \ Normal ` Q \ Abrupt ` A" proof (cases "isFault t") case True with execn_to_execn_mark_guards_Fault [OF exec ] obtain f' where "\\\mark_guards f c,Normal s\ =n\ Fault f'" by (fastforce elim: isFaultE) from cnvalidD [OF valid [rule_format] ctxt this P] have False by auto thus ?thesis .. next case False from execn_to_execn_mark_guards [OF exec False] obtain f' where "\\\mark_guards f c,Normal s\ =n\ t" by auto from cnvalidD [OF valid [rule_format] ctxt this P] show ?thesis by auto qed qed lemma MarkGuardsD: assumes deriv: "\,\\\<^bsub>/{}\<^esub> P mark_guards f c Q,A" shows "\,\\\<^bsub>/{}\<^esub> P c Q,A" apply (rule hoare_complete') apply (rule allI) apply (rule MarkGuardsD_sound) apply (iprover intro: hoare_cnvalid [OF deriv]) done lemma MergeGuardsI_sound: assumes valid: "\n. \,\\n:\<^bsub>/F\<^esub> P c Q,A" shows "\,\\n:\<^bsub>/F\<^esub> P merge_guards c Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec_merge: "\\\merge_guards c,Normal s\ =n\ t" from execn_merge_guards_to_execn [OF exec_merge] have exec: "\\\c,Normal s\ =n\ t" . assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" from cnvalidD [OF valid [rule_format] ctxt exec P t_notin_F] show "t \ Normal ` Q \ Abrupt ` A". qed lemma MergeGuardsI: assumes deriv: "\,\\\<^bsub>/F\<^esub> P c Q,A" shows "\,\\\<^bsub>/F\<^esub> P merge_guards c Q,A" apply (rule hoare_complete') apply (rule allI) apply (rule MergeGuardsI_sound) apply (iprover intro: hoare_cnvalid [OF deriv]) done lemma MergeGuardsD_sound: assumes valid: "\n. \,\\n:\<^bsub>/F\<^esub> P merge_guards c Q,A" shows "\,\\n:\<^bsub>/F\<^esub> P c Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\c,Normal s\ =n\ t" from execn_to_execn_merge_guards [OF exec] have exec_merge: "\\\merge_guards c,Normal s\ =n\ t". assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" from cnvalidD [OF valid [rule_format] ctxt exec_merge P t_notin_F] show "t \ Normal ` Q \ Abrupt ` A". qed lemma MergeGuardsD: assumes deriv: "\,\\\<^bsub>/F\<^esub> P merge_guards c Q,A" shows "\,\\\<^bsub>/F\<^esub> P c Q,A" apply (rule hoare_complete') apply (rule allI) apply (rule MergeGuardsD_sound) apply (iprover intro: hoare_cnvalid [OF deriv]) done lemma SubsetGuards_sound: assumes c_c': "c \\<^sub>g c'" assumes valid: "\n. \,\\n:\<^bsub>/{}\<^esub> P c' Q,A" shows "\,\\n:\<^bsub>/{}\<^esub> P c Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\n:\<^bsub>/{}\<^esub> P (Call p) Q,A" assume exec: "\\\c,Normal s\ =n\ t" from execn_to_execn_subseteq_guards [OF c_c' exec] obtain t' where exec_c': "\\\c',Normal s\ =n\ t'" and t'_noFault: "\ isFault t' \ t' = t" by blast assume P: "s \ P" assume t_noFault: "t \ Fault ` {}" from cnvalidD [OF valid [rule_format] ctxt exec_c' P] t'_noFault t_noFault show "t \ Normal ` Q \ Abrupt ` A" by auto qed lemma SubsetGuards: assumes c_c': "c \\<^sub>g c'" assumes deriv: "\,\\\<^bsub>/{}\<^esub> P c' Q,A" shows "\,\\\<^bsub>/{}\<^esub> P c Q,A" apply (rule hoare_complete') apply (rule allI) apply (rule SubsetGuards_sound [OF c_c']) apply (iprover intro: hoare_cnvalid [OF deriv]) done lemma NormalizeD_sound: assumes valid: "\n. \,\\n:\<^bsub>/F\<^esub> P (normalize c) Q,A" shows "\,\\n:\<^bsub>/F\<^esub> P c Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\c,Normal s\ =n\ t" hence exec_norm: "\\\normalize c,Normal s\ =n\ t" by (rule execn_to_execn_normalize) assume P: "s \ P" assume noFault: "t \ Fault ` F" from cnvalidD [OF valid [rule_format] ctxt exec_norm P noFault] show "t \ Normal ` Q \ Abrupt ` A". qed lemma NormalizeD: assumes deriv: "\,\\\<^bsub>/F\<^esub> P (normalize c) Q,A" shows "\,\\\<^bsub>/F\<^esub> P c Q,A" apply (rule hoare_complete') apply (rule allI) apply (rule NormalizeD_sound) apply (iprover intro: hoare_cnvalid [OF deriv]) done lemma NormalizeI_sound: assumes valid: "\n. \,\\n:\<^bsub>/F\<^esub> P c Q,A" shows "\,\\n:\<^bsub>/F\<^esub> P (normalize c) Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\n:\<^bsub>/F\<^esub> P (Call p) Q,A" assume "\\\normalize c,Normal s\ =n\ t" hence exec: "\\\c,Normal s\ =n\ t" by (rule execn_normalize_to_execn) assume P: "s \ P" assume noFault: "t \ Fault ` F" from cnvalidD [OF valid [rule_format] ctxt exec P noFault] show "t \ Normal ` Q \ Abrupt ` A". qed lemma NormalizeI: assumes deriv: "\,\\\<^bsub>/F\<^esub> P c Q,A" shows "\,\\\<^bsub>/F\<^esub> P (normalize c) Q,A" apply (rule hoare_complete') apply (rule allI) apply (rule NormalizeI_sound) apply (iprover intro: hoare_cnvalid [OF deriv]) done subsubsection \Restricting the Procedure Environment\ lemma nvalid_restrict_to_nvalid: assumes valid_c: "\|\<^bsub>M\<^esub>\n:\<^bsub>/F\<^esub> P c Q,A" shows "\\n:\<^bsub>/F\<^esub> P c Q,A" proof (rule nvalidI) fix s t assume exec: "\\\c,Normal s\ =n\ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" show "t \ Normal ` Q \ Abrupt ` A" proof - from execn_to_execn_restrict [OF exec] obtain t' where exec_res: "\|\<^bsub>M\<^esub>\\c,Normal s\ =n\ t'" and t_Fault: "\f. t = Fault f \ t' \ {Fault f, Stuck}" and t'_notStuck: "t'\Stuck \ t'=t" by blast from t_Fault t_notin_F t'_notStuck have "t' \ Fault ` F" by (cases t') auto with valid_c exec_res P have "t' \ Normal ` Q \ Abrupt ` A" by (auto simp add: nvalid_def) with t'_notStuck show ?thesis by auto qed qed lemma valid_restrict_to_valid: assumes valid_c: "\|\<^bsub>M\<^esub>\\<^bsub>/F\<^esub> P c Q,A" shows "\\\<^bsub>/F\<^esub> P c Q,A" proof (rule validI) fix s t assume exec: "\\\c,Normal s\ \ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" show "t \ Normal ` Q \ Abrupt ` A" proof - from exec_to_exec_restrict [OF exec] obtain t' where exec_res: "\|\<^bsub>M\<^esub>\\c,Normal s\ \ t'" and t_Fault: "\f. t = Fault f \ t' \ {Fault f, Stuck}" and t'_notStuck: "t'\Stuck \ t'=t" by blast from t_Fault t_notin_F t'_notStuck have "t' \ Fault ` F" by (cases t') auto with valid_c exec_res P have "t' \ Normal ` Q \ Abrupt ` A" by (auto simp add: valid_def) with t'_notStuck show ?thesis by auto qed qed lemma augment_procs: assumes deriv_c: "\|\<^bsub>M\<^esub>,{}\\<^bsub>/F\<^esub> P c Q,A" shows "\,{}\\<^bsub>/F\<^esub> P c Q,A" apply (rule hoare_complete) apply (rule valid_restrict_to_valid) apply (insert hoare_sound [OF deriv_c]) by (simp add: cvalid_def) lemma augment_Faults: assumes deriv_c: "\,{}\\<^bsub>/F\<^esub> P c Q,A" assumes F: "F \ F'" shows "\,{}\\<^bsub>/F'\<^esub> P c Q,A" apply (rule hoare_complete) apply (rule valid_augment_Faults [OF _ F]) apply (insert hoare_sound [OF deriv_c]) by (simp add: cvalid_def) end diff --git a/thys/Simpl/HoareTotal.thy b/thys/Simpl/HoareTotal.thy --- a/thys/Simpl/HoareTotal.thy +++ b/thys/Simpl/HoareTotal.thy @@ -1,1231 +1,1430 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: HoareTotal.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. +Copyright (c) 2022 Apple Inc. All rights reserved. -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA *) section \Derived Hoare Rules for Total Correctness\ theory HoareTotal imports HoareTotalProps begin lemma conseq_no_aux: "\\,\ \\<^sub>t\<^bsub>/F\<^esub> P' c Q',A'; \s. s \ P \ (s\P' \ (Q' \ Q)\ (A' \ A))\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (rule conseq [where P'="\Z. P'" and Q'="\Z. Q'" and A'="\Z. A'"]) auto text \If for example a specification for a "procedure pointer" parameter is in the precondition we can extract it with this rule\ lemma conseq_exploit_pre: "\\s \ P. \,\ \\<^sub>t\<^bsub>/F\<^esub> ({s} \ P) c Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" apply (rule Conseq) apply clarify apply (rule_tac x="{s} \ P" in exI) apply (rule_tac x="Q" in exI) apply (rule_tac x="A" in exI) by simp lemma conseq:"\\Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z); \s. s \ P \ (\ Z. s\P' Z \ (Q' Z \ Q)\ (A' Z \ A))\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (rule Conseq') blast lemma Lem:"\\Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z); P \ {s. \ Z. s\P' Z \ (Q' Z \ Q) \ (A' Z \ A)}\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (lem x c) Q,A" apply (unfold lem_def) apply (erule conseq) apply blast done lemma LemAnno: assumes conseq: "P \ {s. \Z. s\P' Z \ (\t. t \ Q' Z \ t \ Q) \ (\t. t \ A' Z \ t \ A)}" assumes lem: "\Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z)" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (lem x c) Q,A" apply (rule Lem [OF lem]) using conseq by blast lemma LemAnnoNoAbrupt: assumes conseq: "P \ {s. \Z. s\P' Z \ (\t. t \ Q' Z \ t \ Q)}" assumes lem: "\Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> (P' Z) c (Q' Z),{}" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (lem x c) Q,{}" apply (rule Lem [OF lem]) using conseq by blast lemma TrivPost: "\Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z) \ \Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> (P' Z) c UNIV,UNIV" apply (rule allI) apply (erule conseq) apply auto done lemma TrivPostNoAbr: "\Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> (P' Z) c (Q' Z),{} \ \Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> (P' Z) c UNIV,{}" apply (rule allI) apply (erule conseq) apply auto done lemma DynComConseq: assumes "P \ {s. \P' Q' A'. \,\\\<^sub>t\<^bsub>/F \<^esub>P' (c s) Q',A' \ P \ P' \ Q' \ Q \ A' \ A}" shows "\,\\\<^sub>t\<^bsub>/F \<^esub>P DynCom c Q,A" using assms apply - apply (rule hoaret.DynCom) apply clarsimp apply (rule hoaret.Conseq) apply clarsimp apply blast done lemma SpecAnno: assumes consequence: "P \ {s. (\ Z. s\P' Z \ (Q' Z \ Q) \ (A' Z \ A))}" assumes spec: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) (c Z) (Q' Z),(A' Z)" assumes bdy_constant: "\Z. c Z = c undefined" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (specAnno P' c Q' A') Q,A" proof - from spec bdy_constant have "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) (c undefined) (Q' Z),(A' Z)" apply - apply (rule allI) apply (erule_tac x=Z in allE) apply (erule_tac x=Z in allE) apply simp done with consequence show ?thesis apply (simp add: specAnno_def) apply (erule conseq) apply blast done qed lemma SpecAnno': "\P \ {s. \ Z. s\P' Z \ (\t. t \ Q' Z \ t \ Q) \ (\t. t \ A' Z \ t \ A)}; \Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) (c Z) (Q' Z),(A' Z); \Z. c Z = c undefined \ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (specAnno P' c Q' A') Q,A" apply (simp only: subset_iff [THEN sym]) apply (erule (1) SpecAnno) apply assumption done lemma SpecAnnoNoAbrupt: "\P \ {s. \ Z. s\P' Z \ (\t. t \ Q' Z \ t \ Q)}; \Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) (c Z) (Q' Z),{}; \Z. c Z = c undefined \ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (specAnno P' c Q' (\s. {})) Q,A" apply (rule SpecAnno') apply auto done lemma Skip: "P \ Q \ \,\\\<^sub>t\<^bsub>/F\<^esub> P Skip Q,A" by (rule hoaret.Skip [THEN conseqPre],simp) lemma Basic: "P \ {s. (f s) \ Q} \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (Basic f) Q,A" by (rule hoaret.Basic [THEN conseqPre]) lemma BasicCond: "\P \ {s. (b s \ f s\Q) \ (\ b s \ g s\Q)}\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P Basic (\s. if b s then f s else g s) Q,A" apply (rule Basic) apply auto done lemma Spec: "P \ {s. (\t. (s,t) \ r \ t \ Q) \ (\t. (s,t) \ r)} \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (Spec r) Q,A" by (rule hoaret.Spec [THEN conseqPre]) lemma SpecIf: "\P \ {s. (b s \ f s \ Q) \ (\ b s \ g s \ Q \ h s \ Q)}\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P Spec (if_rel b f g h) Q,A" apply (rule Spec) apply (auto simp add: if_rel_def) done lemma Seq [trans, intro?]: "\\,\\\<^sub>t\<^bsub>/F\<^esub> P c\<^sub>1 R,A; \,\\\<^sub>t\<^bsub>/F\<^esub> R c\<^sub>2 Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P Seq c\<^sub>1 c\<^sub>2 Q,A" by (rule hoaret.Seq) lemma SeqSwap: "\\,\\\<^sub>t\<^bsub>/F\<^esub> R c2 Q,A; \,\\\<^sub>t\<^bsub>/F\<^esub> P c1 R,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P Seq c1 c2 Q,A" by (rule Seq) lemma BSeq: "\\,\\\<^sub>t\<^bsub>/F\<^esub> P c\<^sub>1 R,A; \,\\\<^sub>t\<^bsub>/F\<^esub> R c\<^sub>2 Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (bseq c\<^sub>1 c\<^sub>2) Q,A" by (unfold bseq_def) (rule Seq) lemma Cond: assumes wp: "P \ {s. (s\b \ s\P\<^sub>1) \ (s\b \ s\P\<^sub>2)}" assumes deriv_c1: "\,\\\<^sub>t\<^bsub>/F\<^esub> P\<^sub>1 c\<^sub>1 Q,A" assumes deriv_c2: "\,\\\<^sub>t\<^bsub>/F\<^esub> P\<^sub>2 c\<^sub>2 Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (Cond b c\<^sub>1 c\<^sub>2) Q,A" proof (rule hoaret.Cond [THEN conseqPre]) from deriv_c1 show "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. (s \ b \ s \ P\<^sub>1) \ (s \ b \ s \ P\<^sub>2)} \ b) c\<^sub>1 Q,A" by (rule conseqPre) blast next from deriv_c2 show "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. (s \ b \ s \ P\<^sub>1) \ (s \ b \ s \ P\<^sub>2)} \ - b) c\<^sub>2 Q,A" by (rule conseqPre) blast qed (insert wp) lemma CondSwap: "\\,\\\<^sub>t\<^bsub>/F\<^esub> P1 c1 Q,A; \,\\\<^sub>t\<^bsub>/F\<^esub> P2 c2 Q,A; P \ {s. (s\b \ s\P1) \ (s\b \ s\P2)}\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (Cond b c1 c2) Q,A" by (rule Cond) lemma Cond': "\P \ {s. (b \ P1) \ (- b \ P2)};\,\\\<^sub>t\<^bsub>/F\<^esub> P1 c1 Q,A; \,\\\<^sub>t\<^bsub>/F\<^esub> P2 c2 Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (Cond b c1 c2) Q,A" by (rule CondSwap) blast+ lemma CondInv: assumes wp: "P \ Q" assumes inv: "Q \ {s. (s\b \ s\P\<^sub>1) \ (s\b \ s\P\<^sub>2)}" assumes deriv_c1: "\,\\\<^sub>t\<^bsub>/F\<^esub> P\<^sub>1 c\<^sub>1 Q,A" assumes deriv_c2: "\,\\\<^sub>t\<^bsub>/F\<^esub> P\<^sub>2 c\<^sub>2 Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (Cond b c\<^sub>1 c\<^sub>2) Q,A" proof - from wp inv have "P \ {s. (s\b \ s\P\<^sub>1) \ (s\b \ s\P\<^sub>2)}" by blast from Cond [OF this deriv_c1 deriv_c2] show ?thesis . qed lemma CondInv': assumes wp: "P \ I" assumes inv: "I \ {s. (s\b \ s\P\<^sub>1) \ (s\b \ s\P\<^sub>2)}" assumes wp': "I \ Q" assumes deriv_c1: "\,\\\<^sub>t\<^bsub>/F\<^esub> P\<^sub>1 c\<^sub>1 I,A" assumes deriv_c2: "\,\\\<^sub>t\<^bsub>/F\<^esub> P\<^sub>2 c\<^sub>2 I,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (Cond b c\<^sub>1 c\<^sub>2) Q,A" proof - from CondInv [OF wp inv deriv_c1 deriv_c2] have "\,\\\<^sub>t\<^bsub>/F\<^esub> P (Cond b c\<^sub>1 c\<^sub>2) I,A" . from conseqPost [OF this wp' subset_refl] show ?thesis . qed lemma switchNil: "P \ Q \ \,\\\<^sub>t\<^bsub>/F \<^esub>P (switch v []) Q,A" by (simp add: Skip) lemma switchCons: "\P \ {s. (v s \ V \ s \ P\<^sub>1) \ (v s \ V \ s \ P\<^sub>2)}; \,\\\<^sub>t\<^bsub>/F \<^esub>P\<^sub>1 c Q,A; \,\\\<^sub>t\<^bsub>/F \<^esub>P\<^sub>2 (switch v vs) Q,A\ \ \,\\\<^sub>t\<^bsub>/F \<^esub>P (switch v ((V,c)#vs)) Q,A" by (simp add: Cond) lemma Guard: "\P \ g \ R; \,\\\<^sub>t\<^bsub>/F\<^esub> R c Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P Guard f g c Q,A" apply (rule HoareTotalDef.Guard [THEN conseqPre, of _ _ _ _ R]) apply (erule conseqPre) apply auto done lemma GuardSwap: "\ \,\\\<^sub>t\<^bsub>/F\<^esub> R c Q,A; P \ g \ R\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P Guard f g c Q,A" by (rule Guard) lemma Guarantee: "\P \ {s. s \ g \ s \ R}; \,\\\<^sub>t\<^bsub>/F\<^esub> R c Q,A; f \ F\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (Guard f g c) Q,A" apply (rule Guarantee [THEN conseqPre, of _ _ _ _ _ "{s. s \ g \ s \ R}"]) apply assumption apply (erule conseqPre) apply auto done lemma GuaranteeSwap: "\ \,\\\<^sub>t\<^bsub>/F\<^esub> R c Q,A; P \ {s. s \ g \ s \ R}; f \ F\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (Guard f g c) Q,A" by (rule Guarantee) lemma GuardStrip: "\P \ R; \,\\\<^sub>t\<^bsub>/F\<^esub> R c Q,A; f \ F\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (Guard f g c) Q,A" apply (rule Guarantee [THEN conseqPre]) apply auto done lemma GuardStripSwap: "\\,\\\<^sub>t\<^bsub>/F\<^esub> R c Q,A; P \ R; f \ F\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (Guard f g c) Q,A" by (rule GuardStrip) lemma GuaranteeStrip: "\P \ R; \,\\\<^sub>t\<^bsub>/F\<^esub> R c Q,A; f \ F\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (guaranteeStrip f g c) Q,A" by (unfold guaranteeStrip_def) (rule GuardStrip) lemma GuaranteeStripSwap: "\\,\\\<^sub>t\<^bsub>/F\<^esub> R c Q,A; P \ R; f \ F\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (guaranteeStrip f g c) Q,A" by (unfold guaranteeStrip_def) (rule GuardStrip) lemma GuaranteeAsGuard: "\P \ g \ R; \,\\\<^sub>t\<^bsub>/F\<^esub> R c Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P guaranteeStrip f g c Q,A" by (unfold guaranteeStrip_def) (rule Guard) lemma GuaranteeAsGuardSwap: "\ \,\\\<^sub>t\<^bsub>/F\<^esub> R c Q,A; P \ g \ R\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P guaranteeStrip f g c Q,A" by (rule GuaranteeAsGuard) lemma GuardsNil: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (guards [] c) Q,A" by simp lemma GuardsCons: "\,\\\<^sub>t\<^bsub>/F\<^esub> P Guard f g (guards gs c) Q,A \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (guards ((f,g)#gs) c) Q,A" by simp lemma GuardsConsGuaranteeStrip: "\,\\\<^sub>t\<^bsub>/F\<^esub> P guaranteeStrip f g (guards gs c) Q,A \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (guards (guaranteeStripPair f g#gs) c) Q,A" by (simp add: guaranteeStripPair_def guaranteeStrip_def) lemma While: assumes P_I: "P \ I" assumes deriv_body: "\\. \,\\\<^sub>t\<^bsub>/F\<^esub> ({\} \ I \ b) c ({t. (t, \) \ V} \ I),A" assumes I_Q: "I \ -b \ Q" assumes wf: "wf V" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (whileAnno b I V c) Q,A" proof - from wf deriv_body P_I I_Q show ?thesis apply (unfold whileAnno_def) apply (erule conseqPrePost [OF HoareTotalDef.While]) apply auto done qed lemma WhileInvPost: assumes P_I: "P \ I" assumes termi_body: "\\. \,\\\<^sub>t\<^bsub>/UNIV\<^esub> ({\} \ I \ b) c ({t. (t, \) \ V} \ P),A" assumes deriv_body: "\,\\\<^bsub>/F\<^esub> (I \ b) c I,A" assumes I_Q: "I \ -b \ Q" assumes wf: "wf V" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (whileAnno b I V c) Q,A" proof - have "\\. \,\\\<^sub>t\<^bsub>/F\<^esub> ({\} \ I \ b) c ({t. (t, \) \ V} \ I),A" proof fix \ from hoare_sound [OF deriv_body] hoaret_sound [OF termi_body [rule_format, of \]] have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({\} \ I \ b) c ({t. (t, \) \ V} \ I),A" by (fastforce simp add: cvalidt_def validt_def cvalid_def valid_def) then show "\,\\\<^sub>t\<^bsub>/F\<^esub> ({\} \ I \ b) c ({t. (t, \) \ V} \ I),A" by (rule hoaret_complete') qed from While [OF P_I this I_Q wf] show ?thesis . qed (* *) lemma "\,\\\<^bsub>/F\<^esub> (P \ b) c Q,A \ \,\\\<^bsub>/F\<^esub> (P \ b) (Seq c (Guard f Q Skip)) Q,A" oops text \@{term "J"} will be instantiated by tactic with @{term "gs' \ I"} for those guards that are not stripped.\ lemma WhileAnnoG: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (guards gs (whileAnno b J V (Seq c (guards gs Skip)))) Q,A \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (whileAnnoG gs b I V c) Q,A" by (simp add: whileAnnoG_def whileAnno_def while_def) text \This form stems from @{term "strip_guards F (whileAnnoG gs b I V c)"}\ lemma WhileNoGuard': assumes P_I: "P \ I" assumes deriv_body: "\\. \,\\\<^sub>t\<^bsub>/F\<^esub> ({\} \ I \ b) c ({t. (t, \) \ V} \ I),A" assumes I_Q: "I \ -b \ Q" assumes wf: "wf V" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (whileAnno b I V (Seq c Skip)) Q,A" apply (rule While [OF P_I _ I_Q wf]) apply (rule allI) apply (rule Seq) apply (rule deriv_body [rule_format]) apply (rule hoaret.Skip) done lemma WhileAnnoFix: assumes consequence: "P \ {s. (\ Z. s\I Z \ (I Z \ -b \ Q)) }" assumes bdy: "\Z \. \,\\\<^sub>t\<^bsub>/F\<^esub> ({\} \ I Z \ b) (c Z) ({t. (t, \) \ V Z} \ I Z),A" assumes bdy_constant: "\Z. c Z = c undefined" assumes wf: "\Z. wf (V Z)" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (whileAnnoFix b I V c) Q,A" proof - from bdy bdy_constant have bdy': "\Z. \\. \,\\\<^sub>t\<^bsub>/F\<^esub> ({\} \ I Z \ b) (c undefined) ({t. (t, \) \ V Z} \ I Z),A" apply - apply (erule_tac x=Z in allE) apply (erule_tac x=Z in allE) apply simp done have "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (I Z) (whileAnnoFix b I V c) (I Z \ -b),A" apply rule apply (unfold whileAnnoFix_def) apply (rule hoaret.While) apply (rule wf [rule_format]) apply (rule bdy') done then show ?thesis apply (rule conseq) using consequence by blast qed lemma WhileAnnoFix': assumes consequence: "P \ {s. (\ Z. s\I Z \ (\t. t \ I Z \ -b \ t \ Q)) }" assumes bdy: "\Z \. \,\\\<^sub>t\<^bsub>/F\<^esub> ({\} \ I Z \ b) (c Z) ({t. (t, \) \ V Z} \ I Z),A" assumes bdy_constant: "\Z. c Z = c undefined" assumes wf: "\Z. wf (V Z)" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (whileAnnoFix b I V c) Q,A" apply (rule WhileAnnoFix [OF _ bdy bdy_constant wf]) using consequence by blast lemma WhileAnnoGFix: assumes whileAnnoFix: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (guards gs (whileAnnoFix b J V (\Z. (Seq (c Z) (guards gs Skip))))) Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (whileAnnoGFix gs b I V c) Q,A" using whileAnnoFix by (simp add: whileAnnoGFix_def whileAnnoFix_def while_def) lemma Bind: assumes adapt: "P \ {s. s \ P' s}" assumes c: "\s. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' s) (c (e s)) Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (bind e c) Q,A" apply (rule conseq [where P'="\Z. {s. s=Z \ s \ P' Z}" and Q'="\Z. Q" and A'="\Z. A"]) apply (rule allI) apply (unfold bind_def) apply (rule HoareTotalDef.DynCom) apply (rule ballI) apply clarsimp apply (rule conseqPre) apply (rule c [rule_format]) apply blast using adapt apply blast done -lemma Block: +lemma Block_exn: assumes adapt: "P \ {s. init s \ P' s}" -assumes bdy: "\s. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' s) bdy {t. return s t \ R s t},{t. return s t \ A}" +assumes bdy: "\s. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' s) bdy {t. return s t \ R s t},{t. result_exn (return s t) t \ A}" assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" -shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (block init bdy return c) Q,A" +shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (block_exn init bdy return result_exn c) Q,A" apply (rule conseq [where P'="\Z. {s. s=Z \ init s \ P' Z}" and Q'="\Z. Q" and A'="\Z. A"]) prefer 2 using adapt apply blast apply (rule allI) -apply (unfold block_def) +apply (unfold block_exn_def) apply (rule HoareTotalDef.DynCom) apply (rule ballI) apply clarsimp apply (rule_tac R="{t. return Z t \ R Z t}" in SeqSwap ) apply (rule_tac P'="\Z'. {t. t=Z' \ return Z t \ R Z t}" and Q'="\Z'. Q" and A'="\Z'. A" in conseq) prefer 2 apply simp apply (rule allI) apply (rule HoareTotalDef.DynCom) apply (clarsimp) apply (rule SeqSwap) apply (rule c [rule_format]) apply (rule Basic) apply clarsimp -apply (rule_tac R="{t. return Z t \ A}" in HoareTotalDef.Catch) +apply (rule_tac R="{t. result_exn (return Z t) t \ A}" in HoareTotalDef.Catch) apply (rule_tac R="{i. i \ P' Z}" in Seq) apply (rule Basic) apply clarsimp apply simp apply (rule bdy [rule_format]) apply (rule SeqSwap) apply (rule Throw) apply (rule Basic) apply simp done +lemma Block: +assumes adapt: "P \ {s. init s \ P' s}" +assumes bdy: "\s. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' s) bdy {t. return s t \ R s t},{t. return s t \ A}" +assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" +shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (block init bdy return c) Q,A" + using adapt bdy c unfolding block_def by (rule Block_exn) + lemma BlockSwap: assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes bdy: "\s. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' s) bdy {t. return s t \ R s t},{t. return s t \ A}" assumes adapt: "P \ {s. init s \ P' s}" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (block init bdy return c) Q,A" using adapt bdy c by (rule Block) -lemma BlockSpec: +lemma Block_exnSwap: +assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" +assumes bdy: "\s. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' s) bdy {t. return s t \ R s t},{t. result_exn (return s t) t \ A}" +assumes adapt: "P \ {s. init s \ P' s}" +shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (block_exn init bdy return result_exn c) Q,A" + using adapt bdy c + by (rule Block_exn) + +lemma Block_exnSpec: assumes adapt: "P \ {s. \Z. init s \ P' Z \ (\t. t \ Q' Z \ return s t \ R s t) \ - (\t. t \ A' Z \ return s t \ A)}" + (\t. t \ A' Z \ result_exn (return s t) t \ A)}" assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes bdy: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) bdy (Q' Z),(A' Z)" - shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (block init bdy return c) Q,A" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (block_exn init bdy return result_exn c) Q,A" apply (rule conseq [where P'="\Z. {s. init s \ P' Z \ (\t. t \ Q' Z \ return s t \ R s t) \ - (\t. t \ A' Z \ return s t \ A)}" and Q'="\Z. Q" and + (\t. t \ A' Z \ result_exn (return s t) t \ A)}" and Q'="\Z. Q" and A'="\Z. A"]) prefer 2 using adapt apply blast apply (rule allI) -apply (unfold block_def) +apply (unfold block_exn_def) apply (rule HoareTotalDef.DynCom) apply (rule ballI) apply clarsimp apply (rule_tac R="{t. return s t \ R s t}" in SeqSwap ) apply (rule_tac P'="\Z'. {t. t=Z' \ return s t \ R s t}" and Q'="\Z'. Q" and A'="\Z'. A" in conseq) prefer 2 apply simp apply (rule allI) apply (rule HoareTotalDef.DynCom) apply (clarsimp) apply (rule SeqSwap) apply (rule c [rule_format]) apply (rule Basic) apply clarsimp -apply (rule_tac R="{t. return s t \ A}" in HoareTotalDef.Catch) +apply (rule_tac R="{t. result_exn (return s t) t \ A}" in HoareTotalDef.Catch) apply (rule_tac R="{i. i \ P' Z}" in Seq) apply (rule Basic) apply clarsimp apply simp apply (rule conseq [OF bdy]) apply clarsimp apply blast apply (rule SeqSwap) apply (rule Throw) apply (rule Basic) apply simp done +lemma BlockSpec: + assumes adapt: "P \ {s. \Z. init s \ P' Z \ + (\t. t \ Q' Z \ return s t \ R s t) \ + (\t. t \ A' Z \ return s t \ A)}" + assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" + assumes bdy: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) bdy (Q' Z),(A' Z)" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (block init bdy return c) Q,A" + using adapt c bdy unfolding block_def by (rule Block_exnSpec) + lemma Throw: "P \ A \ \,\\\<^sub>t\<^bsub>/F\<^esub> P Throw Q,A" by (rule hoaret.Throw [THEN conseqPre]) lemmas Catch = hoaret.Catch lemma CatchSwap: "\\,\\\<^sub>t\<^bsub>/F\<^esub> R c\<^sub>2 Q,A; \,\\\<^sub>t\<^bsub>/F\<^esub> P c\<^sub>1 Q,R\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P Catch c\<^sub>1 c\<^sub>2 Q,A" by (rule hoaret.Catch) lemma raise: "P \ {s. f s \ A} \ \,\\\<^sub>t\<^bsub>/F\<^esub> P raise f Q,A" apply (simp add: raise_def) apply (rule Seq) apply (rule Basic) apply (assumption) apply (rule Throw) apply (rule subset_refl) done lemma condCatch: "\\,\\\<^sub>t\<^bsub>/F\<^esub> P c\<^sub>1 Q,((b \ R) \ (-b \ A));\,\\\<^sub>t\<^bsub>/F\<^esub> R c\<^sub>2 Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P condCatch c\<^sub>1 b c\<^sub>2 Q,A" apply (simp add: condCatch_def) apply (rule Catch) apply assumption apply (rule CondSwap) apply (assumption) apply (rule hoaret.Throw) apply blast done lemma condCatchSwap: "\\,\\\<^sub>t\<^bsub>/F\<^esub> R c\<^sub>2 Q,A; \,\\\<^sub>t\<^bsub>/F\<^esub> P c\<^sub>1 Q,((b \ R) \ (-b \ A))\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P condCatch c\<^sub>1 b c\<^sub>2 Q,A" by (rule condCatch) +lemma Proc_exnSpec: + assumes adapt: "P \ {s. \Z. init s \ P' Z \ + (\t. t \ Q' Z \ return s t \ R s t) \ + (\t. t \ A' Z \ result_exn (return s t) t \ A)}" + assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" + assumes p: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) Call p (Q' Z),(A' Z)" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" +using adapt c p +apply (unfold call_exn_def) + by (rule Block_exnSpec) lemma ProcSpec: assumes adapt: "P \ {s. \Z. init s \ P' Z \ (\t. t \ Q' Z \ return s t \ R s t) \ (\t. t \ A' Z \ return s t \ A)}" assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes p: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) Call p (Q' Z),(A' Z)" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return c) Q,A" using adapt c p apply (unfold call_def) by (rule BlockSpec) -lemma ProcSpec': +lemma Proc_exnSpec': assumes adapt: "P \ {s. \Z. init s \ P' Z \ (\t \ Q' Z. return s t \ R s t) \ - (\t \ A' Z. return s t \ A)}" + (\t \ A' Z. result_exn (return s t) t \ A)}" assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes p: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) Call p (Q' Z),(A' Z)" - shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return c) Q,A" -apply (rule ProcSpec [OF _ c p]) + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" +apply (rule Proc_exnSpec [OF _ c p]) apply (insert adapt) apply clarsimp apply (drule (1) subsetD) apply (clarsimp) apply (rule_tac x=Z in exI) apply blast done +lemma ProcSpec': + assumes adapt: "P \ {s. \Z. init s \ P' Z \ + (\t \ Q' Z. return s t \ R s t) \ + (\t \ A' Z. return s t \ A)}" + assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" + assumes p: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) Call p (Q' Z),(A' Z)" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return c) Q,A" + using adapt c p unfolding call_call_exn by (rule Proc_exnSpec') + + +lemma Proc_exnSpecNoAbrupt: + assumes adapt: "P \ {s. \Z. init s \ P' Z \ + (\t. t \ Q' Z \ return s t \ R s t)}" + assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" + assumes p: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) Call p (Q' Z),{}" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" +apply (rule Proc_exnSpec [OF _ c p]) +using adapt +apply simp + done lemma ProcSpecNoAbrupt: assumes adapt: "P \ {s. \Z. init s \ P' Z \ (\t. t \ Q' Z \ return s t \ R s t)}" assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes p: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) Call p (Q' Z),{}" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return c) Q,A" apply (rule ProcSpec [OF _ c p]) using adapt apply simp done lemma FCall: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return (\s t. c (result t))) Q,A \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (fcall init p return result c) Q,A" by (simp add: fcall_def) lemma ProcRec: assumes deriv_bodies: "\p\Procs. \\ Z. \,\\(\q\Procs. \Z. {(P q Z \ {s. ((s,q), \,p) \ r},q,Q q Z,A q Z)}) \\<^sub>t\<^bsub>/F\<^esub> ({\} \ P p Z) (the (\ p)) (Q p Z),(A p Z)" assumes wf: "wf r" assumes Procs_defined: "Procs \ dom \" shows "\p\Procs. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub>(P p Z) Call p (Q p Z),(A p Z)" by (intro strip) (rule HoareTotalDef.CallRec' [OF _ Procs_defined wf deriv_bodies], simp_all) lemma ProcRec': assumes ctxt: "\'=(\\ p. \\(\q\Procs. \Z. {(P q Z \ {s. ((s,q), \,p) \ r},q,Q q Z,A q Z)}))" assumes deriv_bodies: "\p\Procs. \\ Z. \,\' \ p\\<^sub>t\<^bsub>/F\<^esub> ({\} \ P p Z) (the (\ p)) (Q p Z),(A p Z)" assumes wf: "wf r" assumes Procs_defined: "Procs \ dom \" shows "\p\Procs. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub>(P p Z) Call p (Q p Z),(A p Z)" using ctxt deriv_bodies apply simp apply (erule ProcRec [OF _ wf Procs_defined]) done lemma ProcRecList: assumes deriv_bodies: "\p\set Procs. \\ Z. \,\\(\q\set Procs. \Z. {(P q Z \ {s. ((s,q), \,p) \ r},q,Q q Z,A q Z)}) \\<^sub>t\<^bsub>/F\<^esub> ({\} \ P p Z) (the (\ p)) (Q p Z),(A p Z)" assumes wf: "wf r" assumes dist: "distinct Procs" assumes Procs_defined: "set Procs \ dom \" shows "\p\set Procs. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub>(P p Z) Call p (Q p Z),(A p Z)" using deriv_bodies wf Procs_defined by (rule ProcRec) lemma ProcRecSpecs: "\\\. \(P,p,Q,A) \ Specs. \,\\ ((\(P,q,Q,A). (P \ {s. ((s,q),(\,p)) \ r},q,Q,A)) ` Specs) \\<^sub>t\<^bsub>/F\<^esub> ({\} \ P) (the (\ p)) Q,A; wf r; \(P,p,Q,A) \ Specs. p \ dom \\ \ \(P,p,Q,A) \ Specs. \,\\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" apply (rule ballI) apply (case_tac x) apply (rename_tac x P p Q A) apply simp apply (rule hoaret.CallRec) apply auto done lemma ProcRec1: assumes deriv_body: "\\ Z. \,\\(\Z. {(P Z \ {s. ((s,p), \,p) \ r},p,Q Z,A Z)}) \\<^sub>t\<^bsub>/F\<^esub> ({\} \ P Z) (the (\ p)) (Q Z),(A Z)" assumes wf: "wf r" assumes p_defined: "p \ dom \" shows "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P Z) Call p (Q Z),(A Z)" proof - from deriv_body wf p_defined have "\p\{p}. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P Z) Call p (Q Z),(A Z)" apply - apply (rule ProcRec [where A="\p. A" and P="\p. P" and Q="\p. Q"]) apply simp_all done thus ?thesis by simp qed lemma ProcNoRec1: assumes deriv_body: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P Z) (the (\ p)) (Q Z),(A Z)" assumes p_defined: "p \ dom \" shows "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P Z) Call p (Q Z),(A Z)" proof - have "\\ Z. \,\\\<^sub>t\<^bsub>/F\<^esub> ({\} \ P Z) (the (\ p)) (Q Z),(A Z)" by (blast intro: conseqPre deriv_body [rule_format]) with p_defined have "\\ Z. \,\\(\Z. {(P Z \ {s. ((s,p), \,p) \ {}}, p,Q Z,A Z)}) \\<^sub>t\<^bsub>/F\<^esub> ({\} \ P Z) (the (\ p)) (Q Z),(A Z)" by (blast intro: hoaret_augment_context) from this show ?thesis by (rule ProcRec1) (auto simp add: p_defined) qed lemma ProcBody: assumes WP: "P \ P'" assumes deriv_body: "\,\\\<^sub>t\<^bsub>/F\<^esub> P' body Q,A" assumes body: "\ p = Some body" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P Call p Q,A" apply (rule conseqPre [OF _ WP]) apply (rule ProcNoRec1 [rule_format, where P="\Z. P'" and Q="\Z. Q" and A="\Z. A"]) apply (insert body) apply simp apply (rule hoaret_augment_context [OF deriv_body]) apply blast apply fastforce done lemma CallBody: assumes adapt: "P \ {s. init s \ P' s}" assumes bdy: "\s. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' s) body {t. return s t \ R s t},{t. return s t \ A}" assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes body: "\ p = Some body" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return c) Q,A" apply (unfold call_def) apply (rule Block [OF adapt _ c]) apply (rule allI) apply (rule ProcBody [where \=\, OF _ bdy [rule_format] body]) apply simp done +lemma Call_exnBody: +assumes adapt: "P \ {s. init s \ P' s}" +assumes bdy: "\s. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' s) body {t. return s t \ R s t},{t. result_exn (return s t) t \ A}" +assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" +assumes body: "\ p = Some body" +shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" +apply (unfold call_exn_def) +apply (rule Block_exn [OF adapt _ c]) +apply (rule allI) +apply (rule ProcBody [where \=\, OF _ bdy [rule_format] body]) +apply simp + done + lemmas ProcModifyReturn = HoareTotalProps.ProcModifyReturn lemmas ProcModifyReturnSameFaults = HoareTotalProps.ProcModifyReturnSameFaults +lemmas Proc_exnModifyReturn = HoareTotalProps.Proc_exnModifyReturn +lemmas Proc_exnModifyReturnSameFaults = HoareTotalProps.Proc_exnModifyReturnSameFaults + lemma ProcModifyReturnNoAbr: assumes spec: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return' c) Q,A" assumes result_conform: "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" assumes modifies_spec: "\\. \,\\\<^bsub>/UNIV\<^esub> {\} Call p (Modif \),{}" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return c) Q,A" by (rule ProcModifyReturn [OF spec result_conform _ modifies_spec]) simp +lemma Proc_exnModifyReturnNoAbr: + assumes spec: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call_exn init p return' result_exn c) Q,A" + assumes result_conform: + "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" + assumes modifies_spec: + "\\. \,\\\<^bsub>/UNIV\<^esub> {\} Call p (Modif \),{}" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" + by (rule Proc_exnModifyReturn [OF spec result_conform _ modifies_spec]) simp + lemma ProcModifyReturnNoAbrSameFaults: assumes spec: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return' c) Q,A" assumes result_conform: "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" assumes modifies_spec: "\\. \,\\\<^bsub>/F\<^esub> {\} Call p (Modif \),{}" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return c) Q,A" by (rule ProcModifyReturnSameFaults [OF spec result_conform _ modifies_spec]) simp +lemma Proc_exnModifyReturnNoAbrSameFaults: + assumes spec: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call_exn init p return' result_exn c) Q,A" + assumes result_conform: + "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" + assumes modifies_spec: + "\\. \,\\\<^bsub>/F\<^esub> {\} Call p (Modif \),{}" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" + by (rule Proc_exnModifyReturnSameFaults [OF spec result_conform _ modifies_spec]) simp -lemma DynProc: +lemma DynProc_exn: assumes adapt: "P \ {s. \Z. init s \ P' s Z \ (\t. t \ Q' s Z \ return s t \ R s t) \ - (\t. t \ A' s Z \ return s t \ A)}" + (\t. t \ A' s Z \ result_exn (return s t) t \ A)}" assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes p: "\s\ P. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' s Z) Call (p s) (Q' s Z),(A' s Z)" - shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P dynCall init p return c Q,A" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P dynCall_exn f UNIV init p return result_exn c Q,A" apply (rule conseq [where P'="\Z. {s. s=Z \ s \ P}" and Q'="\Z. Q" and A'="\Z. A"]) prefer 2 using adapt apply blast apply (rule allI) -apply (unfold dynCall_def call_def block_def) +apply (unfold dynCall_exn_def maybe_guard_UNIV call_exn_def block_exn_def guards.simps) apply (rule HoareTotalDef.DynCom) apply clarsimp apply (rule HoareTotalDef.DynCom) apply clarsimp apply (frule in_mono [rule_format, OF adapt]) apply clarsimp apply (rename_tac Z') apply (rule_tac R="Q' Z Z'" in Seq) apply (rule CatchSwap) apply (rule SeqSwap) apply (rule Throw) apply (rule subset_refl) apply (rule Basic) apply (rule subset_refl) apply (rule_tac R="{i. i \ P' Z Z'}" in Seq) apply (rule Basic) apply clarsimp apply simp apply (rule_tac Q'="Q' Z Z'" and A'="A' Z Z'" in conseqPost) using p apply clarsimp apply simp apply clarsimp apply (rule_tac P'="\Z''. {t. t=Z'' \ return Z t \ R Z t}" and Q'="\Z''. Q" and A'="\Z''. A" in conseq) prefer 2 apply simp apply (rule allI) apply (rule HoareTotalDef.DynCom) apply clarsimp apply (rule SeqSwap) apply (rule c [rule_format]) apply (rule Basic) apply clarsimp -done + done + +lemma DynProc_exn_guards_cons: + assumes p: "\,\\\<^sub>t\<^bsub>/F\<^esub> P dynCall_exn f UNIV init p return result_exn c Q,A" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> (g \ P) dynCall_exn f g init p return result_exn c Q,A" + using p apply (clarsimp simp add: dynCall_exn_def maybe_guard_def) + apply (rule Guard) + apply (rule subset_refl) + apply assumption + done + +lemma DynProc: + assumes adapt: "P \ {s. \Z. init s \ P' s Z \ + (\t. t \ Q' s Z \ return s t \ R s t) \ + (\t. t \ A' s Z \ return s t \ A)}" + assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" + assumes p: "\s\ P. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' s Z) Call (p s) (Q' s Z),(A' s Z)" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P dynCall init p return c Q,A" + using adapt c p unfolding dynCall_dynCall_exn by (rule DynProc_exn) + +lemma DynProc_exn': + assumes adapt: "P \ {s. \Z. init s \ P' s Z \ + (\t \ Q' s Z. return s t \ R s t) \ + (\t \ A' s Z. result_exn (return s t) t \ A)}" + assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" + assumes p: "\s\ P. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' s Z) Call (p s) (Q' s Z),(A' s Z)" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P dynCall_exn f UNIV init p return result_exn c Q,A" +proof - + from adapt have "P \ {s. \Z. init s \ P' s Z \ + (\t. t \ Q' s Z \ return s t \ R s t) \ + (\t. t \ A' s Z \ result_exn (return s t) t \ A)}" + by blast + from this c p show ?thesis + by (rule DynProc_exn) +qed lemma DynProc': assumes adapt: "P \ {s. \Z. init s \ P' s Z \ (\t \ Q' s Z. return s t \ R s t) \ (\t \ A' s Z. return s t \ A)}" assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes p: "\s\ P. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' s Z) Call (p s) (Q' s Z),(A' s Z)" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P dynCall init p return c Q,A" -proof - - from adapt have "P \ {s. \Z. init s \ P' s Z \ - (\t. t \ Q' s Z \ return s t \ R s t) \ - (\t. t \ A' s Z \ return s t \ A)}" - by blast - from this c p show ?thesis - by (rule DynProc) -qed + using adapt c p unfolding dynCall_dynCall_exn by (rule DynProc_exn') -lemma DynProcStaticSpec: +lemma DynProc_exnStaticSpec: assumes adapt: "P \ {s. s \ S \ (\Z. init s \ P' Z \ (\\. \ \ Q' Z \ return s \ \ R s \) \ - (\\. \ \ A' Z \ return s \ \ A))}" + (\\. \ \ A' Z \ result_exn (return s \) \ \ A))}" assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes spec: "\s\S. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) Call (p s) (Q' Z),(A' Z)" -shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" +shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall_exn f UNIV init p return result_exn c) Q,A" proof - from adapt have P_S: "P \ S" by blast - have "\,\\\<^sub>t\<^bsub>/F\<^esub> (P \ S) (dynCall init p return c) Q,A" - apply (rule DynProc [where P'="\s Z. P' Z" and Q'="\s Z. Q' Z" + have "\,\\\<^sub>t\<^bsub>/F\<^esub> (P \ S) (dynCall_exn f UNIV init p return result_exn c) Q,A" + apply (rule DynProc_exn [where P'="\s Z. P' Z" and Q'="\s Z. Q' Z" and A'="\s Z. A' Z", OF _ c]) apply clarsimp apply (frule in_mono [rule_format, OF adapt]) apply clarsimp using spec apply clarsimp done thus ?thesis by (rule conseqPre) (insert P_S,blast) qed +lemma DynProcStaticSpec: +assumes adapt: "P \ {s. s \ S \ (\Z. init s \ P' Z \ + (\\. \ \ Q' Z \ return s \ \ R s \) \ + (\\. \ \ A' Z \ return s \ \ A))}" +assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" +assumes spec: "\s\S. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) Call (p s) (Q' Z),(A' Z)" +shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using adapt c spec unfolding dynCall_dynCall_exn by (rule DynProc_exnStaticSpec) + +lemma DynProc_exnProcPar: +assumes adapt: "P \ {s. p s = q \ (\Z. init s \ P' Z \ + (\\. \ \ Q' Z \ return s \ \ R s \) \ + (\\. \ \ A' Z \ result_exn (return s \) \ \ A))}" +assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" +assumes spec: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) Call q (Q' Z),(A' Z)" +shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall_exn f UNIV init p return result_exn c) Q,A" + apply (rule DynProc_exnStaticSpec [where S="{s. p s = q}",simplified, OF adapt c]) + using spec + apply simp + done lemma DynProcProcPar: assumes adapt: "P \ {s. p s = q \ (\Z. init s \ P' Z \ (\\. \ \ Q' Z \ return s \ \ R s \) \ (\\. \ \ A' Z \ return s \ \ A))}" assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes spec: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) Call q (Q' Z),(A' Z)" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" apply (rule DynProcStaticSpec [where S="{s. p s = q}",simplified, OF adapt c]) using spec apply simp done - -lemma DynProcProcParNoAbrupt: +lemma DynProc_exnProcParNoAbrupt: assumes adapt: "P \ {s. p s = q \ (\Z. init s \ P' Z \ (\\. \ \ Q' Z \ return s \ \ R s \))}" assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes spec: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) Call q (Q' Z),{}" -shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" +shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall_exn f UNIV init p return result_exn c) Q,A" proof - have "P \ {s. p s = q \ (\ Z. init s \ P' Z \ (\t. t \ Q' Z \ return s t \ R s t) \ - (\t. t \ {} \ return s t \ A))}" + (\t. t \ {} \ result_exn (return s t) t \ A))}" (is "P \ ?P'") proof fix s assume P: "s\P" with adapt obtain Z where Pre: "p s = q \ init s \ P' Z" and adapt_Norm: "\\. \ \ Q' Z \ return s \ \ R s \" by blast from adapt_Norm have "\t. t \ Q' Z \ return s t \ R s t" by auto then show "s\?P'" using Pre by blast qed note P = this show ?thesis apply - - apply (rule DynProcStaticSpec [where S="{s. p s = q}",simplified, OF P c]) + apply (rule DynProc_exnStaticSpec [where S="{s. p s = q}",simplified, OF P c]) apply (insert spec) apply auto done qed -lemma DynProcModifyReturnNoAbr: - assumes to_prove: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return' c) Q,A" +lemma DynProcProcParNoAbrupt: +assumes adapt: "P \ {s. p s = q \ (\Z. init s \ P' Z \ + (\\. \ \ Q' Z \ return s \ \ R s \))}" +assumes c: "\s t. \,\\\<^sub>t\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" +assumes spec: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) Call q (Q' Z),{}" +shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using adapt c spec unfolding dynCall_dynCall_exn by (rule DynProc_exnProcParNoAbrupt) + +lemma DynProc_exnModifyReturnNoAbr: + assumes to_prove: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall_exn f g init p return' result_exn c) Q,A" assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) \ return' s t = return s t" assumes modif_clause: "\s \ P. \\. \,\\\<^bsub>/UNIV\<^esub> {\} Call (p s) (Modif \),{}" - shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" proof - from ret_nrm_modif have "\s t. t \ (Modif (init s)) \ return' s t = return s t" by iprover then have ret_nrm_modif': "\s t. t \ (Modif (init s)) \ return' s t = return s t" by simp have ret_abr_modif': "\s t. t \ {} - \ return' s t = return s t" + \ result_exn (return' s t) t = result_exn (return s t) t" by simp from to_prove ret_nrm_modif' ret_abr_modif' modif_clause show ?thesis - by (rule dynProcModifyReturn) + by (rule dynProc_exnModifyReturn) qed -lemma ProcDynModifyReturnNoAbrSameFaults: +lemma DynProcModifyReturnNoAbr: assumes to_prove: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return' c) Q,A" assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) \ return' s t = return s t" assumes modif_clause: + "\s \ P. \\. \,\\\<^bsub>/UNIV\<^esub> {\} Call (p s) (Modif \),{}" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using to_prove ret_nrm_modif modif_clause unfolding dynCall_dynCall_exn + by (rule DynProc_exnModifyReturnNoAbr) + +lemma ProcDyn_exnModifyReturnNoAbrSameFaults: + assumes to_prove: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall_exn f g init p return' result_exn c) Q,A" + assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) + \ return' s t = return s t" + assumes modif_clause: "\s \ P. \\. \,\\\<^bsub>/F\<^esub> {\} (Call (p s)) (Modif \),{}" - shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" proof - from ret_nrm_modif have "\s t. t \ (Modif (init s)) \ return' s t = return s t" by iprover then have ret_nrm_modif': "\s t. t \ (Modif (init s)) \ return' s t = return s t" by simp have ret_abr_modif': "\s t. t \ {} - \ return' s t = return s t" + \ result_exn (return' s t) t = result_exn (return s t) t" by simp from to_prove ret_nrm_modif' ret_abr_modif' modif_clause show ?thesis - by (rule dynProcModifyReturnSameFaults) + by (rule dynProc_exnModifyReturnSameFaults) +qed + + +lemma ProcDynModifyReturnNoAbrSameFaults: + assumes to_prove: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return' c) Q,A" + assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) + \ return' s t = return s t" + assumes modif_clause: + "\s \ P. \\. \,\\\<^bsub>/F\<^esub> {\} (Call (p s)) (Modif \),{}" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using to_prove ret_nrm_modif modif_clause unfolding dynCall_dynCall_exn + by (rule ProcDyn_exnModifyReturnNoAbrSameFaults) + +lemma Proc_exnProcParModifyReturn: + assumes q: "P \ {s. p s = q} \ P'" + \ \@{thm[source] DynProcProcPar} introduces the same constraint as first conjunction in + @{term P'}, so the vcg can simplify it.\ + assumes to_prove: "\,\\\<^sub>t\<^bsub>/F\<^esub> P' (dynCall_exn f g init p return' result_exn c) Q,A" + assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) + \ return' s t = return s t" + assumes ret_abr_modif: "\s t. t \ (ModifAbr (init s)) + \ result_exn (return' s t) t = result_exn (return s t) t" + assumes modif_clause: + "\\. \,\\\<^bsub>/UNIV\<^esub> {\} (Call q) (Modif \),(ModifAbr \)" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" +proof - + from to_prove have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall_exn f g init p return' result_exn c) Q,A" + by (rule conseqPre) blast + from this ret_nrm_modif + ret_abr_modif + have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall_exn f g init p return result_exn c) Q,A" + by (rule dynProc_exnModifyReturn) (insert modif_clause,auto) + from this q show ?thesis + by (rule conseqPre) qed lemma ProcProcParModifyReturn: assumes q: "P \ {s. p s = q} \ P'" \ \@{thm[source] DynProcProcPar} introduces the same constraint as first conjunction in @{term P'}, so the vcg can simplify it.\ assumes to_prove: "\,\\\<^sub>t\<^bsub>/F\<^esub> P' (dynCall init p return' c) Q,A" assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) \ return' s t = return s t" assumes ret_abr_modif: "\s t. t \ (ModifAbr (init s)) \ return' s t = return s t" assumes modif_clause: "\\. \,\\\<^bsub>/UNIV\<^esub> {\} (Call q) (Modif \),(ModifAbr \)" - shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using q to_prove ret_nrm_modif ret_abr_modif modif_clause unfolding dynCall_dynCall_exn + by (rule Proc_exnProcParModifyReturn) + +lemma Proc_exnProcParModifyReturnSameFaults: + assumes q: "P \ {s. p s = q} \ P'" + \ \@{thm[source] DynProcProcPar} introduces the same constraint as first conjunction in + @{term P'}, so the vcg can simplify it.\ + assumes to_prove: "\,\\\<^sub>t\<^bsub>/F\<^esub> P' (dynCall_exn f g init p return' result_exn c) Q,A" + assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) + \ return' s t = return s t" + assumes ret_abr_modif: "\s t. t \ (ModifAbr (init s)) + \ result_exn (return' s t) t = result_exn (return s t) t" + assumes modif_clause: + "\\. \,\\\<^bsub>/F\<^esub> {\} Call q (Modif \),(ModifAbr \)" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" proof - - from to_prove have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall init p return' c) Q,A" + from to_prove + have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall_exn f g init p return' result_exn c) Q,A" by (rule conseqPre) blast from this ret_nrm_modif ret_abr_modif - have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall init p return c) Q,A" - by (rule dynProcModifyReturn) (insert modif_clause,auto) + have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall_exn f g init p return result_exn c) Q,A" + by (rule dynProc_exnModifyReturnSameFaults) (insert modif_clause,auto) from this q show ?thesis by (rule conseqPre) qed - lemma ProcProcParModifyReturnSameFaults: assumes q: "P \ {s. p s = q} \ P'" \ \@{thm[source] DynProcProcPar} introduces the same constraint as first conjunction in @{term P'}, so the vcg can simplify it.\ assumes to_prove: "\,\\\<^sub>t\<^bsub>/F\<^esub> P' (dynCall init p return' c) Q,A" assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) \ return' s t = return s t" assumes ret_abr_modif: "\s t. t \ (ModifAbr (init s)) \ return' s t = return s t" assumes modif_clause: "\\. \,\\\<^bsub>/F\<^esub> {\} Call q (Modif \),(ModifAbr \)" - shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using q to_prove ret_nrm_modif ret_abr_modif modif_clause unfolding dynCall_dynCall_exn + by (rule Proc_exnProcParModifyReturnSameFaults) + +lemma Proc_exnProcParModifyReturnNoAbr: + assumes q: "P \ {s. p s = q} \ P'" + \ \@{thm[source] DynProcProcParNoAbrupt} introduces the same constraint as + first conjunction in @{term P'}, so the vcg can simplify it.\ + assumes to_prove: "\,\\\<^sub>t\<^bsub>/F\<^esub> P' (dynCall_exn f g init p return' result_exn c) Q,A" + assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) + \ return' s t = return s t" + assumes modif_clause: + "\\. \,\\\<^bsub>/UNIV\<^esub> {\} (Call q) (Modif \),{}" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" proof - - from to_prove - have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall init p return' c) Q,A" + from to_prove have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall_exn f g init p return' result_exn c) Q,A" by (rule conseqPre) blast from this ret_nrm_modif - ret_abr_modif - have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall init p return c) Q,A" - by (rule dynProcModifyReturnSameFaults) (insert modif_clause,auto) + have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall_exn f g init p return result_exn c) Q,A" + by (rule DynProc_exnModifyReturnNoAbr) (insert modif_clause,auto) from this q show ?thesis by (rule conseqPre) qed lemma ProcProcParModifyReturnNoAbr: assumes q: "P \ {s. p s = q} \ P'" \ \@{thm[source] DynProcProcParNoAbrupt} introduces the same constraint as first conjunction in @{term P'}, so the vcg can simplify it.\ assumes to_prove: "\,\\\<^sub>t\<^bsub>/F\<^esub> P' (dynCall init p return' c) Q,A" assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) \ return' s t = return s t" assumes modif_clause: "\\. \,\\\<^bsub>/UNIV\<^esub> {\} (Call q) (Modif \),{}" - shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using q to_prove ret_nrm_modif modif_clause unfolding dynCall_dynCall_exn + by (rule Proc_exnProcParModifyReturnNoAbr) + + +lemma Proc_exnProcParModifyReturnNoAbrSameFaults: + assumes q: "P \ {s. p s = q} \ P'" + \ \@{thm[source] DynProcProcParNoAbrupt} introduces the same constraint as + first conjunction in @{term P'}, so the vcg can simplify it.\ + assumes to_prove: "\,\\\<^sub>t\<^bsub>/F\<^esub> P' (dynCall_exn f g init p return' result_exn c) Q,A" + assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) + \ return' s t = return s t" + assumes modif_clause: + "\\. \,\\\<^bsub>/F\<^esub> {\} (Call q) (Modif \),{}" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" proof - - from to_prove have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall init p return' c) Q,A" + from to_prove have + "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall_exn f g init p return' result_exn c) Q,A" by (rule conseqPre) blast from this ret_nrm_modif - have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall init p return c) Q,A" - by (rule DynProcModifyReturnNoAbr) (insert modif_clause,auto) + have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall_exn f g init p return result_exn c) Q,A" + by (rule ProcDyn_exnModifyReturnNoAbrSameFaults) (insert modif_clause,auto) from this q show ?thesis by (rule conseqPre) qed - lemma ProcProcParModifyReturnNoAbrSameFaults: assumes q: "P \ {s. p s = q} \ P'" \ \@{thm[source] DynProcProcParNoAbrupt} introduces the same constraint as first conjunction in @{term P'}, so the vcg can simplify it.\ assumes to_prove: "\,\\\<^sub>t\<^bsub>/F\<^esub> P' (dynCall init p return' c) Q,A" assumes ret_nrm_modif: "\s t. t \ (Modif (init s)) \ return' s t = return s t" assumes modif_clause: "\\. \,\\\<^bsub>/F\<^esub> {\} (Call q) (Modif \),{}" - shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" -proof - - from to_prove have - "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall init p return' c) Q,A" - by (rule conseqPre) blast - from this ret_nrm_modif - have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. p s = q} \ P') (dynCall init p return c) Q,A" - by (rule ProcDynModifyReturnNoAbrSameFaults) (insert modif_clause,auto) - from this q show ?thesis - by (rule conseqPre) -qed + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using q to_prove ret_nrm_modif modif_clause unfolding dynCall_dynCall_exn + by (rule Proc_exnProcParModifyReturnNoAbrSameFaults) + lemma MergeGuards_iff: "\,\\\<^sub>t\<^bsub>/F\<^esub> P merge_guards c Q,A = \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (auto intro: MergeGuardsI MergeGuardsD) lemma CombineStrip': assumes deriv: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c' Q,A" assumes deriv_strip_triv: "\,{}\\<^bsub>/{}\<^esub> P c'' UNIV,UNIV" assumes c'': "c''= mark_guards False (strip_guards (-F) c')" assumes c: "merge_guards c = merge_guards (mark_guards False c')" shows "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c Q,A" proof - from deriv_strip_triv have deriv_strip: "\,\\\<^bsub>/{}\<^esub> P c'' UNIV,UNIV" by (auto intro: hoare_augment_context) from deriv_strip [simplified c''] have "\,\\\<^bsub>/{}\<^esub> P (strip_guards (- F) c') UNIV,UNIV" by (rule HoarePartialProps.MarkGuardsD) with deriv have "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c' Q,A" by (rule CombineStrip) hence "\,\\\<^sub>t\<^bsub>/{}\<^esub> P mark_guards False c' Q,A" by (rule MarkGuardsI) hence "\,\\\<^sub>t\<^bsub>/{}\<^esub> P merge_guards (mark_guards False c') Q,A" by (rule MergeGuardsI) hence "\,\\\<^sub>t\<^bsub>/{}\<^esub> P merge_guards c Q,A" by (simp add: c) thus ?thesis by (rule MergeGuardsD) qed lemma CombineStrip'': assumes deriv: "\,\\\<^sub>t\<^bsub>/{True}\<^esub> P c' Q,A" assumes deriv_strip_triv: "\,{}\\<^bsub>/{}\<^esub> P c'' UNIV,UNIV" assumes c'': "c''= mark_guards False (strip_guards ({False}) c')" assumes c: "merge_guards c = merge_guards (mark_guards False c')" shows "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c Q,A" apply (rule CombineStrip' [OF deriv deriv_strip_triv _ c]) apply (insert c'') apply (subgoal_tac "- {True} = {False}") apply auto done lemma AsmUN: "(\Z. {(P Z, p, Q Z,A Z)}) \ \ \ \Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P Z) (Call p) (Q Z),(A Z)" by (blast intro: hoaret.Asm) lemma hoaret_to_hoarep': "\Z. \,{}\\<^sub>t\<^bsub>/F\<^esub> (P Z) p (Q Z),(A Z) \ \Z. \,{}\\<^bsub>/F\<^esub> (P Z) p (Q Z),(A Z)" by (iprover intro: total_to_partial) lemma augment_context': "\\ \ \'; \Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P Z) p (Q Z),(A Z)\ \ \Z. \,\'\\<^sub>t\<^bsub>/F\<^esub> (P Z) p (Q Z),(A Z)" by (iprover intro: hoaret_augment_context) lemma augment_emptyFaults: "\\Z. \,{}\\<^sub>t\<^bsub>/{}\<^esub> (P Z) p (Q Z),(A Z)\ \ \Z. \,{}\\<^sub>t\<^bsub>/F\<^esub> (P Z) p (Q Z),(A Z)" by (blast intro: augment_Faults) lemma augment_FaultsUNIV: "\\Z. \,{}\\<^sub>t\<^bsub>/F\<^esub> (P Z) p (Q Z),(A Z)\ \ \Z. \,{}\\<^sub>t\<^bsub>/UNIV\<^esub> (P Z) p (Q Z),(A Z)" by (blast intro: augment_Faults) lemma PostConjI [trans]: "\\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A; \,\\\<^sub>t\<^bsub>/F\<^esub> P c R,B\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c (Q \ R),(A \ B)" by (rule PostConjI) lemma PostConjI' : "\\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A; \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c R,B\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c (Q \ R),(A \ B)" by (rule PostConjI) iprover+ lemma PostConjE [consumes 1]: assumes conj: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c (Q \ R),(A \ B)" assumes E: "\\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A; \,\\\<^sub>t\<^bsub>/F\<^esub> P c R,B\ \ S" shows "S" proof - from conj have "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (rule conseqPost) blast+ moreover from conj have "\,\\\<^sub>t\<^bsub>/F\<^esub> P c R,B" by (rule conseqPost) blast+ ultimately show "S" by (rule E) qed subsubsection \Rules for Single-Step Proof \label{sec:hoare-isar}\ text \ We are now ready to introduce a set of Hoare rules to be used in single-step structured proofs in Isabelle/Isar. \medskip Assertions of Hoare Logic may be manipulated in calculational proofs, with the inclusion expressed in terms of sets or predicates. Reversed order is supported as well. \ lemma annotateI [trans]: "\\,\\\<^sub>t\<^bsub>/F\<^esub> P anno Q,A; c = anno\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (simp) lemma annotate_normI: assumes deriv_anno: "\,\\\<^sub>t\<^bsub>/F\<^esub>P anno Q,A" assumes norm_eq: "normalize c = normalize anno" shows "\,\\\<^sub>t\<^bsub>/F\<^esub>P c Q,A" proof - from HoareTotalProps.NormalizeI [OF deriv_anno] norm_eq have "\,\\\<^sub>t\<^bsub>/F \<^esub>P normalize c Q,A" by simp from NormalizeD [OF this] show ?thesis . qed lemma annotateWhile: "\\,\\\<^sub>t\<^bsub>/F\<^esub> P (whileAnnoG gs b I V c) Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (while gs b c) Q,A" by (simp add: whileAnnoG_def) lemma reannotateWhile: "\\,\\\<^sub>t\<^bsub>/F\<^esub> P (whileAnnoG gs b I V c) Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (whileAnnoG gs b J V c) Q,A" by (simp add: whileAnnoG_def) lemma reannotateWhileNoGuard: "\\,\\\<^sub>t\<^bsub>/F\<^esub> P (whileAnno b I V c) Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (whileAnno b J V c) Q,A" by (simp add: whileAnno_def) lemma [trans] : "P' \ P \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A \ \,\\\<^sub>t\<^bsub>/F\<^esub> P' c Q,A" by (rule conseqPre) lemma [trans]: "Q \ Q' \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q',A" by (rule conseqPost) blast+ lemma [trans]: "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. P s} c Q,A \ (\s. P' s \ P s) \ \,\\\<^sub>t\<^bsub>/F\<^esub> {s. P' s} c Q,A" by (rule conseqPre) auto lemma [trans]: "(\s. P' s \ P s) \ \,\\\<^sub>t\<^bsub>/F\<^esub> {s. P s} c Q,A \ \,\\\<^sub>t\<^bsub>/F\<^esub> {s. P' s} c Q,A" by (rule conseqPre) auto lemma [trans]: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c {s. Q s},A \ (\s. Q s \ Q' s) \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c {s. Q' s},A" by (rule conseqPost) auto lemma [trans]: "(\s. Q s \ Q' s) \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c {s. Q s},A \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c {s. Q' s},A" by (rule conseqPost) auto lemma [intro?]: "\,\\\<^sub>t\<^bsub>/F\<^esub> P Skip P,A" by (rule Skip) auto lemma CondInt [trans,intro?]: "\\,\\\<^sub>t\<^bsub>/F\<^esub> (P \ b) c1 Q,A; \,\\\<^sub>t\<^bsub>/F\<^esub> (P \ - b) c2 Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (Cond b c1 c2) Q,A" by (rule Cond) auto lemma CondConj [trans, intro?]: "\\,\\\<^sub>t\<^bsub>/F\<^esub> {s. P s \ b s} c1 Q,A; \,\\\<^sub>t\<^bsub>/F\<^esub> {s. P s \ \ b s} c2 Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> {s. P s} (Cond {s. b s} c1 c2) Q,A" by (rule Cond) auto end diff --git a/thys/Simpl/HoareTotalDef.thy b/thys/Simpl/HoareTotalDef.thy --- a/thys/Simpl/HoareTotalDef.thy +++ b/thys/Simpl/HoareTotalDef.thy @@ -1,374 +1,353 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: HoareTotalDef.thy - Author: Norbert Schirmer, TU Muenchen - -Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (C) 2006-2008 Norbert Schirmer *) section \Hoare Logic for Total Correctness\ theory HoareTotalDef imports HoarePartialDef Termination begin subsection \Validity of Hoare Tuples: \\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A\\ definition validt :: "[('s,'p,'f) body,'f set,'s assn,('s,'p,'f) com,'s assn,'s assn] \ bool" ("_\\<^sub>t\<^bsub>'/_\<^esub>/ _ _ _,_" [61,60,1000, 20, 1000,1000] 60) where "\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A \ \\\<^bsub>/F\<^esub> P c Q,A \ (\s \ Normal ` P. \\c\s)" definition cvalidt:: "[('s,'p,'f) body,('s,'p) quadruple set,'f set, 's assn,('s,'p,'f) com,'s assn,'s assn] \ bool" ("_,_\\<^sub>t\<^bsub>'/_\<^esub>/ _ _ _,_" [61,60, 60,1000, 20, 1000,1000] 60) where "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A \ (\(P,p,Q,A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A) \ \ \\<^sub>t\<^bsub>/F\<^esub> P c Q,A" notation (ASCII) validt ("_|=t'/_/ _ _ _,_" [61,60,1000, 20, 1000,1000] 60) and cvalidt ("_,_|=t'/_ / _ _ _,_" [61,60,60,1000, 20, 1000,1000] 60) subsection \Properties of Validity\ lemma validtI: "\\s t. \\\\c,Normal s\ \ t;s \ P;t \ Fault ` F\ \ t \ Normal ` Q \ Abrupt ` A; \s. s \ P \ \\ c\(Normal s) \ \ \\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (auto simp add: validt_def valid_def) lemma cvalidtI: "\\s t. \\(P,p,Q,A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A;\\\c,Normal s\ \ t;s \ P; t \ Fault ` F\ \ t \ Normal ` Q \ Abrupt ` A; \s. \\(P,p,Q,A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A; s\P\ \ \\c\(Normal s)\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (auto simp add: cvalidt_def validt_def valid_def) lemma cvalidt_postD: "\\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A; \(P,p,Q,A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A;\\\c,Normal s \ \ t; s \ P;t \ Fault ` F\ \ t \ Normal ` Q \ Abrupt ` A" by (simp add: cvalidt_def validt_def valid_def) lemma cvalidt_termD: "\\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A; \(P,p,Q,A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A;s \ P\ \ \\c\(Normal s)" by (simp add: cvalidt_def validt_def valid_def) lemma validt_augment_Faults: assumes valid:"\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" assumes F': "F \ F'" shows "\\\<^sub>t\<^bsub>/F'\<^esub> P c Q,A" using valid F' by (auto intro: valid_augment_Faults simp add: validt_def) subsection \The Hoare Rules: \\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A\\ inductive "hoaret"::"[('s,'p,'f) body,('s,'p) quadruple set,'f set, 's assn,('s,'p,'f) com,'s assn,'s assn] => bool" ("(3_,_/\\<^sub>t\<^bsub>'/_\<^esub> (_/ (_)/ _,_))" [61,60,60,1000,20,1000,1000]60) for \::"('s,'p,'f) body" where Skip: "\,\\\<^sub>t\<^bsub>/F\<^esub> Q Skip Q,A" | Basic: "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. f s \ Q} (Basic f) Q,A" | Spec: "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. (\t. (s,t) \ r \ t \ Q) \ (\t. (s,t) \ r)} (Spec r) Q,A" | Seq: "\\,\\\<^sub>t\<^bsub>/F\<^esub> P c\<^sub>1 R,A; \,\\\<^sub>t\<^bsub>/F\<^esub> R c\<^sub>2 Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P Seq c\<^sub>1 c\<^sub>2 Q,A" | Cond: "\\,\\\<^sub>t\<^bsub>/F\<^esub> (P \ b) c\<^sub>1 Q,A; \,\\\<^sub>t\<^bsub>/F\<^esub> (P \ - b) c\<^sub>2 Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (Cond b c\<^sub>1 c\<^sub>2) Q,A" | While: "\wf r; \\. \,\\\<^sub>t\<^bsub>/F\<^esub> ({\} \ P \ b) c ({t. (t,\)\r} \ P),A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (While b c) (P \ - b),A" | Guard: "\,\\\<^sub>t\<^bsub>/F\<^esub> (g \ P) c Q,A \ \,\\\<^sub>t\<^bsub>/F\<^esub> (g \ P) Guard f g c Q,A" | Guarantee: "\f \ F; \,\\\<^sub>t\<^bsub>/F\<^esub> (g \ P) c Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (Guard f g c) Q,A" | CallRec: "\(P,p,Q,A) \ Specs; wf r; Specs_wf = (\p \. (\(P,q,Q,A). (P \ {s. ((s,q),(\,p)) \ r},q,Q,A)) ` Specs); \(P,p,Q,A)\ Specs. p \ dom \ \ (\\. \,\ \ Specs_wf p \\\<^sub>t\<^bsub>/F\<^esub> ({\} \ P) (the (\ p)) Q,A) \ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" | DynCom: "\s \ P. \,\\\<^sub>t\<^bsub>/F\<^esub> P (c s) Q,A \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (DynCom c) Q,A" | Throw: "\,\\\<^sub>t\<^bsub>/F\<^esub> A Throw Q,A" | Catch: "\\,\\\<^sub>t\<^bsub>/F\<^esub> P c\<^sub>1 Q,R; \,\\\<^sub>t\<^bsub>/F\<^esub> R c\<^sub>2 Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P Catch c\<^sub>1 c\<^sub>2 Q,A" | Conseq: "\s \ P. \P' Q' A'. \,\\\<^sub>t\<^bsub>/F\<^esub> P' c Q',A' \ s \ P' \ Q' \ Q \ A' \ A \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" | Asm: "(P,p,Q,A) \ \ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" | ExFalso: "\\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A; \ \\\<^sub>t\<^bsub>/F\<^esub> P c Q,A\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" \ \This is a hack rule that enables us to derive completeness for an arbitrary context \\\, from completeness for an empty context.\ text \Does not work, because of rule ExFalso, the context \\\ is to blame. A weaker version with empty context can be derived from soundness later on.\ lemma hoaret_to_hoarep: assumes hoaret: "\,\\\<^sub>t\<^bsub>/F\<^esub> P p Q,A" shows "\,\\\<^bsub>/F\<^esub> P p Q,A" using hoaret proof (induct) case Skip thus ?case by (rule hoarep.intros) next case Basic thus ?case by (rule hoarep.intros) next case Seq thus ?case by - (rule hoarep.intros) next case Cond thus ?case by - (rule hoarep.intros) next case (While r \ F P b c A) hence "\\. \,\\\<^bsub>/F\<^esub> ({\} \ P \ b) c ({t. (t, \) \ r} \ P),A" by iprover hence "\,\\\<^bsub>/F\<^esub> (P \ b) c P,A" by (rule HoarePartialDef.conseq) blast then show "\,\\\<^bsub>/F\<^esub> P While b c (P \ - b),A" by (rule hoarep.While) next case Guard thus ?case by - (rule hoarep.intros) (*next case (CallRec A F P Procs Q Z \ p r) hence hyp: "\p\Procs. \\ Z. \,\ \ (\q\Procs. \Z. {(P q Z \ {s. ((s, q), \, p) \ r}, Call q, Q q Z,A q Z)})\\<^bsub>/F\<^esub> ({\} \ P p Z) (the (\ p)) (Q p Z),(A p Z)" by blast have "\p\Procs. \Z. \,\ \ (\q\Procs. \Z. {(P q Z, Call q, Q q Z,A q Z)})\\<^bsub>/F\<^esub> (P p Z) (the (\ p)) (Q p Z),(A p Z)" proof (intro ballI allI) fix p Z assume "p \ Procs" with hyp have hyp': "\ \. \,\ \ (\q\Procs. \Z. {(P q Z \ {s. ((s, q), \, p) \ r}, Call q, Q q Z, A q Z)})\\<^bsub>/F\<^esub> ({\} \ P p Z) (the (\ p)) (Q p Z),(A p Z)" by blast have "\\. \,\ \ (\q\Procs. \Z. {(P q Z, Call q, Q q Z,A q Z)})\\<^bsub>/F\<^esub> ({\} \ P p Z) (the (\ p)) (Q p Z),(A p Z)" (is "\\. \,?\'\\<^bsub>/F\<^esub> ({\} \ P p Z) (the (\ p)) (Q p Z),(A p Z)") proof (rule allI, rule WeakenContext [OF hyp'],clarify) fix \ P' c Q' A' assume "(P', c, Q', A') \ \ \ (\q\Procs. \Z. {(P q Z \ {s. ((s, q), \, p) \ r}, Call q, Q q Z, A q Z)})" (is "(P', c, Q', A') \ \ \ ?Spec") then show "\,?\'\\<^bsub>/F\<^esub> P' c Q',A'" proof (cases rule: UnE [consumes 1]) assume "(P',c,Q',A') \ \" then show ?thesis by (blast intro: HoarePartialDef.Asm) next assume "(P',c,Q',A') \ ?Spec" then show ?thesis proof (clarify) fix q Z assume q: "q \ Procs" show "\,?\'\\<^bsub>/F\<^esub> (P q Z \ {s. ((s, q), \, p) \ r}) Call q (Q q Z),(A q Z)" proof - from q have "\,?\'\\<^bsub>/F\<^esub> (P q Z) Call q (Q q Z),(A q Z)" by - (rule HoarePartialDef.Asm,blast) thus ?thesis by (rule HoarePartialDef.conseqPre) blast qed qed qed qed then show "\,\ \ (\q\Procs. \Z. {(P q Z, Call q, Q q Z,A q Z)}) \\<^bsub>/F\<^esub> (P p Z) (the (\ p)) (Q p Z),(A p Z)" by (rule HoarePartialDef.conseq) blast qed thus ?case by - (rule hoarep.CallRec)*) next case DynCom thus ?case by (blast intro: hoarep.DynCom) next case Throw thus ?case by - (rule hoarep.Throw) next case Catch thus ?case by - (rule hoarep.Catch) next case Conseq thus ?case by - (rule hoarep.Conseq,blast) next case Asm thus ?case by (rule HoarePartialDef.Asm) next case (ExFalso \ F P c Q A) assume "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" hence "\,\\\<^bsub>/F\<^esub> P c Q,A" oops lemma hoaret_augment_context: assumes deriv: "\,\\\<^sub>t\<^bsub>/F\<^esub> P p Q,A" shows "\\'. \ \ \' \ \,\'\\<^sub>t\<^bsub>/F\<^esub> P p Q,A" using deriv proof (induct) case (CallRec P p Q A Specs r Specs_wf \ F \') have aug: "\ \ \'" by fact then have h: "\\ p. \ \ Specs_wf p \ \ \' \ Specs_wf p \" by blast have "\(P,p,Q,A)\Specs. p \ dom \ \ (\\. \,\ \ Specs_wf p \\\<^sub>t\<^bsub>/F\<^esub> ({\} \ P) (the (\ p)) Q,A \ (\x. \ \ Specs_wf p \ \ x \ \,x\\<^sub>t\<^bsub>/F\<^esub> ({\} \ P) (the (\ p)) Q,A))" by fact hence "\(P,p,Q,A)\Specs. p \ dom \ \ (\\. \,\'\ Specs_wf p \ \\<^sub>t\<^bsub>/F\<^esub> ({\} \ P) (the (\ p)) Q,A)" apply (clarify) apply (rename_tac P p Q A) apply (drule (1) bspec) apply (clarsimp) apply (erule_tac x=\ in allE) apply clarify apply (erule_tac x="\' \ Specs_wf p \" in allE) apply (insert aug) apply auto done with CallRec show ?case by - (rule hoaret.CallRec) next case DynCom thus ?case by (blast intro: hoaret.DynCom) next case (Conseq P \ F c Q A \') from Conseq have "\s \ P. (\P' Q' A'. (\,\' \\<^sub>t\<^bsub>/F\<^esub> P' c Q',A') \ s \ P'\ Q' \ Q \ A' \ A)" by blast with Conseq show ?case by - (rule hoaret.Conseq) next case (ExFalso \ F P c Q A \') have "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" "\ \\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" "\ \ \'" by fact+ then show ?case by (fastforce intro: hoaret.ExFalso simp add: cvalidt_def) qed (blast intro: hoaret.intros)+ subsection \Some Derived Rules\ lemma Conseq': "\s. s \ P \ (\P' Q' A'. (\ Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z)) \ (\Z. s \ P' Z \ (Q' Z \ Q) \ (A' Z \ A))) \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" apply (rule Conseq) apply (rule ballI) apply (erule_tac x=s in allE) apply (clarify) apply (rule_tac x="P' Z" in exI) apply (rule_tac x="Q' Z" in exI) apply (rule_tac x="A' Z" in exI) apply blast done lemma conseq:"\\Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> (P' Z) c (Q' Z),(A' Z); \s. s \ P \ (\ Z. s\P' Z \ (Q' Z \ Q)\ (A' Z \ A))\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (rule Conseq) blast theorem conseqPrePost: "\,\\\<^sub>t\<^bsub>/F\<^esub> P' c Q',A' \ P \ P' \ Q' \ Q \ A' \ A \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (rule conseq [where ?P'="\Z. P'" and ?Q'="\Z. Q'"]) auto lemma conseqPre: "\,\\\<^sub>t\<^bsub>/F\<^esub> P' c Q,A \ P \ P' \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (rule conseq) auto lemma conseqPost: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q',A'\ Q' \ Q \ A' \ A \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (rule conseq) auto lemma Spec_wf_conv: "(\(P, q, Q, A). (P \ {s. ((s, q), \, p) \ r}, q, Q, A)) ` (\p\Procs. \Z. {(P p Z, p, Q p Z, A p Z)}) = (\q\Procs. \Z. {(P q Z \ {s. ((s, q), \, p) \ r}, q, Q q Z, A q Z)})" by (auto intro!: image_eqI) lemma CallRec': "\p\Procs; Procs \ dom \; wf r; \p\Procs. \\ Z. \,\\(\q\Procs. \Z. {((P q Z) \ {s. ((s,q),(\,p)) \ r},q,Q q Z,(A q Z))}) \\<^sub>t\<^bsub>/F\<^esub> ({\} \ (P p Z)) (the (\ p)) (Q p Z),(A p Z)\ \ \,\\\<^sub>t\<^bsub>/F\<^esub> (P p Z) (Call p) (Q p Z),(A p Z)" apply (rule CallRec [where Specs="\p\Procs. \Z. {((P p Z),p,Q p Z,A p Z)}" and r=r]) apply blast apply assumption apply (rule refl) apply (clarsimp) apply (rename_tac p') apply (rule conjI) apply blast apply (intro allI) apply (rename_tac Z \) apply (drule_tac x=p' in bspec, assumption) apply (erule_tac x=\ in allE) apply (erule_tac x=Z in allE) apply (fastforce simp add: Spec_wf_conv) done end diff --git a/thys/Simpl/HoareTotalProps.thy b/thys/Simpl/HoareTotalProps.thy --- a/thys/Simpl/HoareTotalProps.thy +++ b/thys/Simpl/HoareTotalProps.thy @@ -1,3496 +1,3634 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: HoarePartial.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (c) 2022 Apple Inc. All rights reserved. *) section \Properties of Total Correctness Hoare Logic\ theory HoareTotalProps imports SmallStep HoareTotalDef HoarePartialProps begin subsection \Soundness\ lemma hoaret_sound: assumes hoare: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" using hoare proof (induct) case (Skip \ F P A) show "\,\ \\<^sub>t\<^bsub>/F\<^esub> P Skip P,A" proof (rule cvalidtI) fix s t assume "\\\Skip,Normal s\ \ t" "s \ P" thus "t \ Normal ` P \ Abrupt ` A" by cases auto next fix s show "\\Skip \ Normal s" by (rule terminates.intros) qed next case (Basic \ F f P A) show "\,\ \\<^sub>t\<^bsub>/F\<^esub> {s. f s \ P} (Basic f) P,A" proof (rule cvalidtI) fix s t assume "\\\Basic f,Normal s\ \ t" "s \ {s. f s \ P}" thus "t \ Normal ` P \ Abrupt ` A" by cases auto next fix s show "\\Basic f \ Normal s" by (rule terminates.intros) qed next case (Spec \ F r Q A) show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. (\t. (s, t) \ r \ t \ Q) \ (\t. (s, t) \ r)} Spec r Q,A" proof (rule cvalidtI) fix s t assume "\\\Spec r ,Normal s\ \ t" "s \ {s. (\t. (s, t) \ r \ t \ Q) \ (\t. (s, t) \ r)}" thus "t \ Normal ` Q \ Abrupt ` A" by cases auto next fix s show "\\Spec r \ Normal s" by (rule terminates.intros) qed next case (Seq \ F P c1 R A c2 Q) have valid_c1: "\,\ \\<^sub>t\<^bsub>/F\<^esub> P c1 R,A" by fact have valid_c2: "\,\ \\<^sub>t\<^bsub>/F\<^esub> R c2 Q,A" by fact show "\,\ \\<^sub>t\<^bsub>/F\<^esub> P Seq c1 c2 Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\Seq c1 c2,Normal s\ \ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" from exec P obtain r where exec_c1: "\\\c1,Normal s\ \ r" and exec_c2: "\\\c2,r\ \ t" by cases auto with t_notin_F have "r \ Fault ` F" by (auto dest: Fault_end) from valid_c1 ctxt exec_c1 P this have r: "r \ Normal ` R \ Abrupt ` A" by (rule cvalidt_postD) show "t\Normal ` Q \ Abrupt ` A" proof (cases r) case (Normal r') with exec_c2 r show "t\Normal ` Q \ Abrupt ` A" apply - apply (rule cvalidt_postD [OF valid_c2 ctxt _ _ t_notin_F]) apply auto done next case (Abrupt r') with exec_c2 have "t=Abrupt r'" by (auto elim: exec_elim_cases) with Abrupt r show ?thesis by auto next case Fault with r show ?thesis by blast next case Stuck with r show ?thesis by blast qed next fix s assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume P: "s\P" show "\\Seq c1 c2 \ Normal s" proof - from valid_c1 ctxt P have "\\c1\ Normal s" by (rule cvalidt_termD) moreover { fix r assume exec_c1: "\\\c1,Normal s\ \ r" have "\\c2 \ r" proof (cases r) case (Normal r') with cvalidt_postD [OF valid_c1 ctxt exec_c1 P] have r: "r\Normal ` R" by auto with cvalidt_termD [OF valid_c2 ctxt] exec_c1 show "\\c2 \ r" by auto qed auto } ultimately show ?thesis by (iprover intro: terminates.intros) qed qed next case (Cond \ F P b c1 Q A c2) have valid_c1: "\,\ \\<^sub>t\<^bsub>/F\<^esub> (P \ b) c1 Q,A" by fact have valid_c2: "\,\ \\<^sub>t\<^bsub>/F\<^esub> (P \ - b) c2 Q,A" by fact show "\,\ \\<^sub>t\<^bsub>/F\<^esub> P Cond b c1 c2 Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\Cond b c1 c2,Normal s\ \ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" show "t \ Normal ` Q \ Abrupt ` A" proof (cases "s\b") case True with exec have "\\\c1,Normal s\ \ t" by cases auto with P True show ?thesis by - (rule cvalidt_postD [OF valid_c1 ctxt _ _ t_notin_F],auto) next case False with exec P have "\\\c2,Normal s\ \ t" by cases auto with P False show ?thesis by - (rule cvalidt_postD [OF valid_c2 ctxt _ _ t_notin_F],auto) qed next fix s assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume P: "s \ P" thus "\\Cond b c1 c2 \ Normal s" using cvalidt_termD [OF valid_c1 ctxt] cvalidt_termD [OF valid_c2 ctxt] by (cases "s \ b") (auto intro: terminates.intros) qed next case (While r \ F P b c A) assume wf: "wf r" have valid_c: "\\. \,\\\<^sub>t\<^bsub>/F\<^esub> ({\} \ P \ b) c ({t. (t, \) \ r} \ P),A" using While.hyps by iprover show "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (While b c) (P \ - b),A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume wprems: "\\\While b c,Normal s\ \ t" "s \ P" "t \ Fault ` F" from wf have "\t. \\\\While b c,Normal s\ \ t; s \ P; t \ Fault ` F\ \ t \ Normal ` (P \ - b) \ Abrupt ` A" proof (induct) fix s t assume hyp: "\s' t. \(s',s)\r; \\\While b c,Normal s'\ \ t; s' \ P; t \ Fault ` F\ \ t \ Normal ` (P \ - b) \ Abrupt ` A" assume exec: "\\\While b c,Normal s\ \ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" from exec show "t \ Normal ` (P \ - b) \ Abrupt ` A" proof (cases) fix s' assume b: "s\b" assume exec_c: "\\\c,Normal s\ \ s'" assume exec_w: "\\\While b c,s'\ \ t" from exec_w t_notin_F have "s' \ Fault ` F" by (auto dest: Fault_end) from exec_c P b valid_c ctxt this have s': "s' \ Normal ` ({s'. (s', s) \ r} \ P) \ Abrupt ` A" by (auto simp add: cvalidt_def validt_def valid_def) show ?thesis proof (cases s') case Normal with exec_w s' t_notin_F show ?thesis by - (rule hyp,auto) next case Abrupt with exec_w have "t=s'" by (auto dest: Abrupt_end) with Abrupt s' show ?thesis by blast next case Fault with exec_w have "t=s'" by (auto dest: Fault_end) with Fault s' show ?thesis by blast next case Stuck with exec_w have "t=s'" by (auto dest: Stuck_end) with Stuck s' show ?thesis by blast qed next assume "s\b" "t=Normal s" with P show ?thesis by simp qed qed with wprems show "t \ Normal ` (P \ - b) \ Abrupt ` A" by blast next fix s assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume "s \ P" with wf show "\\While b c \ Normal s" proof (induct) fix s assume hyp: "\s'. \(s',s)\r; s' \ P\ \ \\While b c \ Normal s'" assume P: "s \ P" show "\\While b c \ Normal s" proof (cases "s \ b") case False with P show ?thesis by (blast intro: terminates.intros) next case True with valid_c P ctxt have "\\c \ Normal s" by (simp add: cvalidt_def validt_def) moreover { fix s' assume exec_c: "\\\c,Normal s\ \ s'" have "\\While b c \ s'" proof (cases s') case (Normal s'') with exec_c P True valid_c ctxt have s': "s' \ Normal ` ({s'. (s', s) \ r} \ P)" by (fastforce simp add: cvalidt_def validt_def valid_def) then show ?thesis by (blast intro: hyp) qed auto } ultimately show ?thesis by (blast intro: terminates.intros) qed qed qed next case (Guard \ F g P c Q A f) have valid_c: "\,\ \\<^sub>t\<^bsub>/F\<^esub> (g \ P) c Q,A" by fact show "\,\ \\<^sub>t\<^bsub>/F\<^esub> (g \ P) Guard f g c Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\Guard f g c,Normal s\ \ t" assume t_notin_F: "t \ Fault ` F" assume P:"s \ (g \ P)" from exec P have "\\\c,Normal s\ \ t" by cases auto from valid_c ctxt this P t_notin_F show "t \ Normal ` Q \ Abrupt ` A" by (rule cvalidt_postD) next fix s assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume P:"s \ (g \ P)" thus "\\Guard f g c \ Normal s" by (auto intro: terminates.intros cvalidt_termD [OF valid_c ctxt]) qed next case (Guarantee f F \ g P c Q A) have valid_c: "\,\ \\<^sub>t\<^bsub>/F\<^esub> (g \ P) c Q,A" by fact have f_F: "f \ F" by fact show "\,\ \\<^sub>t\<^bsub>/F\<^esub> P Guard f g c Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\Guard f g c,Normal s\ \ t" assume t_notin_F: "t \ Fault ` F" assume P:"s \ P" from exec f_F t_notin_F have g: "s \ g" by cases auto with P have P': "s \ g \ P" by blast from exec g have "\\\c,Normal s\ \ t" by cases auto from valid_c ctxt this P' t_notin_F show "t \ Normal ` Q \ Abrupt ` A" by (rule cvalidt_postD) next fix s assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume P:"s \ P" thus "\\Guard f g c \ Normal s" by (auto intro: terminates.intros cvalidt_termD [OF valid_c ctxt]) qed next case (CallRec P p Q A Specs r Specs_wf \ F) have p: "(P,p,Q,A) \ Specs" by fact have wf: "wf r" by fact have Specs_wf: "Specs_wf = (\p \. (\(P,q,Q,A). (P \ {s. ((s, q),\,p) \ r},q,Q,A)) ` Specs)" by fact from CallRec.hyps have valid_body: "\(P, p, Q, A)\Specs. p \ dom \ \ (\\. \,\ \ Specs_wf p \\\<^sub>t\<^bsub>/F\<^esub> ({\} \ P) the (\ p) Q,A)" by auto show "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" proof - { fix \p assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" from wf have "\\ p P Q A. \\p = (\,p); (P,p,Q,A) \ Specs\ \ \\\<^sub>t\<^bsub>/F\<^esub> ({\} \ P) (the (\ (p))) Q,A" proof (induct \p rule: wf_induct [rule_format, consumes 1, case_names WF]) case (WF \p \ p P Q A) have \p: "\p = (\, p)" by fact have p: "(P, p, Q, A) \ Specs" by fact { fix q P' Q' A' assume q: "(P',q,Q',A') \ Specs" have "\\\<^sub>t\<^bsub>/F\<^esub> (P' \ {s. ((s,q), \,p) \ r}) (Call q) Q',A'" proof (rule validtI) fix s t assume exec_q: "\\\Call q,Normal s\ \ t" assume Pre: "s \ P' \ {s. ((s,q), \,p) \ r}" assume t_notin_F: "t \ Fault ` F" from Pre q \p have valid_bdy: "\\\<^sub>t\<^bsub>/F\<^esub> ({s} \ P') the (\ q) Q',A'" by - (rule WF.hyps, auto) from Pre q have Pre': "s \ {s} \ P'" by auto from exec_q show "t \ Normal ` Q' \ Abrupt ` A'" proof (cases) fix bdy assume bdy: "\ q = Some bdy" assume exec_bdy: "\\\bdy,Normal s\ \ t" from valid_bdy [simplified bdy option.sel] t_notin_F exec_bdy Pre' have "t \ Normal ` Q' \ Abrupt ` A'" by (auto simp add: validt_def valid_def) with Pre q show ?thesis by auto next assume "\ q = None" with q valid_body have False by auto thus ?thesis .. qed next fix s assume Pre: "s \ P' \ {s. ((s,q), \,p) \ r}" from Pre q \p have valid_bdy: "\ \\<^sub>t\<^bsub>/F\<^esub> ({s} \ P') (the (\ q)) Q',A'" by - (rule WF.hyps, auto) from Pre q have Pre': "s \ {s} \ P'" by auto from valid_bdy ctxt Pre' have "\\the (\ q) \ Normal s" by (auto simp add: validt_def) with valid_body q show "\\Call q\ Normal s" by (fastforce intro: terminates.Call) qed } hence "\(P, p, Q, A)\Specs_wf p \. \\\<^sub>t\<^bsub>/F\<^esub> P Call p Q,A" by (auto simp add: cvalidt_def Specs_wf) with ctxt have "\(P, p, Q, A)\\ \ Specs_wf p \. \\\<^sub>t\<^bsub>/F\<^esub> P Call p Q,A" by auto with p valid_body show "\ \\<^sub>t\<^bsub>/F\<^esub> ({\} \ P) (the (\ p)) Q,A" by (simp add: cvalidt_def) blast qed } note lem = this have valid_body': "\\. \(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A \ \(P,p,Q,A)\Specs. \\\<^sub>t\<^bsub>/F\<^esub> ({\} \ P) (the (\ p)) Q,A" by (auto intro: lem) show "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec_call: "\\\Call p,Normal s\ \ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" from exec_call show "t \ Normal ` Q \ Abrupt ` A" proof (cases) fix bdy assume bdy: "\ p = Some bdy" assume exec_body: "\\\bdy,Normal s\ \ t" from exec_body bdy p P t_notin_F valid_body' [of "s", OF ctxt] ctxt have "t \ Normal ` Q \ Abrupt ` A" apply (simp only: cvalidt_def validt_def valid_def) apply (drule (1) bspec) apply auto done with p P show ?thesis by simp next assume "\ p = None" with p valid_body have False by auto thus ?thesis by simp qed next fix s assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume P: "s \ P" show "\\Call p \ Normal s" proof - from ctxt P p valid_body' [of "s",OF ctxt] have "\\(the (\ p)) \ Normal s" by (auto simp add: cvalidt_def validt_def) with valid_body p show ?thesis by (fastforce intro: terminates.Call) qed qed qed next case (DynCom P \ F c Q A) hence valid_c: "\s\P. \,\\\<^sub>t\<^bsub>/F\<^esub> P (c s) Q,A" by simp show "\,\\\<^sub>t\<^bsub>/F\<^esub> P DynCom c Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\DynCom c,Normal s\ \ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" from exec show "t \ Normal ` Q \ Abrupt ` A" proof (cases) assume "\\\c s,Normal s\ \ t" from cvalidt_postD [OF valid_c [rule_format, OF P] ctxt this P t_notin_F] show ?thesis . qed next fix s assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume P: "s \ P" show "\\DynCom c \ Normal s" proof - from cvalidt_termD [OF valid_c [rule_format, OF P] ctxt P] have "\\c s \ Normal s" . thus ?thesis by (rule terminates.intros) qed qed next case (Throw \ F A Q) show "\,\ \\<^sub>t\<^bsub>/F\<^esub> A Throw Q,A" proof (rule cvalidtI) fix s t assume "\\\Throw,Normal s\ \ t" "s \ A" then show "t \ Normal ` Q \ Abrupt ` A" by cases simp next fix s show "\\Throw \ Normal s" by (rule terminates.intros) qed next case (Catch \ F P c\<^sub>1 Q R c\<^sub>2 A) have valid_c1: "\,\ \\<^sub>t\<^bsub>/F\<^esub> P c\<^sub>1 Q,R" by fact have valid_c2: "\,\ \\<^sub>t\<^bsub>/F\<^esub> R c\<^sub>2 Q,A" by fact show "\,\ \\<^sub>t\<^bsub>/F\<^esub> P Catch c\<^sub>1 c\<^sub>2 Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\Catch c\<^sub>1 c\<^sub>2,Normal s\ \ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" from exec show "t \ Normal ` Q \ Abrupt ` A" proof (cases) fix s' assume exec_c1: "\\\c\<^sub>1,Normal s\ \ Abrupt s'" assume exec_c2: "\\\c\<^sub>2,Normal s'\ \ t" from cvalidt_postD [OF valid_c1 ctxt exec_c1 P] have "Abrupt s' \ Abrupt ` R" by auto with cvalidt_postD [OF valid_c2 ctxt] exec_c2 t_notin_F show ?thesis by fastforce next assume exec_c1: "\\\c\<^sub>1,Normal s\ \ t" assume notAbr: "\ isAbr t" from cvalidt_postD [OF valid_c1 ctxt exec_c1 P] t_notin_F have "t \ Normal ` Q \ Abrupt ` R" . with notAbr show ?thesis by auto qed next fix s assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume P: "s \ P" show "\\Catch c\<^sub>1 c\<^sub>2 \ Normal s" proof - from valid_c1 ctxt P have "\\c\<^sub>1\ Normal s" by (rule cvalidt_termD) moreover { fix r assume exec_c1: "\\\c\<^sub>1,Normal s\ \ Abrupt r" from cvalidt_postD [OF valid_c1 ctxt exec_c1 P] have r: "Abrupt r\Normal ` Q \ Abrupt ` R" by auto hence "Abrupt r\Abrupt ` R" by fast with cvalidt_termD [OF valid_c2 ctxt] exec_c1 have "\\c\<^sub>2 \ Normal r" by fast } ultimately show ?thesis by (iprover intro: terminates.intros) qed qed next case (Conseq P \ F c Q A) hence adapt: "\s \ P. (\P' Q' A'. (\,\ \\<^sub>t\<^bsub>/F\<^esub> P' c Q',A') \ s \ P'\ Q' \ Q \ A' \ A)" by blast show "\,\ \\<^sub>t\<^bsub>/F\<^esub> P c Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\c,Normal s\ \ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" show "t \ Normal ` Q \ Abrupt ` A" proof - from adapt [rule_format, OF P] obtain P' and Q' and A' where valid_P'_Q': "\,\ \\<^sub>t\<^bsub>/F\<^esub> P' c Q',A'" and weaken: "s \ P'" "Q' \ Q" "A'\ A" by blast from exec valid_P'_Q' ctxt t_notin_F have P'_Q': "Normal s \ Normal ` P' \ t \ Normal ` Q' \ Abrupt ` A'" by (unfold cvalidt_def validt_def valid_def) blast hence "s \ P' \ t \ Normal ` Q' \ Abrupt ` A'" by blast with weaken show ?thesis by blast qed next fix s assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume P: "s \ P" show "\\c \ Normal s" proof - from P adapt obtain P' and Q' and A' where "\,\ \\<^sub>t\<^bsub>/F\<^esub> P' c Q',A'" "s \ P'" by blast with ctxt show ?thesis by (simp add: cvalidt_def validt_def) qed qed next case (Asm P p Q A \ F) assume "(P, p, Q, A) \ \" then show "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" by (auto simp add: cvalidt_def ) next case ExFalso thus ?case by iprover qed lemma hoaret_sound': "\,{}\\<^sub>t\<^bsub>/F\<^esub> P c Q,A \ \\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" apply (drule hoaret_sound) apply (simp add: cvalidt_def) done theorem total_to_partial: assumes total: "\,{}\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" shows "\,{}\\<^bsub>/F\<^esub> P c Q,A" proof - from total have "\,{}\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (rule hoaret_sound) hence "\\\<^bsub>/F\<^esub> P c Q,A" by (simp add: cvalidt_def validt_def cvalid_def) thus ?thesis by (rule hoare_complete) qed subsection \Completeness\ lemma MGT_valid: "\\\<^sub>t\<^bsub>/F \<^esub>{s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\c\Normal s} c {t. \\\c,Normal Z\ \ Normal t}, {t. \\\c,Normal Z\ \ Abrupt t}" proof (rule validtI) fix s t assume "\\\c,Normal s\ \ t" "s \ {s. s = Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\c\Normal s}" "t \ Fault ` F" thus "t \ Normal ` {t. \\\c,Normal Z\ \ Normal t} \ Abrupt ` {t. \\\c,Normal Z\ \ Abrupt t}" apply (cases t) apply (auto simp add: final_notin_def) done next fix s assume "s \ {s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\c\Normal s}" thus "\\c\Normal s" by blast qed text \The consequence rule where the existential @{term Z} is instantiated to @{term s}. Usefull in proof of \MGT_lemma\.\ lemma ConseqMGT: assumes modif: "\Z::'a. \,\ \\<^sub>t\<^bsub>/F\<^esub> (P' Z::'a assn) c (Q' Z),(A' Z)" assumes impl: "\s. s \ P \ s \ P' s \ (\t. t \ Q' s \ t \ Q) \ (\t. t \ A' s \ t \ A)" shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P c Q,A" using impl by - (rule conseq [OF modif],blast) lemma MGT_implies_complete: assumes MGT: "\Z. \,{}\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\c\Normal s} c {t. \\\c,Normal Z\ \ Normal t}, {t. \\\c,Normal Z\ \ Abrupt t}" assumes valid: "\ \\<^sub>t\<^bsub>/F\<^esub> P c Q,A" shows "\,{} \\<^sub>t\<^bsub>/F\<^esub> P c Q,A" using MGT apply (rule ConseqMGT) apply (insert valid) apply (auto simp add: validt_def valid_def intro!: final_notinI) done lemma conseq_extract_state_indep_prop: assumes state_indep_prop:"\s \ P. R" assumes to_show: "R \ \,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" apply (rule Conseq) apply (clarify) apply (rule_tac x="P" in exI) apply (rule_tac x="Q" in exI) apply (rule_tac x="A" in exI) using state_indep_prop to_show by blast lemma MGT_lemma: assumes MGT_Calls: "\p \ dom \. \Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\(Call p)\Normal s} (Call p) {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t}" shows "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\c\Normal s} c {t. \\\c,Normal Z\ \ Normal t},{t. \\\c,Normal Z\ \ Abrupt t}" proof (induct c) case Skip show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s = Z \ \\\Skip,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Skip \ Normal s} Skip {t. \\\Skip,Normal Z\ \ Normal t},{t. \\\Skip,Normal Z\ \ Abrupt t}" by (rule hoaret.Skip [THEN conseqPre]) (auto elim: exec_elim_cases simp add: final_notin_def intro: exec.intros terminates.intros) next case (Basic f) show "\,\\\<^sub>t\<^bsub>/F \<^esub>{s. s=Z \ \\\Basic f,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Basic f \ Normal s} Basic f {t. \\\Basic f,Normal Z\ \ Normal t}, {t. \\\Basic f,Normal Z\ \ Abrupt t}" by (rule hoaret.Basic [THEN conseqPre]) (auto elim: exec_elim_cases simp add: final_notin_def intro: exec.intros terminates.intros) next case (Spec r) show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Spec r,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Spec r \ Normal s} Spec r {t. \\\Spec r,Normal Z\ \ Normal t}, {t. \\\Spec r,Normal Z\ \ Abrupt t}" apply (rule hoaret.Spec [THEN conseqPre]) apply (clarsimp simp add: final_notin_def) apply (case_tac "\t. (Z,t) \ r") apply (auto elim: exec_elim_cases simp add: final_notin_def intro: exec.intros) done next case (Seq c1 c2) have hyp_c1: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c1,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\c1\Normal s} c1 {t. \\\c1,Normal Z\ \ Normal t}, {t. \\\c1,Normal Z\ \ Abrupt t}" using Seq.hyps by iprover have hyp_c2: "\ Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\c2\Normal s} c2 {t. \\\c2,Normal Z\ \ Normal t}, {t. \\\c2,Normal Z\ \ Abrupt t}" using Seq.hyps by iprover from hyp_c1 have "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Seq c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Seq c1 c2\Normal s} c1 {t. \\\c1,Normal Z\ \ Normal t \ \\\c2,Normal t\ \\({Stuck} \ Fault ` (-F)) \ \\c2\Normal t}, {t. \\\Seq c1 c2,Normal Z\ \ Abrupt t}" by (rule ConseqMGT) (auto dest: Seq_NoFaultStuckD1 [simplified] Seq_NoFaultStuckD2 [simplified] elim: terminates_Normal_elim_cases intro: exec.intros) thus "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Seq c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Seq c1 c2\Normal s} Seq c1 c2 {t. \\\Seq c1 c2,Normal Z\ \ Normal t}, {t. \\\Seq c1 c2,Normal Z\ \ Abrupt t}" proof (rule hoaret.Seq ) show "\,\\\<^sub>t\<^bsub>/F\<^esub> {t. \\\c1,Normal Z\ \ Normal t \ \\\c2,Normal t\ \\({Stuck} \ Fault ` (-F)) \ \\c2 \ Normal t} c2 {t. \\\Seq c1 c2,Normal Z\ \ Normal t}, {t. \\\Seq c1 c2,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [OF hyp_c2],safe) fix r t assume "\\\c1,Normal Z\ \ Normal r" "\\\c2,Normal r\ \ Normal t" then show "\\\Seq c1 c2,Normal Z\ \ Normal t" by (rule exec.intros) next fix r t assume "\\\c1,Normal Z\ \ Normal r" "\\\c2,Normal r\ \ Abrupt t" then show "\\\Seq c1 c2,Normal Z\ \ Abrupt t" by (rule exec.intros) qed qed next case (Cond b c1 c2) have "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c1,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\c1\Normal s} c1 {t. \\\c1,Normal Z\ \ Normal t}, {t. \\\c1,Normal Z\ \ Abrupt t}" using Cond.hyps by iprover hence "\,\ \\<^sub>t\<^bsub>/F\<^esub> ({s. s=Z \ \\\Cond b c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\(Cond b c1 c2)\Normal s}\b) c1 {t. \\\Cond b c1 c2,Normal Z\ \ Normal t}, {t. \\\Cond b c1 c2,Normal Z\ \ Abrupt t}" by (rule ConseqMGT) (fastforce simp add: final_notin_def intro: exec.CondTrue elim: terminates_Normal_elim_cases) moreover have "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\c2\Normal s} c2 {t. \\\c2,Normal Z\ \ Normal t}, {t. \\\c2,Normal Z\ \ Abrupt t}" using Cond.hyps by iprover hence "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. s=Z \ \\\Cond b c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\(Cond b c1 c2)\Normal s}\-b) c2 {t. \\\Cond b c1 c2,Normal Z\ \ Normal t}, {t. \\\Cond b c1 c2,Normal Z\ \ Abrupt t}" by (rule ConseqMGT) (fastforce simp add: final_notin_def intro: exec.CondFalse elim: terminates_Normal_elim_cases) ultimately show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Cond b c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\(Cond b c1 c2)\Normal s} (Cond b c1 c2) {t. \\\Cond b c1 c2,Normal Z\ \ Normal t}, {t. \\\Cond b c1 c2,Normal Z\ \ Abrupt t}" by (rule hoaret.Cond) next case (While b c) let ?unroll = "({(s,t). s\b \ \\\c,Normal s\ \ Normal t})\<^sup>*" let ?P' = "\Z. {t. (Z,t)\?unroll \ (\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)) \ \\(While b c)\Normal t}" let ?A = "\Z. {t. \\\While b c,Normal Z\ \ Abrupt t}" let ?r = "{(t,s). \\(While b c)\Normal s \ s\b \ \\\c,Normal s\ \ Normal t}" show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\While b c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\(While b c)\Normal s} (While b c) {t. \\\While b c,Normal Z\ \ Normal t}, {t. \\\While b c,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [where ?P'="\ Z. ?P' Z" and ?Q'="\ Z. ?P' Z \ - b"]) have wf_r: "wf ?r" by (rule wf_terminates_while) show "\ Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (?P' Z) (While b c) (?P' Z \ - b),(?A Z)" proof (rule allI, rule hoaret.While [OF wf_r]) fix Z from While have hyp_c: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub>{s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\c\Normal s} c {t. \\\c,Normal Z\ \ Normal t}, {t. \\\c,Normal Z\ \ Abrupt t}" by iprover show "\\. \,\\\<^sub>t\<^bsub>/F\<^esub> ({\} \ ?P' Z \ b) c ({t. (t, \) \ ?r} \ ?P' Z),(?A Z)" proof (rule allI, rule ConseqMGT [OF hyp_c]) fix \ s assume "s\ {\} \ {t. (Z, t) \ ?unroll \ (\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)) \ \\(While b c)\Normal t} \ b" then obtain s_eq_\: "s=\" and Z_s_unroll: "(Z,s) \ ?unroll" and noabort:"\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)" and while_term: "\\(While b c)\Normal s" and s_in_b: "s\b" by blast show "s \ {t. t = s \ \\\c,Normal t\ \\({Stuck} \ Fault ` (-F)) \ \\c\Normal t} \ (\t. t \ {t. \\\c,Normal s\ \ Normal t} \ t \ {t. (t,\) \ ?r} \ {t. (Z, t) \ ?unroll \ (\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)) \ \\(While b c)\Normal t}) \ (\t. t \ {t. \\\c,Normal s\ \ Abrupt t} \ t \ {t. \\\While b c,Normal Z\ \ Abrupt t})" (is "?C1 \ ?C2 \ ?C3") proof (intro conjI) from Z_s_unroll noabort s_in_b while_term show ?C1 by (blast elim: terminates_Normal_elim_cases) next { fix t assume s_t: "\\\c,Normal s\ \ Normal t" with s_eq_\ while_term s_in_b have "(t,\) \ ?r" by blast moreover from Z_s_unroll s_t s_in_b have "(Z, t) \ ?unroll" by (blast intro: rtrancl_into_rtrancl) moreover from while_term s_t s_in_b have "\\(While b c)\Normal t" by (blast elim: terminates_Normal_elim_cases) moreover note noabort ultimately have "(t,\) \ ?r \ (Z, t) \ ?unroll \ (\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)) \ \\(While b c)\Normal t" by iprover } then show ?C2 by blast next { fix t assume s_t: "\\\c,Normal s\ \ Abrupt t" from Z_s_unroll noabort s_t s_in_b have "\\\While b c,Normal Z\ \ Abrupt t" by blast } thus ?C3 by simp qed qed qed next fix s assume P: "s \ {s. s=Z \ \\\While b c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\While b c \ Normal s}" hence WhileNoFault: "\\\While b c,Normal Z\ \\({Stuck} \ Fault ` (-F))" by auto show "s \ ?P' s \ (\t. t\(?P' s \ - b)\ t\{t. \\\While b c,Normal Z\ \ Normal t})\ (\t. t\?A s \ t\?A Z)" proof (intro conjI) { fix e assume "(Z,e) \ ?unroll" "e \ b" from this WhileNoFault have "\\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)" (is "?Prop Z e") proof (induct rule: converse_rtrancl_induct [consumes 1]) assume e_in_b: "e \ b" assume WhileNoFault: "\\\While b c,Normal e\ \\({Stuck} \ Fault ` (-F))" with e_in_b WhileNoFault have cNoFault: "\\\c,Normal e\ \\({Stuck} \ Fault ` (-F))" by (auto simp add: final_notin_def intro: exec.intros) moreover { fix u assume "\\\c,Normal e\ \ Abrupt u" with e_in_b have "\\\While b c,Normal e\ \ Abrupt u" by (blast intro: exec.intros) } ultimately show "?Prop e e" by iprover next fix Z r assume e_in_b: "e\b" assume WhileNoFault: "\\\While b c,Normal Z\ \\({Stuck} \ Fault ` (-F))" assume hyp: "\e\b;\\\While b c,Normal r\ \\({Stuck} \ Fault ` (-F))\ \ ?Prop r e" assume Z_r: "(Z, r) \ {(Z, r). Z \ b \ \\\c,Normal Z\ \ Normal r}" with WhileNoFault have "\\\While b c,Normal r\ \\({Stuck} \ Fault ` (-F))" by (auto simp add: final_notin_def intro: exec.intros) from hyp [OF e_in_b this] obtain cNoFault: "\\\c,Normal e\ \\({Stuck} \ Fault ` (-F))" and Abrupt_r: "\u. \\\c,Normal e\ \ Abrupt u \ \\\While b c,Normal r\ \ Abrupt u" by simp { fix u assume "\\\c,Normal e\ \ Abrupt u" with Abrupt_r have "\\\While b c,Normal r\ \ Abrupt u" by simp moreover from Z_r obtain "Z \ b" "\\\c,Normal Z\ \ Normal r" by simp ultimately have "\\\While b c,Normal Z\ \ Abrupt u" by (blast intro: exec.intros) } with cNoFault show "?Prop Z e" by iprover qed } with P show "s \ ?P' s" by blast next { fix t assume "termination": "t \ b" assume "(Z, t) \ ?unroll" hence "\\\While b c,Normal Z\ \ Normal t" proof (induct rule: converse_rtrancl_induct [consumes 1]) from "termination" show "\\\While b c,Normal t\ \ Normal t" by (blast intro: exec.WhileFalse) next fix Z r assume first_body: "(Z, r) \ {(s, t). s \ b \ \\\c,Normal s\ \ Normal t}" assume "(r, t) \ ?unroll" assume rest_loop: "\\\While b c, Normal r\ \ Normal t" show "\\\While b c,Normal Z\ \ Normal t" proof - from first_body obtain "Z \ b" "\\\c,Normal Z\ \ Normal r" by fast moreover from rest_loop have "\\\While b c,Normal r\ \ Normal t" by fast ultimately show "\\\While b c,Normal Z\ \ Normal t" by (rule exec.WhileTrue) qed qed } with P show "(\t. t\(?P' s \ - b) \t\{t. \\\While b c,Normal Z\ \ Normal t})" by blast next from P show "\t. t\?A s \ t\?A Z" by simp qed qed next case (Call p) from noStuck_Call have "\s \ {s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\ Normal s}. p \ dom \" by (fastforce simp add: final_notin_def) then show ?case proof (rule conseq_extract_state_indep_prop) assume p_defined: "p \ dom \" with MGT_Calls show "\,\\\<^sub>t\<^bsub>/F \<^esub>{s. s=Z \ \\\Call p ,Normal s\ \\({Stuck} \ Fault ` (-F))\ \\Call p\Normal s} (Call p) {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t}" by (auto) qed next case (DynCom c) have hyp: "\s'. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s = Z \ \\\c s',Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\c s'\Normal s} c s' {t. \\\c s',Normal Z\ \ Normal t},{t. \\\c s',Normal Z\ \ Abrupt t}" using DynCom by simp have hyp': "\,\\\<^sub>t\<^bsub>/F \<^esub>{s. s = Z \ \\\DynCom c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\DynCom c\Normal s} (c Z) {t. \\\DynCom c,Normal Z\ \ Normal t},{t. \\\DynCom c,Normal Z\ \ Abrupt t}" by (rule ConseqMGT [OF hyp]) (fastforce simp add: final_notin_def intro: exec.intros elim: terminates_Normal_elim_cases) show "\,\\\<^sub>t\<^bsub>/F \<^esub>{s. s=Z \ \\\DynCom c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\DynCom c\Normal s} DynCom c {t. \\\DynCom c,Normal Z\ \ Normal t}, {t. \\\DynCom c,Normal Z\ \ Abrupt t}" apply (rule hoaret.DynCom) apply (clarsimp) apply (rule hyp' [simplified]) done next case (Guard f g c) have hyp_c: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\c\Normal s} c {t. \\\c,Normal Z\ \ Normal t}, {t. \\\c,Normal Z\ \ Abrupt t}" using Guard by iprover show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s = Z \ \\\Guard f g c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Guard f g c\ Normal s} Guard f g c {t. \\\Guard f g c ,Normal Z\ \ Normal t}, {t. \\\Guard f g c,Normal Z\ \ Abrupt t}" proof (cases "f \ F") case True from hyp_c have "\,\\\<^sub>t\<^bsub>/F \<^esub>(g \ {s. s=Z \ \\\Guard f g c,Normal s\ \\({Stuck}\ Fault ` (-F))\ \\Guard f g c\ Normal s}) c {t. \\\Guard f g c,Normal Z\ \ Normal t}, {t. \\\Guard f g c,Normal Z\ \ Abrupt t}" apply (rule ConseqMGT) apply (insert True) apply (auto simp add: final_notin_def intro: exec.intros elim: terminates_Normal_elim_cases) done from True this show ?thesis by (rule conseqPre [OF Guarantee]) auto next case False from hyp_c have "\,\\\<^sub>t\<^bsub>/F \<^esub>(g \ {s. s \ g \ s=Z \ \\\Guard f g c,Normal s\ \\({Stuck}\ Fault ` (-F))\ \\Guard f g c\ Normal s} ) c {t. \\\Guard f g c,Normal Z\ \ Normal t}, {t. \\\Guard f g c,Normal Z\ \ Abrupt t}" apply (rule ConseqMGT) apply clarify apply (frule Guard_noFaultStuckD [OF _ False]) apply (auto simp add: final_notin_def intro: exec.intros elim: terminates_Normal_elim_cases) done then show ?thesis apply (rule conseqPre [OF hoaret.Guard]) apply clarify apply (frule Guard_noFaultStuckD [OF _ False]) apply auto done qed next case Throw show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s = Z \ \\\Throw,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Throw \ Normal s} Throw {t. \\\Throw,Normal Z\ \ Normal t}, {t. \\\Throw,Normal Z\ \ Abrupt t}" by (rule conseqPre [OF hoaret.Throw]) (blast intro: exec.intros terminates.intros) next case (Catch c\<^sub>1 c\<^sub>2) have "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s = Z \ \\\c\<^sub>1,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\c\<^sub>1 \ Normal s} c\<^sub>1 {t. \\\c\<^sub>1,Normal Z\ \ Normal t}, {t. \\\c\<^sub>1,Normal Z\ \ Abrupt t}" using Catch.hyps by iprover hence "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s = Z \ \\\Catch c\<^sub>1 c\<^sub>2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Catch c\<^sub>1 c\<^sub>2 \ Normal s} c\<^sub>1 {t. \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Normal t}, {t. \\\c\<^sub>1,Normal Z\ \ Abrupt t \ \\c\<^sub>2 \ Normal t \ \\\c\<^sub>2,Normal t\ \\({Stuck} \ Fault ` (-F))}" by (rule ConseqMGT) (fastforce intro: exec.intros terminates.intros elim: terminates_Normal_elim_cases simp add: final_notin_def) moreover have "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c\<^sub>2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\c\<^sub>2 \ Normal s} c\<^sub>2 {t. \\\c\<^sub>2,Normal Z\ \ Normal t}, {t. \\\c\<^sub>2,Normal Z\ \ Abrupt t}" using Catch.hyps by iprover hence "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. \\\c\<^sub>1,Normal Z\ \Abrupt s \ \\c\<^sub>2 \ Normal s \ \\\c\<^sub>2,Normal s\ \\({Stuck} \ Fault ` (-F))} c\<^sub>2 {t. \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Normal t}, {t. \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Abrupt t}" by (rule ConseqMGT) (fastforce intro: exec.intros terminates.intros simp add: noFault_def') ultimately show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s = Z \ \\\Catch c\<^sub>1 c\<^sub>2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Catch c\<^sub>1 c\<^sub>2 \ Normal s} Catch c\<^sub>1 c\<^sub>2 {t. \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Normal t}, {t. \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Abrupt t}" by (rule hoaret.Catch ) qed lemma Call_lemma': assumes Call_hyp: "\q\dom \. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub>{s. s=Z \ \\\Call q,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call q\Normal s \ ((s,q),(\,p)) \ termi_call_steps \} (Call q) {t. \\\Call q,Normal Z\ \ Normal t}, {t. \\\Call q,Normal Z\ \ Abrupt t}" shows "\Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ c \ redexes c')} c {t. \\\c,Normal Z\ \ Normal t}, {t. \\\c,Normal Z\ \ Abrupt t}" proof (induct c) case Skip show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s = Z \ \\\Skip,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p \ Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ Skip \ redexes c')} Skip {t. \\\Skip,Normal Z\ \ Normal t}, {t. \\\Skip,Normal Z\ \ Abrupt t}" by (rule hoaret.Skip [THEN conseqPre]) (blast intro: exec.Skip) next case (Basic f) show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Basic f,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ Basic f \ redexes c')} Basic f {t. \\\Basic f,Normal Z\ \ Normal t}, {t. \\\Basic f,Normal Z\ \ Abrupt t}" by (rule hoaret.Basic [THEN conseqPre]) (blast intro: exec.Basic) next case (Spec r) show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Spec r,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ Spec r \ redexes c')} Spec r {t. \\\Spec r,Normal Z\ \ Normal t}, {t. \\\Spec r,Normal Z\ \ Abrupt t}" apply (rule hoaret.Spec [THEN conseqPre]) apply (clarsimp) apply (case_tac "\t. (Z,t) \ r") apply (auto elim: exec_elim_cases simp add: final_notin_def intro: exec.intros) done next case (Seq c1 c2) have hyp_c1: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c1,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ c1 \ redexes c')} c1 {t. \\\c1,Normal Z\ \ Normal t}, {t. \\\c1,Normal Z\ \ Abrupt t}" using Seq.hyps by iprover have hyp_c2: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ c2 \ redexes c')} c2 {t. \\\c2,Normal Z\ \ Normal t}, {t. \\\c2,Normal Z\ \ Abrupt t}" using Seq.hyps (2) by iprover have c1: "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Seq c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ Seq c1 c2 \ redexes c')} c1 {t. \\\c1,Normal Z\ \ Normal t \ \\\c2,Normal t\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal t) \ c2 \ redexes c')}, {t. \\\Seq c1 c2,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [OF hyp_c1],clarify,safe) assume "\\\Seq c1 c2,Normal Z\ \\({Stuck} \ Fault ` (-F))" thus "\\\c1,Normal Z\ \\({Stuck} \ Fault ` (-F))" by (blast dest: Seq_NoFaultStuckD1) next fix c' assume steps_c': "\\ (Call p, Normal \) \\<^sup>+ (c', Normal Z)" assume red: "Seq c1 c2 \ redexes c'" from redexes_subset [OF red] steps_c' show "\c'. \\ (Call p, Normal \) \\<^sup>+ (c', Normal Z) \ c1 \ redexes c'" by (auto iff: root_in_redexes) next fix t assume "\\\Seq c1 c2,Normal Z\ \\({Stuck} \ Fault ` (-F))" "\\\c1,Normal Z\ \ Normal t" thus "\\\c2,Normal t\ \\({Stuck} \ Fault ` (-F))" by (blast dest: Seq_NoFaultStuckD2) next fix c' t assume steps_c': "\\ (Call p, Normal \) \\<^sup>+ (c', Normal Z)" assume red: "Seq c1 c2 \ redexes c'" assume exec_c1: "\\ \c1,Normal Z\ \ Normal t" show "\c'. \\ (Call p, Normal \) \\<^sup>+ (c', Normal t) \ c2 \ redexes c'" proof - note steps_c' also from exec_impl_steps_Normal [OF exec_c1] have "\\ (c1, Normal Z) \\<^sup>* (Skip, Normal t)". from steps_redexes_Seq [OF this red] obtain c'' where steps_c'': "\\ (c', Normal Z) \\<^sup>* (c'', Normal t)" and Skip: "Seq Skip c2 \ redexes c''" by blast note steps_c'' also have step_Skip: "\\ (Seq Skip c2,Normal t) \ (c2,Normal t)" by (rule step.SeqSkip) from step_redexes [OF step_Skip Skip] obtain c''' where step_c''': "\\ (c'', Normal t) \ (c''', Normal t)" and c2: "c2 \ redexes c'''" by blast note step_c''' finally show ?thesis using c2 by blast qed next fix t assume "\\\c1,Normal Z\ \ Abrupt t" thus "\\\Seq c1 c2,Normal Z\ \ Abrupt t" by (blast intro: exec.intros) qed show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Seq c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ Seq c1 c2 \ redexes c')} Seq c1 c2 {t. \\\Seq c1 c2,Normal Z\ \ Normal t}, {t. \\\Seq c1 c2,Normal Z\ \ Abrupt t}" by (rule hoaret.Seq [OF c1 ConseqMGT [OF hyp_c2]]) (blast intro: exec.intros) next case (Cond b c1 c2) have hyp_c1: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c1,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ c1 \ redexes c')} c1 {t. \\\c1,Normal Z\ \ Normal t}, {t. \\\c1,Normal Z\ \ Abrupt t}" using Cond.hyps by iprover have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. s=Z \ \\\Cond b c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ Cond b c1 c2 \ redexes c')} \ b) c1 {t. \\\Cond b c1 c2,Normal Z\ \ Normal t}, {t. \\\Cond b c1 c2,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [OF hyp_c1],safe) assume "Z \ b" "\\\Cond b c1 c2,Normal Z\ \\({Stuck} \ Fault ` (-F))" thus "\\\c1,Normal Z\ \\({Stuck} \ Fault ` (-F))" by (auto simp add: final_notin_def intro: exec.CondTrue) next fix c' assume b: "Z \ b" assume steps_c': "\\ (Call p, Normal \) \\<^sup>+ (c', Normal Z)" assume redex_c': "Cond b c1 c2 \ redexes c'" show "\c'. \\ (Call p, Normal \) \\<^sup>+ (c', Normal Z) \ c1 \ redexes c'" proof - note steps_c' also from b have "\\(Cond b c1 c2, Normal Z) \ (c1, Normal Z)" by (rule step.CondTrue) from step_redexes [OF this redex_c'] obtain c'' where step_c'': "\\ (c', Normal Z) \ (c'', Normal Z)" and c1: "c1 \ redexes c''" by blast note step_c'' finally show ?thesis using c1 by blast qed next fix t assume "Z \ b" "\\\c1,Normal Z\ \ Normal t" thus "\\\Cond b c1 c2,Normal Z\ \ Normal t" by (blast intro: exec.CondTrue) next fix t assume "Z \ b" "\\\c1,Normal Z\ \ Abrupt t" thus "\\\Cond b c1 c2,Normal Z\ \ Abrupt t" by (blast intro: exec.CondTrue) qed moreover have hyp_c2: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ c2 \ redexes c')} c2 {t. \\\c2,Normal Z\ \ Normal t}, {t. \\\c2,Normal Z\ \ Abrupt t}" using Cond.hyps by iprover have "\,\\\<^sub>t\<^bsub>/F\<^esub> ({s. s=Z \ \\\Cond b c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c', Normal s) \ Cond b c1 c2 \ redexes c')} \ -b) c2 {t. \\\Cond b c1 c2,Normal Z\ \ Normal t}, {t. \\\Cond b c1 c2,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [OF hyp_c2],safe) assume "Z \ b" "\\\Cond b c1 c2,Normal Z\ \\({Stuck} \ Fault ` (-F))" thus "\\\c2,Normal Z\ \\({Stuck} \ Fault ` (-F))" by (auto simp add: final_notin_def intro: exec.CondFalse) next fix c' assume b: "Z \ b" assume steps_c': "\\ (Call p, Normal \) \\<^sup>+ (c', Normal Z)" assume redex_c': "Cond b c1 c2 \ redexes c'" show "\c'. \\ (Call p, Normal \) \\<^sup>+ (c', Normal Z) \ c2 \ redexes c'" proof - note steps_c' also from b have "\\(Cond b c1 c2, Normal Z) \ (c2, Normal Z)" by (rule step.CondFalse) from step_redexes [OF this redex_c'] obtain c'' where step_c'': "\\ (c', Normal Z) \ (c'', Normal Z)" and c1: "c2 \ redexes c''" by blast note step_c'' finally show ?thesis using c1 by blast qed next fix t assume "Z \ b" "\\\c2,Normal Z\ \ Normal t" thus "\\\Cond b c1 c2,Normal Z\ \ Normal t" by (blast intro: exec.CondFalse) next fix t assume "Z \ b" "\\\c2,Normal Z\ \ Abrupt t" thus "\\\Cond b c1 c2,Normal Z\ \ Abrupt t" by (blast intro: exec.CondFalse) qed ultimately show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Cond b c1 c2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ Cond b c1 c2 \ redexes c')} (Cond b c1 c2) {t. \\\Cond b c1 c2,Normal Z\ \ Normal t}, {t. \\\Cond b c1 c2,Normal Z\ \ Abrupt t}" by (rule hoaret.Cond) next case (While b c) let ?unroll = "({(s,t). s\b \ \\\c,Normal s\ \ Normal t})\<^sup>*" let ?P' = "\Z. {t. (Z,t)\?unroll \ (\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal t) \ While b c \ redexes c')}" let ?A = "\Z. {t. \\\While b c,Normal Z\ \ Abrupt t}" let ?r = "{(t,s). \\(While b c)\Normal s \ s\b \ \\\c,Normal s\ \ Normal t}" show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\While b c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \)\\<^sup>+(c',Normal s) \ While b c \ redexes c')} (While b c) {t. \\\While b c,Normal Z\ \ Normal t}, {t. \\\While b c,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [where ?P'="\ Z. ?P' Z" and ?Q'="\ Z. ?P' Z \ - b"]) have wf_r: "wf ?r" by (rule wf_terminates_while) show "\ Z. \,\\\<^sub>t\<^bsub>/F\<^esub> (?P' Z) (While b c) (?P' Z \ - b),(?A Z)" proof (rule allI, rule hoaret.While [OF wf_r]) fix Z from While have hyp_c: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ c \ redexes c')} c {t. \\\c,Normal Z\ \ Normal t}, {t. \\\c,Normal Z\ \ Abrupt t}" by iprover show "\\. \,\\\<^sub>t\<^bsub>/F\<^esub> ({\} \ ?P' Z \ b) c ({t. (t, \) \ ?r} \ ?P' Z),(?A Z)" proof (rule allI, rule ConseqMGT [OF hyp_c]) fix \ s assume asm: "s\ {\} \ {t. (Z, t) \ ?unroll \ (\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)) \ \\Call p\ Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal t) \ While b c \ redexes c')} \ b" then obtain c' where s_eq_\: "s=\" and Z_s_unroll: "(Z,s) \ ?unroll" and noabort:"\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)" and termi: "\\Call p \ Normal \" and reach: "\\(Call p,Normal \) \\<^sup>+ (c',Normal s)" and red_c': "While b c \ redexes c'" and s_in_b: "s\b" by blast obtain c'' where reach_c: "\\(Call p,Normal \) \\<^sup>+ (c'',Normal s)" "Seq c (While b c) \ redexes c''" proof - note reach also from s_in_b have "\\(While b c,Normal s) \ (Seq c (While b c),Normal s)" by (rule step.WhileTrue) from step_redexes [OF this red_c'] obtain c'' where step: "\\ (c', Normal s) \ (c'', Normal s)" and red_c'': "Seq c (While b c) \ redexes c''" by blast note step finally show ?thesis using red_c'' by (blast intro: that) qed from reach termi have "\\c' \ Normal s" by (rule steps_preserves_termination') from redexes_preserves_termination [OF this red_c'] have termi_while: "\\While b c \ Normal s" . show "s \ {t. t = s \ \\\c,Normal t\ \\({Stuck} \ Fault ` (-F)) \ \\Call p \ Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal t) \ c \ redexes c')} \ (\t. t \ {t. \\\c,Normal s\ \ Normal t} \ t \ {t. (t,\) \ ?r} \ {t. (Z, t) \ ?unroll \ (\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)) \ \\Call p \ Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal t) \ While b c \ redexes c')}) \ (\t. t \ {t. \\\c,Normal s\ \ Abrupt t} \ t \ {t. \\\While b c,Normal Z\ \ Abrupt t})" (is "?C1 \ ?C2 \ ?C3") proof (intro conjI) from Z_s_unroll noabort s_in_b termi reach_c show ?C1 apply clarsimp apply (drule redexes_subset) apply simp apply (blast intro: root_in_redexes) done next { fix t assume s_t: "\\\c,Normal s\ \ Normal t" with s_eq_\ termi_while s_in_b have "(t,\) \ ?r" by blast moreover from Z_s_unroll s_t s_in_b have "(Z, t) \ ?unroll" by (blast intro: rtrancl_into_rtrancl) moreover obtain c'' where reach_c'': "\\(Call p,Normal \) \\<^sup>+ (c'',Normal t)" "(While b c) \ redexes c''" proof - note reach_c (1) also from s_in_b have "\\(While b c,Normal s)\ (Seq c (While b c),Normal s)" by (rule step.WhileTrue) have "\\ (Seq c (While b c), Normal s) \\<^sup>+ (While b c, Normal t)" proof - from exec_impl_steps_Normal [OF s_t] have "\\ (c, Normal s) \\<^sup>* (Skip, Normal t)". hence "\\ (Seq c (While b c), Normal s) \\<^sup>* (Seq Skip (While b c), Normal t)" by (rule SeqSteps) auto moreover have "\\(Seq Skip (While b c), Normal t)\(While b c, Normal t)" by (rule step.SeqSkip) ultimately show ?thesis by (rule rtranclp_into_tranclp1) qed from steps_redexes' [OF this reach_c (2)] obtain c''' where step: "\\ (c'', Normal s) \\<^sup>+ (c''', Normal t)" and red_c'': "While b c \ redexes c'''" by blast note step finally show ?thesis using red_c'' by (blast intro: that) qed moreover note noabort termi ultimately have "(t,\) \ ?r \ (Z, t) \ ?unroll \ (\e. (Z,e)\?unroll \ e\b \ \\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)) \ \\Call p \ Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c', Normal t) \ While b c \ redexes c')" by iprover } then show ?C2 by blast next { fix t assume s_t: "\\\c,Normal s\ \ Abrupt t" from Z_s_unroll noabort s_t s_in_b have "\\\While b c,Normal Z\ \ Abrupt t" by blast } thus ?C3 by simp qed qed qed next fix s assume P: "s \ {s. s=Z \ \\\While b c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ While b c \ redexes c')}" hence WhileNoFault: "\\\While b c,Normal Z\ \\({Stuck} \ Fault ` (-F))" by auto show "s \ ?P' s \ (\t. t\(?P' s \ - b)\ t\{t. \\\While b c,Normal Z\ \ Normal t})\ (\t. t\?A s \ t\?A Z)" proof (intro conjI) { fix e assume "(Z,e) \ ?unroll" "e \ b" from this WhileNoFault have "\\\c,Normal e\ \\({Stuck} \ Fault ` (-F)) \ (\u. \\\c,Normal e\ \Abrupt u \ \\\While b c,Normal Z\ \ Abrupt u)" (is "?Prop Z e") proof (induct rule: converse_rtrancl_induct [consumes 1]) assume e_in_b: "e \ b" assume WhileNoFault: "\\\While b c,Normal e\ \\({Stuck} \ Fault ` (-F))" with e_in_b WhileNoFault have cNoFault: "\\\c,Normal e\ \\({Stuck} \ Fault ` (-F))" by (auto simp add: final_notin_def intro: exec.intros) moreover { fix u assume "\\\c,Normal e\ \ Abrupt u" with e_in_b have "\\\While b c,Normal e\ \ Abrupt u" by (blast intro: exec.intros) } ultimately show "?Prop e e" by iprover next fix Z r assume e_in_b: "e\b" assume WhileNoFault: "\\\While b c,Normal Z\ \\({Stuck} \ Fault ` (-F))" assume hyp: "\e\b;\\\While b c,Normal r\ \\({Stuck} \ Fault ` (-F))\ \ ?Prop r e" assume Z_r: "(Z, r) \ {(Z, r). Z \ b \ \\\c,Normal Z\ \ Normal r}" with WhileNoFault have "\\\While b c,Normal r\ \\({Stuck} \ Fault ` (-F))" by (auto simp add: final_notin_def intro: exec.intros) from hyp [OF e_in_b this] obtain cNoFault: "\\\c,Normal e\ \\({Stuck} \ Fault ` (-F))" and Abrupt_r: "\u. \\\c,Normal e\ \ Abrupt u \ \\\While b c,Normal r\ \ Abrupt u" by simp { fix u assume "\\\c,Normal e\ \ Abrupt u" with Abrupt_r have "\\\While b c,Normal r\ \ Abrupt u" by simp moreover from Z_r obtain "Z \ b" "\\\c,Normal Z\ \ Normal r" by simp ultimately have "\\\While b c,Normal Z\ \ Abrupt u" by (blast intro: exec.intros) } with cNoFault show "?Prop Z e" by iprover qed } with P show "s \ ?P' s" by blast next { fix t assume "termination": "t \ b" assume "(Z, t) \ ?unroll" hence "\\\While b c,Normal Z\ \ Normal t" proof (induct rule: converse_rtrancl_induct [consumes 1]) from "termination" show "\\\While b c,Normal t\ \ Normal t" by (blast intro: exec.WhileFalse) next fix Z r assume first_body: "(Z, r) \ {(s, t). s \ b \ \\\c,Normal s\ \ Normal t}" assume "(r, t) \ ?unroll" assume rest_loop: "\\\While b c, Normal r\ \ Normal t" show "\\\While b c,Normal Z\ \ Normal t" proof - from first_body obtain "Z \ b" "\\\c,Normal Z\ \ Normal r" by fast moreover from rest_loop have "\\\While b c,Normal r\ \ Normal t" by fast ultimately show "\\\While b c,Normal Z\ \ Normal t" by (rule exec.WhileTrue) qed qed } with P show "\t. t\(?P' s \ - b) \t\{t. \\\While b c,Normal Z\ \ Normal t}" by blast next from P show "\t. t\?A s \ t\?A Z" by simp qed qed next case (Call q) let ?P = "{s. s=Z \ \\\Call q ,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ Call q \ redexes c')}" from noStuck_Call have "\s \ ?P. q \ dom \" by (fastforce simp add: final_notin_def) then show ?case proof (rule conseq_extract_state_indep_prop) assume q_defined: "q \ dom \" from Call_hyp have "\q\dom \. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub>{s. s=Z \ \\\Call q,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call q\Normal s \ ((s,q),(\,p)) \ termi_call_steps \} (Call q) {t. \\\Call q,Normal Z\ \ Normal t}, {t. \\\Call q,Normal Z\ \ Abrupt t}" by (simp add: exec_Call_body' noFaultStuck_Call_body' [simplified] terminates_Normal_Call_body) from Call_hyp q_defined have Call_hyp': "\Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Call q,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call q\Normal s \ ((s,q),(\,p)) \ termi_call_steps \} (Call q) {t. \\\Call q,Normal Z\ \ Normal t}, {t. \\\Call q,Normal Z\ \ Abrupt t}" by auto show "\,\\\<^sub>t\<^bsub>/F\<^esub> ?P (Call q) {t. \\\Call q ,Normal Z\ \ Normal t}, {t. \\\Call q ,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [OF Call_hyp'],safe) fix c' assume termi: "\\Call p \ Normal \" assume steps_c': "\\ (Call p, Normal \) \\<^sup>+ (c', Normal Z)" assume red_c': "Call q \ redexes c'" show "\\Call q \ Normal Z" proof - from steps_preserves_termination' [OF steps_c' termi] have "\\c' \ Normal Z" . from redexes_preserves_termination [OF this red_c'] show ?thesis . qed next fix c' assume termi: "\\Call p \ Normal \" assume steps_c': "\\ (Call p, Normal \) \\<^sup>+ (c', Normal Z)" assume red_c': "Call q \ redexes c'" from redex_redexes [OF this] have "redex c' = Call q" by auto with termi steps_c' show "((Z, q), \, p) \ termi_call_steps \" by (auto simp add: termi_call_steps_def) qed qed next case (DynCom c) have hyp: "\s'. \Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s = Z \ \\\c s',Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p \ Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ c s' \ redexes c')} (c s') {t. \\\c s',Normal Z\ \ Normal t},{t. \\\c s',Normal Z\ \ Abrupt t}" using DynCom by simp have hyp': "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\DynCom c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p \ Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ DynCom c \ redexes c')} (c Z) {t. \\\DynCom c,Normal Z\ \ Normal t},{t. \\\DynCom c,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [OF hyp],safe) assume "\\\DynCom c,Normal Z\ \\({Stuck} \ Fault ` (-F))" then show "\\\c Z,Normal Z\ \\({Stuck} \ Fault ` (-F))" by (fastforce simp add: final_notin_def intro: exec.intros) next fix c' assume steps: "\\ (Call p, Normal \) \\<^sup>+ (c', Normal Z)" assume c': "DynCom c \ redexes c'" have "\\ (DynCom c, Normal Z) \ (c Z,Normal Z)" by (rule step.DynCom) from step_redexes [OF this c'] obtain c'' where step: "\\ (c', Normal Z) \ (c'', Normal Z)" and c'': "c Z \ redexes c''" by blast note steps also note step finally show "\c'. \\ (Call p, Normal \) \\<^sup>+ (c', Normal Z) \ c Z \ redexes c'" using c'' by blast next fix t assume "\\\c Z,Normal Z\ \ Normal t" thus "\\\DynCom c,Normal Z\ \ Normal t" by (auto intro: exec.intros) next fix t assume "\\\c Z,Normal Z\ \ Abrupt t" thus "\\\DynCom c,Normal Z\ \ Abrupt t" by (auto intro: exec.intros) qed show ?case apply (rule hoaret.DynCom) apply safe apply (rule hyp') done next case (Guard f g c) have hyp_c: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\c,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ c \ redexes c')} c {t. \\\c,Normal Z\ \ Normal t}, {t. \\\c,Normal Z\ \ Abrupt t}" using Guard.hyps by iprover show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Guard f g c ,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ Guard f g c \ redexes c')} Guard f g c {t. \\\Guard f g c ,Normal Z\ \ Normal t}, {t. \\\Guard f g c ,Normal Z\ \ Abrupt t}" proof (cases "f \ F") case True have "\,\\\<^sub>t\<^bsub>/F\<^esub> (g \ {s. s=Z \ \\\Guard f g c ,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ Guard f g c \ redexes c')}) c {t. \\\Guard f g c ,Normal Z\ \ Normal t}, {t. \\\Guard f g c ,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [OF hyp_c], safe) assume "\\\Guard f g c ,Normal Z\ \\({Stuck} \ Fault ` (-F))" "Z\g" thus "\\\c,Normal Z\ \\({Stuck} \ Fault ` (-F))" by (auto simp add: final_notin_def intro: exec.intros) next fix c' assume steps: "\\ (Call p, Normal \) \\<^sup>+ (c', Normal Z)" assume c': "Guard f g c \ redexes c'" assume "Z \ g" from this have "\\(Guard f g c,Normal Z) \ (c,Normal Z)" by (rule step.Guard) from step_redexes [OF this c'] obtain c'' where step: "\\ (c', Normal Z) \ (c'', Normal Z)" and c'': "c \ redexes c''" by blast note steps also note step finally show "\c'. \\ (Call p, Normal \) \\<^sup>+ (c', Normal Z) \ c \ redexes c'" using c'' by blast next fix t assume "\\\Guard f g c ,Normal Z\ \\({Stuck} \ Fault ` (-F))" "\\\c,Normal Z\ \ Normal t" "Z \ g" thus "\\\Guard f g c ,Normal Z\ \ Normal t" by (auto simp add: final_notin_def intro: exec.intros ) next fix t assume "\\\Guard f g c ,Normal Z\ \\({Stuck} \ Fault ` (-F))" "\\\c,Normal Z\ \ Abrupt t" "Z \ g" thus "\\\Guard f g c ,Normal Z\ \ Abrupt t" by (auto simp add: final_notin_def intro: exec.intros ) qed from True this show ?thesis by (rule conseqPre [OF Guarantee]) auto next case False have "\,\\\<^sub>t\<^bsub>/F\<^esub> (g \ {s. s=Z \ \\\Guard f g c ,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ Guard f g c \ redexes c')}) c {t. \\\Guard f g c ,Normal Z\ \ Normal t}, {t. \\\Guard f g c ,Normal Z\ \ Abrupt t}" proof (rule ConseqMGT [OF hyp_c], safe) assume "\\\Guard f g c ,Normal Z\ \\({Stuck} \ Fault ` (-F))" thus "\\\c,Normal Z\ \\({Stuck} \ Fault ` (-F))" using False by (cases "Z\ g") (auto simp add: final_notin_def intro: exec.intros) next fix c' assume steps: "\\ (Call p, Normal \) \\<^sup>+ (c', Normal Z)" assume c': "Guard f g c \ redexes c'" assume "Z \ g" from this have "\\(Guard f g c,Normal Z) \ (c,Normal Z)" by (rule step.Guard) from step_redexes [OF this c'] obtain c'' where step: "\\ (c', Normal Z) \ (c'', Normal Z)" and c'': "c \ redexes c''" by blast note steps also note step finally show "\c'. \\ (Call p, Normal \) \\<^sup>+ (c', Normal Z) \ c \ redexes c'" using c'' by blast next fix t assume "\\\Guard f g c ,Normal Z\ \\({Stuck} \ Fault ` (-F))" "\\\c,Normal Z\ \ Normal t" thus "\\\Guard f g c ,Normal Z\ \ Normal t" using False by (cases "Z\ g") (auto simp add: final_notin_def intro: exec.intros ) next fix t assume "\\\Guard f g c ,Normal Z\ \\({Stuck} \ Fault ` (-F))" "\\\c,Normal Z\ \ Abrupt t" thus "\\\Guard f g c ,Normal Z\ \ Abrupt t" using False by (cases "Z\ g") (auto simp add: final_notin_def intro: exec.intros ) qed then show ?thesis apply (rule conseqPre [OF hoaret.Guard]) apply clarify apply (frule Guard_noFaultStuckD [OF _ False]) apply auto done qed next case Throw show "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Throw,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal \ \ (\c'. \\(Call p, Normal \) \\<^sup>+ (c',Normal s) \ Throw \ redexes c')} Throw {t. \\\Throw,Normal Z\ \ Normal t}, {t. \\\Throw,Normal Z\ \ Abrupt t}" by (rule conseqPre [OF hoaret.Throw]) (blast intro: exec.intros terminates.intros) next case (Catch c\<^sub>1 c\<^sub>2) have hyp_c1: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s= Z \ \\\c\<^sub>1,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p \ Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ c\<^sub>1 \ redexes c')} c\<^sub>1 {t. \\\c\<^sub>1,Normal Z\ \ Normal t},{t. \\\c\<^sub>1,Normal Z\ \ Abrupt t}" using Catch.hyps by iprover have hyp_c2: "\Z. \,\\\<^sub>t\<^bsub>/F\<^esub> {s. s= Z \ \\\c\<^sub>2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\ Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal s) \ c\<^sub>2 \ redexes c')} c\<^sub>2 {t. \\\c\<^sub>2,Normal Z\ \ Normal t},{t. \\\c\<^sub>2,Normal Z\ \ Abrupt t}" using Catch.hyps by iprover have "\,\\\<^sub>t\<^bsub>/F\<^esub> {s. s = Z \ \\\Catch c\<^sub>1 c\<^sub>2,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\ Normal \ \ (\c'. \\(Call p,Normal \)\\<^sup>+(c',Normal s) \ Catch c\<^sub>1 c\<^sub>2 \ redexes c')} c\<^sub>1 {t. \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Normal t}, {t. \\\c\<^sub>1,Normal Z\ \ Abrupt t \ \\\c\<^sub>2,Normal t\ \\({Stuck} \ Fault`(-F)) \ \\Call p \ Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal t) \ c\<^sub>2 \ redexes c')}" proof (rule ConseqMGT [OF hyp_c1],clarify,safe) assume "\\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \\({Stuck} \ Fault ` (-F))" thus "\\\c\<^sub>1,Normal Z\ \\({Stuck} \ Fault ` (-F))" by (fastforce simp add: final_notin_def intro: exec.intros) next fix c' assume steps: "\\ (Call p, Normal \) \\<^sup>+ (c', Normal Z)" assume c': "Catch c\<^sub>1 c\<^sub>2 \ redexes c'" from steps redexes_subset [OF this] show "\c'. \\ (Call p, Normal \) \\<^sup>+ (c', Normal Z) \ c\<^sub>1 \ redexes c'" by (auto iff: root_in_redexes) next fix t assume "\\\c\<^sub>1,Normal Z\ \ Normal t" thus "\\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Normal t" by (auto intro: exec.intros) next fix t assume "\\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \\({Stuck} \ Fault ` (-F))" "\\\c\<^sub>1,Normal Z\ \ Abrupt t" thus "\\\c\<^sub>2,Normal t\ \\({Stuck} \ Fault ` (-F))" by (auto simp add: final_notin_def intro: exec.intros) next fix c' t assume steps_c': "\\ (Call p, Normal \) \\<^sup>+ (c', Normal Z)" assume red: "Catch c\<^sub>1 c\<^sub>2 \ redexes c'" assume exec_c\<^sub>1: "\\ \c\<^sub>1,Normal Z\ \ Abrupt t" show "\c'. \\ (Call p, Normal \) \\<^sup>+ (c', Normal t) \ c\<^sub>2 \ redexes c'" proof - note steps_c' also from exec_impl_steps_Normal_Abrupt [OF exec_c\<^sub>1] have "\\ (c\<^sub>1, Normal Z) \\<^sup>* (Throw, Normal t)". from steps_redexes_Catch [OF this red] obtain c'' where steps_c'': "\\ (c', Normal Z) \\<^sup>* (c'', Normal t)" and Catch: "Catch Throw c\<^sub>2 \ redexes c''" by blast note steps_c'' also have step_Catch: "\\ (Catch Throw c\<^sub>2,Normal t) \ (c\<^sub>2,Normal t)" by (rule step.CatchThrow) from step_redexes [OF step_Catch Catch] obtain c''' where step_c''': "\\ (c'', Normal t) \ (c''', Normal t)" and c2: "c\<^sub>2 \ redexes c'''" by blast note step_c''' finally show ?thesis using c2 by blast qed qed moreover have "\,\\\<^sub>t\<^bsub>/F\<^esub> {t. \\\c\<^sub>1,Normal Z\ \ Abrupt t \ \\\c\<^sub>2,Normal t\ \\({Stuck} \ Fault ` (-F)) \ \\Call p \ Normal \ \ (\c'. \\(Call p,Normal \) \\<^sup>+ (c',Normal t) \ c\<^sub>2 \ redexes c')} c\<^sub>2 {t. \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Normal t}, {t. \\\Catch c\<^sub>1 c\<^sub>2,Normal Z\ \ Abrupt t}" by (rule ConseqMGT [OF hyp_c2]) (fastforce intro: exec.intros) ultimately show ?case by (rule hoaret.Catch) qed text \To prove a procedure implementation correct it suffices to assume only the procedure specifications of procedures that actually occur during evaluation of the body. \ lemma Call_lemma: assumes A: "\q \ dom \. \Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Call q,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call q\Normal s \ ((s,q),(\,p)) \ termi_call_steps \} (Call q) {t. \\\Call q,Normal Z\ \ Normal t}, {t. \\\Call q,Normal Z\ \ Abrupt t}" assumes pdef: "p \ dom \" shows "\Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> ({\} \ {s. s=Z \\\\the (\ p),Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal s}) the (\ p) {t. \\\the (\ p),Normal Z\ \ Normal t}, {t. \\\the (\ p),Normal Z\ \ Abrupt t}" apply (rule conseqPre) apply (rule Call_lemma' [OF A]) using pdef apply (fastforce intro: terminates.intros tranclp.r_into_trancl [of "(step \)", OF step.Call] root_in_redexes) done lemma Call_lemma_switch_Call_body: assumes call: "\q \ dom \. \Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Call q,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call q\Normal s \ ((s,q),(\,p)) \ termi_call_steps \} (Call q) {t. \\\Call q,Normal Z\ \ Normal t}, {t. \\\Call q,Normal Z\ \ Abrupt t}" assumes p_defined: "p \ dom \" shows "\Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> ({\} \ {s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal s}) the (\ p) {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t}" apply (simp only: exec_Call_body' [OF p_defined] noFaultStuck_Call_body' [OF p_defined] terminates_Normal_Call_body [OF p_defined]) apply (rule conseqPre) apply (rule Call_lemma') apply (rule call) using p_defined apply (fastforce intro: terminates.intros tranclp.r_into_trancl [of "(step \)", OF step.Call] root_in_redexes) done lemma MGT_Call: "\p \ dom \. \Z. \,\ \\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\(Call p)\Normal s} (Call p) {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t}" apply (intro ballI allI) apply (rule CallRec' [where Procs="dom \" and P="\p Z. {s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal s}" and Q="\p Z. {t. \\\Call p,Normal Z\ \ Normal t}" and A="\p Z. {t. \\\Call p,Normal Z\ \ Abrupt t}" and r="termi_call_steps \" ]) apply simp apply simp apply (rule wf_termi_call_steps) apply (intro ballI allI) apply simp apply (rule Call_lemma_switch_Call_body [rule_format, simplified]) apply (rule hoaret.Asm) apply fastforce apply assumption done lemma CollInt_iff: "{s. P s} \ {s. Q s} = {s. P s \ Q s}" by auto lemma image_Un_conv: "f ` (\p\dom \. \Z. {x p Z}) = (\p\dom \. \Z. {f (x p Z)})" by (auto iff: not_None_eq) text \Another proof of \MGT_Call\, maybe a little more readable\ lemma "\p \ dom \. \Z. \,{} \\<^sub>t\<^bsub>/F\<^esub> {s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\(Call p)\Normal s} (Call p) {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t}" proof - { fix p Z \ assume defined: "p \ dom \" define Specs where "Specs = (\p\dom \. \Z. {({s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal s}, p, {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t})})" define Specs_wf where "Specs_wf p \ = (\(P,q,Q,A). (P \ {s. ((s,q),\,p) \ termi_call_steps \}, q, Q, A)) ` Specs" for p \ have "\,Specs_wf p \ \\<^sub>t\<^bsub>/F\<^esub>({\} \ {s. s = Z \ \\\the (\ p),Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\the (\ p)\Normal s}) (the (\ p)) {t. \\\the (\ p),Normal Z\ \ Normal t}, {t. \\\the (\ p),Normal Z\ \ Abrupt t}" apply (rule Call_lemma [rule_format, OF _ defined]) apply (rule hoaret.Asm) apply (clarsimp simp add: Specs_wf_def Specs_def image_Un_conv) apply (rule_tac x=q in bexI) apply (rule_tac x=Z in exI) apply (clarsimp simp add: CollInt_iff) apply auto done hence "\,Specs_wf p \ \\<^sub>t\<^bsub>/F\<^esub>({\} \ {s. s = Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal s}) (the (\ p)) {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t}" by (simp only: exec_Call_body' [OF defined] noFaultStuck_Call_body' [OF defined] terminates_Normal_Call_body [OF defined]) } note bdy=this show ?thesis apply (intro ballI allI) apply (rule hoaret.CallRec [where Specs="(\p\dom \. \Z. {({s. s=Z \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) \ \\Call p\Normal s}, p, {t. \\\Call p,Normal Z\ \ Normal t}, {t. \\\Call p,Normal Z\ \ Abrupt t})})", OF _ wf_termi_call_steps [of \] refl]) apply fastforce apply clarify apply (rule conjI) apply fastforce apply (rule allI) apply (simp (no_asm_use) only : Un_empty_left) apply (rule bdy) apply auto done qed theorem hoaret_complete: "\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A \ \,{}\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (iprover intro: MGT_implies_complete MGT_lemma [OF MGT_Call]) lemma hoaret_complete': assumes cvalid: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" proof (cases "\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A") case True hence "\,{}\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (rule hoaret_complete) thus "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (rule hoaret_augment_context) simp next case False with cvalid show ?thesis by (rule ExFalso) qed subsection \And Now: Some Useful Rules\ subsubsection \Modify Return\ -lemma ProcModifyReturn_sound: - assumes valid_call: "\,\ \\<^sub>t\<^bsub>/F\<^esub> P call init p return' c Q,A" +lemma Proc_exnModifyReturn_sound: + assumes valid_call: "\,\ \\<^sub>t\<^bsub>/F\<^esub> P call_exn init p return' result_exn c Q,A" assumes valid_modif: "\\. \,\ \\<^bsub>/UNIV\<^esub> {\} (Call p) (Modif \),(ModifAbr \)" assumes res_modif: "\s t. t \ Modif (init s) \ return' s t = return s t" assumes ret_modifAbr: - "\s t. t \ ModifAbr (init s) \ return' s t = return s t" - shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (call init p return c) Q,A" + "\s t. t \ ModifAbr (init s) \ result_exn (return' s t) t = result_exn (return s t) t" + shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" hence "\(P, p, Q, A)\\. \ \\<^bsub>/F\<^esub> P (Call p) Q,A" by (auto simp add: validt_def) then have ctxt': "\(P, p, Q, A)\\. \ \\<^bsub>/UNIV\<^esub> P (Call p) Q,A" by (auto intro: valid_augment_Faults) - assume exec: "\\\call init p return c,Normal s\ \ t" + assume exec: "\\\call_exn init p return result_exn c,Normal s\ \ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" from exec show "t \ Normal ` Q \ Abrupt ` A" - proof (cases rule: exec_call_Normal_elim) + proof (cases rule: exec_call_exn_Normal_elim) fix bdy t' assume bdy: "\ p = Some bdy" assume exec_body: "\\\bdy,Normal (init s)\ \ Normal t'" assume exec_c: "\\\c s t',Normal (return s t')\ \ t" from exec_body bdy have "\\\(Call p ),Normal (init s)\ \ Normal t'" by (auto simp add: intro: exec.intros) from cvalidD [OF valid_modif [rule_format, of "init s"] ctxt' this] P have "t' \ Modif (init s)" by auto with res_modif have "Normal (return' s t') = Normal (return s t')" by simp with exec_body exec_c bdy - have "\\\call init p return' c,Normal s\ \ t" - by (auto intro: exec_call) + have "\\\call_exn init p return' result_exn c,Normal s\ \ t" + by (auto intro: exec_call_exn) from cvalidt_postD [OF valid_call ctxt this] P t_notin_F show ?thesis by simp next fix bdy t' assume bdy: "\ p = Some bdy" assume exec_body: "\\\bdy,Normal (init s)\ \ Abrupt t'" - assume t: "t = Abrupt (return s t')" + assume t: "t = Abrupt (result_exn (return s t') t')" also from exec_body bdy have "\\\(Call p),Normal (init s)\ \ Abrupt t'" by (auto simp add: intro: exec.intros) from cvalidD [OF valid_modif [rule_format, of "init s"] ctxt' this] P have "t' \ ModifAbr (init s)" by auto - with ret_modifAbr have "Abrupt (return s t') = Abrupt (return' s t')" + with ret_modifAbr have "Abrupt (result_exn (return s t') t') = Abrupt (result_exn (return' s t') t')" by simp - finally have "t = Abrupt (return' s t')" . + finally have "t = Abrupt (result_exn (return' s t') t')" . with exec_body bdy - have "\\\call init p return' c,Normal s\ \ t" - by (auto intro: exec_callAbrupt) + have "\\\call_exn init p return' result_exn c,Normal s\ \ t" + by (auto intro: exec_call_exnAbrupt) from cvalidt_postD [OF valid_call ctxt this] P t_notin_F show ?thesis by simp next fix bdy f assume bdy: "\ p = Some bdy" assume "\\\bdy,Normal (init s)\ \ Fault f" and t: "t = Fault f" - with bdy have "\\\call init p return' c ,Normal s\ \ t" - by (auto intro: exec_callFault) + with bdy have "\\\call_exn init p return' result_exn c ,Normal s\ \ t" + by (auto intro: exec_call_exnFault) from cvalidt_postD [OF valid_call ctxt this P] t t_notin_F show ?thesis by simp next fix bdy assume bdy: "\ p = Some bdy" assume "\\\bdy,Normal (init s)\ \ Stuck" "t = Stuck" - with bdy have "\\\call init p return' c ,Normal s\ \ t" - by (auto intro: exec_callStuck) + with bdy have "\\\call_exn init p return' result_exn c ,Normal s\ \ t" + by (auto intro: exec_call_exnStuck) from valid_call ctxt this P t_notin_F show ?thesis by (rule cvalidt_postD) next assume "\ p = None" "t=Stuck" - hence "\\\call init p return' c ,Normal s\ \ t" - by (auto intro: exec_callUndefined) + hence "\\\call_exn init p return' result_exn c ,Normal s\ \ t" + by (auto intro: exec_call_exnUndefined) from valid_call ctxt this P t_notin_F show ?thesis by (rule cvalidt_postD) qed next fix s assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" hence "\(P, p, Q, A)\\. \ \\<^bsub>/F\<^esub> P (Call p) Q,A" by (auto simp add: validt_def) then have ctxt': "\(P, p, Q, A)\\. \ \\<^bsub>/UNIV\<^esub> P (Call p) Q,A" by (auto intro: valid_augment_Faults) assume P: "s \ P" from valid_call ctxt P - have call: "\\call init p return' c\ Normal s" + have call: "\\call_exn init p return' result_exn c\ Normal s" by (rule cvalidt_termD) - show "\\call init p return c \ Normal s" - proof (cases "p \ dom \") - case True - with call obtain bdy where - bdy: "\ p = Some bdy" and termi_bdy: "\\bdy \ Normal (init s)" and - termi_c: "\t. \\\bdy,Normal (init s)\ \ Normal t \ - \\c s t \ Normal (return' s t)" - by cases auto - { - fix t - assume exec_bdy: "\\\bdy,Normal (init s)\ \ Normal t" - hence "\\c s t \ Normal (return s t)" - proof - - from exec_bdy bdy - have "\\\(Call p ),Normal (init s)\ \ Normal t" - by (auto simp add: intro: exec.intros) - from cvalidD [OF valid_modif [rule_format, of "init s"] ctxt' this] P - res_modif - have "return' s t = return s t" - by auto - with termi_c exec_bdy show ?thesis by auto - qed - } - with bdy termi_bdy - show ?thesis - by (iprover intro: terminates_call) - next - case False - thus ?thesis - by (auto intro: terminates_callUndefined) - qed -qed - -lemma ProcModifyReturn: - assumes spec: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return' c) Q,A" - assumes res_modif: - "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" - assumes ret_modifAbr: - "\s t. t \ ModifAbr (init s) \ (return' s t) = (return s t)" - assumes modifies_spec: - "\\. \,\\\<^bsub>/UNIV\<^esub> {\} (Call p) (Modif \),(ModifAbr \)" - shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return c) Q,A" -apply (rule hoaret_complete') -apply (rule ProcModifyReturn_sound [where Modif=Modif and ModifAbr=ModifAbr, - OF _ _ res_modif ret_modifAbr]) -apply (rule hoaret_sound [OF spec]) -using modifies_spec -apply (blast intro: hoare_sound) -done - -lemma ProcModifyReturnSameFaults_sound: - assumes valid_call: "\,\ \\<^sub>t\<^bsub>/F\<^esub> P call init p return' c Q,A" - assumes valid_modif: - "\\. \,\ \\<^bsub>/F\<^esub> {\} Call p (Modif \),(ModifAbr \)" - assumes res_modif: - "\s t. t \ Modif (init s) \ return' s t = return s t" - assumes ret_modifAbr: - "\s t. t \ ModifAbr (init s) \ return' s t = return s t" - shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (call init p return c) Q,A" -proof (rule cvalidtI) - fix s t - assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" - hence ctxt': "\(P, p, Q, A)\\. \ \\<^bsub>/F\<^esub> P (Call p) Q,A" - by (auto simp add: validt_def) - assume exec: "\\\call init p return c,Normal s\ \ t" - assume P: "s \ P" - assume t_notin_F: "t \ Fault ` F" - from exec - show "t \ Normal ` Q \ Abrupt ` A" - proof (cases rule: exec_call_Normal_elim) - fix bdy t' - assume bdy: "\ p = Some bdy" - assume exec_body: "\\\bdy,Normal (init s)\ \ Normal t'" - assume exec_c: "\\\c s t',Normal (return s t')\ \ t" - from exec_body bdy - have "\\\(Call p) ,Normal (init s)\ \ Normal t'" - by (auto simp add: intro: exec.intros) - from cvalidD [OF valid_modif [rule_format, of "init s"] ctxt' this] P - have "t' \ Modif (init s)" - by auto - with res_modif have "Normal (return' s t') = Normal (return s t')" - by simp - with exec_body exec_c bdy - have "\\\call init p return' c,Normal s\ \ t" - by (auto intro: exec_call) - from cvalidt_postD [OF valid_call ctxt this] P t_notin_F - show ?thesis - by simp - next - fix bdy t' - assume bdy: "\ p = Some bdy" - assume exec_body: "\\\bdy,Normal (init s)\ \ Abrupt t'" - assume t: "t = Abrupt (return s t')" - also - from exec_body bdy - have "\\\Call p ,Normal (init s)\ \ Abrupt t'" - by (auto simp add: intro: exec.intros) - from cvalidD [OF valid_modif [rule_format, of "init s"] ctxt' this] P - have "t' \ ModifAbr (init s)" - by auto - with ret_modifAbr have "Abrupt (return s t') = Abrupt (return' s t')" - by simp - finally have "t = Abrupt (return' s t')" . - with exec_body bdy - have "\\\call init p return' c,Normal s\ \ t" - by (auto intro: exec_callAbrupt) - from cvalidt_postD [OF valid_call ctxt this] P t_notin_F - show ?thesis - by simp - next - fix bdy f - assume bdy: "\ p = Some bdy" - assume "\\\bdy,Normal (init s)\ \ Fault f" and - t: "t = Fault f" - with bdy have "\\\call init p return' c ,Normal s\ \ t" - by (auto intro: exec_callFault) - from cvalidt_postD [OF valid_call ctxt this P] t t_notin_F - show ?thesis - by simp - next - fix bdy - assume bdy: "\ p = Some bdy" - assume "\\\bdy,Normal (init s)\ \ Stuck" - "t = Stuck" - with bdy have "\\\call init p return' c,Normal s\ \ t" - by (auto intro: exec_callStuck) - from valid_call ctxt this P t_notin_F - show ?thesis - by (rule cvalidt_postD) - next - assume "\ p = None" "t=Stuck" - hence "\\\call init p return' c,Normal s\ \ t" - by (auto intro: exec_callUndefined) - from valid_call ctxt this P t_notin_F - show ?thesis - by (rule cvalidt_postD) - qed -next - fix s - assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" - hence ctxt': "\(P, p, Q, A)\\. \ \\<^bsub>/F\<^esub> P (Call p) Q,A" - by (auto simp add: validt_def) - assume P: "s \ P" - from valid_call ctxt P - have call: "\\call init p return' c\ Normal s" - by (rule cvalidt_termD) - show "\\call init p return c \ Normal s" + show "\\call_exn init p return result_exn c \ Normal s" proof (cases "p \ dom \") case True with call obtain bdy where bdy: "\ p = Some bdy" and termi_bdy: "\\bdy \ Normal (init s)" and termi_c: "\t. \\\bdy,Normal (init s)\ \ Normal t \ \\c s t \ Normal (return' s t)" by cases auto { fix t assume exec_bdy: "\\\bdy,Normal (init s)\ \ Normal t" hence "\\c s t \ Normal (return s t)" proof - from exec_bdy bdy have "\\\(Call p ),Normal (init s)\ \ Normal t" by (auto simp add: intro: exec.intros) from cvalidD [OF valid_modif [rule_format, of "init s"] ctxt' this] P res_modif have "return' s t = return s t" by auto with termi_c exec_bdy show ?thesis by auto qed } with bdy termi_bdy show ?thesis - by (iprover intro: terminates_call) + by (iprover intro: terminates_call_exn) next case False thus ?thesis - by (auto intro: terminates_callUndefined) + by (auto intro: terminates_call_exnUndefined) qed qed +lemma ProcModifyReturn_sound: + assumes valid_call: "\,\ \\<^sub>t\<^bsub>/F\<^esub> P call init p return' c Q,A" + assumes valid_modif: + "\\. \,\ \\<^bsub>/UNIV\<^esub> {\} (Call p) (Modif \),(ModifAbr \)" + assumes res_modif: + "\s t. t \ Modif (init s) \ return' s t = return s t" + assumes ret_modifAbr: + "\s t. t \ ModifAbr (init s) \ return' s t = return s t" +shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (call init p return c) Q,A" + using valid_call valid_modif res_modif ret_modifAbr unfolding call_call_exn + by (rule Proc_exnModifyReturn_sound) + + +lemma Proc_exnModifyReturn: + assumes spec: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call_exn init p return' result_exn c) Q,A" + assumes res_modif: + "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" + assumes ret_modifAbr: + "\s t. t \ ModifAbr (init s) \ (result_exn (return' s t) t) = (result_exn (return s t) t)" + assumes modifies_spec: + "\\. \,\\\<^bsub>/UNIV\<^esub> {\} (Call p) (Modif \),(ModifAbr \)" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" +apply (rule hoaret_complete') +apply (rule Proc_exnModifyReturn_sound [where Modif=Modif and ModifAbr=ModifAbr, + OF _ _ res_modif ret_modifAbr]) +apply (rule hoaret_sound [OF spec]) +using modifies_spec +apply (blast intro: hoare_sound) +done + +lemma ProcModifyReturn: + assumes spec: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return' c) Q,A" + assumes res_modif: + "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" + assumes ret_modifAbr: + "\s t. t \ ModifAbr (init s) \ (return' s t) = (return s t)" + assumes modifies_spec: + "\\. \,\\\<^bsub>/UNIV\<^esub> {\} (Call p) (Modif \),(ModifAbr \)" +shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return c) Q,A" + using spec res_modif ret_modifAbr modifies_spec unfolding call_call_exn by (rule Proc_exnModifyReturn) + +lemma Proc_exnModifyReturnSameFaults_sound: + assumes valid_call: "\,\ \\<^sub>t\<^bsub>/F\<^esub> P call_exn init p return' result_exn c Q,A" + assumes valid_modif: + "\\. \,\ \\<^bsub>/F\<^esub> {\} Call p (Modif \),(ModifAbr \)" + assumes res_modif: + "\s t. t \ Modif (init s) \ return' s t = return s t" + assumes ret_modifAbr: + "\s t. t \ ModifAbr (init s) \ result_exn (return' s t) t = result_exn (return s t) t" + shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" +proof (rule cvalidtI) + fix s t + assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" + hence ctxt': "\(P, p, Q, A)\\. \ \\<^bsub>/F\<^esub> P (Call p) Q,A" + by (auto simp add: validt_def) + assume exec: "\\\call_exn init p return result_exn c,Normal s\ \ t" + assume P: "s \ P" + assume t_notin_F: "t \ Fault ` F" + from exec + show "t \ Normal ` Q \ Abrupt ` A" + proof (cases rule: exec_call_exn_Normal_elim) + fix bdy t' + assume bdy: "\ p = Some bdy" + assume exec_body: "\\\bdy,Normal (init s)\ \ Normal t'" + assume exec_c: "\\\c s t',Normal (return s t')\ \ t" + from exec_body bdy + have "\\\(Call p) ,Normal (init s)\ \ Normal t'" + by (auto simp add: intro: exec.intros) + from cvalidD [OF valid_modif [rule_format, of "init s"] ctxt' this] P + have "t' \ Modif (init s)" + by auto + with res_modif have "Normal (return' s t') = Normal (return s t')" + by simp + with exec_body exec_c bdy + have "\\\call_exn init p return' result_exn c,Normal s\ \ t" + by (auto intro: exec_call_exn) + from cvalidt_postD [OF valid_call ctxt this] P t_notin_F + show ?thesis + by simp + next + fix bdy t' + assume bdy: "\ p = Some bdy" + assume exec_body: "\\\bdy,Normal (init s)\ \ Abrupt t'" + assume t: "t = Abrupt (result_exn (return s t') t')" + also + from exec_body bdy + have "\\\Call p ,Normal (init s)\ \ Abrupt t'" + by (auto simp add: intro: exec.intros) + from cvalidD [OF valid_modif [rule_format, of "init s"] ctxt' this] P + have "t' \ ModifAbr (init s)" + by auto + with ret_modifAbr have "Abrupt (result_exn (return s t') t') = Abrupt (result_exn (return' s t') t')" + by simp + finally have "t = Abrupt (result_exn (return' s t') t')" . + with exec_body bdy + have "\\\call_exn init p return' result_exn c,Normal s\ \ t" + by (auto intro: exec_call_exnAbrupt) + from cvalidt_postD [OF valid_call ctxt this] P t_notin_F + show ?thesis + by simp + next + fix bdy f + assume bdy: "\ p = Some bdy" + assume "\\\bdy,Normal (init s)\ \ Fault f" and + t: "t = Fault f" + with bdy have "\\\call_exn init p return' result_exn c ,Normal s\ \ t" + by (auto intro: exec_call_exnFault) + from cvalidt_postD [OF valid_call ctxt this P] t t_notin_F + show ?thesis + by simp + next + fix bdy + assume bdy: "\ p = Some bdy" + assume "\\\bdy,Normal (init s)\ \ Stuck" + "t = Stuck" + with bdy have "\\\call_exn init p return' result_exn c,Normal s\ \ t" + by (auto intro: exec_call_exnStuck) + from valid_call ctxt this P t_notin_F + show ?thesis + by (rule cvalidt_postD) + next + assume "\ p = None" "t=Stuck" + hence "\\\call_exn init p return' result_exn c,Normal s\ \ t" + by (auto intro: exec_call_exnUndefined) + from valid_call ctxt this P t_notin_F + show ?thesis + by (rule cvalidt_postD) + qed +next + fix s + assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" + hence ctxt': "\(P, p, Q, A)\\. \ \\<^bsub>/F\<^esub> P (Call p) Q,A" + by (auto simp add: validt_def) + assume P: "s \ P" + from valid_call ctxt P + have call: "\\call_exn init p return' result_exn c\ Normal s" + by (rule cvalidt_termD) + show "\\call_exn init p return result_exn c \ Normal s" + proof (cases "p \ dom \") + case True + with call obtain bdy where + bdy: "\ p = Some bdy" and termi_bdy: "\\bdy \ Normal (init s)" and + termi_c: "\t. \\\bdy,Normal (init s)\ \ Normal t \ + \\c s t \ Normal (return' s t)" + by cases auto + { + fix t + assume exec_bdy: "\\\bdy,Normal (init s)\ \ Normal t" + hence "\\c s t \ Normal (return s t)" + proof - + from exec_bdy bdy + have "\\\(Call p ),Normal (init s)\ \ Normal t" + by (auto simp add: intro: exec.intros) + from cvalidD [OF valid_modif [rule_format, of "init s"] ctxt' this] P + res_modif + have "return' s t = return s t" + by auto + with termi_c exec_bdy show ?thesis by auto + qed + } + with bdy termi_bdy + show ?thesis + by (iprover intro: terminates_call_exn) + next + case False + thus ?thesis + by (auto intro: terminates_call_exnUndefined) + qed +qed + +lemma ProcModifyReturnSameFaults_sound: + assumes valid_call: "\,\ \\<^sub>t\<^bsub>/F\<^esub> P call init p return' c Q,A" + assumes valid_modif: + "\\. \,\ \\<^bsub>/F\<^esub> {\} Call p (Modif \),(ModifAbr \)" + assumes res_modif: + "\s t. t \ Modif (init s) \ return' s t = return s t" + assumes ret_modifAbr: + "\s t. t \ ModifAbr (init s) \ return' s t = return s t" +shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (call init p return c) Q,A" + using valid_call valid_modif res_modif ret_modifAbr unfolding call_call_exn + by (rule Proc_exnModifyReturnSameFaults_sound) + +lemma Proc_exnModifyReturnSameFaults: + assumes spec: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call_exn init p return' result_exn c) Q,A" + assumes res_modif: + "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" + assumes ret_modifAbr: + "\s t. t \ ModifAbr (init s) \ result_exn (return' s t) t = result_exn (return s t) t" + assumes modifies_spec: + "\\. \,\\\<^bsub>/F\<^esub> {\} (Call p) (Modif \),(ModifAbr \)" + shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call_exn init p return result_exn c) Q,A" +apply (rule hoaret_complete') +apply (rule Proc_exnModifyReturnSameFaults_sound [where Modif=Modif and ModifAbr=ModifAbr, + OF _ _ res_modif ret_modifAbr]) +apply (rule hoaret_sound [OF spec]) +using modifies_spec +apply (blast intro: hoare_sound) +done + + lemma ProcModifyReturnSameFaults: assumes spec: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return' c) Q,A" assumes res_modif: "\s t. t \ Modif (init s) \ (return' s t) = (return s t)" assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) \ (return' s t) = (return s t)" assumes modifies_spec: "\\. \,\\\<^bsub>/F\<^esub> {\} (Call p) (Modif \),(ModifAbr \)" - shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return c) Q,A" -apply (rule hoaret_complete') -apply (rule ProcModifyReturnSameFaults_sound [where Modif=Modif and ModifAbr=ModifAbr, - OF _ _ res_modif ret_modifAbr]) -apply (rule hoaret_sound [OF spec]) -using modifies_spec -apply (blast intro: hoare_sound) -done +shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (call init p return c) Q,A" + using spec res_modif ret_modifAbr modifies_spec unfolding call_call_exn + by (rule Proc_exnModifyReturnSameFaults) + subsubsection \DynCall\ - -lemma dynProcModifyReturn_sound: -assumes valid_call: "\,\ \\<^sub>t\<^bsub>/F\<^esub> P dynCall init p return' c Q,A" +lemma dynProc_exnModifyReturn_sound: +assumes valid_call: "\,\ \\<^sub>t\<^bsub>/F\<^esub> P dynCall_exn f g init p return' result_exn c Q,A" assumes valid_modif: "\s\P. \\. \,\ \\<^bsub>/UNIV\<^esub> {\} (Call (p s)) (Modif \),(ModifAbr \)" assumes ret_modif: "\s t. t \ Modif (init s) \ return' s t = return s t" -assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) \ return' s t = return s t" -shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" +assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) \ result_exn (return' s t) t = result_exn (return s t) t" +shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" hence "\(P, p, Q, A)\\. \ \\<^bsub>/F\<^esub> P (Call p) Q,A" by (auto simp add: validt_def) then have ctxt': "\(P, p, Q, A)\\. \ \\<^bsub>/UNIV\<^esub> P (Call p) Q,A" by (auto intro: valid_augment_Faults) - assume exec: "\\\dynCall init p return c,Normal s\ \ t" + assume exec: "\\\dynCall_exn f g init p return result_exn c,Normal s\ \ t" assume t_notin_F: "t \ Fault ` F" assume P: "s \ P" with valid_modif have valid_modif': "\\. \,\\\<^bsub>/UNIV\<^esub> {\} (Call (p s)) (Modif \),(ModifAbr \)" by blast from exec - have "\\\call init (p s) return c,Normal s\ \ t" - by (cases rule: exec_dynCall_Normal_elim) + have "\\\maybe_guard f g (call_exn init (p s) return result_exn c),Normal s\ \ t" + by (cases rule: exec_dynCall_exn_Normal_elim) then show "t \ Normal ` Q \ Abrupt ` A" - proof (cases rule: exec_call_Normal_elim) - fix bdy t' - assume bdy: "\ (p s) = Some bdy" - assume exec_body: "\\\bdy,Normal (init s)\ \ Normal t'" - assume exec_c: "\\\c s t',Normal (return s t')\ \ t" - from exec_body bdy - have "\\\Call (p s),Normal (init s)\ \ Normal t'" - by (auto simp add: intro: exec.Call) - from cvalidD [OF valid_modif' [rule_format, of "init s"] ctxt' this] P - have "t' \ Modif (init s)" - by auto - with ret_modif have "Normal (return' s t') = - Normal (return s t')" - by simp - with exec_body exec_c bdy - have "\\\call init (p s) return' c,Normal s\ \ t" - by (auto intro: exec_call) - hence "\\\dynCall init p return' c,Normal s\ \ t" - by (rule exec_dynCall) - from cvalidt_postD [OF valid_call ctxt this] P t_notin_F - show ?thesis - by simp + proof (cases rule: exec_maybe_guard_Normal_elim_cases) + case noFault + from noFault have guards_ok: "s \ g" by simp + from noFault have "\\ \call_exn init (p s) return result_exn c,Normal s\ \ t" by simp + then show "t \ Normal ` Q \ Abrupt ` A" + proof (cases rule: exec_call_exn_Normal_elim) + fix bdy t' + assume bdy: "\ (p s) = Some bdy" + assume exec_body: "\\\bdy,Normal (init s)\ \ Normal t'" + assume exec_c: "\\\c s t',Normal (return s t')\ \ t" + from exec_body bdy + have "\\\Call (p s),Normal (init s)\ \ Normal t'" + by (auto simp add: intro: exec.Call) + from cvalidD [OF valid_modif' [rule_format, of "init s"] ctxt' this] P + have "t' \ Modif (init s)" + by auto + with ret_modif have "Normal (return' s t') = + Normal (return s t')" + by simp + with exec_body exec_c bdy + have "\\\call_exn init (p s) return' result_exn c,Normal s\ \ t" + by (auto intro: exec_call_exn) + from exec_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ \ t" + by (simp add: dynCall_exn_def exec_maybe_guard_DynCom) + from cvalidt_postD [OF valid_call ctxt this] P t_notin_F + show ?thesis + by simp + next + fix bdy t' + assume bdy: "\ (p s) = Some bdy" + assume exec_body: "\\\bdy,Normal (init s)\ \ Abrupt t'" + assume t: "t = Abrupt (result_exn (return s t') t')" + also from exec_body bdy + have "\\\Call (p s) ,Normal (init s)\ \ Abrupt t'" + by (auto simp add: intro: exec.intros) + from cvalidD [OF valid_modif' [rule_format, of "init s"] ctxt' this] P + have "t' \ ModifAbr (init s)" + by auto + with ret_modifAbr have "Abrupt (result_exn (return s t') t') = Abrupt (result_exn (return' s t') t')" + by simp + finally have "t = Abrupt (result_exn (return' s t') t')" . + with exec_body bdy + have "\\\call_exn init (p s) return' result_exn c,Normal s\ \ t" + by (auto intro: exec_call_exnAbrupt) + from exec_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ \ t" + by (simp add: dynCall_exn_def exec_maybe_guard_DynCom) + from cvalidt_postD [OF valid_call ctxt this] P t_notin_F + show ?thesis + by simp + next + fix bdy f' + assume bdy: "\ (p s) = Some bdy" + assume "\\\bdy,Normal (init s)\ \ Fault f'" and + t: "t = Fault f'" + with bdy have "\\\call_exn init (p s) return' result_exn c ,Normal s\ \ t" + by (auto intro: exec_call_exnFault) + from exec_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ \ t" + by (simp add: dynCall_exn_def exec_maybe_guard_DynCom) + from cvalidt_postD [OF valid_call ctxt this P] t t_notin_F + show ?thesis + by blast + next + fix bdy + assume bdy: "\ (p s) = Some bdy" + assume "\\\bdy,Normal (init s)\ \ Stuck" + "t = Stuck" + with bdy have "\\\call_exn init (p s) return' result_exn c ,Normal s\ \ t" + by (auto intro: exec_call_exnStuck) + from exec_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ \ t" + by (simp add: dynCall_exn_def exec_maybe_guard_DynCom) + from valid_call ctxt this P t_notin_F + show ?thesis + by (rule cvalidt_postD) + next + fix bdy + assume "\ (p s) = None" "t=Stuck" + hence "\\\call_exn init (p s) return' result_exn c ,Normal s\ \ t" + by (auto intro: exec_call_exnUndefined) + from exec_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ \ t" + by (simp add: dynCall_exn_def exec_maybe_guard_DynCom) + from valid_call ctxt this P t_notin_F + show ?thesis + by (rule cvalidt_postD) + qed next - fix bdy t' - assume bdy: "\ (p s) = Some bdy" - assume exec_body: "\\\bdy,Normal (init s)\ \ Abrupt t'" - assume t: "t = Abrupt (return s t')" - also from exec_body bdy - have "\\\Call (p s) ,Normal (init s)\ \ Abrupt t'" - by (auto simp add: intro: exec.intros) - from cvalidD [OF valid_modif' [rule_format, of "init s"] ctxt' this] P - have "t' \ ModifAbr (init s)" - by auto - with ret_modifAbr have "Abrupt (return s t') = Abrupt (return' s t')" - by simp - finally have "t = Abrupt (return' s t')" . - with exec_body bdy - have "\\\call init (p s) return' c,Normal s\ \ t" - by (auto intro: exec_callAbrupt) - hence "\\\dynCall init p return' c,Normal s\ \ t" - by (rule exec_dynCall) + case (someFault) + then obtain guards_fail:"s \ g" + and t: "t = Fault f" by simp + from exec_maybe_guard_Fault [OF guards_fail] t + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ \ t" + by (simp add: dynCall_exn_def exec_guards_DynCom) from cvalidt_postD [OF valid_call ctxt this] P t_notin_F - show ?thesis - by simp - next - fix bdy f - assume bdy: "\ (p s) = Some bdy" - assume "\\\bdy,Normal (init s)\ \ Fault f" and - t: "t = Fault f" - with bdy have "\\\call init (p s) return' c ,Normal s\ \ t" - by (auto intro: exec_callFault) - hence "\\\dynCall init p return' c,Normal s\ \ t" - by (rule exec_dynCall) - from cvalidt_postD [OF valid_call ctxt this P] t t_notin_F - show ?thesis - by blast - next - fix bdy - assume bdy: "\ (p s) = Some bdy" - assume "\\\bdy,Normal (init s)\ \ Stuck" - "t = Stuck" - with bdy have "\\\call init (p s) return' c ,Normal s\ \ t" - by (auto intro: exec_callStuck) - hence "\\\dynCall init p return' c,Normal s\ \ t" - by (rule exec_dynCall) - from valid_call ctxt this P t_notin_F - show ?thesis - by (rule cvalidt_postD) - next - fix bdy - assume "\ (p s) = None" "t=Stuck" - hence "\\\call init (p s) return' c ,Normal s\ \ t" - by (auto intro: exec_callUndefined) - hence "\\\dynCall init p return' c,Normal s\ \ t" - by (rule exec_dynCall) - from valid_call ctxt this P t_notin_F - show ?thesis - by (rule cvalidt_postD) + show ?thesis by simp qed next fix s assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" hence "\(P, p, Q, A)\\. \ \\<^bsub>/F\<^esub> P (Call p) Q,A" by (auto simp add: validt_def) then have ctxt': "\(P, p, Q, A)\\. \ \\<^bsub>/UNIV\<^esub> P (Call p) Q,A" by (auto intro: valid_augment_Faults) assume P: "s \ P" from valid_call ctxt P - have "\\dynCall init p return' c\ Normal s" + have "\\dynCall_exn f g init p return' result_exn c\ Normal s" by (rule cvalidt_termD) - hence call: "\\call init (p s) return' c\ Normal s" - by cases - have "\\call init (p s) return c \ Normal s" - proof (cases "p s \ dom \") - case True - with call obtain bdy where - bdy: "\ (p s) = Some bdy" and termi_bdy: "\\bdy \ Normal (init s)" and - termi_c: "\t. \\\bdy,Normal (init s)\ \ Normal t \ - \\c s t \ Normal (return' s t)" - by cases auto - { - fix t - assume exec_bdy: "\\\bdy,Normal (init s)\ \ Normal t" - hence "\\c s t \ Normal (return s t)" - proof - - from exec_bdy bdy - have "\\\Call (p s),Normal (init s)\ \ Normal t" - by (auto simp add: intro: exec.intros) - from cvalidD [OF valid_modif [rule_format, of s "init s"] ctxt' this] P - ret_modif - have "return' s t = return s t" - by auto - with termi_c exec_bdy show ?thesis by auto - qed - } - with bdy termi_bdy - show ?thesis - by (iprover intro: terminates_call) + thus "\\dynCall_exn f g init p return result_exn c \ Normal s" + proof (cases rule: terminates_dynCall_exn_elim) + case noFault + then obtain guards_ok: "s \ g" + and call: "\\call_exn init (p s) return' result_exn c\ Normal s" + by auto + have "\\call_exn init (p s) return result_exn c \ Normal s" + proof (cases "p s \ dom \") + case True + with call obtain bdy where + bdy: "\ (p s) = Some bdy" and termi_bdy: "\\bdy \ Normal (init s)" and + termi_c: "\t. \\\bdy,Normal (init s)\ \ Normal t \ + \\c s t \ Normal (return' s t)" + by cases auto + { + fix t + assume exec_bdy: "\\\bdy,Normal (init s)\ \ Normal t" + hence "\\c s t \ Normal (return s t)" + proof - + from exec_bdy bdy + have "\\\Call (p s),Normal (init s)\ \ Normal t" + by (auto simp add: intro: exec.intros) + from cvalidD [OF valid_modif [rule_format, of s "init s"] ctxt' this] P + ret_modif + have "return' s t = return s t" + by auto + with termi_c exec_bdy show ?thesis by auto + qed + } + with bdy termi_bdy + show ?thesis + by (iprover intro: terminates_call_exn) + next + case False + thus ?thesis + by (auto intro: terminates_call_exnUndefined) + qed + thus "\\dynCall_exn f g init p return result_exn c \ Normal s" + by (iprover intro: terminates_dynCall_exn) next - case False + case (someFault) + then have guard_fail: "s \ g" by simp thus ?thesis - by (auto intro: terminates_callUndefined) + by (simp add: terminates_maybe_guard_Fault dynCall_exn_def) qed - thus "\\dynCall init p return c \ Normal s" - by (iprover intro: terminates_dynCall) qed +lemma dynProcModifyReturn_sound: +assumes valid_call: "\,\ \\<^sub>t\<^bsub>/F\<^esub> P dynCall init p return' c Q,A" +assumes valid_modif: + "\s\P. \\. \,\ \\<^bsub>/UNIV\<^esub> {\} (Call (p s)) (Modif \),(ModifAbr \)" +assumes ret_modif: + "\s t. t \ Modif (init s) \ return' s t = return s t" +assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) \ return' s t = return s t" +shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using valid_call valid_modif ret_modif ret_modifAbr unfolding dynCall_dynCall_exn + by (rule dynProc_exnModifyReturn_sound) + + +lemma dynProc_exnModifyReturn: +assumes dyn_call: "\,\\\<^sub>t\<^bsub>/F\<^esub> P dynCall_exn f g init p return' result_exn c Q,A" +assumes ret_modif: + "\s t. t \ Modif (init s) + \ return' s t = return s t" +assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) + \ result_exn (return' s t) t = result_exn (return s t) t" +assumes modif: + "\s \ P. \\. + \,\\\<^bsub>/UNIV\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" +shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" +apply (rule hoaret_complete') +apply (rule dynProc_exnModifyReturn_sound + [where Modif=Modif and ModifAbr=ModifAbr, + OF hoaret_sound [OF dyn_call] _ ret_modif ret_modifAbr]) +apply (intro ballI allI) +apply (rule hoare_sound [OF modif [rule_format]]) +apply assumption + done + lemma dynProcModifyReturn: assumes dyn_call: "\,\\\<^sub>t\<^bsub>/F\<^esub> P dynCall init p return' c Q,A" assumes ret_modif: "\s t. t \ Modif (init s) \ return' s t = return s t" assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) \ return' s t = return s t" assumes modif: "\s \ P. \\. \,\\\<^bsub>/UNIV\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" -shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" -apply (rule hoaret_complete') -apply (rule dynProcModifyReturn_sound - [where Modif=Modif and ModifAbr=ModifAbr, - OF hoaret_sound [OF dyn_call] _ ret_modif ret_modifAbr]) -apply (intro ballI allI) -apply (rule hoare_sound [OF modif [rule_format]]) -apply assumption -done + shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using dyn_call ret_modif ret_modifAbr modif unfolding dynCall_dynCall_exn + by (rule dynProc_exnModifyReturn) + +lemma dynProc_exnModifyReturnSameFaults_sound: +assumes valid_call: "\,\ \\<^sub>t\<^bsub>/F\<^esub> P dynCall_exn f g init p return' result_exn c Q,A" +assumes valid_modif: + "\s\P. \\. \,\ \\<^bsub>/F\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" +assumes ret_modif: + "\s t. t \ Modif (init s) \ return' s t = return s t" +assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) \ result_exn (return' s t) t = result_exn (return s t) t" +shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" +proof (rule cvalidtI) + fix s t + assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" + hence ctxt': "\(P, p, Q, A)\\. \ \\<^bsub>/F\<^esub> P (Call p) Q,A" + by (auto simp add: validt_def) + assume exec: "\\\dynCall_exn f g init p return result_exn c,Normal s\ \ t" + assume t_notin_F: "t \ Fault ` F" + assume P: "s \ P" + with valid_modif + have valid_modif': + "\\. \,\\\<^bsub>/F\<^esub> {\} (Call (p s)) (Modif \),(ModifAbr \)" + by blast + from exec + have "\\\maybe_guard f g (call_exn init (p s) return result_exn c),Normal s\ \ t" + by (cases rule: exec_dynCall_exn_Normal_elim) + then show "t \ Normal ` Q \ Abrupt ` A" + proof (cases rule: exec_maybe_guard_Normal_elim_cases) + case noFault + from noFault have guards_ok: "s \ g" by simp + from noFault have "\\ \call_exn init (p s) return result_exn c,Normal s\ \ t" by simp + then show "t \ Normal ` Q \ Abrupt ` A" + proof (cases rule: exec_call_exn_Normal_elim) + fix bdy t' + assume bdy: "\ (p s) = Some bdy" + assume exec_body: "\\\bdy,Normal (init s)\ \ Normal t'" + assume exec_c: "\\\c s t',Normal (return s t')\ \ t" + from exec_body bdy + have "\\\Call (p s),Normal (init s)\ \ Normal t'" + by (auto simp add: intro: exec.intros) + from cvalidD [OF valid_modif' [rule_format, of "init s"] ctxt' this] P + have "t' \ Modif (init s)" + by auto + with ret_modif have "Normal (return' s t') = + Normal (return s t')" + by simp + with exec_body exec_c bdy + have "\\\call_exn init (p s) return' result_exn c,Normal s\ \ t" + by (auto intro: exec_call_exn) + from exec_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ \ t" + by (simp add: dynCall_exn_def exec_maybe_guard_DynCom) + from cvalidt_postD [OF valid_call ctxt this] P t_notin_F + show ?thesis + by simp + next + fix bdy t' + assume bdy: "\ (p s) = Some bdy" + assume exec_body: "\\\bdy,Normal (init s)\ \ Abrupt t'" + assume t: "t = Abrupt (result_exn (return s t') t')" + also from exec_body bdy + have "\\\Call (p s) ,Normal (init s)\ \ Abrupt t'" + by (auto simp add: intro: exec.intros) + from cvalidD [OF valid_modif' [rule_format, of "init s"] ctxt' this] P + have "t' \ ModifAbr (init s)" + by auto + with ret_modifAbr have "Abrupt (result_exn (return s t') t') = Abrupt (result_exn (return' s t') t')" + by simp + finally have "t = Abrupt (result_exn (return' s t') t')" . + with exec_body bdy + have "\\\call_exn init (p s) return' result_exn c,Normal s\ \ t" + by (auto intro: exec_call_exnAbrupt) + from exec_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ \ t" + by (simp add: dynCall_exn_def exec_maybe_guard_DynCom) + from cvalidt_postD [OF valid_call ctxt this] P t_notin_F + show ?thesis + by simp + next + fix bdy f' + assume bdy: "\ (p s) = Some bdy" + assume "\\\bdy,Normal (init s)\ \ Fault f'" and + t: "t = Fault f'" + with bdy have "\\\call_exn init (p s) return' result_exn c ,Normal s\ \ t" + by (auto intro: exec_call_exnFault) + from exec_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ \ t" + by (simp add: dynCall_exn_def exec_maybe_guard_DynCom) + from cvalidt_postD [OF valid_call ctxt this P] t t_notin_F + show ?thesis + by simp + next + fix bdy + assume bdy: "\ (p s) = Some bdy" + assume "\\\bdy,Normal (init s)\ \ Stuck" + "t = Stuck" + with bdy have "\\\call_exn init (p s) return' result_exn c ,Normal s\ \ t" + by (auto intro: exec_call_exnStuck) + from exec_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ \ t" + by (simp add: dynCall_exn_def exec_maybe_guard_DynCom) + from valid_call ctxt this P t_notin_F + show ?thesis + by (rule cvalidt_postD) + next + fix bdy + assume "\ (p s) = None" "t=Stuck" + hence "\\\call_exn init (p s) return' result_exn c ,Normal s\ \ t" + by (auto intro: exec_call_exnUndefined) + from exec_maybe_guard_noFault [OF this guards_ok] + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ \ t" + by (simp add: dynCall_exn_def exec_maybe_guard_DynCom) + from valid_call ctxt this P t_notin_F + show ?thesis + by (rule cvalidt_postD) + qed + next + case (someFault) + then obtain guards_fail:"s \ g" + and t: "t = Fault f" by simp + from exec_maybe_guard_Fault [OF guards_fail] t + have "\\\dynCall_exn f g init p return' result_exn c,Normal s\ \ t" + by (simp add: dynCall_exn_def exec_guards_DynCom) + from cvalidt_postD [OF valid_call ctxt this] P t_notin_F + show ?thesis by simp + qed +next + fix s + assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" + hence ctxt': "\(P, p, Q, A)\\. \ \\<^bsub>/F\<^esub> P (Call p) Q,A" + by (auto simp add: validt_def) + assume P: "s \ P" + from valid_call ctxt P + have "\\dynCall_exn f g init p return' result_exn c\ Normal s" + by (rule cvalidt_termD) + thus "\\dynCall_exn f g init p return result_exn c \ Normal s" + proof (cases rule: terminates_dynCall_exn_elim) + case noFault + then obtain guards_ok: "s \ g" + and call: "\\call_exn init (p s) return' result_exn c\ Normal s" + by auto + have "\\call_exn init (p s) return result_exn c \ Normal s" + proof (cases "p s \ dom \") + case True + with call obtain bdy where + bdy: "\ (p s) = Some bdy" and termi_bdy: "\\bdy \ Normal (init s)" and + termi_c: "\t. \\\bdy,Normal (init s)\ \ Normal t \ + \\c s t \ Normal (return' s t)" + by cases auto + { + fix t + assume exec_bdy: "\\\bdy,Normal (init s)\ \ Normal t" + hence "\\c s t \ Normal (return s t)" + proof - + from exec_bdy bdy + have "\\\Call (p s),Normal (init s)\ \ Normal t" + by (auto simp add: intro: exec.intros) + from cvalidD [OF valid_modif [rule_format, of s "init s"] ctxt' this] P + ret_modif + have "return' s t = return s t" + by auto + with termi_c exec_bdy show ?thesis by auto + qed + } + with bdy termi_bdy + show ?thesis + by (iprover intro: terminates_call_exn) + next + case False + thus ?thesis + by (auto intro: terminates_call_exnUndefined) + qed + thus "\\dynCall_exn f g init p return result_exn c \ Normal s" + by (iprover intro: terminates_dynCall_exn) + next + case (someFault) + then have guard_fail: "s \ g" by simp + thus ?thesis + by (simp add: terminates_maybe_guard_Fault dynCall_exn_def) + qed +qed lemma dynProcModifyReturnSameFaults_sound: assumes valid_call: "\,\ \\<^sub>t\<^bsub>/F\<^esub> P dynCall init p return' c Q,A" assumes valid_modif: "\s\P. \\. \,\ \\<^bsub>/F\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" assumes ret_modif: "\s t. t \ Modif (init s) \ return' s t = return s t" assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) \ return' s t = return s t" shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" -proof (rule cvalidtI) - fix s t - assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" - hence ctxt': "\(P, p, Q, A)\\. \ \\<^bsub>/F\<^esub> P (Call p) Q,A" - by (auto simp add: validt_def) - assume exec: "\\\dynCall init p return c,Normal s\ \ t" - assume t_notin_F: "t \ Fault ` F" - assume P: "s \ P" - with valid_modif - have valid_modif': - "\\. \,\\\<^bsub>/F\<^esub> {\} (Call (p s)) (Modif \),(ModifAbr \)" - by blast - from exec - have "\\\call init (p s) return c,Normal s\ \ t" - by (cases rule: exec_dynCall_Normal_elim) - then show "t \ Normal ` Q \ Abrupt ` A" - proof (cases rule: exec_call_Normal_elim) - fix bdy t' - assume bdy: "\ (p s) = Some bdy" - assume exec_body: "\\\bdy,Normal (init s)\ \ Normal t'" - assume exec_c: "\\\c s t',Normal (return s t')\ \ t" - from exec_body bdy - have "\\\Call (p s),Normal (init s)\ \ Normal t'" - by (auto simp add: intro: exec.intros) - from cvalidD [OF valid_modif' [rule_format, of "init s"] ctxt' this] P - have "t' \ Modif (init s)" - by auto - with ret_modif have "Normal (return' s t') = - Normal (return s t')" - by simp - with exec_body exec_c bdy - have "\\\call init (p s) return' c,Normal s\ \ t" - by (auto intro: exec_call) - hence "\\\dynCall init p return' c,Normal s\ \ t" - by (rule exec_dynCall) - from cvalidt_postD [OF valid_call ctxt this] P t_notin_F - show ?thesis - by simp - next - fix bdy t' - assume bdy: "\ (p s) = Some bdy" - assume exec_body: "\\\bdy,Normal (init s)\ \ Abrupt t'" - assume t: "t = Abrupt (return s t')" - also from exec_body bdy - have "\\\Call (p s) ,Normal (init s)\ \ Abrupt t'" - by (auto simp add: intro: exec.intros) - from cvalidD [OF valid_modif' [rule_format, of "init s"] ctxt' this] P - have "t' \ ModifAbr (init s)" - by auto - with ret_modifAbr have "Abrupt (return s t') = Abrupt (return' s t')" - by simp - finally have "t = Abrupt (return' s t')" . - with exec_body bdy - have "\\\call init (p s) return' c,Normal s\ \ t" - by (auto intro: exec_callAbrupt) - hence "\\\dynCall init p return' c,Normal s\ \ t" - by (rule exec_dynCall) - from cvalidt_postD [OF valid_call ctxt this] P t_notin_F - show ?thesis - by simp - next - fix bdy f - assume bdy: "\ (p s) = Some bdy" - assume "\\\bdy,Normal (init s)\ \ Fault f" and - t: "t = Fault f" - with bdy have "\\\call init (p s) return' c ,Normal s\ \ t" - by (auto intro: exec_callFault) - hence "\\\dynCall init p return' c,Normal s\ \ t" - by (rule exec_dynCall) - from cvalidt_postD [OF valid_call ctxt this P] t t_notin_F - show ?thesis - by simp - next - fix bdy - assume bdy: "\ (p s) = Some bdy" - assume "\\\bdy,Normal (init s)\ \ Stuck" - "t = Stuck" - with bdy have "\\\call init (p s) return' c ,Normal s\ \ t" - by (auto intro: exec_callStuck) - hence "\\\dynCall init p return' c,Normal s\ \ t" - by (rule exec_dynCall) - from valid_call ctxt this P t_notin_F - show ?thesis - by (rule cvalidt_postD) - next - fix bdy - assume "\ (p s) = None" "t=Stuck" - hence "\\\call init (p s) return' c ,Normal s\ \ t" - by (auto intro: exec_callUndefined) - hence "\\\dynCall init p return' c,Normal s\ \ t" - by (rule exec_dynCall) - from valid_call ctxt this P t_notin_F - show ?thesis - by (rule cvalidt_postD) - qed -next - fix s - assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" - hence ctxt': "\(P, p, Q, A)\\. \ \\<^bsub>/F\<^esub> P (Call p) Q,A" - by (auto simp add: validt_def) - assume P: "s \ P" - from valid_call ctxt P - have "\\dynCall init p return' c\ Normal s" - by (rule cvalidt_termD) - hence call: "\\call init (p s) return' c\ Normal s" - by cases - have "\\call init (p s) return c \ Normal s" - proof (cases "p s \ dom \") - case True - with call obtain bdy where - bdy: "\ (p s) = Some bdy" and termi_bdy: "\\bdy \ Normal (init s)" and - termi_c: "\t. \\\bdy,Normal (init s)\ \ Normal t \ - \\c s t \ Normal (return' s t)" - by cases auto - { - fix t - assume exec_bdy: "\\\bdy,Normal (init s)\ \ Normal t" - hence "\\c s t \ Normal (return s t)" - proof - - from exec_bdy bdy - have "\\\Call (p s),Normal (init s)\ \ Normal t" - by (auto simp add: intro: exec.intros) - from cvalidD [OF valid_modif [rule_format, of s "init s"] ctxt' this] P - ret_modif - have "return' s t = return s t" - by auto - with termi_c exec_bdy show ?thesis by auto - qed - } - with bdy termi_bdy - show ?thesis - by (iprover intro: terminates_call) - next - case False - thus ?thesis - by (auto intro: terminates_callUndefined) - qed - thus "\\dynCall init p return c \ Normal s" - by (iprover intro: terminates_dynCall) -qed + using valid_call valid_modif ret_modif ret_modifAbr unfolding dynCall_dynCall_exn + by (rule dynProc_exnModifyReturnSameFaults_sound) + +lemma dynProc_exnModifyReturnSameFaults: +assumes dyn_call: "\,\\\<^sub>t\<^bsub>/F\<^esub> P dynCall_exn f g init p return' result_exn c Q,A" +assumes ret_modif: + "\s t. t \ Modif (init s) \ return' s t = return s t" +assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) \ result_exn (return' s t) t = result_exn (return s t) t" +assumes modif: + "\s \ P. \\. \,\\\<^bsub>/F\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" +shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (dynCall_exn f g init p return result_exn c) Q,A" +apply (rule hoaret_complete') +apply (rule dynProc_exnModifyReturnSameFaults_sound + [where Modif=Modif and ModifAbr=ModifAbr, + OF hoaret_sound [OF dyn_call] _ ret_modif ret_modifAbr]) +apply (intro ballI allI) +apply (rule hoare_sound [OF modif [rule_format]]) +apply assumption + done lemma dynProcModifyReturnSameFaults: assumes dyn_call: "\,\\\<^sub>t\<^bsub>/F\<^esub> P dynCall init p return' c Q,A" assumes ret_modif: "\s t. t \ Modif (init s) \ return' s t = return s t" assumes ret_modifAbr: "\s t. t \ ModifAbr (init s) \ return' s t = return s t" assumes modif: "\s \ P. \\. \,\\\<^bsub>/F\<^esub> {\} Call (p s) (Modif \),(ModifAbr \)" -shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" -apply (rule hoaret_complete') -apply (rule dynProcModifyReturnSameFaults_sound - [where Modif=Modif and ModifAbr=ModifAbr, - OF hoaret_sound [OF dyn_call] _ ret_modif ret_modifAbr]) -apply (intro ballI allI) -apply (rule hoare_sound [OF modif [rule_format]]) -apply assumption -done + shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P (dynCall init p return c) Q,A" + using dyn_call ret_modif ret_modifAbr modif unfolding dynCall_dynCall_exn + by (rule dynProc_exnModifyReturnSameFaults) + subsubsection \Conjunction of Postcondition\ lemma PostConjI_sound: assumes valid_Q: "\,\ \\<^sub>t\<^bsub>/F\<^esub> P c Q,A" assumes valid_R: "\,\ \\<^sub>t\<^bsub>/F\<^esub> P c R,B" shows "\,\ \\<^sub>t\<^bsub>/F\<^esub> P c (Q \ R),(A \ B)" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\c,Normal s\ \ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" from valid_Q ctxt exec P t_notin_F have "t \ Normal ` Q \ Abrupt ` A" by (rule cvalidt_postD) moreover from valid_R ctxt exec P t_notin_F have "t \ Normal ` R \ Abrupt ` B" by (rule cvalidt_postD) ultimately show "t \ Normal ` (Q \ R) \ Abrupt ` (A \ B)" by blast next fix s assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume P: "s \ P" from valid_Q ctxt P show "\\c \ Normal s" by (rule cvalidt_termD) qed lemma PostConjI: assumes deriv_Q: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" assumes deriv_R: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c R,B" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c (Q \ R),(A \ B)" apply (rule hoaret_complete') apply (rule PostConjI_sound) apply (rule hoaret_sound [OF deriv_Q]) apply (rule hoaret_sound [OF deriv_R]) done lemma Merge_PostConj_sound: assumes validF: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" assumes validG: "\,\\\<^sub>t\<^bsub>/G\<^esub> P' c R,X" assumes F_G: "F \ G" assumes P_P': "P \ P'" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c (Q \ R),(A \ X)" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" with F_G have ctxt': "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/G\<^esub> P (Call p) Q,A" by (auto intro: validt_augment_Faults) assume exec: "\\\c,Normal s\ \ t" assume P: "s \ P" with P_P' have P': "s \ P'" by auto assume t_noFault: "t \ Fault ` F" show "t \ Normal ` (Q \ R) \ Abrupt ` (A \ X)" proof - from cvalidt_postD [OF validF [rule_format] ctxt exec P t_noFault] have t: "t \ Normal ` Q \ Abrupt ` A". then have "t \ Fault ` G" by auto from cvalidt_postD [OF validG [rule_format] ctxt' exec P' this] have "t \ Normal ` R \ Abrupt ` X" . with t show ?thesis by auto qed next fix s assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume P: "s \ P" from validF ctxt P show "\\c \ Normal s" by (rule cvalidt_termD) qed lemma Merge_PostConj: assumes validF: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" assumes validG: "\,\\\<^sub>t\<^bsub>/G\<^esub> P' c R,X" assumes F_G: "F \ G" assumes P_P': "P \ P'" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c (Q \ R),(A \ X)" apply (rule hoaret_complete') apply (rule Merge_PostConj_sound [OF _ _ F_G P_P']) using validF apply (blast intro:hoaret_sound) using validG apply (blast intro:hoaret_sound) done subsubsection \Guards and Guarantees\ lemma SplitGuards_sound: assumes valid_c1: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c\<^sub>1 Q,A" assumes valid_c2: "\,\\\<^bsub>/F\<^esub> P c\<^sub>2 UNIV,UNIV" assumes c: "(c\<^sub>1 \\<^sub>g c\<^sub>2) = Some c" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" hence ctxt': "\(P, p, Q, A)\\. \ \\<^bsub>/F\<^esub> P (Call p) Q,A" by (auto simp add: validt_def) assume exec: "\\\c,Normal s\ \ t" assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" show "t \ Normal ` Q \ Abrupt ` A" proof (cases t) case Normal with inter_guards_exec_noFault [OF c exec] have "\\\c\<^sub>1,Normal s\ \ t" by simp from valid_c1 ctxt this P t_notin_F show ?thesis by (rule cvalidt_postD) next case Abrupt with inter_guards_exec_noFault [OF c exec] have "\\\c\<^sub>1,Normal s\ \ t" by simp from valid_c1 ctxt this P t_notin_F show ?thesis by (rule cvalidt_postD) next case (Fault f) assume t: "t=Fault f" with exec inter_guards_exec_Fault [OF c] have "\\\c\<^sub>1,Normal s\ \ Fault f \ \\\c\<^sub>2,Normal s\ \ Fault f" by auto then show ?thesis proof (cases rule: disjE [consumes 1]) assume "\\\c\<^sub>1,Normal s\ \ Fault f" from cvalidt_postD [OF valid_c1 ctxt this P] t t_notin_F show ?thesis by blast next assume "\\\c\<^sub>2,Normal s\ \ Fault f" from cvalidD [OF valid_c2 ctxt' this P] t t_notin_F show ?thesis by blast qed next case Stuck with inter_guards_exec_noFault [OF c exec] have "\\\c\<^sub>1,Normal s\ \ t" by simp from valid_c1 ctxt this P t_notin_F show ?thesis by (rule cvalidt_postD) qed next fix s assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume P: "s \ P" show "\\c \ Normal s" proof - from valid_c1 ctxt P have "\\c\<^sub>1 \ Normal s" by (rule cvalidt_termD) with c show ?thesis by (rule inter_guards_terminates) qed qed lemma SplitGuards: assumes c: "(c\<^sub>1 \\<^sub>g c\<^sub>2) = Some c" assumes deriv_c1: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c\<^sub>1 Q,A" assumes deriv_c2: "\,\\\<^bsub>/F\<^esub> P c\<^sub>2 UNIV,UNIV" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" apply (rule hoaret_complete') apply (rule SplitGuards_sound [OF _ _ c]) apply (rule hoaret_sound [OF deriv_c1]) apply (rule hoare_sound [OF deriv_c2]) done lemma CombineStrip_sound: assumes valid: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" assumes valid_strip: "\,\\\<^bsub>/{}\<^esub> P (strip_guards (-F) c) UNIV,UNIV" shows "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/{}\<^esub> P (Call p) Q,A" hence ctxt': "\(P, p, Q, A)\\. \\\<^bsub>/{}\<^esub> P (Call p) Q,A" by (auto simp add: validt_def) from ctxt have ctxt'': "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" by (auto intro: valid_augment_Faults simp add: validt_def) assume exec: "\\\c,Normal s\ \ t" assume P: "s \ P" assume t_noFault: "t \ Fault ` {}" show "t \ Normal ` Q \ Abrupt ` A" proof (cases t) case (Normal t') from cvalidt_postD [OF valid ctxt'' exec P] Normal show ?thesis by auto next case (Abrupt t') from cvalidt_postD [OF valid ctxt'' exec P] Abrupt show ?thesis by auto next case (Fault f) show ?thesis proof (cases "f \ F") case True hence "f \ -F" by simp with exec Fault have "\\\strip_guards (-F) c,Normal s\ \ Fault f" by (auto intro: exec_to_exec_strip_guards_Fault) from cvalidD [OF valid_strip ctxt' this P] Fault have False by auto thus ?thesis .. next case False with cvalidt_postD [OF valid ctxt'' exec P] Fault show ?thesis by auto qed next case Stuck from cvalidt_postD [OF valid ctxt'' exec P] Stuck show ?thesis by auto qed next fix s assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/{}\<^esub> P (Call p) Q,A" hence ctxt': "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" by (auto intro: valid_augment_Faults simp add: validt_def) assume P: "s \ P" show "\\c \ Normal s" proof - from valid ctxt' P show "\\c \ Normal s" by (rule cvalidt_termD) qed qed lemma CombineStrip: assumes deriv: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" assumes deriv_strip: "\,\\\<^bsub>/{}\<^esub> P (strip_guards (-F) c) UNIV,UNIV" shows "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c Q,A" apply (rule hoaret_complete') apply (rule CombineStrip_sound) apply (iprover intro: hoaret_sound [OF deriv]) apply (iprover intro: hoare_sound [OF deriv_strip]) done lemma GuardsFlip_sound: assumes valid: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" assumes validFlip: "\,\\\<^bsub>/-F\<^esub> P c UNIV,UNIV" shows "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/{}\<^esub> P (Call p) Q,A" from ctxt have ctxt': "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" by (auto intro: valid_augment_Faults simp add: validt_def) from ctxt have ctxtFlip: "\(P, p, Q, A)\\. \\\<^bsub>/-F\<^esub> P (Call p) Q,A" by (auto intro: valid_augment_Faults simp add: validt_def) assume exec: "\\\c,Normal s\ \ t" assume P: "s \ P" assume t_noFault: "t \ Fault ` {}" show "t \ Normal ` Q \ Abrupt ` A" proof (cases t) case (Normal t') from cvalidt_postD [OF valid ctxt' exec P] Normal show ?thesis by auto next case (Abrupt t') from cvalidt_postD [OF valid ctxt' exec P] Abrupt show ?thesis by auto next case (Fault f) show ?thesis proof (cases "f \ F") case True hence "f \ -F" by simp with cvalidD [OF validFlip ctxtFlip exec P] Fault have False by auto thus ?thesis .. next case False with cvalidt_postD [OF valid ctxt' exec P] Fault show ?thesis by auto qed next case Stuck from cvalidt_postD [OF valid ctxt' exec P] Stuck show ?thesis by auto qed next fix s assume ctxt: "\(P, p, Q, A)\\. \ \\<^sub>t\<^bsub>/{}\<^esub> P (Call p) Q,A" hence ctxt': "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" by (auto intro: valid_augment_Faults simp add: validt_def) assume P: "s \ P" show "\\c \ Normal s" proof - from valid ctxt' P show "\\c \ Normal s" by (rule cvalidt_termD) qed qed lemma GuardsFlip: assumes deriv: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" assumes derivFlip: "\,\\\<^bsub>/-F\<^esub> P c UNIV,UNIV" shows "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c Q,A" apply (rule hoaret_complete') apply (rule GuardsFlip_sound) apply (iprover intro: hoaret_sound [OF deriv]) apply (iprover intro: hoare_sound [OF derivFlip]) done lemma MarkGuardsI_sound: assumes valid: "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c Q,A" shows "\,\\\<^sub>t\<^bsub>/{}\<^esub> P mark_guards f c Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/{}\<^esub> P (Call p) Q,A" assume exec: "\\\mark_guards f c,Normal s\ \ t" from exec_mark_guards_to_exec [OF exec] obtain t' where exec_c: "\\\c,Normal s\ \ t'" and t'_noFault: "\ isFault t' \ t' = t" by blast assume P: "s \ P" assume t_noFault: "t \ Fault ` {}" show "t \ Normal ` Q \ Abrupt ` A" proof - from cvalidt_postD [OF valid [rule_format] ctxt exec_c P] have "t' \ Normal ` Q \ Abrupt ` A" by blast with t'_noFault show ?thesis by auto qed next fix s assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/{}\<^esub> P (Call p) Q,A" assume P: "s \ P" from cvalidt_termD [OF valid ctxt P] have "\\c \ Normal s". thus "\\mark_guards f c \ Normal s" by (rule terminates_to_terminates_mark_guards) qed lemma MarkGuardsI: assumes deriv: "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c Q,A" shows "\,\\\<^sub>t\<^bsub>/{}\<^esub> P mark_guards f c Q,A" apply (rule hoaret_complete') apply (rule MarkGuardsI_sound) apply (iprover intro: hoaret_sound [OF deriv]) done lemma MarkGuardsD_sound: assumes valid: "\,\\\<^sub>t\<^bsub>/{}\<^esub> P mark_guards f c Q,A" shows "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/{}\<^esub> P (Call p) Q,A" assume exec: "\\\c,Normal s\ \ t" assume P: "s \ P" assume t_noFault: "t \ Fault ` {}" show "t \ Normal ` Q \ Abrupt ` A" proof (cases "isFault t") case True with exec_to_exec_mark_guards_Fault exec obtain f' where "\\\mark_guards f c,Normal s\ \ Fault f'" by (fastforce elim: isFaultE) from cvalidt_postD [OF valid [rule_format] ctxt this P] have False by auto thus ?thesis .. next case False from exec_to_exec_mark_guards [OF exec False] obtain f' where "\\\mark_guards f c,Normal s\ \ t" by auto from cvalidt_postD [OF valid [rule_format] ctxt this P] show ?thesis by auto qed next fix s assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/{}\<^esub> P (Call p) Q,A" assume P: "s \ P" from cvalidt_termD [OF valid ctxt P] have "\\mark_guards f c \ Normal s". thus "\\c \ Normal s" by (rule terminates_mark_guards_to_terminates) qed lemma MarkGuardsD: assumes deriv: "\,\\\<^sub>t\<^bsub>/{}\<^esub> P mark_guards f c Q,A" shows "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c Q,A" apply (rule hoaret_complete') apply (rule MarkGuardsD_sound) apply (iprover intro: hoaret_sound [OF deriv]) done lemma MergeGuardsI_sound: assumes valid: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P merge_guards c Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec_merge: "\\\merge_guards c,Normal s\ \ t" from exec_merge_guards_to_exec [OF exec_merge] have exec: "\\\c,Normal s\ \ t" . assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" from cvalidt_postD [OF valid [rule_format] ctxt exec P t_notin_F] show "t \ Normal ` Q \ Abrupt ` A". next fix s assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume P: "s \ P" from cvalidt_termD [OF valid ctxt P] have "\\c \ Normal s". thus "\\merge_guards c \ Normal s" by (rule terminates_to_terminates_merge_guards) qed lemma MergeGuardsI: assumes deriv: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P merge_guards c Q,A" apply (rule hoaret_complete') apply (rule MergeGuardsI_sound) apply (iprover intro: hoaret_sound [OF deriv]) done lemma MergeGuardsD_sound: assumes valid: "\,\\\<^sub>t\<^bsub>/F\<^esub> P merge_guards c Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\c,Normal s\ \ t" from exec_to_exec_merge_guards [OF exec] have exec_merge: "\\\merge_guards c,Normal s\ \ t". assume P: "s \ P" assume t_notin_F: "t \ Fault ` F" from cvalidt_postD [OF valid [rule_format] ctxt exec_merge P t_notin_F] show "t \ Normal ` Q \ Abrupt ` A". next fix s assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume P: "s \ P" from cvalidt_termD [OF valid ctxt P] have "\\merge_guards c \ Normal s". thus "\\c \ Normal s" by (rule terminates_merge_guards_to_terminates) qed lemma MergeGuardsD: assumes deriv: "\,\\\<^sub>t\<^bsub>/F\<^esub> P merge_guards c Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" apply (rule hoaret_complete') apply (rule MergeGuardsD_sound) apply (iprover intro: hoaret_sound [OF deriv]) done lemma SubsetGuards_sound: assumes c_c': "c \\<^sub>g c'" assumes valid: "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c' Q,A" shows "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/{}\<^esub> P (Call p) Q,A" assume exec: "\\\c,Normal s\ \ t" from exec_to_exec_subseteq_guards [OF c_c' exec] obtain t' where exec_c': "\\\c',Normal s\ \ t'" and t'_noFault: "\ isFault t' \ t' = t" by blast assume P: "s \ P" assume t_noFault: "t \ Fault ` {}" from cvalidt_postD [OF valid [rule_format] ctxt exec_c' P] t'_noFault t_noFault show "t \ Normal ` Q \ Abrupt ` A" by auto next fix s assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/{}\<^esub> P (Call p) Q,A" assume P: "s \ P" from cvalidt_termD [OF valid ctxt P] have termi_c': "\\c' \ Normal s". from cvalidt_postD [OF valid ctxt _ P] have noFault_c': "\\\c',Normal s\ \\Fault ` UNIV" by (auto simp add: final_notin_def) from termi_c' c_c' noFault_c' show "\\c \ Normal s" by (rule terminates_fewer_guards) qed lemma SubsetGuards: assumes c_c': "c \\<^sub>g c'" assumes deriv: "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c' Q,A" shows "\,\\\<^sub>t\<^bsub>/{}\<^esub> P c Q,A" apply (rule hoaret_complete') apply (rule SubsetGuards_sound [OF c_c']) apply (iprover intro: hoaret_sound [OF deriv]) done lemma NormalizeD_sound: assumes valid: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (normalize c) Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume exec: "\\\c,Normal s\ \ t" hence exec_norm: "\\\normalize c,Normal s\ \ t" by (rule exec_to_exec_normalize) assume P: "s \ P" assume noFault: "t \ Fault ` F" from cvalidt_postD [OF valid [rule_format] ctxt exec_norm P noFault] show "t \ Normal ` Q \ Abrupt ` A". next fix s assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume P: "s \ P" from cvalidt_termD [OF valid ctxt P] have "\\normalize c \ Normal s". thus "\\c \ Normal s" by (rule terminates_normalize_to_terminates) qed lemma NormalizeD: assumes deriv: "\,\\\<^sub>t\<^bsub>/F\<^esub> P (normalize c) Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" apply (rule hoaret_complete') apply (rule NormalizeD_sound) apply (iprover intro: hoaret_sound [OF deriv]) done lemma NormalizeI_sound: assumes valid: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (normalize c) Q,A" proof (rule cvalidtI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume "\\\normalize c,Normal s\ \ t" hence exec: "\\\c,Normal s\ \ t" by (rule exec_normalize_to_exec) assume P: "s \ P" assume noFault: "t \ Fault ` F" from cvalidt_postD [OF valid [rule_format] ctxt exec P noFault] show "t \ Normal ` Q \ Abrupt ` A". next fix s assume ctxt: "\(P, p, Q, A)\\. \\\<^sub>t\<^bsub>/F\<^esub> P (Call p) Q,A" assume P: "s \ P" from cvalidt_termD [OF valid ctxt P] have "\\ c \ Normal s". thus "\\normalize c \ Normal s" by (rule terminates_to_terminates_normalize) qed lemma NormalizeI: assumes deriv: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P (normalize c) Q,A" apply (rule hoaret_complete') apply (rule NormalizeI_sound) apply (iprover intro: hoaret_sound [OF deriv]) done subsubsection \Restricting the Procedure Environment\ lemma validt_restrict_to_validt: assumes validt_c: "\|\<^bsub>M\<^esub>\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" shows "\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" proof - from validt_c have valid_c: "\|\<^bsub>M\<^esub>\\<^bsub>/F\<^esub> P c Q,A" by (simp add: validt_def) hence "\\\<^bsub>/F\<^esub> P c Q,A" by (rule valid_restrict_to_valid) moreover { fix s assume P: "s \ P" have "\\c\Normal s" proof - from P validt_c have "\|\<^bsub>M\<^esub>\c\Normal s" by (auto simp add: validt_def) moreover from P valid_c have "\|\<^bsub>M\<^esub>\\c,Normal s\ \\{Stuck}" by (auto simp add: valid_def final_notin_def) ultimately show ?thesis by (rule terminates_restrict_to_terminates) qed } ultimately show ?thesis by (auto simp add: validt_def) qed lemma augment_procs: assumes deriv_c: "\|\<^bsub>M\<^esub>,{}\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" shows "\,{}\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" apply (rule hoaret_complete) apply (rule validt_restrict_to_validt) apply (insert hoaret_sound [OF deriv_c]) by (simp add: cvalidt_def) subsubsection \Miscellaneous\ lemma augment_Faults: assumes deriv_c: "\,{}\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" assumes F: "F \ F'" shows "\,{}\\<^sub>t\<^bsub>/F'\<^esub> P c Q,A" apply (rule hoaret_complete) apply (rule validt_augment_Faults [OF _ F]) apply (insert hoaret_sound [OF deriv_c]) by (simp add: cvalidt_def) lemma TerminationPartial_sound: assumes "termination": "\s \ P. \\c\Normal s" assumes partial_corr: "\,\\\<^bsub>/F\<^esub> P c Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" using "termination" partial_corr by (auto simp add: cvalidt_def validt_def cvalid_def) lemma TerminationPartial: assumes partial_deriv: "\,\\\<^bsub>/F\<^esub> P c Q,A" assumes "termination": "\s \ P. \\c\Normal s" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" apply (rule hoaret_complete') apply (rule TerminationPartial_sound [OF "termination"]) apply (rule hoare_sound [OF partial_deriv]) done lemma TerminationPartialStrip: assumes partial_deriv: "\,\\\<^bsub>/F\<^esub> P c Q,A" assumes "termination": "\s \ P. strip F' \\strip_guards F' c\Normal s" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" proof - from "termination" have "\s \ P. \\c\Normal s" by (auto intro: terminates_strip_guards_to_terminates terminates_strip_to_terminates) with partial_deriv show ?thesis by (rule TerminationPartial) qed lemma SplitTotalPartial: assumes termi: "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q',A'" assumes part: "\,\\\<^bsub>/F\<^esub> P c Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" proof - from hoaret_sound [OF termi] hoare_sound [OF part] have "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (fastforce simp add: cvalidt_def validt_def cvalid_def valid_def) thus ?thesis by (rule hoaret_complete') qed lemma SplitTotalPartial': assumes termi: "\,\\\<^sub>t\<^bsub>/UNIV\<^esub> P c Q',A'" assumes part: "\,\\\<^bsub>/F\<^esub> P c Q,A" shows "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" proof - from hoaret_sound [OF termi] hoare_sound [OF part] have "\,\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" by (fastforce simp add: cvalidt_def validt_def cvalid_def valid_def) thus ?thesis by (rule hoaret_complete') qed end diff --git a/thys/Simpl/Language.thy b/thys/Simpl/Language.thy --- a/thys/Simpl/Language.thy +++ b/thys/Simpl/Language.thy @@ -1,1150 +1,1242 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: Language.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (c) 2022 Apple Inc. All rights reserved. *) section \The Simpl Syntax\ theory Language imports "HOL-Library.Old_Recdef" begin subsection \The Core Language\ text \We use a shallow embedding of boolean expressions as well as assertions as sets of states. \ type_synonym 's bexp = "'s set" type_synonym 's assn = "'s set" datatype (dead 's, 'p, 'f) com = Skip | Basic "'s \ 's" | Spec "('s \ 's) set" | Seq "('s ,'p, 'f) com" "('s,'p, 'f) com" | Cond "'s bexp" "('s,'p,'f) com" "('s,'p,'f) com" | While "'s bexp" "('s,'p,'f) com" | Call "'p" | DynCom "'s \ ('s,'p,'f) com" | Guard "'f" "'s bexp" "('s,'p,'f) com" | Throw | Catch "('s,'p,'f) com" "('s,'p,'f) com" subsection \Derived Language Constructs\ definition raise:: "('s \ 's) \ ('s,'p,'f) com" where "raise f = Seq (Basic f) Throw" definition condCatch:: "('s,'p,'f) com \ 's bexp \ ('s,'p,'f) com \ ('s,'p,'f) com" where "condCatch c\<^sub>1 b c\<^sub>2 = Catch c\<^sub>1 (Cond b c\<^sub>2 Throw)" definition bind:: "('s \ 'v) \ ('v \ ('s,'p,'f) com) \ ('s,'p,'f) com" where "bind e c = DynCom (\s. c (e s))" definition bseq:: "('s,'p,'f) com \ ('s,'p,'f) com \ ('s,'p,'f) com" where "bseq = Seq" definition + block_exn:: "['s\'s,('s,'p,'f) com,'s\'s\'s,'s\'s\'s,'s\'s\('s,'p,'f) com]\('s,'p,'f) com" +where + "block_exn init bdy return result_exn c = + DynCom (\s. (Seq (Catch (Seq (Basic init) bdy) (Seq (Basic (\t. result_exn (return s t) t)) Throw)) + (DynCom (\t. Seq (Basic (return s)) (c s t)))) + )" + +definition + call_exn:: "('s\'s) \ 'p \ ('s \ 's \ 's)\ ('s \ 's \ 's) \('s\'s\('s,'p,'f) com)\('s,'p,'f)com" where + "call_exn init p return result_exn c = block_exn init (Call p) return result_exn c" + +primrec guards:: "('f \ 's set ) list \ ('s,'p,'f) com \ ('s,'p,'f) com" +where +"guards [] c = c" | +"guards (g#gs) c = Guard (fst g) (snd g) (guards gs c)" + + +definition maybe_guard:: "'f \ 's set \ ('s,'p,'f) com \ ('s,'p,'f) com" +where +"maybe_guard f g c = (if g = UNIV then c else Guard f g c)" + +lemma maybe_guard_UNIV [simp]: "maybe_guard f UNIV c = c" + by (simp add: maybe_guard_def) + + +definition + dynCall_exn:: "'f \ 's set \ ('s \ 's) \ ('s \ 'p) \ + ('s \ 's \ 's) \ ('s \ 's \ 's) \ ('s \ 's \ ('s,'p,'f) com) \ ('s,'p,'f) com" where + "dynCall_exn f g init p return result_exn c = + maybe_guard f g (DynCom (\s. call_exn init (p s) return result_exn c))" + +definition block:: "['s\'s,('s,'p,'f) com,'s\'s\'s,'s\'s\('s,'p,'f) com]\('s,'p,'f) com" where - "block init bdy return c = - DynCom (\s. (Seq (Catch (Seq (Basic init) bdy) (Seq (Basic (return s)) Throw)) - (DynCom (\t. Seq (Basic (return s)) (c s t)))) - )" + "block init bdy return c = block_exn init bdy return (\s t. s) c" + definition call:: "('s\'s) \ 'p \ ('s \ 's \ 's)\('s\'s\('s,'p,'f) com)\('s,'p,'f)com" where "call init p return c = block init (Call p) return c" definition dynCall:: "('s \ 's) \ ('s \ 'p) \ ('s \ 's \ 's) \ ('s \ 's \ ('s,'p,'f) com) \ ('s,'p,'f) com" where "dynCall init p return c = DynCom (\s. call init (p s) return c)" + + definition fcall:: "('s\'s) \ 'p \ ('s \ 's \ 's)\('s \ 'v) \ ('v\('s,'p,'f) com) \('s,'p,'f)com" where "fcall init p return result c = call init p return (\s t. c (result t))" + definition lem:: "'x \ ('s,'p,'f)com \('s,'p,'f)com" where "lem x c = c" primrec switch:: "('s \ 'v) \ ('v set \ ('s,'p,'f) com) list \ ('s,'p,'f) com" where "switch v [] = Skip" | "switch v (Vc#vs) = Cond {s. v s \ fst Vc} (snd Vc) (switch v vs)" definition guaranteeStrip:: "'f \ 's set \ ('s,'p,'f) com \ ('s,'p,'f) com" where "guaranteeStrip f g c = Guard f g c" definition guaranteeStripPair:: "'f \ 's set \ ('f \ 's set)" where "guaranteeStripPair f g = (f,g)" -primrec guards:: "('f \ 's set ) list \ ('s,'p,'f) com \ ('s,'p,'f) com" -where -"guards [] c = c" | -"guards (g#gs) c = Guard (fst g) (snd g) (guards gs c)" definition while:: "('f \ 's set) list \ 's bexp \ ('s,'p,'f) com \ ('s, 'p, 'f) com" where "while gs b c = guards gs (While b (Seq c (guards gs Skip)))" definition whileAnno:: "'s bexp \ 's assn \ ('s \ 's) assn \ ('s,'p,'f) com \ ('s,'p,'f) com" where "whileAnno b I V c = While b c" definition whileAnnoG:: "('f \ 's set) list \ 's bexp \ 's assn \ ('s \ 's) assn \ ('s,'p,'f) com \ ('s,'p,'f) com" where "whileAnnoG gs b I V c = while gs b c" definition specAnno:: "('a \ 's assn) \ ('a \ ('s,'p,'f) com) \ ('a \ 's assn) \ ('a \ 's assn) \ ('s,'p,'f) com" where "specAnno P c Q A = (c undefined)" definition whileAnnoFix:: "'s bexp \ ('a \ 's assn) \ ('a \ ('s \ 's) assn) \ ('a \ ('s,'p,'f) com) \ ('s,'p,'f) com" where "whileAnnoFix b I V c = While b (c undefined)" definition whileAnnoGFix:: "('f \ 's set) list \ 's bexp \ ('a \ 's assn) \ ('a \ ('s \ 's) assn) \ ('a \ ('s,'p,'f) com) \ ('s,'p,'f) com" where "whileAnnoGFix gs b I V c = while gs b (c undefined)" definition if_rel::"('s \ bool) \ ('s \ 's) \ ('s \ 's) \ ('s \ 's) \ ('s \ 's) set" where "if_rel b f g h = {(s,t). if b s then t = f s else t = g s \ t = h s}" lemma fst_guaranteeStripPair: "fst (guaranteeStripPair f g) = f" by (simp add: guaranteeStripPair_def) lemma snd_guaranteeStripPair: "snd (guaranteeStripPair f g) = g" by (simp add: guaranteeStripPair_def) +lemma call_call_exn: "call init p return result = call_exn init p return (\s t. s) result" + by (simp add: call_def call_exn_def block_def) + +lemma dynCall_dynCall_exn: "dynCall init p return result = dynCall_exn undefined UNIV init p return (\s t. s) result" + by (simp add: dynCall_def dynCall_exn_def call_call_exn) + + subsection \Operations on Simpl-Syntax\ subsubsection \Normalisation of Sequential Composition: \sequence\, \flatten\ and \normalize\\ primrec flatten:: "('s,'p,'f) com \ ('s,'p,'f) com list" where "flatten Skip = [Skip]" | "flatten (Basic f) = [Basic f]" | "flatten (Spec r) = [Spec r]" | "flatten (Seq c\<^sub>1 c\<^sub>2) = flatten c\<^sub>1 @ flatten c\<^sub>2" | "flatten (Cond b c\<^sub>1 c\<^sub>2) = [Cond b c\<^sub>1 c\<^sub>2]" | "flatten (While b c) = [While b c]" | "flatten (Call p) = [Call p]" | "flatten (DynCom c) = [DynCom c]" | "flatten (Guard f g c) = [Guard f g c]" | "flatten Throw = [Throw]" | "flatten (Catch c\<^sub>1 c\<^sub>2) = [Catch c\<^sub>1 c\<^sub>2]" primrec sequence:: "(('s,'p,'f) com \ ('s,'p,'f) com \ ('s,'p,'f) com) \ ('s,'p,'f) com list \ ('s,'p,'f) com" where "sequence seq [] = Skip" | "sequence seq (c#cs) = (case cs of [] \ c | _ \ seq c (sequence seq cs))" primrec normalize:: "('s,'p,'f) com \ ('s,'p,'f) com" where "normalize Skip = Skip" | "normalize (Basic f) = Basic f" | "normalize (Spec r) = Spec r" | "normalize (Seq c\<^sub>1 c\<^sub>2) = sequence Seq ((flatten (normalize c\<^sub>1)) @ (flatten (normalize c\<^sub>2)))" | "normalize (Cond b c\<^sub>1 c\<^sub>2) = Cond b (normalize c\<^sub>1) (normalize c\<^sub>2)" | "normalize (While b c) = While b (normalize c)" | "normalize (Call p) = Call p" | "normalize (DynCom c) = DynCom (\s. (normalize (c s)))" | "normalize (Guard f g c) = Guard f g (normalize c)" | "normalize Throw = Throw" | "normalize (Catch c\<^sub>1 c\<^sub>2) = Catch (normalize c\<^sub>1) (normalize c\<^sub>2)" lemma flatten_nonEmpty: "flatten c \ []" by (induct c) simp_all lemma flatten_single: "\c \ set (flatten c'). flatten c = [c]" apply (induct c') apply simp apply simp apply simp apply (simp (no_asm_use) ) apply blast apply (simp (no_asm_use) ) apply (simp (no_asm_use) ) apply simp apply (simp (no_asm_use)) apply (simp (no_asm_use)) apply simp apply (simp (no_asm_use)) done lemma flatten_sequence_id: "\cs\[];\c \ set cs. flatten c = [c]\ \ flatten (sequence Seq cs) = cs" apply (induct cs) apply simp apply (case_tac cs) apply simp apply auto done lemma flatten_app: "flatten (sequence Seq (flatten c1 @ flatten c2)) = flatten c1 @ flatten c2" apply (rule flatten_sequence_id) apply (simp add: flatten_nonEmpty) apply (simp) apply (insert flatten_single) apply blast done lemma flatten_sequence_flatten: "flatten (sequence Seq (flatten c)) = flatten c" apply (induct c) apply (auto simp add: flatten_app) done lemma sequence_flatten_normalize: "sequence Seq (flatten (normalize c)) = normalize c" apply (induct c) apply (auto simp add: flatten_app) done lemma flatten_normalize: "\x xs. flatten (normalize c) = x#xs \ (case xs of [] \ normalize c = x | (x'#xs') \ normalize c= Seq x (sequence Seq xs))" proof (induct c) case (Seq c1 c2) have "flatten (normalize (Seq c1 c2)) = x # xs" by fact hence "flatten (sequence Seq (flatten (normalize c1) @ flatten (normalize c2))) = x#xs" by simp hence x_xs: "flatten (normalize c1) @ flatten (normalize c2) = x # xs" by (simp add: flatten_app) show ?case proof (cases "flatten (normalize c1)") case Nil with flatten_nonEmpty show ?thesis by auto next case (Cons x1 xs1) note Cons_x1_xs1 = this with x_xs obtain x_x1: "x=x1" and xs_rest: "xs=xs1@flatten (normalize c2)" by auto show ?thesis proof (cases xs1) case Nil from Seq.hyps (1) [OF Cons_x1_xs1] Nil have "normalize c1 = x1" by simp with Cons_x1_xs1 Nil x_x1 xs_rest show ?thesis apply (cases "flatten (normalize c2)") apply (fastforce simp add: flatten_nonEmpty) apply simp done next case Cons from Seq.hyps (1) [OF Cons_x1_xs1] Cons have "normalize c1 = Seq x1 (sequence Seq xs1)" by simp with Cons_x1_xs1 Nil x_x1 xs_rest show ?thesis apply (cases "flatten (normalize c2)") apply (fastforce simp add: flatten_nonEmpty) apply (simp split: list.splits) done qed qed qed (auto) lemma flatten_raise [simp]: "flatten (raise f) = [Basic f, Throw]" by (simp add: raise_def) lemma flatten_condCatch [simp]: "flatten (condCatch c1 b c2) = [condCatch c1 b c2]" by (simp add: condCatch_def) lemma flatten_bind [simp]: "flatten (bind e c) = [bind e c]" by (simp add: bind_def) lemma flatten_bseq [simp]: "flatten (bseq c1 c2) = flatten c1 @ flatten c2" by (simp add: bseq_def) +lemma flatten_block_exn [simp]: + "flatten (block_exn init bdy return result_exn result) = [block_exn init bdy return result_exn result]" + by (simp add: block_exn_def) + lemma flatten_block [simp]: "flatten (block init bdy return result) = [block init bdy return result]" by (simp add: block_def) lemma flatten_call [simp]: "flatten (call init p return result) = [call init p return result]" by (simp add: call_def) lemma flatten_dynCall [simp]: "flatten (dynCall init p return result) = [dynCall init p return result]" by (simp add: dynCall_def) +lemma flatten_call_exn [simp]: "flatten (call_exn init p return result_exn result) = [call_exn init p return result_exn result]" + by (simp add: call_exn_def) + +lemma flatten_dynCall_exn [simp]: "flatten (dynCall_exn f g init p return result_exn result) = [dynCall_exn f g init p return result_exn result]" + by (simp add: dynCall_exn_def maybe_guard_def) + lemma flatten_fcall [simp]: "flatten (fcall init p return result c) = [fcall init p return result c]" by (simp add: fcall_def) lemma flatten_switch [simp]: "flatten (switch v Vcs) = [switch v Vcs]" by (cases Vcs) auto lemma flatten_guaranteeStrip [simp]: "flatten (guaranteeStrip f g c) = [guaranteeStrip f g c]" by (simp add: guaranteeStrip_def) lemma flatten_while [simp]: "flatten (while gs b c) = [while gs b c]" apply (simp add: while_def) apply (induct gs) apply auto done lemma flatten_whileAnno [simp]: "flatten (whileAnno b I V c) = [whileAnno b I V c]" by (simp add: whileAnno_def) lemma flatten_whileAnnoG [simp]: "flatten (whileAnnoG gs b I V c) = [whileAnnoG gs b I V c]" by (simp add: whileAnnoG_def) lemma flatten_specAnno [simp]: "flatten (specAnno P c Q A) = flatten (c undefined)" by (simp add: specAnno_def) lemmas flatten_simps = flatten.simps flatten_raise flatten_condCatch flatten_bind flatten_block flatten_call flatten_dynCall flatten_fcall flatten_switch flatten_guaranteeStrip flatten_while flatten_whileAnno flatten_whileAnnoG flatten_specAnno lemma normalize_raise [simp]: "normalize (raise f) = raise f" by (simp add: raise_def) lemma normalize_condCatch [simp]: "normalize (condCatch c1 b c2) = condCatch (normalize c1) b (normalize c2)" by (simp add: condCatch_def) lemma normalize_bind [simp]: "normalize (bind e c) = bind e (\v. normalize (c v))" by (simp add: bind_def) lemma normalize_bseq [simp]: "normalize (bseq c1 c2) = sequence bseq ((flatten (normalize c1)) @ (flatten (normalize c2)))" by (simp add: bseq_def) -lemma normalize_block [simp]: "normalize (block init bdy return c) = - block init (normalize bdy) return (\s t. normalize (c s t))" - apply (simp add: block_def) +lemma normalize_block_exn [simp]: "normalize (block_exn init bdy return result_exn c) = + block_exn init (normalize bdy) return result_exn (\s t. normalize (c s t))" + apply (simp add: block_exn_def) apply (rule ext) apply (simp) apply (cases "flatten (normalize bdy)") apply (simp add: flatten_nonEmpty) apply (rule conjI) apply simp apply (drule flatten_normalize) apply (case_tac list) apply simp apply simp apply (rule ext) apply (case_tac "flatten (normalize (c s sa))") apply (simp add: flatten_nonEmpty) apply simp apply (thin_tac "flatten (normalize bdy) = P" for P) apply (drule flatten_normalize) apply (case_tac lista) apply simp apply simp done +lemma normalize_block [simp]: "normalize (block init bdy return c) = + block init (normalize bdy) return (\s t. normalize (c s t))" + by (simp add: block_def) + lemma normalize_call [simp]: "normalize (call init p return c) = call init p return (\i t. normalize (c i t))" by (simp add: call_def) +lemma normalize_call_exn [simp]: + "normalize (call_exn init p return result_exn c) = call_exn init p return result_exn (\i t. normalize (c i t))" + by (simp add: call_exn_def) + lemma normalize_dynCall [simp]: "normalize (dynCall init p return c) = dynCall init p return (\s t. normalize (c s t))" by (simp add: dynCall_def) +lemma normalize_guards [simp]: + "normalize (guards gs c) = guards gs (normalize c)" + by (induct gs) auto + +lemma normalize_dynCall_exn [simp]: + "normalize (dynCall_exn f g init p return result_exn c) = + dynCall_exn f g init p return result_exn (\s t. normalize (c s t))" + by (simp add: dynCall_exn_def maybe_guard_def) + lemma normalize_fcall [simp]: "normalize (fcall init p return result c) = fcall init p return result (\v. normalize (c v))" by (simp add: fcall_def) lemma normalize_switch [simp]: "normalize (switch v Vcs) = switch v (map (\(V,c). (V,normalize c)) Vcs)" apply (induct Vcs) apply auto done lemma normalize_guaranteeStrip [simp]: "normalize (guaranteeStrip f g c) = guaranteeStrip f g (normalize c)" by (simp add: guaranteeStrip_def) -lemma normalize_guards [simp]: - "normalize (guards gs c) = guards gs (normalize c)" - by (induct gs) auto text \Sequencial composition with guards in the body is not preserved by normalize\ lemma normalize_while [simp]: "normalize (while gs b c) = guards gs (While b (sequence Seq (flatten (normalize c) @ flatten (guards gs Skip))))" by (simp add: while_def) lemma normalize_whileAnno [simp]: "normalize (whileAnno b I V c) = whileAnno b I V (normalize c)" by (simp add: whileAnno_def) lemma normalize_whileAnnoG [simp]: "normalize (whileAnnoG gs b I V c) = guards gs (While b (sequence Seq (flatten (normalize c) @ flatten (guards gs Skip))))" by (simp add: whileAnnoG_def) lemma normalize_specAnno [simp]: "normalize (specAnno P c Q A) = specAnno P (\s. normalize (c undefined)) Q A" by (simp add: specAnno_def) lemmas normalize_simps = normalize.simps normalize_raise normalize_condCatch normalize_bind normalize_block normalize_call normalize_dynCall normalize_fcall normalize_switch normalize_guaranteeStrip normalize_guards normalize_while normalize_whileAnno normalize_whileAnnoG normalize_specAnno subsubsection \Stripping Guards: \strip_guards\\ primrec strip_guards:: "'f set \ ('s,'p,'f) com \ ('s,'p,'f) com" where "strip_guards F Skip = Skip" | "strip_guards F (Basic f) = Basic f" | "strip_guards F (Spec r) = Spec r" | "strip_guards F (Seq c\<^sub>1 c\<^sub>2) = (Seq (strip_guards F c\<^sub>1) (strip_guards F c\<^sub>2))" | "strip_guards F (Cond b c\<^sub>1 c\<^sub>2) = Cond b (strip_guards F c\<^sub>1) (strip_guards F c\<^sub>2)" | "strip_guards F (While b c) = While b (strip_guards F c)" | "strip_guards F (Call p) = Call p" | "strip_guards F (DynCom c) = DynCom (\s. (strip_guards F (c s)))" | "strip_guards F (Guard f g c) = (if f \ F then strip_guards F c else Guard f g (strip_guards F c))" | "strip_guards F Throw = Throw" | "strip_guards F (Catch c\<^sub>1 c\<^sub>2) = Catch (strip_guards F c\<^sub>1) (strip_guards F c\<^sub>2)" definition strip:: "'f set \ ('p \ ('s,'p,'f) com option) \ ('p \ ('s,'p,'f) com option)" where "strip F \ = (\p. map_option (strip_guards F) (\ p))" lemma strip_simp [simp]: "(strip F \) p = map_option (strip_guards F) (\ p)" by (simp add: strip_def) lemma dom_strip: "dom (strip F \) = dom \" by (auto) lemma strip_guards_idem: "strip_guards F (strip_guards F c) = strip_guards F c" by (induct c) auto lemma strip_idem: "strip F (strip F \) = strip F \" apply (rule ext) apply (case_tac "\ x") apply (auto simp add: strip_guards_idem strip_def) done lemma strip_guards_raise [simp]: "strip_guards F (raise f) = raise f" by (simp add: raise_def) lemma strip_guards_condCatch [simp]: "strip_guards F (condCatch c1 b c2) = condCatch (strip_guards F c1) b (strip_guards F c2)" by (simp add: condCatch_def) lemma strip_guards_bind [simp]: "strip_guards F (bind e c) = bind e (\v. strip_guards F (c v))" by (simp add: bind_def) lemma strip_guards_bseq [simp]: "strip_guards F (bseq c1 c2) = bseq (strip_guards F c1) (strip_guards F c2)" by (simp add: bseq_def) +lemma strip_guards_block_exn [simp]: + "strip_guards F (block_exn init bdy return result_exn c) = + block_exn init (strip_guards F bdy) return result_exn (\s t. strip_guards F (c s t))" + by (simp add: block_exn_def) + lemma strip_guards_block [simp]: "strip_guards F (block init bdy return c) = block init (strip_guards F bdy) return (\s t. strip_guards F (c s t))" by (simp add: block_def) lemma strip_guards_call [simp]: "strip_guards F (call init p return c) = call init p return (\s t. strip_guards F (c s t))" by (simp add: call_def) +lemma strip_guards_call_exn [simp]: + "strip_guards F (call_exn init p return result_exn c) = + call_exn init p return result_exn (\s t. strip_guards F (c s t))" + by (simp add: call_exn_def) + lemma strip_guards_dynCall [simp]: "strip_guards F (dynCall init p return c) = dynCall init p return (\s t. strip_guards F (c s t))" by (simp add: dynCall_def) +lemma strip_guards_guards [simp]: "strip_guards F (guards gs c) = + guards (filter (\(f,g). f \ F) gs) (strip_guards F c)" + by (induct gs) auto + +lemma strip_guards_dynCall_exn [simp]: + "strip_guards F (dynCall_exn f g init p return result_exn c) = + dynCall_exn f (if f \ F then UNIV else g) init p return result_exn (\s t. strip_guards F (c s t))" + by (simp add: dynCall_exn_def maybe_guard_def) + lemma strip_guards_fcall [simp]: "strip_guards F (fcall init p return result c) = fcall init p return result (\v. strip_guards F (c v))" by (simp add: fcall_def) lemma strip_guards_switch [simp]: "strip_guards F (switch v Vc) = switch v (map (\(V,c). (V,strip_guards F c)) Vc)" by (induct Vc) auto lemma strip_guards_guaranteeStrip [simp]: "strip_guards F (guaranteeStrip f g c) = (if f \ F then strip_guards F c else guaranteeStrip f g (strip_guards F c))" by (simp add: guaranteeStrip_def) lemma guaranteeStripPair_split_conv [simp]: "case_prod c (guaranteeStripPair f g) = c f g" by (simp add: guaranteeStripPair_def) -lemma strip_guards_guards [simp]: "strip_guards F (guards gs c) = - guards (filter (\(f,g). f \ F) gs) (strip_guards F c)" - by (induct gs) auto + lemma strip_guards_while [simp]: "strip_guards F (while gs b c) = while (filter (\(f,g). f \ F) gs) b (strip_guards F c)" by (simp add: while_def) lemma strip_guards_whileAnno [simp]: "strip_guards F (whileAnno b I V c) = whileAnno b I V (strip_guards F c)" by (simp add: whileAnno_def while_def) lemma strip_guards_whileAnnoG [simp]: "strip_guards F (whileAnnoG gs b I V c) = whileAnnoG (filter (\(f,g). f \ F) gs) b I V (strip_guards F c)" by (simp add: whileAnnoG_def) lemma strip_guards_specAnno [simp]: "strip_guards F (specAnno P c Q A) = specAnno P (\s. strip_guards F (c undefined)) Q A" by (simp add: specAnno_def) lemmas strip_guards_simps = strip_guards.simps strip_guards_raise strip_guards_condCatch strip_guards_bind strip_guards_bseq strip_guards_block strip_guards_dynCall strip_guards_fcall strip_guards_switch strip_guards_guaranteeStrip guaranteeStripPair_split_conv strip_guards_guards strip_guards_while strip_guards_whileAnno strip_guards_whileAnnoG strip_guards_specAnno subsubsection \Marking Guards: \mark_guards\\ primrec mark_guards:: "'f \ ('s,'p,'g) com \ ('s,'p,'f) com" where "mark_guards f Skip = Skip" | "mark_guards f (Basic g) = Basic g" | "mark_guards f (Spec r) = Spec r" | "mark_guards f (Seq c\<^sub>1 c\<^sub>2) = (Seq (mark_guards f c\<^sub>1) (mark_guards f c\<^sub>2))" | "mark_guards f (Cond b c\<^sub>1 c\<^sub>2) = Cond b (mark_guards f c\<^sub>1) (mark_guards f c\<^sub>2)" | "mark_guards f (While b c) = While b (mark_guards f c)" | "mark_guards f (Call p) = Call p" | "mark_guards f (DynCom c) = DynCom (\s. (mark_guards f (c s)))" | "mark_guards f (Guard f' g c) = Guard f g (mark_guards f c)" | "mark_guards f Throw = Throw" | "mark_guards f (Catch c\<^sub>1 c\<^sub>2) = Catch (mark_guards f c\<^sub>1) (mark_guards f c\<^sub>2)" lemma mark_guards_raise: "mark_guards f (raise g) = raise g" by (simp add: raise_def) lemma mark_guards_condCatch [simp]: "mark_guards f (condCatch c1 b c2) = condCatch (mark_guards f c1) b (mark_guards f c2)" by (simp add: condCatch_def) lemma mark_guards_bind [simp]: "mark_guards f (bind e c) = bind e (\v. mark_guards f (c v))" by (simp add: bind_def) lemma mark_guards_bseq [simp]: "mark_guards f (bseq c1 c2) = bseq (mark_guards f c1) (mark_guards f c2)" by (simp add: bseq_def) +lemma mark_guards_block_exn [simp]: + "mark_guards f (block_exn init bdy return result_exn c) = + block_exn init (mark_guards f bdy) return result_exn (\s t. mark_guards f (c s t))" + by (simp add: block_exn_def) + lemma mark_guards_block [simp]: "mark_guards f (block init bdy return c) = block init (mark_guards f bdy) return (\s t. mark_guards f (c s t))" by (simp add: block_def) lemma mark_guards_call [simp]: "mark_guards f (call init p return c) = call init p return (\s t. mark_guards f (c s t))" by (simp add: call_def) +lemma mark_guards_call_exn [simp]: + "mark_guards f (call_exn init p return result_exn c) = + call_exn init p return result_exn (\s t. mark_guards f (c s t))" + by (simp add: call_exn_def) + lemma mark_guards_dynCall [simp]: "mark_guards f (dynCall init p return c) = dynCall init p return (\s t. mark_guards f (c s t))" by (simp add: dynCall_def) +lemma mark_guards_guards [simp]: + "mark_guards f (guards gs c) = guards (map (\(f',g). (f,g)) gs) (mark_guards f c)" + by (induct gs) auto + +lemma mark_guards_dynCall_exn [simp]: + "mark_guards f (dynCall_exn f' g init p return result_exn c) = + dynCall_exn f g init p return result_exn (\s t. mark_guards f (c s t))" + by (simp add: dynCall_exn_def maybe_guard_def) + lemma mark_guards_fcall [simp]: "mark_guards f (fcall init p return result c) = fcall init p return result (\v. mark_guards f (c v))" by (simp add: fcall_def) lemma mark_guards_switch [simp]: "mark_guards f (switch v vs) = switch v (map (\(V,c). (V,mark_guards f c)) vs)" by (induct vs) auto lemma mark_guards_guaranteeStrip [simp]: "mark_guards f (guaranteeStrip f' g c) = guaranteeStrip f g (mark_guards f c)" by (simp add: guaranteeStrip_def) -lemma mark_guards_guards [simp]: - "mark_guards f (guards gs c) = guards (map (\(f',g). (f,g)) gs) (mark_guards f c)" - by (induct gs) auto lemma mark_guards_while [simp]: "mark_guards f (while gs b c) = while (map (\(f',g). (f,g)) gs) b (mark_guards f c)" by (simp add: while_def) lemma mark_guards_whileAnno [simp]: "mark_guards f (whileAnno b I V c) = whileAnno b I V (mark_guards f c)" by (simp add: whileAnno_def while_def) lemma mark_guards_whileAnnoG [simp]: "mark_guards f (whileAnnoG gs b I V c) = whileAnnoG (map (\(f',g). (f,g)) gs) b I V (mark_guards f c)" by (simp add: whileAnno_def whileAnnoG_def while_def) lemma mark_guards_specAnno [simp]: "mark_guards f (specAnno P c Q A) = specAnno P (\s. mark_guards f (c undefined)) Q A" by (simp add: specAnno_def) lemmas mark_guards_simps = mark_guards.simps mark_guards_raise mark_guards_condCatch mark_guards_bind mark_guards_bseq mark_guards_block mark_guards_dynCall mark_guards_fcall mark_guards_switch mark_guards_guaranteeStrip guaranteeStripPair_split_conv mark_guards_guards mark_guards_while mark_guards_whileAnno mark_guards_whileAnnoG mark_guards_specAnno definition is_Guard:: "('s,'p,'f) com \ bool" where "is_Guard c = (case c of Guard f g c' \ True | _ \ False)" lemma is_Guard_basic_simps [simp]: + "is_Guard (guards (pg# pgs) c) = True" "is_Guard Skip = False" "is_Guard (Basic f) = False" "is_Guard (Spec r) = False" "is_Guard (Seq c1 c2) = False" "is_Guard (Cond b c1 c2) = False" "is_Guard (While b c) = False" "is_Guard (Call p) = False" "is_Guard (DynCom C) = False" "is_Guard (Guard F g c) = True" "is_Guard (Throw) = False" "is_Guard (Catch c1 c2) = False" "is_Guard (raise f) = False" "is_Guard (condCatch c1 b c2) = False" "is_Guard (bind e cv) = False" "is_Guard (bseq c1 c2) = False" + "is_Guard (block_exn init bdy return result_exn cont) = False" "is_Guard (block init bdy return cont) = False" "is_Guard (call init p return cont) = False" "is_Guard (dynCall init P return cont) = False" + "is_Guard (call_exn init p return result_exn cont) = False" + "is_Guard (dynCall_exn f UNIV init P return result_exn cont) = False" "is_Guard (fcall init p return result cont') = False" "is_Guard (whileAnno b I V c) = False" "is_Guard (guaranteeStrip F g c) = True" by (auto simp add: is_Guard_def raise_def condCatch_def bind_def bseq_def - block_def call_def dynCall_def fcall_def whileAnno_def guaranteeStrip_def) + block_def block_exn_def call_def dynCall_def call_exn_def dynCall_exn_def + fcall_def whileAnno_def guaranteeStrip_def) lemma is_Guard_switch [simp]: "is_Guard (switch v Vc) = False" by (induct Vc) auto lemmas is_Guard_simps = is_Guard_basic_simps is_Guard_switch primrec dest_Guard:: "('s,'p,'f) com \ ('f \ 's set \ ('s,'p,'f) com)" where "dest_Guard (Guard f g c) = (f,g,c)" lemma dest_Guard_guaranteeStrip [simp]: "dest_Guard (guaranteeStrip f g c) = (f,g,c)" by (simp add: guaranteeStrip_def) lemmas dest_Guard_simps = dest_Guard.simps dest_Guard_guaranteeStrip subsubsection \Merging Guards: \merge_guards\\ primrec merge_guards:: "('s,'p,'f) com \ ('s,'p,'f) com" where "merge_guards Skip = Skip" | "merge_guards (Basic g) = Basic g" | "merge_guards (Spec r) = Spec r" | "merge_guards (Seq c\<^sub>1 c\<^sub>2) = (Seq (merge_guards c\<^sub>1) (merge_guards c\<^sub>2))" | "merge_guards (Cond b c\<^sub>1 c\<^sub>2) = Cond b (merge_guards c\<^sub>1) (merge_guards c\<^sub>2)" | "merge_guards (While b c) = While b (merge_guards c)" | "merge_guards (Call p) = Call p" | "merge_guards (DynCom c) = DynCom (\s. (merge_guards (c s)))" | (*"merge_guards (Guard f g c) = (case (merge_guards c) of Guard f' g' c' \ if f=f' then Guard f (g \ g') c' else Guard f g (Guard f' g' c') | _ \ Guard f g (merge_guards c))"*) (* the following version works better with derived language constructs *) "merge_guards (Guard f g c) = (let c' = (merge_guards c) in if is_Guard c' then let (f',g',c'') = dest_Guard c' in if f=f' then Guard f (g \ g') c'' else Guard f g (Guard f' g' c'') else Guard f g c')" | "merge_guards Throw = Throw" | "merge_guards (Catch c\<^sub>1 c\<^sub>2) = Catch (merge_guards c\<^sub>1) (merge_guards c\<^sub>2)" lemma merge_guards_res_Skip: "merge_guards c = Skip \ c = Skip" by (cases c) (auto split: com.splits if_split_asm simp add: is_Guard_def Let_def) lemma merge_guards_res_Basic: "merge_guards c = Basic f \ c = Basic f" by (cases c) (auto split: com.splits if_split_asm simp add: is_Guard_def Let_def) lemma merge_guards_res_Spec: "merge_guards c = Spec r \ c = Spec r" by (cases c) (auto split: com.splits if_split_asm simp add: is_Guard_def Let_def) lemma merge_guards_res_Seq: "merge_guards c = Seq c1 c2 \ \c1' c2'. c = Seq c1' c2' \ merge_guards c1' = c1 \ merge_guards c2' = c2" by (cases c) (auto split: com.splits if_split_asm simp add: is_Guard_def Let_def) lemma merge_guards_res_Cond: "merge_guards c = Cond b c1 c2 \ \c1' c2'. c = Cond b c1' c2' \ merge_guards c1' = c1 \ merge_guards c2' = c2" by (cases c) (auto split: com.splits if_split_asm simp add: is_Guard_def Let_def) lemma merge_guards_res_While: "merge_guards c = While b c' \ \c''. c = While b c'' \ merge_guards c'' = c'" by (cases c) (auto split: com.splits if_split_asm simp add: is_Guard_def Let_def) lemma merge_guards_res_Call: "merge_guards c = Call p \ c = Call p" by (cases c) (auto split: com.splits if_split_asm simp add: is_Guard_def Let_def) lemma merge_guards_res_DynCom: "merge_guards c = DynCom c' \ \c''. c = DynCom c'' \ (\s. (merge_guards (c'' s))) = c'" by (cases c) (auto split: com.splits if_split_asm simp add: is_Guard_def Let_def) lemma merge_guards_res_Throw: "merge_guards c = Throw \ c = Throw" by (cases c) (auto split: com.splits if_split_asm simp add: is_Guard_def Let_def) lemma merge_guards_res_Catch: "merge_guards c = Catch c1 c2 \ \c1' c2'. c = Catch c1' c2' \ merge_guards c1' = c1 \ merge_guards c2' = c2" by (cases c) (auto split: com.splits if_split_asm simp add: is_Guard_def Let_def) lemma merge_guards_res_Guard: "merge_guards c = Guard f g c' \ \c'' f' g'. c = Guard f' g' c''" by (cases c) (auto split: com.splits if_split_asm simp add: is_Guard_def Let_def) lemmas merge_guards_res_simps = merge_guards_res_Skip merge_guards_res_Basic merge_guards_res_Spec merge_guards_res_Seq merge_guards_res_Cond merge_guards_res_While merge_guards_res_Call merge_guards_res_DynCom merge_guards_res_Throw merge_guards_res_Catch merge_guards_res_Guard +lemma merge_guards_guards_empty: "merge_guards (guards [] c) = merge_guards c" + by (simp) + lemma merge_guards_raise: "merge_guards (raise g) = raise g" by (simp add: raise_def) lemma merge_guards_condCatch [simp]: "merge_guards (condCatch c1 b c2) = condCatch (merge_guards c1) b (merge_guards c2)" by (simp add: condCatch_def) lemma merge_guards_bind [simp]: "merge_guards (bind e c) = bind e (\v. merge_guards (c v))" by (simp add: bind_def) lemma merge_guards_bseq [simp]: "merge_guards (bseq c1 c2) = bseq (merge_guards c1) (merge_guards c2)" by (simp add: bseq_def) +lemma merge_guards_block_exn [simp]: + "merge_guards (block_exn init bdy return result_exn c) = + block_exn init (merge_guards bdy) return result_exn (\s t. merge_guards (c s t))" + by (simp add: block_exn_def) + lemma merge_guards_block [simp]: "merge_guards (block init bdy return c) = block init (merge_guards bdy) return (\s t. merge_guards (c s t))" by (simp add: block_def) lemma merge_guards_call [simp]: "merge_guards (call init p return c) = call init p return (\s t. merge_guards (c s t))" by (simp add: call_def) +lemma merge_guards_call_exn [simp]: + "merge_guards (call_exn init p return result_exn c) = + call_exn init p return result_exn (\s t. merge_guards (c s t))" + by (simp add: call_exn_def) + lemma merge_guards_dynCall [simp]: "merge_guards (dynCall init p return c) = dynCall init p return (\s t. merge_guards (c s t))" by (simp add: dynCall_def) lemma merge_guards_fcall [simp]: "merge_guards (fcall init p return result c) = fcall init p return result (\v. merge_guards (c v))" by (simp add: fcall_def) lemma merge_guards_switch [simp]: "merge_guards (switch v vs) = switch v (map (\(V,c). (V,merge_guards c)) vs)" by (induct vs) auto lemma merge_guards_guaranteeStrip [simp]: "merge_guards (guaranteeStrip f g c) = (let c' = (merge_guards c) in if is_Guard c' then let (f',g',c') = dest_Guard c' in if f=f' then Guard f (g \ g') c' else Guard f g (Guard f' g' c') else Guard f g c')" by (simp add: guaranteeStrip_def) lemma merge_guards_whileAnno [simp]: "merge_guards (whileAnno b I V c) = whileAnno b I V (merge_guards c)" by (simp add: whileAnno_def while_def) lemma merge_guards_specAnno [simp]: "merge_guards (specAnno P c Q A) = specAnno P (\s. merge_guards (c undefined)) Q A" by (simp add: specAnno_def) text \@{term "merge_guards"} for guard-lists as in @{const guards}, @{const while} and @{const whileAnnoG} may have funny effects since the guard-list has to be merged with the body statement too.\ lemmas merge_guards_simps = merge_guards.simps merge_guards_raise merge_guards_condCatch merge_guards_bind merge_guards_bseq merge_guards_block merge_guards_dynCall merge_guards_fcall merge_guards_switch + merge_guards_block_exn merge_guards_call_exn merge_guards_guaranteeStrip merge_guards_whileAnno merge_guards_specAnno primrec noguards:: "('s,'p,'f) com \ bool" where "noguards Skip = True" | "noguards (Basic f) = True" | "noguards (Spec r ) = True" | "noguards (Seq c\<^sub>1 c\<^sub>2) = (noguards c\<^sub>1 \ noguards c\<^sub>2)" | "noguards (Cond b c\<^sub>1 c\<^sub>2) = (noguards c\<^sub>1 \ noguards c\<^sub>2)" | "noguards (While b c) = (noguards c)" | "noguards (Call p) = True" | "noguards (DynCom c) = (\s. noguards (c s))" | "noguards (Guard f g c) = False" | "noguards Throw = True" | "noguards (Catch c\<^sub>1 c\<^sub>2) = (noguards c\<^sub>1 \ noguards c\<^sub>2)" lemma noguards_strip_guards: "noguards (strip_guards UNIV c)" by (induct c) auto primrec nothrows:: "('s,'p,'f) com \ bool" where "nothrows Skip = True" | "nothrows (Basic f) = True" | "nothrows (Spec r) = True" | "nothrows (Seq c\<^sub>1 c\<^sub>2) = (nothrows c\<^sub>1 \ nothrows c\<^sub>2)" | "nothrows (Cond b c\<^sub>1 c\<^sub>2) = (nothrows c\<^sub>1 \ nothrows c\<^sub>2)" | "nothrows (While b c) = nothrows c" | "nothrows (Call p) = True" | "nothrows (DynCom c) = (\s. nothrows (c s))" | "nothrows (Guard f g c) = nothrows c" | "nothrows Throw = False" | "nothrows (Catch c\<^sub>1 c\<^sub>2) = (nothrows c\<^sub>1 \ nothrows c\<^sub>2)" subsubsection \Intersecting Guards: \c\<^sub>1 \\<^sub>g c\<^sub>2\\ inductive_set com_rel ::"(('s,'p,'f) com \ ('s,'p,'f) com) set" where "(c1, Seq c1 c2) \ com_rel" | "(c2, Seq c1 c2) \ com_rel" | "(c1, Cond b c1 c2) \ com_rel" | "(c2, Cond b c1 c2) \ com_rel" | "(c, While b c) \ com_rel" | "(c x, DynCom c) \ com_rel" | "(c, Guard f g c) \ com_rel" | "(c1, Catch c1 c2) \ com_rel" | "(c2, Catch c1 c2) \ com_rel" inductive_cases com_rel_elim_cases: "(c, Skip) \ com_rel" "(c, Basic f) \ com_rel" "(c, Spec r) \ com_rel" "(c, Seq c1 c2) \ com_rel" "(c, Cond b c1 c2) \ com_rel" "(c, While b c1) \ com_rel" "(c, Call p) \ com_rel" "(c, DynCom c1) \ com_rel" "(c, Guard f g c1) \ com_rel" "(c, Throw) \ com_rel" "(c, Catch c1 c2) \ com_rel" lemma wf_com_rel: "wf com_rel" apply (rule wfUNIVI) apply (induct_tac x) apply (erule allE, erule mp, (rule allI impI)+, erule com_rel_elim_cases) apply (erule allE, erule mp, (rule allI impI)+, erule com_rel_elim_cases) apply (erule allE, erule mp, (rule allI impI)+, erule com_rel_elim_cases) apply (erule allE, erule mp, (rule allI impI)+, erule com_rel_elim_cases, simp,simp) apply (erule allE, erule mp, (rule allI impI)+, erule com_rel_elim_cases, simp,simp) apply (erule allE, erule mp, (rule allI impI)+, erule com_rel_elim_cases,simp) apply (erule allE, erule mp, (rule allI impI)+, erule com_rel_elim_cases) apply (erule allE, erule mp, (rule allI impI)+, erule com_rel_elim_cases,simp) apply (erule allE, erule mp, (rule allI impI)+, erule com_rel_elim_cases,simp) apply (erule allE, erule mp, (rule allI impI)+, erule com_rel_elim_cases) apply (erule allE, erule mp, (rule allI impI)+, erule com_rel_elim_cases,simp,simp) done consts inter_guards:: "('s,'p,'f) com \ ('s,'p,'f) com \ ('s,'p,'f) com option" abbreviation inter_guards_syntax :: "('s,'p,'f) com \ ('s,'p,'f) com \ ('s,'p,'f) com option" ("_ \\<^sub>g _" [20,20] 19) where "c \\<^sub>g d == inter_guards (c,d)" recdef inter_guards "inv_image com_rel fst" "(Skip \\<^sub>g Skip) = Some Skip" "(Basic f1 \\<^sub>g Basic f2) = (if f1 = f2 then Some (Basic f1) else None)" "(Spec r1 \\<^sub>g Spec r2) = (if r1 = r2 then Some (Spec r1) else None)" "(Seq a1 a2 \\<^sub>g Seq b1 b2) = (case a1 \\<^sub>g b1 of None \ None | Some c1 \ (case a2 \\<^sub>g b2 of None \ None | Some c2 \ Some (Seq c1 c2)))" "(Cond cnd1 t1 e1 \\<^sub>g Cond cnd2 t2 e2) = (if cnd1 = cnd2 then (case t1 \\<^sub>g t2 of None \ None | Some t \ (case e1 \\<^sub>g e2 of None \ None | Some e \ Some (Cond cnd1 t e))) else None)" "(While cnd1 c1 \\<^sub>g While cnd2 c2) = (if cnd1 = cnd2 then (case c1 \\<^sub>g c2 of None \ None | Some c \ Some (While cnd1 c)) else None)" "(Call p1 \\<^sub>g Call p2) = (if p1 = p2 then Some (Call p1) else None)" "(DynCom P1 \\<^sub>g DynCom P2) = (if (\s. (P1 s \\<^sub>g P2 s) \ None) then Some (DynCom (\s. the (P1 s \\<^sub>g P2 s))) else None)" "(Guard m1 g1 c1 \\<^sub>g Guard m2 g2 c2) = (if m1 = m2 then (case c1 \\<^sub>g c2 of None \ None | Some c \ Some (Guard m1 (g1 \ g2) c)) else None)" "(Throw \\<^sub>g Throw) = Some Throw" "(Catch a1 a2 \\<^sub>g Catch b1 b2) = (case a1 \\<^sub>g b1 of None \ None | Some c1 \ (case a2 \\<^sub>g b2 of None \ None | Some c2 \ Some (Catch c1 c2)))" "(c \\<^sub>g d) = None" (hints cong add: option.case_cong if_cong recdef_wf: wf_com_rel simp: com_rel.intros) lemma inter_guards_strip_eq: "\c. (c1 \\<^sub>g c2) = Some c \ (strip_guards UNIV c = strip_guards UNIV c1) \ (strip_guards UNIV c = strip_guards UNIV c2)" apply (induct c1 c2 rule: inter_guards.induct) prefer 8 apply (simp split: if_split_asm) apply hypsubst apply simp apply (rule ext) apply (erule_tac x=s in allE, erule exE) apply (erule_tac x=s in allE) apply fastforce apply (fastforce split: option.splits if_split_asm)+ done lemma inter_guards_sym: "\c. (c1 \\<^sub>g c2) = Some c \ (c2 \\<^sub>g c1) = Some c" apply (induct c1 c2 rule: inter_guards.induct) apply (simp_all) prefer 7 apply (simp split: if_split_asm add: not_None_eq) apply (rule conjI) apply (clarsimp) apply (rule ext) apply (erule_tac x=s in allE)+ apply fastforce apply fastforce apply (fastforce split: option.splits if_split_asm)+ done lemma inter_guards_Skip: "(Skip \\<^sub>g c2) = Some c = (c2=Skip \ c=Skip)" by (cases c2) auto lemma inter_guards_Basic: "((Basic f) \\<^sub>g c2) = Some c = (c2=Basic f \ c=Basic f)" by (cases c2) auto lemma inter_guards_Spec: "((Spec r) \\<^sub>g c2) = Some c = (c2=Spec r \ c=Spec r)" by (cases c2) auto lemma inter_guards_Seq: "(Seq a1 a2 \\<^sub>g c2) = Some c = (\b1 b2 d1 d2. c2=Seq b1 b2 \ (a1 \\<^sub>g b1) = Some d1 \ (a2 \\<^sub>g b2) = Some d2 \ c=Seq d1 d2)" by (cases c2) (auto split: option.splits) lemma inter_guards_Cond: "(Cond cnd t1 e1 \\<^sub>g c2) = Some c = (\t2 e2 t e. c2=Cond cnd t2 e2 \ (t1 \\<^sub>g t2) = Some t \ (e1 \\<^sub>g e2) = Some e \ c=Cond cnd t e)" by (cases c2) (auto split: option.splits) lemma inter_guards_While: "(While cnd bdy1 \\<^sub>g c2) = Some c = (\bdy2 bdy. c2 =While cnd bdy2 \ (bdy1 \\<^sub>g bdy2) = Some bdy \ c=While cnd bdy)" by (cases c2) (auto split: option.splits if_split_asm) lemma inter_guards_Call: "(Call p \\<^sub>g c2) = Some c = (c2=Call p \ c=Call p)" by (cases c2) (auto split: if_split_asm) lemma inter_guards_DynCom: "(DynCom f1 \\<^sub>g c2) = Some c = (\f2. c2=DynCom f2 \ (\s. ((f1 s) \\<^sub>g (f2 s)) \ None) \ c=DynCom (\s. the ((f1 s) \\<^sub>g (f2 s))))" by (cases c2) (auto split: if_split_asm) lemma inter_guards_Guard: "(Guard f g1 bdy1 \\<^sub>g c2) = Some c = (\g2 bdy2 bdy. c2=Guard f g2 bdy2 \ (bdy1 \\<^sub>g bdy2) = Some bdy \ c=Guard f (g1 \ g2) bdy)" by (cases c2) (auto split: option.splits) lemma inter_guards_Throw: "(Throw \\<^sub>g c2) = Some c = (c2=Throw \ c=Throw)" by (cases c2) auto lemma inter_guards_Catch: "(Catch a1 a2 \\<^sub>g c2) = Some c = (\b1 b2 d1 d2. c2=Catch b1 b2 \ (a1 \\<^sub>g b1) = Some d1 \ (a2 \\<^sub>g b2) = Some d2 \ c=Catch d1 d2)" by (cases c2) (auto split: option.splits) lemmas inter_guards_simps = inter_guards_Skip inter_guards_Basic inter_guards_Spec inter_guards_Seq inter_guards_Cond inter_guards_While inter_guards_Call inter_guards_DynCom inter_guards_Guard inter_guards_Throw inter_guards_Catch subsubsection \Subset on Guards: \c\<^sub>1 \\<^sub>g c\<^sub>2\\ inductive subseteq_guards :: "('s,'p,'f) com \ ('s,'p,'f) com \ bool" ("_ \\<^sub>g _" [20,20] 19) where "Skip \\<^sub>g Skip" | "f1 = f2 \ Basic f1 \\<^sub>g Basic f2" | "r1 = r2 \ Spec r1 \\<^sub>g Spec r2" | "a1 \\<^sub>g b1 \ a2 \\<^sub>g b2 \ Seq a1 a2 \\<^sub>g Seq b1 b2" | "cnd1 = cnd2 \ t1 \\<^sub>g t2 \ e1 \\<^sub>g e2 \ Cond cnd1 t1 e1 \\<^sub>g Cond cnd2 t2 e2" | "cnd1 = cnd2 \ c1 \\<^sub>g c2 \ While cnd1 c1 \\<^sub>g While cnd2 c2" | "p1 = p2 \ Call p1 \\<^sub>g Call p2" | "(\s. P1 s \\<^sub>g P2 s) \ DynCom P1 \\<^sub>g DynCom P2" | "m1 = m2 \ g1 = g2 \ c1 \\<^sub>g c2 \ Guard m1 g1 c1 \\<^sub>g Guard m2 g2 c2" | "c1 \\<^sub>g c2 \ c1 \\<^sub>g Guard m2 g2 c2" | "Throw \\<^sub>g Throw" | "a1 \\<^sub>g b1 \ a2 \\<^sub>g b2 \ Catch a1 a2 \\<^sub>g Catch b1 b2" lemma subseteq_guards_Skip: "c = Skip" if "c \\<^sub>g Skip" using that by cases lemma subseteq_guards_Basic: "c = Basic f" if "c \\<^sub>g Basic f" using that by cases simp lemma subseteq_guards_Spec: "c = Spec r" if "c \\<^sub>g Spec r" using that by cases simp lemma subseteq_guards_Seq: "\c1' c2'. c = Seq c1' c2' \ (c1' \\<^sub>g c1) \ (c2' \\<^sub>g c2)" if "c \\<^sub>g Seq c1 c2" using that by cases simp lemma subseteq_guards_Cond: "\c1' c2'. c=Cond b c1' c2' \ (c1' \\<^sub>g c1) \ (c2' \\<^sub>g c2)" if "c \\<^sub>g Cond b c1 c2" using that by cases simp lemma subseteq_guards_While: "\c''. c=While b c'' \ (c'' \\<^sub>g c')" if "c \\<^sub>g While b c'" using that by cases simp lemma subseteq_guards_Call: "c = Call p" if "c \\<^sub>g Call p" using that by cases simp lemma subseteq_guards_DynCom: "\C'. c=DynCom C' \ (\s. C' s \\<^sub>g C s)" if "c \\<^sub>g DynCom C" using that by cases simp lemma subseteq_guards_Guard: "(c \\<^sub>g c') \ (\c''. c = Guard f g c'' \ (c'' \\<^sub>g c'))" if "c \\<^sub>g Guard f g c'" using that by cases simp_all lemma subseteq_guards_Throw: "c = Throw" if "c \\<^sub>g Throw" using that by cases lemma subseteq_guards_Catch: "\c1' c2'. c = Catch c1' c2' \ (c1' \\<^sub>g c1) \ (c2' \\<^sub>g c2)" if "c \\<^sub>g Catch c1 c2" using that by cases simp lemmas subseteq_guardsD = subseteq_guards_Skip subseteq_guards_Basic subseteq_guards_Spec subseteq_guards_Seq subseteq_guards_Cond subseteq_guards_While subseteq_guards_Call subseteq_guards_DynCom subseteq_guards_Guard subseteq_guards_Throw subseteq_guards_Catch lemma subseteq_guards_Guard': "\f' b' c'. d = Guard f' b' c'" if "Guard f b c \\<^sub>g d" using that by cases auto lemma subseteq_guards_refl: "c \\<^sub>g c" by (induct c) (auto intro: subseteq_guards.intros) (* Antisymmetry and transitivity should hold as well\ *) end diff --git a/thys/Simpl/Semantic.thy b/thys/Simpl/Semantic.thy --- a/thys/Simpl/Semantic.thy +++ b/thys/Simpl/Semantic.thy @@ -1,4683 +1,5194 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: Semantic.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (c) 2022 Apple Inc. All rights reserved. *) section \Big-Step Semantics for Simpl\ theory Semantic imports Language begin notation restrict_map ("_|\<^bsub>_\<^esub>" [90, 91] 90) datatype ('s,'f) xstate = Normal 's | Abrupt 's | Fault 'f | Stuck definition isAbr::"('s,'f) xstate \ bool" where "isAbr S = (\s. S=Abrupt s)" lemma isAbr_simps [simp]: "isAbr (Normal s) = False" "isAbr (Abrupt s) = True" "isAbr (Fault f) = False" "isAbr Stuck = False" by (auto simp add: isAbr_def) lemma isAbrE [consumes 1, elim?]: "\isAbr S; \s. S=Abrupt s \ P\ \ P" by (auto simp add: isAbr_def) lemma not_isAbrD: "\ isAbr s \ (\s'. s=Normal s') \ s = Stuck \ (\f. s=Fault f)" by (cases s) auto definition isFault:: "('s,'f) xstate \ bool" where "isFault S = (\f. S=Fault f)" lemma isFault_simps [simp]: "isFault (Normal s) = False" "isFault (Abrupt s) = False" "isFault (Fault f) = True" "isFault Stuck = False" by (auto simp add: isFault_def) lemma isFaultE [consumes 1, elim?]: "\isFault s; \f. s=Fault f \ P\ \ P" by (auto simp add: isFault_def) lemma not_isFault_iff: "(\ isFault t) = (\f. t \ Fault f)" by (auto elim: isFaultE) (* ************************************************************************* *) subsection \Big-Step Execution: \\\\c, s\ \ t\\ (* ************************************************************************* *) text \The procedure environment\ type_synonym ('s,'p,'f) body = "'p \ ('s,'p,'f) com option" inductive "exec"::"[('s,'p,'f) body,('s,'p,'f) com,('s,'f) xstate,('s,'f) xstate] \ bool" ("_\ \_,_\ \ _" [60,20,98,98] 89) for \::"('s,'p,'f) body" where Skip: "\\\Skip,Normal s\ \ Normal s" | Guard: "\s\g; \\\c,Normal s\ \ t\ \ \\\Guard f g c,Normal s\ \ t" | GuardFault: "s\g \ \\\Guard f g c,Normal s\ \ Fault f" | FaultProp [intro,simp]: "\\\c,Fault f\ \ Fault f" | Basic: "\\\Basic f,Normal s\ \ Normal (f s)" | Spec: "(s,t) \ r \ \\\Spec r,Normal s\ \ Normal t" | SpecStuck: "\t. (s,t) \ r \ \\\Spec r,Normal s\ \ Stuck" | Seq: "\\\\c\<^sub>1,Normal s\ \ s'; \\\c\<^sub>2,s'\ \ t\ \ \\\Seq c\<^sub>1 c\<^sub>2,Normal s\ \ t" | CondTrue: "\s \ b; \\\c\<^sub>1,Normal s\ \ t\ \ \\\Cond b c\<^sub>1 c\<^sub>2,Normal s\ \ t" | CondFalse: "\s \ b; \\\c\<^sub>2,Normal s\ \ t\ \ \\\Cond b c\<^sub>1 c\<^sub>2,Normal s\ \ t" | WhileTrue: "\s \ b; \\\c,Normal s\ \ s'; \\\While b c,s'\ \ t\ \ \\\While b c,Normal s\ \ t" | WhileFalse: "\s \ b\ \ \\\While b c,Normal s\ \ Normal s" | Call: "\\ p=Some bdy;\\\bdy,Normal s\ \ t\ \ \\\Call p,Normal s\ \ t" | CallUndefined: "\\ p=None\ \ \\\Call p,Normal s\ \ Stuck" | StuckProp [intro,simp]: "\\\c,Stuck\ \ Stuck" | DynCom: "\\\\(c s),Normal s\ \ t\ \ \\\DynCom c,Normal s\ \ t" | Throw: "\\\Throw,Normal s\ \ Abrupt s" | AbruptProp [intro,simp]: "\\\c,Abrupt s\ \ Abrupt s" | CatchMatch: "\\\\c\<^sub>1,Normal s\ \ Abrupt s'; \\\c\<^sub>2,Normal s'\ \ t\ \ \\\Catch c\<^sub>1 c\<^sub>2,Normal s\ \ t" | CatchMiss: "\\\\c\<^sub>1,Normal s\ \ t; \isAbr t\ \ \\\Catch c\<^sub>1 c\<^sub>2,Normal s\ \ t" inductive_cases exec_elim_cases [cases set]: "\\\c,Fault f\ \ t" "\\\c,Stuck\ \ t" "\\\c,Abrupt s\ \ t" "\\\Skip,s\ \ t" "\\\Seq c1 c2,s\ \ t" "\\\Guard f g c,s\ \ t" "\\\Basic f,s\ \ t" "\\\Spec r,s\ \ t" "\\\Cond b c1 c2,s\ \ t" "\\\While b c,s\ \ t" "\\\Call p,s\ \ t" "\\\DynCom c,s\ \ t" "\\\Throw,s\ \ t" "\\\Catch c1 c2,s\ \ t" inductive_cases exec_Normal_elim_cases [cases set]: "\\\c,Fault f\ \ t" "\\\c,Stuck\ \ t" "\\\c,Abrupt s\ \ t" "\\\Skip,Normal s\ \ t" "\\\Guard f g c,Normal s\ \ t" "\\\Basic f,Normal s\ \ t" "\\\Spec r,Normal s\ \ t" "\\\Seq c1 c2,Normal s\ \ t" "\\\Cond b c1 c2,Normal s\ \ t" "\\\While b c,Normal s\ \ t" "\\\Call p,Normal s\ \ t" "\\\DynCom c,Normal s\ \ t" "\\\Throw,Normal s\ \ t" "\\\Catch c1 c2,Normal s\ \ t" + +lemma exec_block_exn: + "\\\\bdy,Normal (init s)\ \ Normal t; \\\c s t,Normal (return s t)\ \ u\ + \ + \\\block_exn init bdy return result_exn c,Normal s\ \ u" +apply (unfold block_exn_def) + by (fastforce intro: exec.intros) + lemma exec_block: "\\\\bdy,Normal (init s)\ \ Normal t; \\\c s t,Normal (return s t)\ \ u\ \ \\\block init bdy return c,Normal s\ \ u" -apply (unfold block_def) -by (fastforce intro: exec.intros) + unfolding block_def + by (rule exec_block_exn) + +lemma exec_block_exnAbrupt: + "\\\\bdy,Normal (init s)\ \ Abrupt t\ + \ + \\\block_exn init bdy return result_exn c,Normal s\ \ Abrupt (result_exn (return s t) t)" +apply (unfold block_exn_def) + by (fastforce intro: exec.intros) lemma exec_blockAbrupt: "\\\\bdy,Normal (init s)\ \ Abrupt t\ \ \\\block init bdy return c,Normal s\ \ Abrupt (return s t)" -apply (unfold block_def) -by (fastforce intro: exec.intros) + unfolding block_def + by (rule exec_block_exnAbrupt) + +lemma exec_block_exnFault: + "\\\\bdy,Normal (init s)\ \ Fault f\ + \ + \\\block_exn init bdy return result_exn c,Normal s\ \ Fault f" +apply (unfold block_exn_def) + by (fastforce intro: exec.intros) lemma exec_blockFault: "\\\\bdy,Normal (init s)\ \ Fault f\ \ \\\block init bdy return c,Normal s\ \ Fault f" -apply (unfold block_def) -by (fastforce intro: exec.intros) + unfolding block_def + by (rule exec_block_exnFault) + +lemma exec_block_exnStuck: + "\\\\bdy,Normal (init s)\ \ Stuck\ + \ + \\\block_exn init bdy return result_exn c,Normal s\ \ Stuck" +apply (unfold block_exn_def) + by (fastforce intro: exec.intros) lemma exec_blockStuck: "\\\\bdy,Normal (init s)\ \ Stuck\ \ \\\block init bdy return c,Normal s\ \ Stuck" -apply (unfold block_def) -by (fastforce intro: exec.intros) + unfolding block_def + by (rule exec_block_exnStuck) lemma exec_call: "\\ p=Some bdy;\\\bdy,Normal (init s)\ \ Normal t; \\\c s t,Normal (return s t)\ \ u\ \ \\\call init p return c,Normal s\ \ u" apply (simp add: call_def) apply (rule exec_block) apply (erule (1) Call) apply assumption done - lemma exec_callAbrupt: "\\ p=Some bdy;\\\bdy,Normal (init s)\ \ Abrupt t\ \ \\\call init p return c,Normal s\ \ Abrupt (return s t)" apply (simp add: call_def) apply (rule exec_blockAbrupt) apply (erule (1) Call) done lemma exec_callFault: "\\ p=Some bdy; \\\bdy,Normal (init s)\ \ Fault f\ \ \\\call init p return c,Normal s\ \ Fault f" apply (simp add: call_def) apply (rule exec_blockFault) apply (erule (1) Call) done lemma exec_callStuck: "\\ p=Some bdy; \\\bdy,Normal (init s)\ \ Stuck\ \ \\\call init p return c,Normal s\ \ Stuck" apply (simp add: call_def) apply (rule exec_blockStuck) apply (erule (1) Call) done lemma exec_callUndefined: "\\ p=None\ \ \\\call init p return c,Normal s\ \ Stuck" apply (simp add: call_def) apply (rule exec_blockStuck) apply (erule CallUndefined) done +lemma exec_call_exn: + "\\ p=Some bdy;\\\bdy,Normal (init s)\ \ Normal t; \\\c s t,Normal (return s t)\ \ u\ + \ + \\\call_exn init p return result_exn c,Normal s\ \ u" +apply (simp add: call_exn_def) +apply (rule exec_block_exn) +apply (erule (1) Call) +apply assumption +done + +lemma exec_call_exnAbrupt: + "\\ p=Some bdy;\\\bdy,Normal (init s)\ \ Abrupt t\ + \ + \\\call_exn init p return result_exn c,Normal s\ \ Abrupt (result_exn (return s t) t)" +apply (simp add: call_exn_def) +apply (rule exec_block_exnAbrupt) +apply (erule (1) Call) +done + +lemma exec_call_exnFault: + "\\ p=Some bdy; \\\bdy,Normal (init s)\ \ Fault f\ + \ + \\\call_exn init p return result_exn c,Normal s\ \ Fault f" +apply (simp add: call_exn_def) +apply (rule exec_block_exnFault) +apply (erule (1) Call) +done + +lemma exec_call_exnStuck: + "\\ p=Some bdy; \\\bdy,Normal (init s)\ \ Stuck\ + \ + \\\call_exn init p return result_exn c,Normal s\ \ Stuck" +apply (simp add: call_exn_def) +apply (rule exec_block_exnStuck) +apply (erule (1) Call) +done + +lemma exec_call_exnUndefined: + "\\ p=None\ + \ + \\\call_exn init p return result_exn c,Normal s\ \ Stuck" +apply (simp add: call_exn_def) +apply (rule exec_block_exnStuck) +apply (erule CallUndefined) + done + lemma Fault_end: assumes exec: "\\\c,s\ \ t" and s: "s=Fault f" shows "t=Fault f" using exec s by (induct) auto lemma Stuck_end: assumes exec: "\\\c,s\ \ t" and s: "s=Stuck" shows "t=Stuck" using exec s by (induct) auto lemma Abrupt_end: assumes exec: "\\\c,s\ \ t" and s: "s=Abrupt s'" shows "t=Abrupt s'" using exec s by (induct) auto lemma exec_Call_body_aux: "\ p=Some bdy \ \\\Call p,s\ \ t = \\\bdy,s\ \ t" apply (rule) apply (fastforce elim: exec_elim_cases ) apply (cases s) apply (cases t) apply (auto intro: exec.intros dest: Fault_end Stuck_end Abrupt_end) done lemma exec_Call_body': "p \ dom \ \ \\\Call p,s\ \ t = \\\the (\ p),s\ \ t" apply clarsimp by (rule exec_Call_body_aux) - -lemma exec_block_Normal_elim [consumes 1]: -assumes exec_block: "\\\block init bdy return c,Normal s\ \ t" +lemma exec_block_exn_Normal_elim [consumes 1]: +assumes exec_block: "\\\block_exn init bdy return result_exn c,Normal s\ \ t" assumes Normal: "\t'. \\\\bdy,Normal (init s)\ \ Normal t'; \\\c s t',Normal (return s t')\ \ t\ \ P" assumes Abrupt: "\t'. \\\\bdy,Normal (init s)\ \ Abrupt t'; - t = Abrupt (return s t')\ + t = Abrupt (result_exn (return s t') t')\ \ P" assumes Fault: "\f. \\\\bdy,Normal (init s)\ \ Fault f; t = Fault f\ \ P" assumes Stuck: "\\\\bdy,Normal (init s)\ \ Stuck; t = Stuck\ \ P" assumes "\\ p = None; t = Stuck\ \ P" shows "P" using exec_block -apply (unfold block_def) +apply (unfold block_exn_def) apply (elim exec_Normal_elim_cases) apply simp_all apply (case_tac s') apply simp_all apply (elim exec_Normal_elim_cases) apply simp apply (drule Abrupt_end) apply simp apply (erule exec_Normal_elim_cases) apply simp apply (rule Abrupt,assumption+) apply (drule Fault_end) apply simp apply (erule exec_Normal_elim_cases) apply simp apply (drule Stuck_end) apply simp apply (erule exec_Normal_elim_cases) apply simp apply (case_tac s') apply simp_all apply (elim exec_Normal_elim_cases) apply simp apply (rule Normal, assumption+) apply (drule Fault_end) apply simp apply (rule Fault,assumption+) apply (drule Stuck_end) apply simp apply (rule Stuck,assumption+) done + +lemma exec_block_Normal_elim [consumes 1]: +assumes exec_block: "\\\block init bdy return c,Normal s\ \ t" +assumes Normal: + "\t'. + \\\\bdy,Normal (init s)\ \ Normal t'; + \\\c s t',Normal (return s t')\ \ t\ + \ P" +assumes Abrupt: + "\t'. + \\\\bdy,Normal (init s)\ \ Abrupt t'; + t = Abrupt (return s t')\ + \ P" +assumes Fault: + "\f. + \\\\bdy,Normal (init s)\ \ Fault f; + t = Fault f\ + \ P" +assumes Stuck: + "\\\\bdy,Normal (init s)\ \ Stuck; + t = Stuck\ + \ P" +assumes + Undef: "\\ p = None; t = Stuck\ \ P" +shows "P" + by (rule exec_block_exn_Normal_elim [OF exec_block [simplified block_def] + Normal Abrupt Fault Stuck Undef ]) + +lemma exec_call_exn_Normal_elim [consumes 1]: +assumes exec_call: "\\\call_exn init p return result_exn c,Normal s\ \ t" +assumes Normal: + "\bdy t'. + \\ p = Some bdy; \\\bdy,Normal (init s)\ \ Normal t'; + \\\c s t',Normal (return s t')\ \ t\ + \ P" +assumes Abrupt: + "\bdy t'. + \\ p = Some bdy; \\\bdy,Normal (init s)\ \ Abrupt t'; + t = Abrupt (result_exn (return s t') t')\ + \ P" +assumes Fault: + "\bdy f. + \\ p = Some bdy; \\\bdy,Normal (init s)\ \ Fault f; + t = Fault f\ + \ P" +assumes Stuck: + "\bdy. + \\ p = Some bdy; \\\bdy,Normal (init s)\ \ Stuck; + t = Stuck\ + \ P" +assumes Undef: + "\\ p = None; t = Stuck\ \ P" +shows "P" + using exec_call + apply (unfold call_exn_def) + apply (cases "\ p") + apply (erule exec_block_exn_Normal_elim) + apply (elim exec_Normal_elim_cases) + apply simp + apply simp + apply (elim exec_Normal_elim_cases) + apply simp + apply simp + apply (elim exec_Normal_elim_cases) + apply simp + apply simp + apply (elim exec_Normal_elim_cases) + apply simp + apply (rule Undef,assumption,assumption) + apply (rule Undef,assumption+) + apply (erule exec_block_exn_Normal_elim) + apply (elim exec_Normal_elim_cases) + apply simp + apply (rule Normal,assumption+) + apply simp + apply (elim exec_Normal_elim_cases) + apply simp + apply (rule Abrupt,assumption+) + apply simp + apply (elim exec_Normal_elim_cases) + apply simp + apply (rule Fault, assumption+) + apply simp + apply (elim exec_Normal_elim_cases) + apply simp + apply (rule Stuck,assumption,assumption,assumption) + apply simp + apply (rule Undef,assumption+) + done + lemma exec_call_Normal_elim [consumes 1]: assumes exec_call: "\\\call init p return c,Normal s\ \ t" assumes Normal: "\bdy t'. \\ p = Some bdy; \\\bdy,Normal (init s)\ \ Normal t'; \\\c s t',Normal (return s t')\ \ t\ \ P" assumes Abrupt: "\bdy t'. \\ p = Some bdy; \\\bdy,Normal (init s)\ \ Abrupt t'; t = Abrupt (return s t')\ \ P" assumes Fault: "\bdy f. \\ p = Some bdy; \\\bdy,Normal (init s)\ \ Fault f; t = Fault f\ \ P" assumes Stuck: "\bdy. \\ p = Some bdy; \\\bdy,Normal (init s)\ \ Stuck; t = Stuck\ \ P" assumes Undef: "\\ p = None; t = Stuck\ \ P" shows "P" - using exec_call - apply (unfold call_def) - apply (cases "\ p") - apply (erule exec_block_Normal_elim) - apply (elim exec_Normal_elim_cases) - apply simp - apply simp - apply (elim exec_Normal_elim_cases) - apply simp - apply simp - apply (elim exec_Normal_elim_cases) - apply simp - apply simp - apply (elim exec_Normal_elim_cases) - apply simp - apply (rule Undef,assumption,assumption) - apply (rule Undef,assumption+) - apply (erule exec_block_Normal_elim) - apply (elim exec_Normal_elim_cases) - apply simp - apply (rule Normal,assumption+) - apply simp - apply (elim exec_Normal_elim_cases) - apply simp - apply (rule Abrupt,assumption+) - apply simp - apply (elim exec_Normal_elim_cases) - apply simp - apply (rule Fault, assumption+) - apply simp - apply (elim exec_Normal_elim_cases) - apply simp - apply (rule Stuck,assumption,assumption,assumption) - apply simp - apply (rule Undef,assumption+) - done + using exec_call [simplified call_call_exn] Normal Abrupt Fault Stuck Undef + by (rule exec_call_exn_Normal_elim) lemma exec_dynCall: "\\\\call init (p s) return c,Normal s\ \ t\ \ \\\dynCall init p return c,Normal s\ \ t" apply (simp add: dynCall_def) by (rule DynCom) +lemma exec_dynCall_exn: + "\\\\call_exn init (p s) return result_exn c,Normal s\ \ t\ + \ + \\\dynCall_exn f UNIV init p return result_exn c,Normal s\ \ t" +apply (simp add: dynCall_exn_def) + by (rule DynCom) + lemma exec_dynCall_Normal_elim: assumes exec: "\\\dynCall init p return c,Normal s\ \ t" assumes call: "\\\call init (p s) return c,Normal s\ \ t \ P" shows "P" using exec apply (simp add: dynCall_def) apply (erule exec_Normal_elim_cases) apply (rule call,assumption) done +lemma exec_guards_Normal_elim_cases [consumes 1, case_names noFault someFault]: + assumes exec_guards: "\\\guards gs c,Normal s\ \ t" + assumes noFault: "\f g. (f, g) \ set gs \ s \ g \ \\\c,Normal s\ \ t \ P" + assumes someFault: "\f g. find (\(f,g). s \ g) gs = Some (f, g) \ t = Fault f \ P" + shows "P" + using exec_guards noFault someFault +proof (induct gs) + case Nil + then show ?case by simp +next + case (Cons pg gs) + obtain f g where pg: "pg = (f, g)" + by (cases pg) auto + show ?thesis + proof (cases "s \ g") + case True + from Cons.prems(1) have exec_gs: "\\ \guards gs c,Normal s\ \ t" + by (simp add: pg) (meson True pg exec_Normal_elim_cases) + + from Cons.hyps [OF exec_gs] Cons.prems(2,3) + show ?thesis + by (simp add: pg True) + next + case False + from Cons.prems(1) have t: "t = Fault f" + by (simp add: pg) (meson False pg exec_Normal_elim_cases) + + from t Cons.prems(3) + show ?thesis + by (simp add: pg False) + qed +qed + +lemma exec_guards_noFault: + assumes exec: "\\\c,Normal s\ \ t" + assumes noFault: "\f g. (f, g) \ set gs \ s \ g" + shows "\\\guards gs c,Normal s\ \ t" + using exec noFault by (induct gs) (auto intro: exec.intros) + +lemma exec_guards_Fault: + assumes Fault: "find (\(f,g). s \ g) gs = Some (f, g)" + shows "\\\guards gs c,Normal s\ \ Fault f" + using Fault by (induct gs) (auto intro: exec.intros split: prod.splits if_split_asm) + +lemma exec_guards_DynCom: + assumes exec_c: "\\\guards gs (c s), Normal s\ \ t" + shows "\\\guards gs (DynCom c), Normal s\ \ t" + using exec_c apply (induct gs) + apply (fastforce intro: exec.intros) + apply simp + by (metis exec.Guard exec.GuardFault exec_Normal_elim_cases) + +lemma exec_guards_DynCom_Normal_elim: + assumes exec: "\\\guards gs (DynCom c), Normal s\ \ t" + assumes call: "\\\guards gs (c s), Normal s\ \ t \ P" + shows "P" + using exec call proof (induct gs) + case Nil + then show ?case + apply simp + apply (erule exec_Normal_elim_cases) + apply simp + done +next + case (Cons g gs) + then show ?case + apply (cases g) + apply simp + apply (erule exec_Normal_elim_cases) + apply simp + apply (meson Guard) + apply simp + apply (meson GuardFault) + done +qed + +lemma exec_maybe_guard_DynCom: + assumes exec_c: "\\\maybe_guard f g (c s), Normal s\ \ t" + shows "\\\maybe_guard f g (DynCom c), Normal s\ \ t" + using exec_c + by (metis DynCom Guard GuardFault exec_Normal_elim_cases(5) maybe_guard_def) + +lemma exec_maybe_guard_Normal_elim_cases [consumes 1, case_names noFault someFault]: + assumes exec_guards: "\\\maybe_guard f g c,Normal s\ \ t" + assumes noFault: "s \ g \ \\\c,Normal s\ \ t \ P" + assumes someFault: "s \ g \ t = Fault f \ P" + shows "P" + using exec_guards noFault someFault + by (metis UNIV_I exec_Normal_elim_cases(5) maybe_guard_def) + +lemma exec_maybe_guard_noFault: + assumes exec: "\\\c,Normal s\ \ t" + assumes noFault: "s \ g" + shows "\\\maybe_guard f g c,Normal s\ \ t" + using exec noFault + by (simp add: Guard maybe_guard_def) + +lemma exec_maybe_guard_Fault: + assumes Fault: "s \ g" + shows "\\\maybe_guard f g c,Normal s\ \ Fault f" + using Fault + by (metis GuardFault iso_tuple_UNIV_I maybe_guard_def) + +lemma exec_maybe_guard_DynCom_Normal_elim: + assumes exec: "\\\maybe_guard f g (DynCom c), Normal s\ \ t" + assumes call: "\\\maybe_guard f g (c s), Normal s\ \ t \ P" + shows "P" + using exec call + apply (cases "g=UNIV") + subgoal + apply simp + apply (erule exec_Normal_elim_cases) + apply simp + done + subgoal + apply (simp add: maybe_guard_def) + apply (erule exec_Normal_elim_cases) + apply (meson Guard exec_Normal_elim_cases(12)) + by (meson GuardFault) + done + + +lemma exec_dynCall_exn_Normal_elim: + assumes exec: "\\\dynCall_exn f g init p return result_exn c,Normal s\ \ t" + assumes call: "\\\maybe_guard f g (call_exn init (p s) return result_exn c),Normal s\ \ t \ P" + shows "P" + using exec + apply (simp add: dynCall_exn_def) + apply (erule exec_maybe_guard_DynCom_Normal_elim) + by (rule call) + lemma exec_Call_body: "\ p=Some bdy \ \\\Call p,s\ \ t = \\\the (\ p),s\ \ t" apply (rule) apply (fastforce elim: exec_elim_cases ) apply (cases s) apply (cases t) apply (fastforce intro: exec.intros dest: Fault_end Abrupt_end Stuck_end)+ done lemma exec_Seq': "\\\\c1,s\ \ s'; \\\c2,s'\ \ s''\ \ \\\Seq c1 c2,s\ \ s''" apply (cases s) apply (fastforce intro: exec.intros) apply (fastforce dest: Abrupt_end) apply (fastforce dest: Fault_end) apply (fastforce dest: Stuck_end) done lemma exec_assoc: "\\\Seq c1 (Seq c2 c3),s\ \ t = \\\Seq (Seq c1 c2) c3,s\ \ t" by (blast elim!: exec_elim_cases intro: exec_Seq' ) (* ************************************************************************* *) subsection \Big-Step Execution with Recursion Limit: \\\\c, s\ =n\ t\\ (* ************************************************************************* *) inductive "execn"::"[('s,'p,'f) body,('s,'p,'f) com,('s,'f) xstate,nat,('s,'f) xstate] \ bool" ("_\ \_,_\ =_\ _" [60,20,98,65,98] 89) for \::"('s,'p,'f) body" where Skip: "\\\Skip,Normal s\ =n\ Normal s" | Guard: "\s\g; \\\c,Normal s\ =n\ t\ \ \\\Guard f g c,Normal s\ =n\ t" | GuardFault: "s\g \ \\\Guard f g c,Normal s\ =n\ Fault f" | FaultProp [intro,simp]: "\\\c,Fault f\ =n\ Fault f" | Basic: "\\\Basic f,Normal s\ =n\ Normal (f s)" | Spec: "(s,t) \ r \ \\\Spec r,Normal s\ =n\ Normal t" | SpecStuck: "\t. (s,t) \ r \ \\\Spec r,Normal s\ =n\ Stuck" | Seq: "\\\\c\<^sub>1,Normal s\ =n\ s'; \\\c\<^sub>2,s'\ =n\ t\ \ \\\Seq c\<^sub>1 c\<^sub>2,Normal s\ =n\ t" | CondTrue: "\s \ b; \\\c\<^sub>1,Normal s\ =n\ t\ \ \\\Cond b c\<^sub>1 c\<^sub>2,Normal s\ =n\ t" | CondFalse: "\s \ b; \\\c\<^sub>2,Normal s\ =n\ t\ \ \\\Cond b c\<^sub>1 c\<^sub>2,Normal s\ =n\ t" | WhileTrue: "\s \ b; \\\c,Normal s\ =n\ s'; \\\While b c,s'\ =n\ t\ \ \\\While b c,Normal s\ =n\ t" | WhileFalse: "\s \ b\ \ \\\While b c,Normal s\ =n\ Normal s" | Call: "\\ p=Some bdy;\\\bdy,Normal s\ =n\ t\ \ \\\Call p ,Normal s\ =Suc n\ t" | CallUndefined: "\\ p=None\ \ \\\Call p ,Normal s\ =Suc n\ Stuck" | StuckProp [intro,simp]: "\\\c,Stuck\ =n\ Stuck" | DynCom: "\\\\(c s),Normal s\ =n\ t\ \ \\\DynCom c,Normal s\ =n\ t" | Throw: "\\\Throw,Normal s\ =n\ Abrupt s" | AbruptProp [intro,simp]: "\\\c,Abrupt s\ =n\ Abrupt s" | CatchMatch: "\\\\c\<^sub>1,Normal s\ =n\ Abrupt s'; \\\c\<^sub>2,Normal s'\ =n\ t\ \ \\\Catch c\<^sub>1 c\<^sub>2,Normal s\ =n\ t" | CatchMiss: "\\\\c\<^sub>1,Normal s\ =n\ t; \isAbr t\ \ \\\Catch c\<^sub>1 c\<^sub>2,Normal s\ =n\ t" inductive_cases execn_elim_cases [cases set]: "\\\c,Fault f\ =n\ t" "\\\c,Stuck\ =n\ t" "\\\c,Abrupt s\ =n\ t" "\\\Skip,s\ =n\ t" "\\\Seq c1 c2,s\ =n\ t" "\\\Guard f g c,s\ =n\ t" "\\\Basic f,s\ =n\ t" "\\\Spec r,s\ =n\ t" "\\\Cond b c1 c2,s\ =n\ t" "\\\While b c,s\ =n\ t" "\\\Call p ,s\ =n\ t" "\\\DynCom c,s\ =n\ t" "\\\Throw,s\ =n\ t" "\\\Catch c1 c2,s\ =n\ t" inductive_cases execn_Normal_elim_cases [cases set]: "\\\c,Fault f\ =n\ t" "\\\c,Stuck\ =n\ t" "\\\c,Abrupt s\ =n\ t" "\\\Skip,Normal s\ =n\ t" "\\\Guard f g c,Normal s\ =n\ t" "\\\Basic f,Normal s\ =n\ t" "\\\Spec r,Normal s\ =n\ t" "\\\Seq c1 c2,Normal s\ =n\ t" "\\\Cond b c1 c2,Normal s\ =n\ t" "\\\While b c,Normal s\ =n\ t" "\\\Call p,Normal s\ =n\ t" "\\\DynCom c,Normal s\ =n\ t" "\\\Throw,Normal s\ =n\ t" "\\\Catch c1 c2,Normal s\ =n\ t" lemma execn_Skip': "\\\Skip,t\ =n\ t" by (cases t) (auto intro: execn.intros) lemma execn_Fault_end: assumes exec: "\\\c,s\ =n\ t" and s: "s=Fault f" shows "t=Fault f" using exec s by (induct) auto lemma execn_Stuck_end: assumes exec: "\\\c,s\ =n\ t" and s: "s=Stuck" shows "t=Stuck" using exec s by (induct) auto lemma execn_Abrupt_end: assumes exec: "\\\c,s\ =n\ t" and s: "s=Abrupt s'" shows "t=Abrupt s'" using exec s by (induct) auto +lemma execn_block_exn: + "\\\\bdy,Normal (init s)\ =n\ Normal t; \\\c s t,Normal (return s t)\ =n\ u\ + \ + \\\block_exn init bdy return result_exn c,Normal s\ =n\ u" +apply (unfold block_exn_def) + by (fastforce intro: execn.intros) + lemma execn_block: "\\\\bdy,Normal (init s)\ =n\ Normal t; \\\c s t,Normal (return s t)\ =n\ u\ \ \\\block init bdy return c,Normal s\ =n\ u" -apply (unfold block_def) -by (fastforce intro: execn.intros) + unfolding block_def + by (rule execn_block_exn) + +lemma execn_block_exnAbrupt: + "\\\\bdy,Normal (init s)\ =n\ Abrupt t\ + \ + \\\block_exn init bdy return result_exn c,Normal s\ =n\ Abrupt (result_exn (return s t) t)" +apply (unfold block_exn_def) + by (fastforce intro: execn.intros) lemma execn_blockAbrupt: "\\\\bdy,Normal (init s)\ =n\ Abrupt t\ \ \\\block init bdy return c,Normal s\ =n\ Abrupt (return s t)" -apply (unfold block_def) -by (fastforce intro: execn.intros) + unfolding block_def + by (rule execn_block_exnAbrupt) + +lemma execn_block_exnFault: + "\\\\bdy,Normal (init s)\ =n\ Fault f\ + \ + \\\block_exn init bdy return result_exn c,Normal s\ =n\ Fault f" +apply (unfold block_exn_def) + by (fastforce intro: execn.intros) lemma execn_blockFault: "\\\\bdy,Normal (init s)\ =n\ Fault f\ \ \\\block init bdy return c,Normal s\ =n\ Fault f" -apply (unfold block_def) -by (fastforce intro: execn.intros) + unfolding block_def +by (rule execn_block_exnFault) + +lemma execn_block_exnStuck: + "\\\\bdy,Normal (init s)\ =n\ Stuck\ + \ + \\\block_exn init bdy return result_exn c,Normal s\ =n\ Stuck" +apply (unfold block_exn_def) + by (fastforce intro: execn.intros) lemma execn_blockStuck: "\\\\bdy,Normal (init s)\ =n\ Stuck\ \ \\\block init bdy return c,Normal s\ =n\ Stuck" -apply (unfold block_def) -by (fastforce intro: execn.intros) + unfolding block_def +by (rule execn_block_exnStuck) lemma execn_call: "\\ p=Some bdy;\\\bdy,Normal (init s)\ =n\ Normal t; \\\c s t,Normal (return s t)\ =Suc n\ u\ \ \\\call init p return c,Normal s\ =Suc n\ u" apply (simp add: call_def) apply (rule execn_block) apply (erule (1) Call) apply assumption done +lemma execn_call_exn: + "\\ p=Some bdy;\\\bdy,Normal (init s)\ =n\ Normal t; + \\\c s t,Normal (return s t)\ =Suc n\ u\ + \ + \\\call_exn init p return result_exn c,Normal s\ =Suc n\ u" +apply (simp add: call_exn_def) +apply (rule execn_block_exn) +apply (erule (1) Call) +apply assumption +done + lemma execn_callAbrupt: "\\ p=Some bdy;\\\bdy,Normal (init s)\ =n\ Abrupt t\ \ \\\call init p return c,Normal s\ =Suc n\ Abrupt (return s t)" apply (simp add: call_def) apply (rule execn_blockAbrupt) apply (erule (1) Call) done +lemma execn_call_exnAbrupt: + "\\ p=Some bdy;\\\bdy,Normal (init s)\ =n\ Abrupt t\ + \ + \\\call_exn init p return result_exn c,Normal s\ =Suc n\ Abrupt (result_exn (return s t) t)" +apply (simp add: call_exn_def) +apply (rule execn_block_exnAbrupt) +apply (erule (1) Call) + done + lemma execn_callFault: "\\ p=Some bdy; \\\bdy,Normal (init s)\ =n\ Fault f\ \ \\\call init p return c,Normal s\ =Suc n\ Fault f" apply (simp add: call_def) apply (rule execn_blockFault) apply (erule (1) Call) done +lemma execn_call_exnFault: + "\\ p=Some bdy; \\\bdy,Normal (init s)\ =n\ Fault f\ + \ + \\\call_exn init p return result_exn c,Normal s\ =Suc n\ Fault f" +apply (simp add: call_exn_def) +apply (rule execn_block_exnFault) +apply (erule (1) Call) +done + lemma execn_callStuck: "\\ p=Some bdy; \\\bdy,Normal (init s)\ =n\ Stuck\ \ \\\call init p return c,Normal s\ =Suc n\ Stuck" apply (simp add: call_def) apply (rule execn_blockStuck) apply (erule (1) Call) done +lemma execn_call_exnStuck: + "\\ p=Some bdy; \\\bdy,Normal (init s)\ =n\ Stuck\ + \ + \\\call_exn init p return result_exn c,Normal s\ =Suc n\ Stuck" +apply (simp add: call_exn_def) +apply (rule execn_block_exnStuck) +apply (erule (1) Call) + done + lemma execn_callUndefined: "\\ p=None\ \ \\\call init p return c,Normal s\ =Suc n\ Stuck" apply (simp add: call_def) apply (rule execn_blockStuck) apply (erule CallUndefined) done +lemma execn_call_exnUndefined: + "\\ p=None\ + \ + \\\call_exn init p return result_exn c,Normal s\ =Suc n\ Stuck" +apply (simp add: call_exn_def) +apply (rule execn_block_exnStuck) +apply (erule CallUndefined) +done + +lemma execn_block_exn_Normal_elim [consumes 1]: +assumes execn_block: "\\\block_exn init bdy return result_exn c,Normal s\ =n\ t" +assumes Normal: + "\t'. + \\\\bdy,Normal (init s)\ =n\ Normal t'; + \\\c s t',Normal (return s t')\ =n\ t\ + \ P" +assumes Abrupt: + "\t'. + \\\\bdy,Normal (init s)\ =n\ Abrupt t'; + t = Abrupt (result_exn (return s t') t')\ + \ P" +assumes Fault: + "\f. + \\\\bdy,Normal (init s)\ =n\ Fault f; + t = Fault f\ + \ P" +assumes Stuck: + "\\\\bdy,Normal (init s)\ =n\ Stuck; + t = Stuck\ + \ P" +assumes Undef: + "\\ p = None; t = Stuck\ \ P" +shows "P" + using execn_block +apply (unfold block_exn_def) +apply (elim execn_Normal_elim_cases) +apply simp_all +apply (case_tac s') +apply simp_all +apply (elim execn_Normal_elim_cases) +apply simp +apply (drule execn_Abrupt_end) apply simp +apply (erule execn_Normal_elim_cases) +apply simp +apply (rule Abrupt,assumption+) +apply (drule execn_Fault_end) apply simp +apply (erule execn_Normal_elim_cases) +apply simp +apply (drule execn_Stuck_end) apply simp +apply (erule execn_Normal_elim_cases) +apply simp +apply (case_tac s') +apply simp_all +apply (elim execn_Normal_elim_cases) +apply simp +apply (rule Normal,assumption+) +apply (drule execn_Fault_end) apply simp +apply (rule Fault,assumption+) +apply (drule execn_Stuck_end) apply simp +apply (rule Stuck,assumption+) +done + lemma execn_block_Normal_elim [consumes 1]: assumes execn_block: "\\\block init bdy return c,Normal s\ =n\ t" assumes Normal: "\t'. \\\\bdy,Normal (init s)\ =n\ Normal t'; \\\c s t',Normal (return s t')\ =n\ t\ \ P" assumes Abrupt: "\t'. \\\\bdy,Normal (init s)\ =n\ Abrupt t'; t = Abrupt (return s t')\ \ P" assumes Fault: "\f. \\\\bdy,Normal (init s)\ =n\ Fault f; t = Fault f\ \ P" assumes Stuck: "\\\\bdy,Normal (init s)\ =n\ Stuck; t = Stuck\ \ P" assumes Undef: "\\ p = None; t = Stuck\ \ P" shows "P" - using execn_block -apply (unfold block_def) -apply (elim execn_Normal_elim_cases) -apply simp_all -apply (case_tac s') -apply simp_all -apply (elim execn_Normal_elim_cases) -apply simp -apply (drule execn_Abrupt_end) apply simp -apply (erule execn_Normal_elim_cases) -apply simp -apply (rule Abrupt,assumption+) -apply (drule execn_Fault_end) apply simp -apply (erule execn_Normal_elim_cases) -apply simp -apply (drule execn_Stuck_end) apply simp -apply (erule execn_Normal_elim_cases) -apply simp -apply (case_tac s') -apply simp_all -apply (elim execn_Normal_elim_cases) -apply simp -apply (rule Normal,assumption+) -apply (drule execn_Fault_end) apply simp -apply (rule Fault,assumption+) -apply (drule execn_Stuck_end) apply simp -apply (rule Stuck,assumption+) -done + using execn_block [unfolded block_def] Normal Abrupt Fault Stuck Undef + by (rule execn_block_exn_Normal_elim) + +lemma execn_call_exn_Normal_elim [consumes 1]: +assumes exec_call: "\\\call_exn init p return result_exn c,Normal s\ =n\ t" +assumes Normal: + "\bdy i t'. + \\ p = Some bdy; \\\bdy,Normal (init s)\ =i\ Normal t'; + \\\c s t',Normal (return s t')\ =Suc i\ t; n = Suc i\ + \ P" +assumes Abrupt: + "\bdy i t'. + \\ p = Some bdy; \\\bdy,Normal (init s)\ =i\ Abrupt t'; n = Suc i; + t = Abrupt (result_exn (return s t') t')\ + \ P" +assumes Fault: + "\bdy i f. + \\ p = Some bdy; \\\bdy,Normal (init s)\ =i\ Fault f; n = Suc i; + t = Fault f\ + \ P" +assumes Stuck: + "\bdy i. + \\ p = Some bdy; \\\bdy,Normal (init s)\ =i\ Stuck; n = Suc i; + t = Stuck\ + \ P" +assumes Undef: + "\i. \\ p = None; n = Suc i; t = Stuck\ \ P" +shows "P" + using exec_call + apply (unfold call_exn_def) + apply (cases n) + apply (simp only: block_exn_def) + apply (fastforce elim: execn_Normal_elim_cases) + apply (cases "\ p") + apply (erule execn_block_exn_Normal_elim) + apply (elim execn_Normal_elim_cases) + apply simp + apply simp + apply (elim execn_Normal_elim_cases) + apply simp + apply simp + apply (elim execn_Normal_elim_cases) + apply simp + apply simp + apply (elim execn_Normal_elim_cases) + apply simp + apply (rule Undef,assumption,assumption,assumption) + apply (rule Undef,assumption+) + apply (erule execn_block_exn_Normal_elim) + apply (elim execn_Normal_elim_cases) + apply simp + apply (rule Normal,assumption+) + apply simp + apply (elim execn_Normal_elim_cases) + apply simp + apply (rule Abrupt,assumption+) + apply simp + apply (elim execn_Normal_elim_cases) + apply simp + apply (rule Fault,assumption+) + apply simp + apply (elim execn_Normal_elim_cases) + apply simp + apply (rule Stuck,assumption,assumption,assumption,assumption) + apply (rule Undef,assumption,assumption,assumption) + apply (rule Undef,assumption+) + done + lemma execn_call_Normal_elim [consumes 1]: assumes exec_call: "\\\call init p return c,Normal s\ =n\ t" assumes Normal: "\bdy i t'. \\ p = Some bdy; \\\bdy,Normal (init s)\ =i\ Normal t'; \\\c s t',Normal (return s t')\ =Suc i\ t; n = Suc i\ \ P" assumes Abrupt: "\bdy i t'. \\ p = Some bdy; \\\bdy,Normal (init s)\ =i\ Abrupt t'; n = Suc i; t = Abrupt (return s t')\ \ P" assumes Fault: "\bdy i f. \\ p = Some bdy; \\\bdy,Normal (init s)\ =i\ Fault f; n = Suc i; t = Fault f\ \ P" assumes Stuck: "\bdy i. \\ p = Some bdy; \\\bdy,Normal (init s)\ =i\ Stuck; n = Suc i; t = Stuck\ \ P" assumes Undef: "\i. \\ p = None; n = Suc i; t = Stuck\ \ P" shows "P" - using exec_call - apply (unfold call_def) - apply (cases n) - apply (simp only: block_def) - apply (fastforce elim: execn_Normal_elim_cases) - apply (cases "\ p") - apply (erule execn_block_Normal_elim) - apply (elim execn_Normal_elim_cases) - apply simp - apply simp - apply (elim execn_Normal_elim_cases) - apply simp - apply simp - apply (elim execn_Normal_elim_cases) - apply simp - apply simp - apply (elim execn_Normal_elim_cases) - apply simp - apply (rule Undef,assumption,assumption,assumption) - apply (rule Undef,assumption+) - apply (erule execn_block_Normal_elim) - apply (elim execn_Normal_elim_cases) - apply simp - apply (rule Normal,assumption+) - apply simp - apply (elim execn_Normal_elim_cases) - apply simp - apply (rule Abrupt,assumption+) - apply simp - apply (elim execn_Normal_elim_cases) - apply simp - apply (rule Fault,assumption+) - apply simp - apply (elim execn_Normal_elim_cases) - apply simp - apply (rule Stuck,assumption,assumption,assumption,assumption) - apply (rule Undef,assumption,assumption,assumption) - apply (rule Undef,assumption+) - done + using exec_call [simplified call_call_exn] Normal Abrupt Fault Stuck Undef + by (rule execn_call_exn_Normal_elim) + lemma execn_dynCall: "\\\\call init (p s) return c,Normal s\ =n\ t\ \ \\\dynCall init p return c,Normal s\ =n\ t" apply (simp add: dynCall_def) by (rule DynCom) +lemma execn_dynCall_exn: + "\\\\call_exn init (p s) return result_exn c,Normal s\ =n\ t\ + \ + \\\dynCall_exn f UNIV init p return result_exn c,Normal s\ =n\ t" +apply (simp add: dynCall_exn_def) + by (rule DynCom) + lemma execn_dynCall_Normal_elim: assumes exec: "\\\dynCall init p return c,Normal s\ =n\ t" assumes "\\\call init (p s) return c,Normal s\ =n\ t \ P" shows "P" using exec apply (simp add: dynCall_def) apply (erule execn_Normal_elim_cases) apply fact done - - - +lemma execn_guards_Normal_elim_cases [consumes 1, case_names noFault someFault]: + assumes exec_guards: "\\\guards gs c,Normal s\ =n\ t" + assumes noFault: "\f g. (f, g) \ set gs \ s \ g \ \\\c,Normal s\ =n\ t \ P" + assumes someFault: "\f g. find (\(f,g). s \ g) gs = Some (f, g) \ t = Fault f \ P" + shows "P" + using exec_guards noFault someFault +proof (induct gs) + case Nil + then show ?case by simp +next + case (Cons pg gs) + obtain f g where pg: "pg = (f, g)" + by (cases pg) auto + show ?thesis + proof (cases "s \ g") + case True + from Cons.prems(1) have exec_gs: "\\ \guards gs c,Normal s\ =n\ t" + by (simp add: pg) (meson True pg execn_Normal_elim_cases) + + from Cons.hyps [OF exec_gs] Cons.prems(2,3) + show ?thesis + by (simp add: pg True) + next + case False + from Cons.prems(1) have t: "t = Fault f" + by (simp add: pg) (meson False pg execn_Normal_elim_cases) + + from t Cons.prems(3) + show ?thesis + by (simp add: pg False) + qed +qed + +lemma execn_maybe_guard_Normal_elim_cases [consumes 1, case_names noFault someFault]: + assumes exec_guards: "\\\maybe_guard f g c,Normal s\ =n\ t" + assumes noFault: "s \ g \ \\\c,Normal s\ =n\ t \ P" + assumes someFault: "s \ g \ t = Fault f \ P" + shows "P" + using exec_guards noFault someFault + by (metis UNIV_I execn_Normal_elim_cases(5) maybe_guard_def) + +lemma execn_guards_noFault: + assumes exec: "\\\c,Normal s\ =n\ t" + assumes noFault: "\f g. (f, g) \ set gs \ s \ g" + shows "\\\guards gs c,Normal s\ =n\ t" + using exec noFault by (induct gs) (auto intro: execn.intros) + +lemma execn_guards_Fault: + assumes Fault: "find (\(f,g). s \ g) gs = Some (f, g)" + shows "\\\guards gs c,Normal s\ =n\ Fault f" + using Fault by (induct gs) (auto intro: execn.intros split: prod.splits if_split_asm) + +lemma execn_maybe_guard_noFault: + assumes exec: "\\\c,Normal s\ =n\ t" + assumes noFault: "s \ g" + shows "\\\maybe_guard f g c,Normal s\ =n\ t" + using exec noFault + by (auto intro: execn.intros simp add: maybe_guard_def) + +lemma execn_maybe_guard_Fault: + assumes Fault: "s \ g" + shows "\\\maybe_guard f g c,Normal s\ =n\ Fault f" + using Fault by (auto simp add: maybe_guard_def intro: execn.intros split: prod.splits if_split_asm) + +lemma execn_guards_DynCom_Normal_elim: + assumes exec: "\\\guards gs (DynCom c), Normal s\ =n\ t" + assumes call: "\\\guards gs (c s), Normal s\ =n\ t \ P" + shows "P" + using exec call proof (induct gs) + case Nil + then show ?case + apply simp + apply (erule execn_Normal_elim_cases) + apply simp + done +next + case (Cons g gs) + then show ?case + apply (cases g) + apply simp + apply (erule execn_Normal_elim_cases) + apply simp + apply (meson execn.Guard) + apply simp + apply (meson execn.GuardFault) + done +qed + +lemma execn_maybe_guard_DynCom_Normal_elim: + assumes exec: "\\\maybe_guard f g (DynCom c), Normal s\ =n\ t" + assumes call: "\\\maybe_guard f g (c s), Normal s\ =n\ t \ P" + shows "P" + using exec call + by (metis execn.Guard execn.GuardFault execn_Normal_elim_cases(12) execn_Normal_elim_cases(5) maybe_guard_def) + +lemma execn_guards_DynCom: + assumes exec_c: "\\\guards gs (c s), Normal s\ =n\ t" + shows "\\\guards gs (DynCom c), Normal s\ =n\ t" + using exec_c apply (induct gs) + apply (fastforce intro: execn.intros) + apply simp + by (metis execn.Guard execn.GuardFault execn_Normal_elim_cases) + +lemma execn_maybe_guard_DynCom: + assumes exec_c: "\\\maybe_guard f g (c s), Normal s\ =n\ t" + shows "\\\maybe_guard f g (DynCom c), Normal s\ =n\ t" + using exec_c + apply (cases "g = UNIV") + subgoal + apply simp + apply (rule execn.intros) + apply simp + done + subgoal + apply (simp add: maybe_guard_def) + by (metis execn.DynCom execn.Guard execn.GuardFault execn_Normal_elim_cases(5)) + done + + +lemma execn_dynCall_exn_Normal_elim: + assumes exec: "\\\dynCall_exn f g init p return result_exn c,Normal s\ =n\ t" + assumes "\\\maybe_guard f g (call_exn init (p s) return result_exn c),Normal s\ =n\ t \ P" + shows "P" + using exec + apply (simp add: dynCall_exn_def) + apply (erule execn_maybe_guard_DynCom_Normal_elim) + apply fact + done lemma execn_Seq': "\\\\c1,s\ =n\ s'; \\\c2,s'\ =n\ s''\ \ \\\Seq c1 c2,s\ =n\ s''" apply (cases s) apply (fastforce intro: execn.intros) apply (fastforce dest: execn_Abrupt_end) apply (fastforce dest: execn_Fault_end) apply (fastforce dest: execn_Stuck_end) done lemma execn_mono: assumes exec: "\\\c,s\ =n\ t" shows "\ m. n \ m \ \\\c,s\ =m\ t" using exec by (induct) (auto intro: execn.intros dest: Suc_le_D) lemma execn_Suc: "\\\c,s\ =n\ t \ \\\c,s\ =Suc n\ t" by (rule execn_mono [OF _ le_refl [THEN le_SucI]]) lemma execn_assoc: "\\\Seq c1 (Seq c2 c3),s\ =n\ t = \\\Seq (Seq c1 c2) c3,s\ =n\ t" by (auto elim!: execn_elim_cases intro: execn_Seq') lemma execn_to_exec: assumes execn: "\\\c,s\ =n\ t" shows "\\\c,s\ \ t" using execn by induct (auto intro: exec.intros) lemma exec_to_execn: assumes execn: "\\\c,s\ \ t" shows "\n. \\\c,s\ =n\ t" using execn proof (induct) case Skip thus ?case by (iprover intro: execn.intros) next case Guard thus ?case by (iprover intro: execn.intros) next case GuardFault thus ?case by (iprover intro: execn.intros) next case FaultProp thus ?case by (iprover intro: execn.intros) next case Basic thus ?case by (iprover intro: execn.intros) next case Spec thus ?case by (iprover intro: execn.intros) next case SpecStuck thus ?case by (iprover intro: execn.intros) next case (Seq c1 s s' c2 s'') then obtain n m where "\\\c1,Normal s\ =n\ s'" "\\\c2,s'\ =m\ s''" by blast then have "\\\c1,Normal s\ =max n m\ s'" "\\\c2,s'\ =max n m\ s''" by (auto elim!: execn_mono intro: max.cobounded1 max.cobounded2) thus ?case by (iprover intro: execn.intros) next case CondTrue thus ?case by (iprover intro: execn.intros) next case CondFalse thus ?case by (iprover intro: execn.intros) next case (WhileTrue s b c s' s'') then obtain n m where "\\\c,Normal s\ =n\ s'" "\\\While b c,s'\ =m\ s''" by blast then have "\\\c,Normal s\ =max n m\ s'" "\\\While b c,s'\ =max n m\ s''" by (auto elim!: execn_mono intro: max.cobounded1 max.cobounded2) with WhileTrue show ?case by (iprover intro: execn.intros) next case WhileFalse thus ?case by (iprover intro: execn.intros) next case Call thus ?case by (iprover intro: execn.intros) next case CallUndefined thus ?case by (iprover intro: execn.intros) next case StuckProp thus ?case by (iprover intro: execn.intros) next case DynCom thus ?case by (iprover intro: execn.intros) next case Throw thus ?case by (iprover intro: execn.intros) next case AbruptProp thus ?case by (iprover intro: execn.intros) next case (CatchMatch c1 s s' c2 s'') then obtain n m where "\\\c1,Normal s\ =n\ Abrupt s'" "\\\c2,Normal s'\ =m\ s''" by blast then have "\\\c1,Normal s\ =max n m\ Abrupt s'" "\\\c2,Normal s'\ =max n m\ s''" by (auto elim!: execn_mono intro: max.cobounded1 max.cobounded2) with CatchMatch.hyps show ?case by (iprover intro: execn.intros) next case CatchMiss thus ?case by (iprover intro: execn.intros) qed theorem exec_iff_execn: "(\\\c,s\ \ t) = (\n. \\\c,s\ =n\ t)" by (iprover intro: exec_to_execn execn_to_exec) definition nfinal_notin:: "('s,'p,'f) body \ ('s,'p,'f) com \ ('s,'f) xstate \ nat \ ('s,'f) xstate set \ bool" ("_\ \_,_\ =_\\_" [60,20,98,65,60] 89) where "\\ \c,s\ =n\\T = (\t. \\ \c,s\ =n\ t \ t\T)" definition final_notin:: "('s,'p,'f) body \ ('s,'p,'f) com \ ('s,'f) xstate \ ('s,'f) xstate set \ bool" ("_\ \_,_\ \\_" [60,20,98,60] 89) where "\\ \c,s\ \\T = (\t. \\ \c,s\ \t \ t\T)" lemma final_notinI: "\\t. \\\c,s\ \ t \ t \ T\ \ \\\c,s\ \\T" by (simp add: final_notin_def) lemma noFaultStuck_Call_body': "p \ dom \ \ \\\Call p,Normal s\ \\({Stuck} \ Fault ` (-F)) = \\\the (\ p),Normal s\ \\({Stuck} \ Fault ` (-F))" by (clarsimp simp add: final_notin_def exec_Call_body) lemma noFault_startn: assumes execn: "\\\c,s\ =n\ t" and t: "t\Fault f" shows "s\Fault f" using execn t by (induct) auto lemma noFault_start: assumes exec: "\\\c,s\ \ t" and t: "t\Fault f" shows "s\Fault f" using exec t by (induct) auto lemma noStuck_startn: assumes execn: "\\\c,s\ =n\ t" and t: "t\Stuck" shows "s\Stuck" using execn t by (induct) auto lemma noStuck_start: assumes exec: "\\\c,s\ \ t" and t: "t\Stuck" shows "s\Stuck" using exec t by (induct) auto lemma noAbrupt_startn: assumes execn: "\\\c,s\ =n\ t" and t: "\t'. t\Abrupt t'" shows "s\Abrupt s'" using execn t by (induct) auto lemma noAbrupt_start: assumes exec: "\\\c,s\ \ t" and t: "\t'. t\Abrupt t'" shows "s\Abrupt s'" using exec t by (induct) auto lemma noFaultn_startD: "\\\c,s\ =n\ Normal t \ s \ Fault f" by (auto dest: noFault_startn) lemma noFaultn_startD': "t\Fault f \ \\\c,s\ =n\ t \ s \ Fault f" by (auto dest: noFault_startn) lemma noFault_startD: "\\\c,s\ \ Normal t \ s \ Fault f" by (auto dest: noFault_start) lemma noFault_startD': "t\Fault f\ \\\c,s\ \ t \ s \ Fault f" by (auto dest: noFault_start) lemma noStuckn_startD: "\\\c,s\ =n\ Normal t \ s \ Stuck" by (auto dest: noStuck_startn) lemma noStuckn_startD': "t\Stuck \ \\\c,s\ =n\ t \ s \ Stuck" by (auto dest: noStuck_startn) lemma noStuck_startD: "\\\c,s\ \ Normal t \ s \ Stuck" by (auto dest: noStuck_start) lemma noStuck_startD': "t\Stuck \ \\\c,s\ \ t \ s \ Stuck" by (auto dest: noStuck_start) lemma noAbruptn_startD: "\\\c,s\ =n\ Normal t \ s \ Abrupt s'" by (auto dest: noAbrupt_startn) lemma noAbrupt_startD: "\\\c,s\ \ Normal t \ s \ Abrupt s'" by (auto dest: noAbrupt_start) lemma noFaultnI: "\\t. \\\c,s\ =n\t \ t\Fault f\ \ \\\c,s\ =n\\{Fault f}" by (simp add: nfinal_notin_def) lemma noFaultnI': assumes contr: "\\\c,s\ =n\ Fault f \ False" shows "\\\c,s\ =n\\{Fault f}" proof (rule noFaultnI) fix t assume "\\\c,s\ =n\ t" with contr show "t \ Fault f" by (cases "t=Fault f") auto qed lemma noFaultn_def': "\\\c,s\ =n\\{Fault f} = (\\\\c,s\ =n\ Fault f)" apply rule apply (fastforce simp add: nfinal_notin_def) apply (fastforce intro: noFaultnI') done lemma noStucknI: "\\t. \\\c,s\ =n\t \ t\Stuck\ \ \\\c,s\ =n\\{Stuck}" by (simp add: nfinal_notin_def) lemma noStucknI': assumes contr: "\\\c,s\ =n\ Stuck \ False" shows "\\\c,s\ =n\\{Stuck}" proof (rule noStucknI) fix t assume "\\\c,s\ =n\ t" with contr show "t \ Stuck" by (cases t) auto qed lemma noStuckn_def': "\\\c,s\ =n\\{Stuck} = (\\\\c,s\ =n\ Stuck)" apply rule apply (fastforce simp add: nfinal_notin_def) apply (fastforce intro: noStucknI') done lemma noFaultI: "\\t. \\\c,s\ \t \ t\Fault f\ \ \\\c,s\ \\{Fault f}" by (simp add: final_notin_def) lemma noFaultI': assumes contr: "\\\c,s\ \ Fault f\ False" shows "\\\c,s\ \\{Fault f}" proof (rule noFaultI) fix t assume "\\\c,s\ \ t" with contr show "t \ Fault f" by (cases "t=Fault f") auto qed lemma noFaultE: "\\\\c,s\ \\{Fault f}; \\\c,s\ \ Fault f\ \ P" by (auto simp add: final_notin_def) lemma noFault_def': "\\\c,s\ \\{Fault f} = (\\\\c,s\ \ Fault f)" apply rule apply (fastforce simp add: final_notin_def) apply (fastforce intro: noFaultI') done lemma noStuckI: "\\t. \\\c,s\ \t \ t\Stuck\ \ \\\c,s\ \\{Stuck}" by (simp add: final_notin_def) lemma noStuckI': assumes contr: "\\\c,s\ \ Stuck \ False" shows "\\\c,s\ \\{Stuck}" proof (rule noStuckI) fix t assume "\\\c,s\ \ t" with contr show "t \ Stuck" by (cases t) auto qed lemma noStuckE: "\\\\c,s\ \\{Stuck}; \\\c,s\ \ Stuck\ \ P" by (auto simp add: final_notin_def) lemma noStuck_def': "\\\c,s\ \\{Stuck} = (\\\\c,s\ \ Stuck)" apply rule apply (fastforce simp add: final_notin_def) apply (fastforce intro: noStuckI') done lemma noFaultn_execD: "\\\\c,s\ =n\\{Fault f}; \\\c,s\ =n\t\ \ t\Fault f" by (simp add: nfinal_notin_def) lemma noFault_execD: "\\\\c,s\ \\{Fault f}; \\\c,s\ \t\ \ t\Fault f" by (simp add: final_notin_def) lemma noFaultn_exec_startD: "\\\\c,s\ =n\\{Fault f}; \\\c,s\ =n\t\ \ s\Fault f" by (auto simp add: nfinal_notin_def dest: noFaultn_startD) lemma noFault_exec_startD: "\\\\c,s\ \\{Fault f}; \\\c,s\ \t\ \ s\Fault f" by (auto simp add: final_notin_def dest: noFault_startD) lemma noStuckn_execD: "\\\\c,s\ =n\\{Stuck}; \\\c,s\ =n\t\ \ t\Stuck" by (simp add: nfinal_notin_def) lemma noStuck_execD: "\\\\c,s\ \\{Stuck}; \\\c,s\ \t\ \ t\Stuck" by (simp add: final_notin_def) lemma noStuckn_exec_startD: "\\\\c,s\ =n\\{Stuck}; \\\c,s\ =n\t\ \ s\Stuck" by (auto simp add: nfinal_notin_def dest: noStuckn_startD) lemma noStuck_exec_startD: "\\\\c,s\ \\{Stuck}; \\\c,s\ \t\ \ s\Stuck" by (auto simp add: final_notin_def dest: noStuck_startD) lemma noFaultStuckn_execD: "\\\\c,s\ =n\\{Fault True,Fault False,Stuck}; \\\c,s\ =n\t\ \ t\{Fault True,Fault False,Stuck}" by (simp add: nfinal_notin_def) lemma noFaultStuck_execD: "\\\\c,s\ \\{Fault True,Fault False,Stuck}; \\\c,s\ \t\ \ t\{Fault True,Fault False,Stuck}" by (simp add: final_notin_def) lemma noFaultStuckn_exec_startD: "\\\\c,s\ =n\\{Fault True, Fault False,Stuck}; \\\c,s\ =n\t\ \ s\{Fault True,Fault False,Stuck}" by (auto simp add: nfinal_notin_def ) lemma noFaultStuck_exec_startD: "\\\\c,s\ \\{Fault True, Fault False,Stuck}; \\\c,s\ \t\ \ s\{Fault True,Fault False,Stuck}" by (auto simp add: final_notin_def ) lemma noStuck_Call: assumes noStuck: "\\\Call p,Normal s\ \\{Stuck}" shows "p \ dom \" proof (cases "p \ dom \") case True thus ?thesis by simp next case False hence "\ p = None" by auto hence "\\\Call p,Normal s\ \Stuck" by (rule exec.CallUndefined) with noStuck show ?thesis by (auto simp add: final_notin_def) qed lemma Guard_noFaultStuckD: assumes "\\\Guard f g c,Normal s\ \\({Stuck} \ Fault ` (-F))" assumes "f \ F" shows "s \ g" using assms by (auto simp add: final_notin_def intro: exec.intros) lemma final_notin_to_finaln: assumes notin: "\\\c,s\ \\T" shows "\\\c,s\ =n\\T" proof (clarsimp simp add: nfinal_notin_def) fix t assume "\\\c,s\ =n\ t" and "t\T" with notin show "False" by (auto intro: execn_to_exec simp add: final_notin_def) qed lemma noFault_Call_body: "\ p=Some bdy\ \\\Call p ,Normal s\ \\{Fault f} = \\\the (\ p),Normal s\ \\{Fault f}" by (simp add: noFault_def' exec_Call_body) lemma noStuck_Call_body: "\ p=Some bdy\ \\\Call p,Normal s\ \\{Stuck} = \\\the (\ p),Normal s\ \\{Stuck}" by (simp add: noStuck_def' exec_Call_body) lemma exec_final_notin_to_execn: "\\\c,s\ \\T \ \\\c,s\ =n\\T" by (auto simp add: final_notin_def nfinal_notin_def dest: execn_to_exec) lemma execn_final_notin_to_exec: "\n. \\\c,s\ =n\\T \ \\\c,s\ \\T" by (auto simp add: final_notin_def nfinal_notin_def dest: exec_to_execn) lemma exec_final_notin_iff_execn: "\\\c,s\ \\T = (\n. \\\c,s\ =n\\T)" by (auto intro: exec_final_notin_to_execn execn_final_notin_to_exec) lemma Seq_NoFaultStuckD2: assumes noabort: "\\\Seq c1 c2,s\ \\({Stuck} \ Fault ` F)" shows "\t. \\\c1,s\ \ t \ t\ ({Stuck} \ Fault ` F) \ \\\c2,t\ \\({Stuck} \ Fault ` F)" using noabort by (auto simp add: final_notin_def intro: exec_Seq') lemma Seq_NoFaultStuckD1: assumes noabort: "\\\Seq c1 c2,s\ \\({Stuck} \ Fault ` F)" shows "\\\c1,s\ \\({Stuck} \ Fault ` F)" proof (rule final_notinI) fix t assume exec_c1: "\\\c1,s\ \ t" show "t \ {Stuck} \ Fault ` F" proof assume "t \ {Stuck} \ Fault ` F" moreover { assume "t = Stuck" with exec_c1 have "\\\Seq c1 c2,s\ \ Stuck" by (auto intro: exec_Seq') with noabort have False by (auto simp add: final_notin_def) hence False .. } moreover { assume "t \ Fault ` F" then obtain f where t: "t=Fault f" and f: "f \ F" by auto from t exec_c1 have "\\\Seq c1 c2,s\ \ Fault f" by (auto intro: exec_Seq') with noabort f have False by (auto simp add: final_notin_def) hence False .. } ultimately show False by auto qed qed lemma Seq_NoFaultStuckD2': assumes noabort: "\\\Seq c1 c2,s\ \\({Stuck} \ Fault ` F)" shows "\t. \\\c1,s\ \ t \ t\ ({Stuck} \ Fault ` F) \ \\\c2,t\ \\({Stuck} \ Fault ` F)" using noabort by (auto simp add: final_notin_def intro: exec_Seq') (* ************************************************************************* *) subsection \Lemmas about @{const "sequence"}, @{const "flatten"} and @{const "normalize"}\ (* ************************************************************************ *) lemma execn_sequence_app: "\s s' t. \\\\sequence Seq xs,Normal s\ =n\ s'; \\\sequence Seq ys,s'\ =n\ t\ \ \\\sequence Seq (xs@ys),Normal s\ =n\ t" proof (induct xs) case Nil thus ?case by (auto elim: execn_Normal_elim_cases) next case (Cons x xs) have exec_x_xs: "\\\sequence Seq (x # xs),Normal s\ =n\ s'" by fact have exec_ys: "\\\sequence Seq ys,s'\ =n\ t" by fact show ?case proof (cases xs) case Nil with exec_x_xs have "\\\x,Normal s\ =n\ s'" by (auto elim: execn_Normal_elim_cases ) with Nil exec_ys show ?thesis by (cases ys) (auto intro: execn.intros elim: execn_elim_cases) next case Cons with exec_x_xs obtain s'' where exec_x: "\\\x,Normal s\ =n\ s''" and exec_xs: "\\\sequence Seq xs,s''\ =n\ s'" by (auto elim: execn_Normal_elim_cases ) show ?thesis proof (cases s'') case (Normal s''') from Cons.hyps [OF exec_xs [simplified Normal] exec_ys] have "\\\sequence Seq (xs @ ys),Normal s'''\ =n\ t" . with Cons exec_x Normal show ?thesis by (auto intro: execn.intros) next case (Abrupt s''') with exec_xs have "s'=Abrupt s'''" by (auto dest: execn_Abrupt_end) with exec_ys have "t=Abrupt s'''" by (auto dest: execn_Abrupt_end) with exec_x Abrupt Cons show ?thesis by (auto intro: execn.intros) next case (Fault f) with exec_xs have "s'=Fault f" by (auto dest: execn_Fault_end) with exec_ys have "t=Fault f" by (auto dest: execn_Fault_end) with exec_x Fault Cons show ?thesis by (auto intro: execn.intros) next case Stuck with exec_xs have "s'=Stuck" by (auto dest: execn_Stuck_end) with exec_ys have "t=Stuck" by (auto dest: execn_Stuck_end) with exec_x Stuck Cons show ?thesis by (auto intro: execn.intros) qed qed qed lemma execn_sequence_appD: "\s t. \\\sequence Seq (xs @ ys),Normal s\ =n\ t \ \s'. \\\sequence Seq xs,Normal s\ =n\ s' \ \\\sequence Seq ys,s'\ =n\ t" proof (induct xs) case Nil thus ?case by (auto intro: execn.intros) next case (Cons x xs) have exec_app: "\\\sequence Seq ((x # xs) @ ys),Normal s\ =n\ t" by fact show ?case proof (cases xs) case Nil with exec_app show ?thesis by (cases ys) (auto elim: execn_Normal_elim_cases intro: execn_Skip') next case Cons with exec_app obtain s' where exec_x: "\\\x,Normal s\ =n\ s'" and exec_xs_ys: "\\\sequence Seq (xs @ ys),s'\ =n\ t" by (auto elim: execn_Normal_elim_cases) show ?thesis proof (cases s') case (Normal s'') from Cons.hyps [OF exec_xs_ys [simplified Normal]] Normal exec_x Cons show ?thesis by (auto intro: execn.intros) next case (Abrupt s'') with exec_xs_ys have "t=Abrupt s''" by (auto dest: execn_Abrupt_end) with Abrupt exec_x Cons show ?thesis by (auto intro: execn.intros) next case (Fault f) with exec_xs_ys have "t=Fault f" by (auto dest: execn_Fault_end) with Fault exec_x Cons show ?thesis by (auto intro: execn.intros) next case Stuck with exec_xs_ys have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck exec_x Cons show ?thesis by (auto intro: execn.intros) qed qed qed lemma execn_sequence_appE [consumes 1]: "\\\\sequence Seq (xs @ ys),Normal s\ =n\ t; \s'. \\\\sequence Seq xs,Normal s\ =n\ s';\\\sequence Seq ys,s'\ =n\ t\ \ P \ \ P" by (auto dest: execn_sequence_appD) lemma execn_to_execn_sequence_flatten: assumes exec: "\\\c,s\ =n\ t" shows "\\\sequence Seq (flatten c),s\ =n\ t" using exec proof induct case (Seq c1 c2 n s s' s'') thus ?case by (auto intro: execn.intros execn_sequence_app) qed (auto intro: execn.intros) lemma execn_to_execn_normalize: assumes exec: "\\\c,s\ =n\ t" shows "\\\normalize c,s\ =n\ t" using exec proof induct case (Seq c1 c2 n s s' s'') thus ?case by (auto intro: execn_to_execn_sequence_flatten execn_sequence_app ) qed (auto intro: execn.intros) lemma execn_sequence_flatten_to_execn: shows "\s t. \\\sequence Seq (flatten c),s\ =n\ t \ \\\c,s\ =n\ t" proof (induct c) case (Seq c1 c2) have exec_seq: "\\\sequence Seq (flatten (Seq c1 c2)),s\ =n\ t" by fact show ?case proof (cases s) case (Normal s') with exec_seq obtain s'' where "\\\sequence Seq (flatten c1),Normal s'\ =n\ s''" and "\\\sequence Seq (flatten c2),s''\ =n\ t" by (auto elim: execn_sequence_appE) with Seq.hyps Normal show ?thesis by (fastforce intro: execn.intros) next case Abrupt with exec_seq show ?thesis by (auto intro: execn.intros dest: execn_Abrupt_end) next case Fault with exec_seq show ?thesis by (auto intro: execn.intros dest: execn_Fault_end) next case Stuck with exec_seq show ?thesis by (auto intro: execn.intros dest: execn_Stuck_end) qed qed auto lemma execn_normalize_to_execn: shows "\s t n. \\\normalize c,s\ =n\ t \ \\\c,s\ =n\ t" proof (induct c) case Skip thus ?case by simp next case Basic thus ?case by simp next case Spec thus ?case by simp next case (Seq c1 c2) have "\\\normalize (Seq c1 c2),s\ =n\ t" by fact hence exec_norm_seq: "\\\sequence Seq (flatten (normalize c1) @ flatten (normalize c2)),s\ =n\ t" by simp show ?case proof (cases s) case (Normal s') with exec_norm_seq obtain s'' where exec_norm_c1: "\\\sequence Seq (flatten (normalize c1)),Normal s'\ =n\ s''" and exec_norm_c2: "\\\sequence Seq (flatten (normalize c2)),s''\ =n\ t" by (auto elim: execn_sequence_appE) from execn_sequence_flatten_to_execn [OF exec_norm_c1] execn_sequence_flatten_to_execn [OF exec_norm_c2] Seq.hyps Normal show ?thesis by (fastforce intro: execn.intros) next case (Abrupt s') with exec_norm_seq have "t=Abrupt s'" by (auto dest: execn_Abrupt_end) with Abrupt show ?thesis by (auto intro: execn.intros) next case (Fault f) with exec_norm_seq have "t=Fault f" by (auto dest: execn_Fault_end) with Fault show ?thesis by (auto intro: execn.intros) next case Stuck with exec_norm_seq have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck show ?thesis by (auto intro: execn.intros) qed next case Cond thus ?case by (auto intro: execn.intros elim!: execn_elim_cases) next case (While b c) have "\\\normalize (While b c),s\ =n\ t" by fact hence exec_norm_w: "\\\While b (normalize c),s\ =n\ t" by simp { fix s t w assume exec_w: "\\\w,s\ =n\ t" have "w=While b (normalize c) \ \\\While b c,s\ =n\ t" using exec_w proof (induct) case (WhileTrue s b' c' n w t) from WhileTrue obtain s_in_b: "s \ b" and exec_c: "\\\normalize c,Normal s\ =n\ w" and hyp_w: "\\\While b c,w\ =n\ t" by simp from While.hyps [OF exec_c] have "\\\c,Normal s\ =n\ w" by simp with hyp_w s_in_b have "\\\While b c,Normal s\ =n\ t" by (auto intro: execn.intros) with WhileTrue show ?case by simp qed (auto intro: execn.intros) } from this [OF exec_norm_w] show ?case by simp next case Call thus ?case by simp next case DynCom thus ?case by (auto intro: execn.intros elim!: execn_elim_cases) next case Guard thus ?case by (auto intro: execn.intros elim!: execn_elim_cases) next case Throw thus ?case by simp next case Catch thus ?case by (fastforce intro: execn.intros elim!: execn_elim_cases) qed lemma execn_normalize_iff_execn: "\\\normalize c,s\ =n\ t = \\\c,s\ =n\ t" by (auto intro: execn_to_execn_normalize execn_normalize_to_execn) lemma exec_sequence_app: assumes exec_xs: "\\\sequence Seq xs,Normal s\ \ s'" assumes exec_ys: "\\\sequence Seq ys,s'\ \ t" shows "\\\sequence Seq (xs@ys),Normal s\ \ t" proof - from exec_to_execn [OF exec_xs] obtain n where execn_xs: "\\\sequence Seq xs,Normal s\ =n\ s'".. from exec_to_execn [OF exec_ys] obtain m where execn_ys: "\\\sequence Seq ys,s'\ =m\ t".. with execn_xs obtain "\\\sequence Seq xs,Normal s\ =max n m\ s'" "\\\sequence Seq ys,s'\ =max n m\ t" by (auto intro: execn_mono max.cobounded1 max.cobounded2) from execn_sequence_app [OF this] have "\\\sequence Seq (xs @ ys),Normal s\ =max n m\ t" . thus ?thesis by (rule execn_to_exec) qed lemma exec_sequence_appD: assumes exec_xs_ys: "\\\sequence Seq (xs @ ys),Normal s\ \ t" shows "\s'. \\\sequence Seq xs,Normal s\ \ s' \ \\\sequence Seq ys,s'\ \ t" proof - from exec_to_execn [OF exec_xs_ys] obtain n where "\\\sequence Seq (xs @ ys),Normal s\ =n\ t".. thus ?thesis by (cases rule: execn_sequence_appE) (auto intro: execn_to_exec) qed lemma exec_sequence_appE [consumes 1]: "\\\\sequence Seq (xs @ ys),Normal s\ \ t; \s'. \\\\sequence Seq xs,Normal s\ \ s';\\\sequence Seq ys,s'\ \ t\ \ P \ \ P" by (auto dest: exec_sequence_appD) lemma exec_to_exec_sequence_flatten: assumes exec: "\\\c,s\ \ t" shows "\\\sequence Seq (flatten c),s\ \ t" proof - from exec_to_execn [OF exec] obtain n where "\\\c,s\ =n\ t".. from execn_to_execn_sequence_flatten [OF this] show ?thesis by (rule execn_to_exec) qed lemma exec_sequence_flatten_to_exec: assumes exec_seq: "\\\sequence Seq (flatten c),s\ \ t" shows "\\\c,s\ \ t" proof - from exec_to_execn [OF exec_seq] obtain n where "\\\sequence Seq (flatten c),s\ =n\ t".. from execn_sequence_flatten_to_execn [OF this] show ?thesis by (rule execn_to_exec) qed lemma exec_to_exec_normalize: assumes exec: "\\\c,s\ \ t" shows "\\\normalize c,s\ \ t" proof - from exec_to_execn [OF exec] obtain n where "\\\c,s\ =n\ t".. hence "\\\normalize c,s\ =n\ t" by (rule execn_to_execn_normalize) thus ?thesis by (rule execn_to_exec) qed lemma exec_normalize_to_exec: assumes exec: "\\\normalize c,s\ \ t" shows "\\\c,s\ \ t" proof - from exec_to_execn [OF exec] obtain n where "\\\normalize c,s\ =n\ t".. hence "\\\c,s\ =n\ t" by (rule execn_normalize_to_execn) thus ?thesis by (rule execn_to_exec) qed lemma exec_normalize_iff_exec: "\\\normalize c,s\ \ t = \\\c,s\ \ t" by (auto intro: exec_to_exec_normalize exec_normalize_to_exec) (* ************************************************************************* *) subsection \Lemmas about @{term "c\<^sub>1 \\<^sub>g c\<^sub>2"}\ (* ************************************************************************ *) lemma execn_to_execn_subseteq_guards: "\c s t n. \c \\<^sub>g c'; \\\c,s\ =n\ t\ \ \t'. \\\c',s\ =n\ t' \ (isFault t \ isFault t') \ (\ isFault t' \ t'=t)" proof (induct c') case Skip thus ?case by (fastforce dest: subseteq_guardsD elim: execn_elim_cases) next case Basic thus ?case by (fastforce dest: subseteq_guardsD elim: execn_elim_cases) next case Spec thus ?case by (fastforce dest: subseteq_guardsD elim: execn_elim_cases) next case (Seq c1' c2') have "c \\<^sub>g Seq c1' c2'" by fact from subseteq_guards_Seq [OF this] obtain c1 c2 where c: "c = Seq c1 c2" and c1_c1': "c1 \\<^sub>g c1'" and c2_c2': "c2 \\<^sub>g c2'" by blast have exec: "\\\c,s\ =n\ t" by fact with c obtain w where exec_c1: "\\\c1,s\ =n\ w" and exec_c2: "\\\c2,w\ =n\ t" by (auto elim: execn_elim_cases) from exec_c1 Seq.hyps c1_c1' obtain w' where exec_c1': "\\\c1',s\ =n\ w'" and w_Fault: "isFault w \ isFault w'" and w'_noFault: "\ isFault w' \ w'=w" by blast show ?case proof (cases "s") case (Fault f) with exec have "t=Fault f" by (auto dest: execn_Fault_end) with Fault show ?thesis by auto next case Stuck with exec have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck show ?thesis by auto next case (Abrupt s') with exec have "t=Abrupt s'" by (auto dest: execn_Abrupt_end) with Abrupt show ?thesis by auto next case (Normal s') show ?thesis proof (cases "isFault w") case True then obtain f where w': "w=Fault f".. moreover with exec_c2 have t: "t=Fault f" by (auto dest: execn_Fault_end) ultimately show ?thesis using Normal w_Fault exec_c1' by (fastforce intro: execn.intros elim: isFaultE) next case False note noFault_w = this show ?thesis proof (cases "isFault w'") case True then obtain f' where w': "w'=Fault f'".. with Normal exec_c1' have exec: "\\\Seq c1' c2',s\ =n\ Fault f'" by (auto intro: execn.intros) then show ?thesis by auto next case False with w'_noFault have w': "w'=w" by simp from Seq.hyps exec_c2 c2_c2' obtain t' where "\\\c2',w\ =n\ t'" and "isFault t \ isFault t'" and "\ isFault t' \ t'=t" by blast with Normal exec_c1' w' show ?thesis by (fastforce intro: execn.intros) qed qed qed next case (Cond b c1' c2') have exec: "\\\c,s\ =n\ t" by fact have "c \\<^sub>g Cond b c1' c2'" by fact from subseteq_guards_Cond [OF this] obtain c1 c2 where c: "c = Cond b c1 c2" and c1_c1': "c1 \\<^sub>g c1'" and c2_c2': "c2 \\<^sub>g c2'" by blast show ?case proof (cases "s") case (Fault f) with exec have "t=Fault f" by (auto dest: execn_Fault_end) with Fault show ?thesis by auto next case Stuck with exec have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck show ?thesis by auto next case (Abrupt s') with exec have "t=Abrupt s'" by (auto dest: execn_Abrupt_end) with Abrupt show ?thesis by auto next case (Normal s') from exec [simplified c Normal] show ?thesis proof (cases) assume s'_in_b: "s' \ b" assume "\\\c1,Normal s'\ =n\ t" with c1_c1' Normal Cond.hyps obtain t' where "\\\c1',Normal s'\ =n\ t'" "isFault t \ isFault t'" "\ isFault t' \ t' = t" by blast with s'_in_b Normal show ?thesis by (fastforce intro: execn.intros) next assume s'_notin_b: "s' \ b" assume "\\\c2,Normal s'\ =n\ t" with c2_c2' Normal Cond.hyps obtain t' where "\\\c2',Normal s'\ =n\ t'" "isFault t \ isFault t'" "\ isFault t' \ t' = t" by blast with s'_notin_b Normal show ?thesis by (fastforce intro: execn.intros) qed qed next case (While b c') have exec: "\\\c,s\ =n\ t" by fact have "c \\<^sub>g While b c'" by fact from subseteq_guards_While [OF this] obtain c'' where c: "c = While b c''" and c''_c': "c'' \\<^sub>g c'" by blast { fix c r w assume exec: "\\\c,r\ =n\ w" assume c: "c=While b c''" have "\w'. \\\While b c',r\ =n\ w' \ (isFault w \ isFault w') \ (\ isFault w' \ w'=w)" using exec c proof (induct) case (WhileTrue r b' ca n u w) have eqs: "While b' ca = While b c''" by fact from WhileTrue have r_in_b: "r \ b" by simp from WhileTrue have exec_c'': "\\\c'',Normal r\ =n\ u" by simp from While.hyps [OF c''_c' exec_c''] obtain u' where exec_c': "\\\c',Normal r\ =n\ u'" and u_Fault: "isFault u \ isFault u' "and u'_noFault: "\ isFault u' \ u' = u" by blast from WhileTrue obtain w' where exec_w: "\\\While b c',u\ =n\ w'" and w_Fault: "isFault w \ isFault w'" and w'_noFault: "\ isFault w' \ w' = w" by blast show ?case proof (cases "isFault u'") case True with exec_c' r_in_b show ?thesis by (fastforce intro: execn.intros elim: isFaultE) next case False with exec_c' r_in_b u'_noFault exec_w w_Fault w'_noFault show ?thesis by (fastforce intro: execn.intros) qed next case WhileFalse thus ?case by (fastforce intro: execn.intros) qed auto } from this [OF exec c] show ?case . next case Call thus ?case by (fastforce dest: subseteq_guardsD elim: execn_elim_cases) next case (DynCom C') have exec: "\\\c,s\ =n\ t" by fact have "c \\<^sub>g DynCom C'" by fact from subseteq_guards_DynCom [OF this] obtain C where c: "c = DynCom C" and C_C': "\s. C s \\<^sub>g C' s" by blast show ?case proof (cases "s") case (Fault f) with exec have "t=Fault f" by (auto dest: execn_Fault_end) with Fault show ?thesis by auto next case Stuck with exec have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck show ?thesis by auto next case (Abrupt s') with exec have "t=Abrupt s'" by (auto dest: execn_Abrupt_end) with Abrupt show ?thesis by auto next case (Normal s') from exec [simplified c Normal] have "\\\C s',Normal s'\ =n\ t" by cases from DynCom.hyps C_C' [rule_format] this obtain t' where "\\\C' s',Normal s'\ =n\ t'" "isFault t \ isFault t'" "\ isFault t' \ t' = t" by blast with Normal show ?thesis by (fastforce intro: execn.intros) qed next case (Guard f' g' c') have exec: "\\\c,s\ =n\ t" by fact have "c \\<^sub>g Guard f' g' c'" by fact hence subset_cases: "(c \\<^sub>g c') \ (\c''. c = Guard f' g' c'' \ (c'' \\<^sub>g c'))" by (rule subseteq_guards_Guard) show ?case proof (cases "s") case (Fault f) with exec have "t=Fault f" by (auto dest: execn_Fault_end) with Fault show ?thesis by auto next case Stuck with exec have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck show ?thesis by auto next case (Abrupt s') with exec have "t=Abrupt s'" by (auto dest: execn_Abrupt_end) with Abrupt show ?thesis by auto next case (Normal s') from subset_cases show ?thesis proof assume c_c': "c \\<^sub>g c'" from Guard.hyps [OF this exec] Normal obtain t' where exec_c': "\\\c',Normal s'\ =n\ t'" and t_Fault: "isFault t \ isFault t'" and t_noFault: "\ isFault t' \ t' = t" by blast with Normal show ?thesis by (cases "s' \ g'") (fastforce intro: execn.intros)+ next assume "\c''. c = Guard f' g' c'' \ (c'' \\<^sub>g c')" then obtain c'' where c: "c = Guard f' g' c''" and c''_c': "c'' \\<^sub>g c'" by blast from c exec Normal have exec_Guard': "\\\Guard f' g' c'',Normal s'\ =n\ t" by simp thus ?thesis proof (cases) assume s'_in_g': "s' \ g'" assume exec_c'': "\\\c'',Normal s'\ =n\ t" from Guard.hyps [OF c''_c' exec_c''] obtain t' where exec_c': "\\\c',Normal s'\ =n\ t'" and t_Fault: "isFault t \ isFault t'" and t_noFault: "\ isFault t' \ t' = t" by blast with Normal s'_in_g' show ?thesis by (fastforce intro: execn.intros) next assume "s' \ g'" "t=Fault f'" with Normal show ?thesis by (fastforce intro: execn.intros) qed qed qed next case Throw thus ?case by (fastforce dest: subseteq_guardsD intro: execn.intros elim: execn_elim_cases) next case (Catch c1' c2') have "c \\<^sub>g Catch c1' c2'" by fact from subseteq_guards_Catch [OF this] obtain c1 c2 where c: "c = Catch c1 c2" and c1_c1': "c1 \\<^sub>g c1'" and c2_c2': "c2 \\<^sub>g c2'" by blast have exec: "\\\c,s\ =n\ t" by fact show ?case proof (cases "s") case (Fault f) with exec have "t=Fault f" by (auto dest: execn_Fault_end) with Fault show ?thesis by auto next case Stuck with exec have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck show ?thesis by auto next case (Abrupt s') with exec have "t=Abrupt s'" by (auto dest: execn_Abrupt_end) with Abrupt show ?thesis by auto next case (Normal s') from exec [simplified c Normal] show ?thesis proof (cases) fix w assume exec_c1: "\\\c1,Normal s'\ =n\ Abrupt w" assume exec_c2: "\\\c2,Normal w\ =n\ t" from Normal exec_c1 c1_c1' Catch.hyps obtain w' where exec_c1': "\\\c1',Normal s'\ =n\ w'" and w'_noFault: "\ isFault w' \ w' = Abrupt w" by blast show ?thesis proof (cases "isFault w'") case True with exec_c1' Normal show ?thesis by (fastforce intro: execn.intros elim: isFaultE) next case False with w'_noFault have w': "w'=Abrupt w" by simp from Normal exec_c2 c2_c2' Catch.hyps obtain t' where "\\\c2',Normal w\ =n\ t'" "isFault t \ isFault t'" "\ isFault t' \ t' = t" by blast with exec_c1' w' Normal show ?thesis by (fastforce intro: execn.intros ) qed next assume exec_c1: "\\\c1,Normal s'\ =n\ t" assume t: "\ isAbr t" from Normal exec_c1 c1_c1' Catch.hyps obtain t' where exec_c1': "\\\c1',Normal s'\ =n\ t'" and t_Fault: "isFault t \ isFault t'" and t'_noFault: "\ isFault t' \ t' = t" by blast show ?thesis proof (cases "isFault t'") case True with exec_c1' Normal show ?thesis by (fastforce intro: execn.intros elim: isFaultE) next case False with exec_c1' Normal t_Fault t'_noFault t show ?thesis by (fastforce intro: execn.intros) qed qed qed qed lemma exec_to_exec_subseteq_guards: assumes c_c': "c \\<^sub>g c'" assumes exec: "\\\c,s\ \ t" shows "\t'. \\\c',s\ \ t' \ (isFault t \ isFault t') \ (\ isFault t' \ t'=t)" proof - from exec_to_execn [OF exec] obtain n where "\\\c,s\ =n\ t" .. from execn_to_execn_subseteq_guards [OF c_c' this] show ?thesis by (blast intro: execn_to_exec) qed (* ************************************************************************* *) subsection \Lemmas about @{const "merge_guards"}\ (* ************************************************************************ *) theorem execn_to_execn_merge_guards: assumes exec_c: "\\\c,s\ =n\ t" shows "\\\merge_guards c,s\ =n\ t " using exec_c proof (induct) case (Guard s g c n t f) have s_in_g: "s \ g" by fact have exec_merge_c: "\\\merge_guards c,Normal s\ =n\ t" by fact show ?case proof (cases "\f' g' c'. merge_guards c = Guard f' g' c'") case False with exec_merge_c s_in_g show ?thesis by (cases "merge_guards c") (auto intro: execn.intros simp add: Let_def) next case True then obtain f' g' c' where merge_guards_c: "merge_guards c = Guard f' g' c'" by iprover show ?thesis proof (cases "f=f'") case False from exec_merge_c s_in_g merge_guards_c False show ?thesis by (auto intro: execn.intros simp add: Let_def) next case True from exec_merge_c s_in_g merge_guards_c True show ?thesis by (fastforce intro: execn.intros elim: execn.cases) qed qed next case (GuardFault s g f c n) have s_notin_g: "s \ g" by fact show ?case proof (cases "\f' g' c'. merge_guards c = Guard f' g' c'") case False with s_notin_g show ?thesis by (cases "merge_guards c") (auto intro: execn.intros simp add: Let_def) next case True then obtain f' g' c' where merge_guards_c: "merge_guards c = Guard f' g' c'" by iprover show ?thesis proof (cases "f=f'") case False from s_notin_g merge_guards_c False show ?thesis by (auto intro: execn.intros simp add: Let_def) next case True from s_notin_g merge_guards_c True show ?thesis by (fastforce intro: execn.intros) qed qed qed (fastforce intro: execn.intros)+ lemma execn_merge_guards_to_execn_Normal: "\s n t. \\\merge_guards c,Normal s\ =n\ t \ \\\c,Normal s\ =n\ t" proof (induct c) case Skip thus ?case by auto next case Basic thus ?case by auto next case Spec thus ?case by auto next case (Seq c1 c2) have "\\\merge_guards (Seq c1 c2),Normal s\ =n\ t" by fact hence exec_merge: "\\\Seq (merge_guards c1) (merge_guards c2),Normal s\ =n\ t" by simp then obtain s' where exec_merge_c1: "\\\merge_guards c1,Normal s\ =n\ s'" and exec_merge_c2: "\\\merge_guards c2,s'\ =n\ t" by cases from exec_merge_c1 have exec_c1: "\\\c1,Normal s\ =n\ s'" by (rule Seq.hyps) show ?case proof (cases s') case (Normal s'') with exec_merge_c2 have "\\\c2,s'\ =n\ t" by (auto intro: Seq.hyps) with exec_c1 show ?thesis by (auto intro: execn.intros) next case (Abrupt s'') with exec_merge_c2 have "t=Abrupt s''" by (auto dest: execn_Abrupt_end) with exec_c1 Abrupt show ?thesis by (auto intro: execn.intros) next case (Fault f) with exec_merge_c2 have "t=Fault f" by (auto dest: execn_Fault_end) with exec_c1 Fault show ?thesis by (auto intro: execn.intros) next case Stuck with exec_merge_c2 have "t=Stuck" by (auto dest: execn_Stuck_end) with exec_c1 Stuck show ?thesis by (auto intro: execn.intros) qed next case Cond thus ?case by (fastforce intro: execn.intros elim: execn_Normal_elim_cases) next case (While b c) { fix c' r w assume exec_c': "\\\c',r\ =n\ w" assume c': "c'=While b (merge_guards c)" have "\\\While b c,r\ =n\ w" using exec_c' c' proof (induct) case (WhileTrue r b' c'' n u w) have eqs: "While b' c'' = While b (merge_guards c)" by fact from WhileTrue have r_in_b: "r \ b" by simp from WhileTrue While.hyps have exec_c: "\\\c,Normal r\ =n\ u" by simp from WhileTrue have exec_w: "\\\While b c,u\ =n\ w" by simp from r_in_b exec_c exec_w show ?case by (rule execn.WhileTrue) next case WhileFalse thus ?case by (auto intro: execn.WhileFalse) qed auto } with While.prems show ?case by (auto) next case Call thus ?case by simp next case DynCom thus ?case by (fastforce intro: execn.intros elim: execn_Normal_elim_cases) next case (Guard f g c) have exec_merge: "\\\merge_guards (Guard f g c),Normal s\ =n\ t" by fact show ?case proof (cases "s \ g") case False with exec_merge have "t=Fault f" by (auto split: com.splits if_split_asm elim: execn_Normal_elim_cases simp add: Let_def is_Guard_def) with False show ?thesis by (auto intro: execn.intros) next case True note s_in_g = this show ?thesis proof (cases "\f' g' c'. merge_guards c = Guard f' g' c'") case False then have "merge_guards (Guard f g c) = Guard f g (merge_guards c)" by (cases "merge_guards c") (auto simp add: Let_def) with exec_merge s_in_g obtain "\\\merge_guards c,Normal s\ =n\ t" by (auto elim: execn_Normal_elim_cases) from Guard.hyps [OF this] s_in_g show ?thesis by (auto intro: execn.intros) next case True then obtain f' g' c' where merge_guards_c: "merge_guards c = Guard f' g' c'" by iprover show ?thesis proof (cases "f=f'") case False with merge_guards_c have "merge_guards (Guard f g c) = Guard f g (merge_guards c)" by (simp add: Let_def) with exec_merge s_in_g obtain "\\\merge_guards c,Normal s\ =n\ t" by (auto elim: execn_Normal_elim_cases) from Guard.hyps [OF this] s_in_g show ?thesis by (auto intro: execn.intros) next case True note f_eq_f' = this with merge_guards_c have merge_guards_Guard: "merge_guards (Guard f g c) = Guard f (g \ g') c'" by simp show ?thesis proof (cases "s \ g'") case True with exec_merge merge_guards_Guard merge_guards_c s_in_g have "\\\merge_guards c,Normal s\ =n\ t" by (auto intro: execn.intros elim: execn_Normal_elim_cases) with Guard.hyps [OF this] s_in_g show ?thesis by (auto intro: execn.intros) next case False with exec_merge merge_guards_Guard have "t=Fault f" by (auto elim: execn_Normal_elim_cases) with merge_guards_c f_eq_f' False have "\\\merge_guards c,Normal s\ =n\ t" by (auto intro: execn.intros) from Guard.hyps [OF this] s_in_g show ?thesis by (auto intro: execn.intros) qed qed qed qed next case Throw thus ?case by simp next case (Catch c1 c2) have "\\\merge_guards (Catch c1 c2),Normal s\ =n\ t" by fact hence "\\\Catch (merge_guards c1) (merge_guards c2),Normal s\ =n\ t" by simp thus ?case by cases (auto intro: execn.intros Catch.hyps) qed theorem execn_merge_guards_to_execn: "\\\merge_guards c,s\ =n\ t \ \\\c, s\ =n\ t" apply (cases s) apply (fastforce intro: execn_merge_guards_to_execn_Normal) apply (fastforce dest: execn_Abrupt_end) apply (fastforce dest: execn_Fault_end) apply (fastforce dest: execn_Stuck_end) done corollary execn_iff_execn_merge_guards: "\\\c, s\ =n\ t = \\\merge_guards c,s\ =n\ t" by (blast intro: execn_merge_guards_to_execn execn_to_execn_merge_guards) theorem exec_iff_exec_merge_guards: "\\\c, s\ \ t = \\\merge_guards c,s\ \ t" by (blast dest: exec_to_execn intro: execn_to_exec intro: execn_to_execn_merge_guards execn_merge_guards_to_execn) corollary exec_to_exec_merge_guards: "\\\c, s\ \ t \ \\\merge_guards c,s\ \ t" by (rule iffD1 [OF exec_iff_exec_merge_guards]) corollary exec_merge_guards_to_exec: "\\\merge_guards c,s\ \ t \ \\\c, s\ \ t" by (rule iffD2 [OF exec_iff_exec_merge_guards]) (* ************************************************************************* *) subsection \Lemmas about @{const "mark_guards"}\ (* ************************************************************************ *) lemma execn_to_execn_mark_guards: assumes exec_c: "\\\c,s\ =n\ t" assumes t_not_Fault: "\ isFault t" shows "\\\mark_guards f c,s\ =n\ t " using exec_c t_not_Fault [simplified not_isFault_iff] by (induct) (auto intro: execn.intros dest: noFaultn_startD') lemma execn_to_execn_mark_guards_Fault: assumes exec_c: "\\\c,s\ =n\ t" shows "\f. \t=Fault f\ \ \f'. \\\mark_guards x c,s\ =n\ Fault f'" using exec_c proof (induct) case Skip thus ?case by auto next case Guard thus ?case by (fastforce intro: execn.intros) next case GuardFault thus ?case by (fastforce intro: execn.intros) next case FaultProp thus ?case by auto next case Basic thus ?case by auto next case Spec thus ?case by auto next case SpecStuck thus ?case by auto next case (Seq c1 s n w c2 t) have exec_c1: "\\\c1,Normal s\ =n\ w" by fact have exec_c2: "\\\c2,w\ =n\ t" by fact have t: "t=Fault f" by fact show ?case proof (cases w) case (Fault f') with exec_c2 t have "f'=f" by (auto dest: execn_Fault_end) with Fault Seq.hyps obtain f'' where "\\\mark_guards x c1,Normal s\ =n\ Fault f''" by auto moreover have "\\\mark_guards x c2,Fault f''\ =n\ Fault f''" by auto ultimately show ?thesis by (auto intro: execn.intros) next case (Normal s') with execn_to_execn_mark_guards [OF exec_c1] have exec_mark_c1: "\\\mark_guards x c1,Normal s\ =n\ w" by simp with Seq.hyps t obtain f' where "\\\mark_guards x c2,w\ =n\ Fault f'" by blast with exec_mark_c1 show ?thesis by (auto intro: execn.intros) next case (Abrupt s') with execn_to_execn_mark_guards [OF exec_c1] have exec_mark_c1: "\\\mark_guards x c1,Normal s\ =n\ w" by simp with Seq.hyps t obtain f' where "\\\mark_guards x c2,w\ =n\ Fault f'" by (auto intro: execn.intros) with exec_mark_c1 show ?thesis by (auto intro: execn.intros) next case Stuck with exec_c2 have "t=Stuck" by (auto dest: execn_Stuck_end) with t show ?thesis by simp qed next case CondTrue thus ?case by (fastforce intro: execn.intros) next case CondFalse thus ?case by (fastforce intro: execn.intros) next case (WhileTrue s b c n w t) have exec_c: "\\\c,Normal s\ =n\ w" by fact have exec_w: "\\\While b c,w\ =n\ t" by fact have t: "t = Fault f" by fact have s_in_b: "s \ b" by fact show ?case proof (cases w) case (Fault f') with exec_w t have "f'=f" by (auto dest: execn_Fault_end) with Fault WhileTrue.hyps obtain f'' where "\\\mark_guards x c,Normal s\ =n\ Fault f''" by auto moreover have "\\\mark_guards x (While b c),Fault f''\ =n\ Fault f''" by auto ultimately show ?thesis using s_in_b by (auto intro: execn.intros) next case (Normal s') with execn_to_execn_mark_guards [OF exec_c] have exec_mark_c: "\\\mark_guards x c,Normal s\ =n\ w" by simp with WhileTrue.hyps t obtain f' where "\\\mark_guards x (While b c),w\ =n\ Fault f'" by blast with exec_mark_c s_in_b show ?thesis by (auto intro: execn.intros) next case (Abrupt s') with execn_to_execn_mark_guards [OF exec_c] have exec_mark_c: "\\\mark_guards x c,Normal s\ =n\ w" by simp with WhileTrue.hyps t obtain f' where "\\\mark_guards x (While b c),w\ =n\ Fault f'" by (auto intro: execn.intros) with exec_mark_c s_in_b show ?thesis by (auto intro: execn.intros) next case Stuck with exec_w have "t=Stuck" by (auto dest: execn_Stuck_end) with t show ?thesis by simp qed next case WhileFalse thus ?case by (fastforce intro: execn.intros) next case Call thus ?case by (fastforce intro: execn.intros) next case CallUndefined thus ?case by simp next case StuckProp thus ?case by simp next case DynCom thus ?case by (fastforce intro: execn.intros) next case Throw thus ?case by simp next case AbruptProp thus ?case by simp next case (CatchMatch c1 s n w c2 t) have exec_c1: "\\\c1,Normal s\ =n\ Abrupt w" by fact have exec_c2: "\\\c2,Normal w\ =n\ t" by fact have t: "t = Fault f" by fact from execn_to_execn_mark_guards [OF exec_c1] have exec_mark_c1: "\\\mark_guards x c1,Normal s\ =n\ Abrupt w" by simp with CatchMatch.hyps t obtain f' where "\\\mark_guards x c2,Normal w\ =n\ Fault f'" by blast with exec_mark_c1 show ?case by (auto intro: execn.intros) next case CatchMiss thus ?case by (fastforce intro: execn.intros) qed lemma execn_mark_guards_to_execn: "\s n t. \\\mark_guards f c,s\ =n\ t \ \t'. \\\c,s\ =n\ t' \ (isFault t \ isFault t') \ (t' = Fault f \ t'=t) \ (isFault t' \ isFault t) \ (\ isFault t' \ t'=t)" proof (induct c) case Skip thus ?case by auto next case Basic thus ?case by auto next case Spec thus ?case by auto next case (Seq c1 c2 s n t) have exec_mark: "\\\mark_guards f (Seq c1 c2),s\ =n\ t" by fact then obtain w where exec_mark_c1: "\\\mark_guards f c1,s\ =n\ w" and exec_mark_c2: "\\\mark_guards f c2,w\ =n\ t" by (auto elim: execn_elim_cases) from Seq.hyps exec_mark_c1 obtain w' where exec_c1: "\\\c1,s\ =n\ w'" and w_Fault: "isFault w \ isFault w'" and w'_Fault_f: "w' = Fault f \ w'=w" and w'_Fault: "isFault w' \ isFault w" and w'_noFault: "\ isFault w' \ w'=w" by blast show ?case proof (cases "s") case (Fault f) with exec_mark have "t=Fault f" by (auto dest: execn_Fault_end) with Fault show ?thesis by auto next case Stuck with exec_mark have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck show ?thesis by auto next case (Abrupt s') with exec_mark have "t=Abrupt s'" by (auto dest: execn_Abrupt_end) with Abrupt show ?thesis by auto next case (Normal s') show ?thesis proof (cases "isFault w") case True then obtain f where w': "w=Fault f".. moreover with exec_mark_c2 have t: "t=Fault f" by (auto dest: execn_Fault_end) ultimately show ?thesis using Normal w_Fault w'_Fault_f exec_c1 by (fastforce intro: execn.intros elim: isFaultE) next case False note noFault_w = this show ?thesis proof (cases "isFault w'") case True then obtain f' where w': "w'=Fault f'".. with Normal exec_c1 have exec: "\\\Seq c1 c2,s\ =n\ Fault f'" by (auto intro: execn.intros) from w'_Fault_f w' noFault_w have "f' \ f" by (cases w) auto moreover from w' w'_Fault exec_mark_c2 have "isFault t" by (auto dest: execn_Fault_end elim: isFaultE) ultimately show ?thesis using exec by auto next case False with w'_noFault have w': "w'=w" by simp from Seq.hyps exec_mark_c2 obtain t' where "\\\c2,w\ =n\ t'" and "isFault t \ isFault t'" and "t' = Fault f \ t'=t" and "isFault t' \ isFault t" and "\ isFault t' \ t'=t" by blast with Normal exec_c1 w' show ?thesis by (fastforce intro: execn.intros) qed qed qed next case (Cond b c1 c2 s n t) have exec_mark: "\\\mark_guards f (Cond b c1 c2),s\ =n\ t" by fact show ?case proof (cases s) case (Fault f) with exec_mark have "t=Fault f" by (auto dest: execn_Fault_end) with Fault show ?thesis by auto next case Stuck with exec_mark have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck show ?thesis by auto next case (Abrupt s') with exec_mark have "t=Abrupt s'" by (auto dest: execn_Abrupt_end) with Abrupt show ?thesis by auto next case (Normal s') show ?thesis proof (cases "s'\ b") case True with Normal exec_mark have "\\\mark_guards f c1 ,Normal s'\ =n\ t" by (auto elim: execn_Normal_elim_cases) with Normal True Cond.hyps obtain t' where "\\\c1,Normal s'\ =n\ t'" "isFault t \ isFault t'" "t' = Fault f \ t'=t" "isFault t' \ isFault t" "\ isFault t' \ t' = t" by blast with Normal True show ?thesis by (blast intro: execn.intros) next case False with Normal exec_mark have "\\\mark_guards f c2 ,Normal s'\ =n\ t" by (auto elim: execn_Normal_elim_cases) with Normal False Cond.hyps obtain t' where "\\\c2,Normal s'\ =n\ t'" "isFault t \ isFault t'" "t' = Fault f \ t'=t" "isFault t' \ isFault t" "\ isFault t' \ t' = t" by blast with Normal False show ?thesis by (blast intro: execn.intros) qed qed next case (While b c s n t) have exec_mark: "\\\mark_guards f (While b c),s\ =n\ t" by fact show ?case proof (cases s) case (Fault f) with exec_mark have "t=Fault f" by (auto dest: execn_Fault_end) with Fault show ?thesis by auto next case Stuck with exec_mark have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck show ?thesis by auto next case (Abrupt s') with exec_mark have "t=Abrupt s'" by (auto dest: execn_Abrupt_end) with Abrupt show ?thesis by auto next case (Normal s') { fix c' r w assume exec_c': "\\\c',r\ =n\ w" assume c': "c'=While b (mark_guards f c)" have "\w'. \\\While b c,r\ =n\ w' \ (isFault w \ isFault w') \ (w' = Fault f \ w'=w) \ (isFault w' \ isFault w) \ (\ isFault w' \ w'=w)" using exec_c' c' proof (induct) case (WhileTrue r b' c'' n u w) have eqs: "While b' c'' = While b (mark_guards f c)" by fact from WhileTrue.hyps eqs have r_in_b: "r\b" by simp from WhileTrue.hyps eqs have exec_mark_c: "\\\mark_guards f c,Normal r\ =n\ u" by simp from WhileTrue.hyps eqs have exec_mark_w: "\\\While b (mark_guards f c),u\ =n\ w" by simp show ?case proof - from WhileTrue.hyps eqs have "\\\mark_guards f c,Normal r\ =n\ u" by simp with While.hyps obtain u' where exec_c: "\\\c,Normal r\ =n\ u'" and u_Fault: "isFault u \ isFault u'" and u'_Fault_f: "u' = Fault f \ u'=u" and u'_Fault: "isFault u' \ isFault u" and u'_noFault: "\ isFault u' \ u'=u" by blast show ?thesis proof (cases "isFault u'") case False with u'_noFault have u': "u'=u" by simp from WhileTrue.hyps eqs obtain w' where "\\\While b c,u\ =n\ w'" "isFault w \ isFault w'" "w' = Fault f \ w'=w" "isFault w' \ isFault w" "\ isFault w' \ w' = w" by blast with u' exec_c r_in_b show ?thesis by (blast intro: execn.WhileTrue) next case True then obtain f' where u': "u'=Fault f'".. with exec_c r_in_b have exec: "\\\While b c,Normal r\ =n\ Fault f'" by (blast intro: execn.intros) from True u'_Fault have "isFault u" by simp then obtain f where u: "u=Fault f".. with exec_mark_w have "w=Fault f" by (auto dest: execn_Fault_end) with exec u' u u'_Fault_f show ?thesis by auto qed qed next case (WhileFalse r b' c'' n) have eqs: "While b' c'' = While b (mark_guards f c)" by fact from WhileFalse.hyps eqs have r_not_in_b: "r\b" by simp show ?case proof - from r_not_in_b have "\\\While b c,Normal r\ =n\ Normal r" by (rule execn.WhileFalse) thus ?thesis by blast qed qed auto } note hyp_while = this show ?thesis proof (cases "s'\b") case False with Normal exec_mark have "t=s" by (auto elim: execn_Normal_elim_cases) with Normal False show ?thesis by (auto intro: execn.intros) next case True note s'_in_b = this with Normal exec_mark obtain r where exec_mark_c: "\\\mark_guards f c,Normal s'\ =n\ r" and exec_mark_w: "\\\While b (mark_guards f c),r\ =n\ t" by (auto elim: execn_Normal_elim_cases) from While.hyps exec_mark_c obtain r' where exec_c: "\\\c,Normal s'\ =n\ r'" and r_Fault: "isFault r \ isFault r'" and r'_Fault_f: "r' = Fault f \ r'=r" and r'_Fault: "isFault r' \ isFault r" and r'_noFault: "\ isFault r' \ r'=r" by blast show ?thesis proof (cases "isFault r'") case False with r'_noFault have r': "r'=r" by simp from hyp_while exec_mark_w obtain t' where "\\\While b c,r\ =n\ t'" "isFault t \ isFault t'" "t' = Fault f \ t'=t" "isFault t' \ isFault t" "\ isFault t' \ t'=t" by blast with r' exec_c Normal s'_in_b show ?thesis by (blast intro: execn.intros) next case True then obtain f' where r': "r'=Fault f'".. hence "\\\While b c,r'\ =n\ Fault f'" by auto with Normal s'_in_b exec_c have exec: "\\\While b c,Normal s'\ =n\ Fault f'" by (auto intro: execn.intros) from True r'_Fault have "isFault r" by simp then obtain f where r: "r=Fault f".. with exec_mark_w have "t=Fault f" by (auto dest: execn_Fault_end) with Normal exec r' r r'_Fault_f show ?thesis by auto qed qed qed next case Call thus ?case by auto next case DynCom thus ?case by (fastforce elim!: execn_elim_cases intro: execn.intros) next case (Guard f' g c s n t) have exec_mark: "\\\mark_guards f (Guard f' g c),s\ =n\ t" by fact show ?case proof (cases s) case (Fault f) with exec_mark have "t=Fault f" by (auto dest: execn_Fault_end) with Fault show ?thesis by auto next case Stuck with exec_mark have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck show ?thesis by auto next case (Abrupt s') with exec_mark have "t=Abrupt s'" by (auto dest: execn_Abrupt_end) with Abrupt show ?thesis by auto next case (Normal s') show ?thesis proof (cases "s'\g") case False with Normal exec_mark have t: "t=Fault f" by (auto elim: execn_Normal_elim_cases) from False have "\\\Guard f' g c,Normal s'\ =n\ Fault f'" by (blast intro: execn.intros) with Normal t show ?thesis by auto next case True with exec_mark Normal have "\\\mark_guards f c,Normal s'\ =n\ t" by (auto elim: execn_Normal_elim_cases) with Guard.hyps obtain t' where "\\\c,Normal s'\ =n\ t'" and "isFault t \ isFault t'" and "t' = Fault f \ t'=t" and "isFault t' \ isFault t" and "\ isFault t' \ t'=t" by blast with Normal True show ?thesis by (blast intro: execn.intros) qed qed next case Throw thus ?case by auto next case (Catch c1 c2 s n t) have exec_mark: "\\\mark_guards f (Catch c1 c2),s\ =n\ t" by fact show ?case proof (cases "s") case (Fault f) with exec_mark have "t=Fault f" by (auto dest: execn_Fault_end) with Fault show ?thesis by auto next case Stuck with exec_mark have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck show ?thesis by auto next case (Abrupt s') with exec_mark have "t=Abrupt s'" by (auto dest: execn_Abrupt_end) with Abrupt show ?thesis by auto next case (Normal s') note s=this with exec_mark have "\\\Catch (mark_guards f c1) (mark_guards f c2),Normal s'\ =n\ t" by simp thus ?thesis proof (cases) fix w assume exec_mark_c1: "\\\mark_guards f c1,Normal s'\ =n\ Abrupt w" assume exec_mark_c2: "\\\mark_guards f c2,Normal w\ =n\ t" from exec_mark_c1 Catch.hyps obtain w' where exec_c1: "\\\c1,Normal s'\ =n\ w'" and w'_Fault_f: "w' = Fault f \ w'=Abrupt w" and w'_Fault: "isFault w' \ isFault (Abrupt w)" and w'_noFault: "\ isFault w' \ w'=Abrupt w" by fastforce show ?thesis proof (cases "w'") case (Fault f') with Normal exec_c1 have "\\\Catch c1 c2,s\ =n\ Fault f'" by (auto intro: execn.intros) with w'_Fault Fault show ?thesis by auto next case Stuck with w'_noFault have False by simp thus ?thesis .. next case (Normal w'') with w'_noFault have False by simp thus ?thesis .. next case (Abrupt w'') with w'_noFault have w'': "w''=w" by simp from exec_mark_c2 Catch.hyps obtain t' where "\\\c2,Normal w\ =n\ t'" "isFault t \ isFault t'" "t' = Fault f \ t'=t" "isFault t' \ isFault t" "\ isFault t' \ t'=t" by blast with w'' Abrupt s exec_c1 show ?thesis by (blast intro: execn.intros) qed next assume t: "\ isAbr t" assume "\\\mark_guards f c1,Normal s'\ =n\ t" with Catch.hyps obtain t' where exec_c1: "\\\c1,Normal s'\ =n\ t'" and t_Fault: "isFault t \ isFault t'" and t'_Fault_f: "t' = Fault f \ t'=t" and t'_Fault: "isFault t' \ isFault t" and t'_noFault: "\ isFault t' \ t'=t" by blast show ?thesis proof (cases "isFault t'") case True then obtain f' where t': "t'=Fault f'".. with exec_c1 have "\\\Catch c1 c2,Normal s'\ =n\ Fault f'" by (auto intro: execn.intros) with t'_Fault_f t'_Fault t' s show ?thesis by auto next case False with t'_noFault have "t'=t" by simp with t exec_c1 s show ?thesis by (blast intro: execn.intros) qed qed qed qed lemma exec_to_exec_mark_guards: assumes exec_c: "\\\c,s\ \ t" assumes t_not_Fault: "\ isFault t" shows "\\\mark_guards f c,s\ \ t " proof - from exec_to_execn [OF exec_c] obtain n where "\\\c,s\ =n\ t" .. from execn_to_execn_mark_guards [OF this t_not_Fault] show ?thesis by (blast intro: execn_to_exec) qed lemma exec_to_exec_mark_guards_Fault: assumes exec_c: "\\\c,s\ \ Fault f" shows "\f'. \\\mark_guards x c,s\ \ Fault f'" proof - from exec_to_execn [OF exec_c] obtain n where "\\\c,s\ =n\ Fault f" .. from execn_to_execn_mark_guards_Fault [OF this] show ?thesis by (blast intro: execn_to_exec) qed lemma exec_mark_guards_to_exec: assumes exec_mark: "\\\mark_guards f c,s\ \ t" shows "\t'. \\\c,s\ \ t' \ (isFault t \ isFault t') \ (t' = Fault f \ t'=t) \ (isFault t' \ isFault t) \ (\ isFault t' \ t'=t)" proof - from exec_to_execn [OF exec_mark] obtain n where "\\\mark_guards f c,s\ =n\ t" .. from execn_mark_guards_to_execn [OF this] show ?thesis by (blast intro: execn_to_exec) qed (* ************************************************************************* *) subsection \Lemmas about @{const "strip_guards"}\ (* ************************************************************************* *) lemma execn_to_execn_strip_guards: assumes exec_c: "\\\c,s\ =n\ t" assumes t_not_Fault: "\ isFault t" shows "\\\strip_guards F c,s\ =n\ t " using exec_c t_not_Fault [simplified not_isFault_iff] by (induct) (auto intro: execn.intros dest: noFaultn_startD') lemma execn_to_execn_strip_guards_Fault: assumes exec_c: "\\\c,s\ =n\ t" shows "\f. \t=Fault f; f \ F\ \ \\\strip_guards F c,s\ =n\ Fault f" using exec_c proof (induct) case Skip thus ?case by auto next case Guard thus ?case by (fastforce intro: execn.intros) next case GuardFault thus ?case by (fastforce intro: execn.intros) next case FaultProp thus ?case by auto next case Basic thus ?case by auto next case Spec thus ?case by auto next case SpecStuck thus ?case by auto next case (Seq c1 s n w c2 t) have exec_c1: "\\\c1,Normal s\ =n\ w" by fact have exec_c2: "\\\c2,w\ =n\ t" by fact have t: "t=Fault f" by fact have notinF: "f \ F" by fact show ?case proof (cases w) case (Fault f') with exec_c2 t have "f'=f" by (auto dest: execn_Fault_end) with Fault notinF Seq.hyps have "\\\strip_guards F c1,Normal s\ =n\ Fault f" by auto moreover have "\\\strip_guards F c2,Fault f\ =n\ Fault f" by auto ultimately show ?thesis by (auto intro: execn.intros) next case (Normal s') with execn_to_execn_strip_guards [OF exec_c1] have exec_strip_c1: "\\\strip_guards F c1,Normal s\ =n\ w" by simp with Seq.hyps t notinF have "\\\strip_guards F c2,w\ =n\ Fault f" by blast with exec_strip_c1 show ?thesis by (auto intro: execn.intros) next case (Abrupt s') with execn_to_execn_strip_guards [OF exec_c1] have exec_strip_c1: "\\\strip_guards F c1,Normal s\ =n\ w" by simp with Seq.hyps t notinF have "\\\strip_guards F c2,w\ =n\ Fault f" by (auto intro: execn.intros) with exec_strip_c1 show ?thesis by (auto intro: execn.intros) next case Stuck with exec_c2 have "t=Stuck" by (auto dest: execn_Stuck_end) with t show ?thesis by simp qed next case CondTrue thus ?case by (fastforce intro: execn.intros) next case CondFalse thus ?case by (fastforce intro: execn.intros) next case (WhileTrue s b c n w t) have exec_c: "\\\c,Normal s\ =n\ w" by fact have exec_w: "\\\While b c,w\ =n\ t" by fact have t: "t = Fault f" by fact have notinF: "f \ F" by fact have s_in_b: "s \ b" by fact show ?case proof (cases w) case (Fault f') with exec_w t have "f'=f" by (auto dest: execn_Fault_end) with Fault notinF WhileTrue.hyps have "\\\strip_guards F c,Normal s\ =n\ Fault f" by auto moreover have "\\\strip_guards F (While b c),Fault f\ =n\ Fault f" by auto ultimately show ?thesis using s_in_b by (auto intro: execn.intros) next case (Normal s') with execn_to_execn_strip_guards [OF exec_c] have exec_strip_c: "\\\strip_guards F c,Normal s\ =n\ w" by simp with WhileTrue.hyps t notinF have "\\\strip_guards F (While b c),w\ =n\ Fault f" by blast with exec_strip_c s_in_b show ?thesis by (auto intro: execn.intros) next case (Abrupt s') with execn_to_execn_strip_guards [OF exec_c] have exec_strip_c: "\\\strip_guards F c,Normal s\ =n\ w" by simp with WhileTrue.hyps t notinF have "\\\strip_guards F (While b c),w\ =n\ Fault f" by (auto intro: execn.intros) with exec_strip_c s_in_b show ?thesis by (auto intro: execn.intros) next case Stuck with exec_w have "t=Stuck" by (auto dest: execn_Stuck_end) with t show ?thesis by simp qed next case WhileFalse thus ?case by (fastforce intro: execn.intros) next case Call thus ?case by (fastforce intro: execn.intros) next case CallUndefined thus ?case by simp next case StuckProp thus ?case by simp next case DynCom thus ?case by (fastforce intro: execn.intros) next case Throw thus ?case by simp next case AbruptProp thus ?case by simp next case (CatchMatch c1 s n w c2 t) have exec_c1: "\\\c1,Normal s\ =n\ Abrupt w" by fact have exec_c2: "\\\c2,Normal w\ =n\ t" by fact have t: "t = Fault f" by fact have notinF: "f \ F" by fact from execn_to_execn_strip_guards [OF exec_c1] have exec_strip_c1: "\\\strip_guards F c1,Normal s\ =n\ Abrupt w" by simp with CatchMatch.hyps t notinF have "\\\strip_guards F c2,Normal w\ =n\ Fault f" by blast with exec_strip_c1 show ?case by (auto intro: execn.intros) next case CatchMiss thus ?case by (fastforce intro: execn.intros) qed lemma execn_to_execn_strip_guards': assumes exec_c: "\\\c,s\ =n\ t" assumes t_not_Fault: "t \ Fault ` F" shows "\\\strip_guards F c,s\ =n\ t" proof (cases t) case (Fault f) with t_not_Fault exec_c show ?thesis by (auto intro: execn_to_execn_strip_guards_Fault) qed (insert exec_c, auto intro: execn_to_execn_strip_guards) lemma execn_strip_guards_to_execn: "\s n t. \\\strip_guards F c,s\ =n\ t \ \t'. \\\c,s\ =n\ t' \ (isFault t \ isFault t') \ (t' \ Fault ` (- F) \ t'=t) \ (\ isFault t' \ t'=t)" proof (induct c) case Skip thus ?case by auto next case Basic thus ?case by auto next case Spec thus ?case by auto next case (Seq c1 c2 s n t) have exec_strip: "\\\strip_guards F (Seq c1 c2),s\ =n\ t" by fact then obtain w where exec_strip_c1: "\\\strip_guards F c1,s\ =n\ w" and exec_strip_c2: "\\\strip_guards F c2,w\ =n\ t" by (auto elim: execn_elim_cases) from Seq.hyps exec_strip_c1 obtain w' where exec_c1: "\\\c1,s\ =n\ w'" and w_Fault: "isFault w \ isFault w'" and w'_Fault: "w' \ Fault ` (- F) \ w'=w" and w'_noFault: "\ isFault w' \ w'=w" by blast show ?case proof (cases "s") case (Fault f) with exec_strip have "t=Fault f" by (auto dest: execn_Fault_end) with Fault show ?thesis by auto next case Stuck with exec_strip have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck show ?thesis by auto next case (Abrupt s') with exec_strip have "t=Abrupt s'" by (auto dest: execn_Abrupt_end) with Abrupt show ?thesis by auto next case (Normal s') show ?thesis proof (cases "isFault w") case True then obtain f where w': "w=Fault f".. moreover with exec_strip_c2 have t: "t=Fault f" by (auto dest: execn_Fault_end) ultimately show ?thesis using Normal w_Fault w'_Fault exec_c1 by (fastforce intro: execn.intros elim: isFaultE) next case False note noFault_w = this show ?thesis proof (cases "isFault w'") case True then obtain f' where w': "w'=Fault f'".. with Normal exec_c1 have exec: "\\\Seq c1 c2,s\ =n\ Fault f'" by (auto intro: execn.intros) from w'_Fault w' noFault_w have "f' \ F" by (cases w) auto with exec show ?thesis by auto next case False with w'_noFault have w': "w'=w" by simp from Seq.hyps exec_strip_c2 obtain t' where "\\\c2,w\ =n\ t'" and "isFault t \ isFault t'" and "t' \ Fault ` (-F) \ t'=t" and "\ isFault t' \ t'=t" by blast with Normal exec_c1 w' show ?thesis by (fastforce intro: execn.intros) qed qed qed next next case (Cond b c1 c2 s n t) have exec_strip: "\\\strip_guards F (Cond b c1 c2),s\ =n\ t" by fact show ?case proof (cases s) case (Fault f) with exec_strip have "t=Fault f" by (auto dest: execn_Fault_end) with Fault show ?thesis by auto next case Stuck with exec_strip have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck show ?thesis by auto next case (Abrupt s') with exec_strip have "t=Abrupt s'" by (auto dest: execn_Abrupt_end) with Abrupt show ?thesis by auto next case (Normal s') show ?thesis proof (cases "s'\ b") case True with Normal exec_strip have "\\\strip_guards F c1 ,Normal s'\ =n\ t" by (auto elim: execn_Normal_elim_cases) with Normal True Cond.hyps obtain t' where "\\\c1,Normal s'\ =n\ t'" "isFault t \ isFault t'" "t' \ Fault ` (-F) \ t'=t" "\ isFault t' \ t' = t" by blast with Normal True show ?thesis by (blast intro: execn.intros) next case False with Normal exec_strip have "\\\strip_guards F c2 ,Normal s'\ =n\ t" by (auto elim: execn_Normal_elim_cases) with Normal False Cond.hyps obtain t' where "\\\c2,Normal s'\ =n\ t'" "isFault t \ isFault t'" "t' \ Fault ` (-F) \ t'=t" "\ isFault t' \ t' = t" by blast with Normal False show ?thesis by (blast intro: execn.intros) qed qed next case (While b c s n t) have exec_strip: "\\\strip_guards F (While b c),s\ =n\ t" by fact show ?case proof (cases s) case (Fault f) with exec_strip have "t=Fault f" by (auto dest: execn_Fault_end) with Fault show ?thesis by auto next case Stuck with exec_strip have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck show ?thesis by auto next case (Abrupt s') with exec_strip have "t=Abrupt s'" by (auto dest: execn_Abrupt_end) with Abrupt show ?thesis by auto next case (Normal s') { fix c' r w assume exec_c': "\\\c',r\ =n\ w" assume c': "c'=While b (strip_guards F c)" have "\w'. \\\While b c,r\ =n\ w' \ (isFault w \ isFault w') \ (w' \ Fault ` (-F) \ w'=w) \ (\ isFault w' \ w'=w)" using exec_c' c' proof (induct) case (WhileTrue r b' c'' n u w) have eqs: "While b' c'' = While b (strip_guards F c)" by fact from WhileTrue.hyps eqs have r_in_b: "r\b" by simp from WhileTrue.hyps eqs have exec_strip_c: "\\\strip_guards F c,Normal r\ =n\ u" by simp from WhileTrue.hyps eqs have exec_strip_w: "\\\While b (strip_guards F c),u\ =n\ w" by simp show ?case proof - from WhileTrue.hyps eqs have "\\\strip_guards F c,Normal r\ =n\ u" by simp with While.hyps obtain u' where exec_c: "\\\c,Normal r\ =n\ u'" and u_Fault: "isFault u \ isFault u'" and u'_Fault: "u' \ Fault ` (-F) \ u'=u" and u'_noFault: "\ isFault u' \ u'=u" by blast show ?thesis proof (cases "isFault u'") case False with u'_noFault have u': "u'=u" by simp from WhileTrue.hyps eqs obtain w' where "\\\While b c,u\ =n\ w'" "isFault w \ isFault w'" "w' \ Fault ` (-F) \ w'=w" "\ isFault w' \ w' = w" by blast with u' exec_c r_in_b show ?thesis by (blast intro: execn.WhileTrue) next case True then obtain f' where u': "u'=Fault f'".. with exec_c r_in_b have exec: "\\\While b c,Normal r\ =n\ Fault f'" by (blast intro: execn.intros) show ?thesis proof (cases "isFault u") case True then obtain f where u: "u=Fault f".. with exec_strip_w have "w=Fault f" by (auto dest: execn_Fault_end) with exec u' u u'_Fault show ?thesis by auto next case False with u'_Fault u' have "f' \ F" by (cases u) auto with exec show ?thesis by auto qed qed qed next case (WhileFalse r b' c'' n) have eqs: "While b' c'' = While b (strip_guards F c)" by fact from WhileFalse.hyps eqs have r_not_in_b: "r\b" by simp show ?case proof - from r_not_in_b have "\\\While b c,Normal r\ =n\ Normal r" by (rule execn.WhileFalse) thus ?thesis by blast qed qed auto } note hyp_while = this show ?thesis proof (cases "s'\b") case False with Normal exec_strip have "t=s" by (auto elim: execn_Normal_elim_cases) with Normal False show ?thesis by (auto intro: execn.intros) next case True note s'_in_b = this with Normal exec_strip obtain r where exec_strip_c: "\\\strip_guards F c,Normal s'\ =n\ r" and exec_strip_w: "\\\While b (strip_guards F c),r\ =n\ t" by (auto elim: execn_Normal_elim_cases) from While.hyps exec_strip_c obtain r' where exec_c: "\\\c,Normal s'\ =n\ r'" and r_Fault: "isFault r \ isFault r'" and r'_Fault: "r' \ Fault ` (-F) \ r'=r" and r'_noFault: "\ isFault r' \ r'=r" by blast show ?thesis proof (cases "isFault r'") case False with r'_noFault have r': "r'=r" by simp from hyp_while exec_strip_w obtain t' where "\\\While b c,r\ =n\ t'" "isFault t \ isFault t'" "t' \ Fault ` (-F) \ t'=t" "\ isFault t' \ t'=t" by blast with r' exec_c Normal s'_in_b show ?thesis by (blast intro: execn.intros) next case True then obtain f' where r': "r'=Fault f'".. hence "\\\While b c,r'\ =n\ Fault f'" by auto with Normal s'_in_b exec_c have exec: "\\\While b c,Normal s'\ =n\ Fault f'" by (auto intro: execn.intros) show ?thesis proof (cases "isFault r") case True then obtain f where r: "r=Fault f".. with exec_strip_w have "t=Fault f" by (auto dest: execn_Fault_end) with Normal exec r' r r'_Fault show ?thesis by auto next case False with r'_Fault r' have "f' \ F" by (cases r) auto with Normal exec show ?thesis by auto qed qed qed qed next case Call thus ?case by auto next case DynCom thus ?case by (fastforce elim!: execn_elim_cases intro: execn.intros) next case (Guard f g c s n t) have exec_strip: "\\\strip_guards F (Guard f g c),s\ =n\ t" by fact show ?case proof (cases s) case (Fault f) with exec_strip have "t=Fault f" by (auto dest: execn_Fault_end) with Fault show ?thesis by auto next case Stuck with exec_strip have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck show ?thesis by auto next case (Abrupt s') with exec_strip have "t=Abrupt s'" by (auto dest: execn_Abrupt_end) with Abrupt show ?thesis by auto next case (Normal s') show ?thesis proof (cases "f\F") case True with exec_strip Normal have exec_strip_c: "\\\strip_guards F c,Normal s'\ =n\ t" by simp with Guard.hyps obtain t' where "\\\c,Normal s'\ =n\ t'" and "isFault t \ isFault t'" and "t' \ Fault ` (-F) \ t'=t" and "\ isFault t' \ t'=t" by blast with Normal True show ?thesis by (cases "s'\ g") (fastforce intro: execn.intros)+ next case False note f_notin_F = this show ?thesis proof (cases "s'\g") case False with Normal exec_strip f_notin_F have t: "t=Fault f" by (auto elim: execn_Normal_elim_cases) from False have "\\\Guard f g c,Normal s'\ =n\ Fault f" by (blast intro: execn.intros) with False Normal t show ?thesis by auto next case True with exec_strip Normal f_notin_F have "\\\strip_guards F c,Normal s'\ =n\ t" by (auto elim: execn_Normal_elim_cases) with Guard.hyps obtain t' where "\\\c,Normal s'\ =n\ t'" and "isFault t \ isFault t'" and "t' \ Fault ` (-F) \ t'=t" and "\ isFault t' \ t'=t" by blast with Normal True show ?thesis by (blast intro: execn.intros) qed qed qed next case Throw thus ?case by auto next case (Catch c1 c2 s n t) have exec_strip: "\\\strip_guards F (Catch c1 c2),s\ =n\ t" by fact show ?case proof (cases "s") case (Fault f) with exec_strip have "t=Fault f" by (auto dest: execn_Fault_end) with Fault show ?thesis by auto next case Stuck with exec_strip have "t=Stuck" by (auto dest: execn_Stuck_end) with Stuck show ?thesis by auto next case (Abrupt s') with exec_strip have "t=Abrupt s'" by (auto dest: execn_Abrupt_end) with Abrupt show ?thesis by auto next case (Normal s') note s=this with exec_strip have "\\\Catch (strip_guards F c1) (strip_guards F c2),Normal s'\ =n\ t" by simp thus ?thesis proof (cases) fix w assume exec_strip_c1: "\\\strip_guards F c1,Normal s'\ =n\ Abrupt w" assume exec_strip_c2: "\\\strip_guards F c2,Normal w\ =n\ t" from exec_strip_c1 Catch.hyps obtain w' where exec_c1: "\\\c1,Normal s'\ =n\ w'" and w'_Fault: "w' \ Fault ` (-F) \ w'=Abrupt w" and w'_noFault: "\ isFault w' \ w'=Abrupt w" by blast show ?thesis proof (cases "w'") case (Fault f') with Normal exec_c1 have "\\\Catch c1 c2,s\ =n\ Fault f'" by (auto intro: execn.intros) with w'_Fault Fault show ?thesis by auto next case Stuck with w'_noFault have False by simp thus ?thesis .. next case (Normal w'') with w'_noFault have False by simp thus ?thesis .. next case (Abrupt w'') with w'_noFault have w'': "w''=w" by simp from exec_strip_c2 Catch.hyps obtain t' where "\\\c2,Normal w\ =n\ t'" "isFault t \ isFault t'" "t' \ Fault ` (-F) \ t'=t" "\ isFault t' \ t'=t" by blast with w'' Abrupt s exec_c1 show ?thesis by (blast intro: execn.intros) qed next assume t: "\ isAbr t" assume "\\\strip_guards F c1,Normal s'\ =n\ t" with Catch.hyps obtain t' where exec_c1: "\\\c1,Normal s'\ =n\ t'" and t_Fault: "isFault t \ isFault t'" and t'_Fault: "t' \ Fault ` (-F) \ t'=t" and t'_noFault: "\ isFault t' \ t'=t" by blast show ?thesis proof (cases "isFault t'") case True then obtain f' where t': "t'=Fault f'".. with exec_c1 have "\\\Catch c1 c2,Normal s'\ =n\ Fault f'" by (auto intro: execn.intros) with t'_Fault t' s show ?thesis by auto next case False with t'_noFault have "t'=t" by simp with t exec_c1 s show ?thesis by (blast intro: execn.intros) qed qed qed qed lemma execn_strip_to_execn: assumes exec_strip: "strip F \\\c,s\ =n\ t" shows "\t'. \\\c,s\ =n\ t' \ (isFault t \ isFault t') \ (t' \ Fault ` (- F) \ t'=t) \ (\ isFault t' \ t'=t)" using exec_strip proof (induct) case Skip thus ?case by (blast intro: execn.intros) next case Guard thus ?case by (blast intro: execn.intros) next case GuardFault thus ?case by (blast intro: execn.intros) next case FaultProp thus ?case by (blast intro: execn.intros) next case Basic thus ?case by (blast intro: execn.intros) next case Spec thus ?case by (blast intro: execn.intros) next case SpecStuck thus ?case by (blast intro: execn.intros) next case Seq thus ?case by (blast intro: execn.intros elim: isFaultE) next case CondTrue thus ?case by (blast intro: execn.intros) next case CondFalse thus ?case by (blast intro: execn.intros) next case WhileTrue thus ?case by (blast intro: execn.intros elim: isFaultE) next case WhileFalse thus ?case by (blast intro: execn.intros) next case Call thus ?case by simp (blast intro: execn.intros dest: execn_strip_guards_to_execn) next case CallUndefined thus ?case by simp (blast intro: execn.intros) next case StuckProp thus ?case by blast next case DynCom thus ?case by (blast intro: execn.intros) next case Throw thus ?case by (blast intro: execn.intros) next case AbruptProp thus ?case by (blast intro: execn.intros) next case (CatchMatch c1 s n r c2 t) then obtain r' t' where exec_c1: "\\\c1,Normal s\ =n\ r'" and r'_Fault: "r' \ Fault ` (-F) \ r' = Abrupt r" and r'_noFault: "\ isFault r' \ r' = Abrupt r" and exec_c2: "\\\c2,Normal r\ =n\ t'" and t_Fault: "isFault t \ isFault t'" and t'_Fault: "t' \ Fault ` (-F) \ t' = t" and t'_noFault: "\ isFault t' \ t' = t" by blast show ?case proof (cases "isFault r'") case True then obtain f' where r': "r'=Fault f'".. with exec_c1 have "\\\Catch c1 c2,Normal s\ =n\ Fault f'" by (auto intro: execn.intros) with r' r'_Fault show ?thesis by (auto intro: execn.intros) next case False with r'_noFault have "r'=Abrupt r" by simp with exec_c1 exec_c2 t_Fault t'_noFault t'_Fault show ?thesis by (blast intro: execn.intros) qed next case CatchMiss thus ?case by (fastforce intro: execn.intros elim: isFaultE) qed lemma exec_strip_guards_to_exec: assumes exec_strip: "\\\strip_guards F c,s\ \ t" shows "\t'. \\\c,s\ \ t' \ (isFault t \ isFault t') \ (t' \ Fault ` (-F) \ t'=t) \ (\ isFault t' \ t'=t)" proof - from exec_strip obtain n where execn_strip: "\\\strip_guards F c,s\ =n\ t" by (auto simp add: exec_iff_execn) then obtain t' where "\\\c,s\ =n\ t'" "isFault t \ isFault t'" "t' \ Fault ` (-F) \ t'=t" "\ isFault t' \ t'=t" by (blast dest: execn_strip_guards_to_execn) thus ?thesis by (blast intro: execn_to_exec) qed lemma exec_strip_to_exec: assumes exec_strip: "strip F \\\c,s\ \ t" shows "\t'. \\\c,s\ \ t' \ (isFault t \ isFault t') \ (t' \ Fault ` (-F) \ t'=t) \ (\ isFault t' \ t'=t)" proof - from exec_strip obtain n where execn_strip: "strip F \\\c,s\ =n\ t" by (auto simp add: exec_iff_execn) then obtain t' where "\\\c,s\ =n\ t'" "isFault t \ isFault t'" "t' \ Fault ` (-F) \ t'=t" "\ isFault t' \ t'=t" by (blast dest: execn_strip_to_execn) thus ?thesis by (blast intro: execn_to_exec) qed lemma exec_to_exec_strip_guards: assumes exec_c: "\\\c,s\ \ t" assumes t_not_Fault: "\ isFault t" shows "\\\strip_guards F c,s\ \ t" proof - from exec_c obtain n where "\\\c,s\ =n\t" by (auto simp add: exec_iff_execn) from this t_not_Fault have "\\\strip_guards F c,s\ =n\ t" by (rule execn_to_execn_strip_guards ) thus "\\\strip_guards F c,s\ \ t" by (rule execn_to_exec) qed lemma exec_to_exec_strip_guards': assumes exec_c: "\\\c,s\ \ t" assumes t_not_Fault: "t \ Fault ` F" shows "\\\strip_guards F c,s\ \ t" proof - from exec_c obtain n where "\\\c,s\ =n\t" by (auto simp add: exec_iff_execn) from this t_not_Fault have "\\\strip_guards F c,s\ =n\ t" by (rule execn_to_execn_strip_guards' ) thus "\\\strip_guards F c,s\ \ t" by (rule execn_to_exec) qed lemma execn_to_execn_strip: assumes exec_c: "\\\c,s\ =n\ t" assumes t_not_Fault: "\ isFault t" shows "strip F \\\c,s\ =n\ t" using exec_c t_not_Fault proof (induct) case (Call p bdy s n s') have bdy: "\ p = Some bdy" by fact from Call have "strip F \\\bdy,Normal s\ =n\ s'" by blast from execn_to_execn_strip_guards [OF this] Call have "strip F \\\strip_guards F bdy,Normal s\ =n\ s'" by simp moreover from bdy have "(strip F \) p = Some (strip_guards F bdy)" by simp ultimately show ?case by (blast intro: execn.intros) next case CallUndefined thus ?case by (auto intro: execn.CallUndefined) qed (auto intro: execn.intros dest: noFaultn_startD' simp add: not_isFault_iff) lemma execn_to_execn_strip': assumes exec_c: "\\\c,s\ =n\ t" assumes t_not_Fault: "t \ Fault ` F" shows "strip F \\\c,s\ =n\ t" using exec_c t_not_Fault proof (induct) case (Call p bdy s n s') have bdy: "\ p = Some bdy" by fact from Call have "strip F \\\bdy,Normal s\ =n\ s'" by blast from execn_to_execn_strip_guards' [OF this] Call have "strip F \\\strip_guards F bdy,Normal s\ =n\ s'" by simp moreover from bdy have "(strip F \) p = Some (strip_guards F bdy)" by simp ultimately show ?case by (blast intro: execn.intros) next case CallUndefined thus ?case by (auto intro: execn.CallUndefined) next case (Seq c1 s n s' c2 t) show ?case proof (cases "isFault s'") case False with Seq show ?thesis by (auto intro: execn.intros simp add: not_isFault_iff) next case True then obtain f' where s': "s'=Fault f'" by (auto simp add: isFault_def) with Seq obtain "t=Fault f'" and "f' \ F" by (force dest: execn_Fault_end) with Seq s' show ?thesis by (auto intro: execn.intros) qed next case (WhileTrue b c s n s' t) show ?case proof (cases "isFault s'") case False with WhileTrue show ?thesis by (auto intro: execn.intros simp add: not_isFault_iff) next case True then obtain f' where s': "s'=Fault f'" by (auto simp add: isFault_def) with WhileTrue obtain "t=Fault f'" and "f' \ F" by (force dest: execn_Fault_end) with WhileTrue s' show ?thesis by (auto intro: execn.intros) qed qed (auto intro: execn.intros) lemma exec_to_exec_strip: assumes exec_c: "\\\c,s\ \ t" assumes t_not_Fault: "\ isFault t" shows "strip F \\\c,s\ \ t" proof - from exec_c obtain n where "\\\c,s\ =n\t" by (auto simp add: exec_iff_execn) from this t_not_Fault have "strip F \\\c,s\ =n\ t" by (rule execn_to_execn_strip) thus "strip F \\\c,s\ \ t" by (rule execn_to_exec) qed lemma exec_to_exec_strip': assumes exec_c: "\\\c,s\ \ t" assumes t_not_Fault: "t \ Fault ` F" shows "strip F \\\c,s\ \ t" proof - from exec_c obtain n where "\\\c,s\ =n\t" by (auto simp add: exec_iff_execn) from this t_not_Fault have "strip F \\\c,s\ =n\ t" by (rule execn_to_execn_strip' ) thus "strip F \\\c,s\ \ t" by (rule execn_to_exec) qed lemma exec_to_exec_strip_guards_Fault: assumes exec_c: "\\\c,s\ \ Fault f" assumes f_notin_F: "f \ F" shows"\\\strip_guards F c,s\ \ Fault f" proof - from exec_c obtain n where "\\\c,s\ =n\Fault f" by (auto simp add: exec_iff_execn) from execn_to_execn_strip_guards_Fault [OF this _ f_notin_F] have "\\\strip_guards F c,s\ =n\ Fault f" by simp thus "\\\strip_guards F c,s\ \ Fault f" by (rule execn_to_exec) qed (* ************************************************************************* *) subsection \Lemmas about @{term "c\<^sub>1 \\<^sub>g c\<^sub>2"}\ (* ************************************************************************* *) lemma inter_guards_execn_Normal_noFault: "\c c2 s t n. \(c1 \\<^sub>g c2) = Some c; \\\c,Normal s\ =n\ t; \ isFault t\ \ \\\c1,Normal s\ =n\ t \ \\\c2,Normal s\ =n\ t" proof (induct c1) case Skip have "(Skip \\<^sub>g c2) = Some c" by fact then obtain c2: "c2=Skip" and c: "c=Skip" by (simp add: inter_guards_Skip) have "\\\c,Normal s\ =n\ t" by fact with c have "t=Normal s" by (auto elim: execn_Normal_elim_cases) with Skip c2 show ?case by (auto intro: execn.intros) next case (Basic f) have "(Basic f \\<^sub>g c2) = Some c" by fact then obtain c2: "c2=Basic f" and c: "c=Basic f" by (simp add: inter_guards_Basic) have "\\\c,Normal s\ =n\ t" by fact with c have "t=Normal (f s)" by (auto elim: execn_Normal_elim_cases) with Basic c2 show ?case by (auto intro: execn.intros) next case (Spec r) have "(Spec r \\<^sub>g c2) = Some c" by fact then obtain c2: "c2=Spec r" and c: "c=Spec r" by (simp add: inter_guards_Spec) have "\\\c,Normal s\ =n\ t" by fact with c have "\\\Spec r,Normal s\ =n\ t" by simp from this Spec c2 show ?case by (cases) (auto intro: execn.intros) next case (Seq a1 a2) have noFault: "\ isFault t" by fact have "(Seq a1 a2 \\<^sub>g c2) = Some c" by fact then obtain b1 b2 d1 d2 where c2: "c2=Seq b1 b2" and d1: "(a1 \\<^sub>g b1) = Some d1" and d2: "(a2 \\<^sub>g b2) = Some d2" and c: "c=Seq d1 d2" by (auto simp add: inter_guards_Seq) have "\\\c,Normal s\ =n\ t" by fact with c obtain s' where exec_d1: "\\\d1,Normal s\ =n\ s'" and exec_d2: "\\\d2,s'\ =n\ t" by (auto elim: execn_Normal_elim_cases) show ?case proof (cases s') case (Fault f') with exec_d2 have "t=Fault f'" by (auto intro: execn_Fault_end) with noFault show ?thesis by simp next case (Normal s'') with d1 exec_d1 Seq.hyps obtain "\\\a1,Normal s\ =n\ Normal s''" and "\\\b1,Normal s\ =n\ Normal s''" by auto moreover from Normal d2 exec_d2 noFault Seq.hyps obtain "\\\a2,Normal s''\ =n\ t" and "\\\b2,Normal s''\ =n\ t" by auto ultimately show ?thesis using Normal c2 by (auto intro: execn.intros) next case (Abrupt s'') with exec_d2 have "t=Abrupt s''" by (auto simp add: execn_Abrupt_end) moreover from Abrupt d1 exec_d1 Seq.hyps obtain "\\\a1,Normal s\ =n\ Abrupt s''" and "\\\b1,Normal s\ =n\ Abrupt s''" by auto moreover obtain "\\\a2,Abrupt s''\ =n\ Abrupt s''" and "\\\b2,Abrupt s''\ =n\ Abrupt s''" by auto ultimately show ?thesis using Abrupt c2 by (auto intro: execn.intros) next case Stuck with exec_d2 have "t=Stuck" by (auto simp add: execn_Stuck_end) moreover from Stuck d1 exec_d1 Seq.hyps obtain "\\\a1,Normal s\ =n\ Stuck" and "\\\b1,Normal s\ =n\ Stuck" by auto moreover obtain "\\\a2,Stuck\ =n\ Stuck" and "\\\b2,Stuck\ =n\ Stuck" by auto ultimately show ?thesis using Stuck c2 by (auto intro: execn.intros) qed next case (Cond b t1 e1) have noFault: "\ isFault t" by fact have "(Cond b t1 e1 \\<^sub>g c2) = Some c" by fact then obtain t2 e2 t3 e3 where c2: "c2=Cond b t2 e2" and t3: "(t1 \\<^sub>g t2) = Some t3" and e3: "(e1 \\<^sub>g e2) = Some e3" and c: "c=Cond b t3 e3" by (auto simp add: inter_guards_Cond) have "\\\c,Normal s\ =n\ t" by fact with c have "\\\Cond b t3 e3,Normal s\ =n\ t" by simp then show ?case proof (cases) assume s_in_b: "s\b" assume "\\\t3,Normal s\ =n\ t" with Cond.hyps t3 noFault obtain "\\\t1,Normal s\ =n\ t" "\\\t2,Normal s\ =n\ t" by auto with s_in_b c2 show ?thesis by (auto intro: execn.intros) next assume s_notin_b: "s\b" assume "\\\e3,Normal s\ =n\ t" with Cond.hyps e3 noFault obtain "\\\e1,Normal s\ =n\ t" "\\\e2,Normal s\ =n\ t" by auto with s_notin_b c2 show ?thesis by (auto intro: execn.intros) qed next case (While b bdy1) have noFault: "\ isFault t" by fact have "(While b bdy1 \\<^sub>g c2) = Some c" by fact then obtain bdy2 bdy where c2: "c2=While b bdy2" and bdy: "(bdy1 \\<^sub>g bdy2) = Some bdy" and c: "c=While b bdy" by (auto simp add: inter_guards_While) have exec_c: "\\\c,Normal s\ =n\ t" by fact { fix s t n w w1 w2 assume exec_w: "\\\w,Normal s\ =n\ t" assume w: "w=While b bdy" assume noFault: "\ isFault t" from exec_w w noFault have "\\\While b bdy1,Normal s\ =n\ t \ \\\While b bdy2,Normal s\ =n\ t" proof (induct) prefer 10 case (WhileTrue s b' bdy' n s' s'') have eqs: "While b' bdy' = While b bdy" by fact from WhileTrue have s_in_b: "s \ b" by simp have noFault_s'': "\ isFault s''" by fact from WhileTrue have exec_bdy: "\\\bdy,Normal s\ =n\ s'" by simp from WhileTrue have exec_w: "\\\While b bdy,s'\ =n\ s''" by simp show ?case proof (cases s') case (Fault f) with exec_w have "s''=Fault f" by (auto intro: execn_Fault_end) with noFault_s'' show ?thesis by simp next case (Normal s''') with exec_bdy bdy While.hyps obtain "\\\bdy1,Normal s\ =n\ Normal s'''" "\\\bdy2,Normal s\ =n\ Normal s'''" by auto moreover from Normal WhileTrue obtain "\\\While b bdy1,Normal s'''\ =n\ s''" "\\\While b bdy2,Normal s'''\ =n\ s''" by simp ultimately show ?thesis using s_in_b Normal by (auto intro: execn.intros) next case (Abrupt s''') with exec_bdy bdy While.hyps obtain "\\\bdy1,Normal s\ =n\ Abrupt s'''" "\\\bdy2,Normal s\ =n\ Abrupt s'''" by auto moreover from Abrupt WhileTrue obtain "\\\While b bdy1,Abrupt s'''\ =n\ s''" "\\\While b bdy2,Abrupt s'''\ =n\ s''" by simp ultimately show ?thesis using s_in_b Abrupt by (auto intro: execn.intros) next case Stuck with exec_bdy bdy While.hyps obtain "\\\bdy1,Normal s\ =n\ Stuck" "\\\bdy2,Normal s\ =n\ Stuck" by auto moreover from Stuck WhileTrue obtain "\\\While b bdy1,Stuck\ =n\ s''" "\\\While b bdy2,Stuck\ =n\ s''" by simp ultimately show ?thesis using s_in_b Stuck by (auto intro: execn.intros) qed next case WhileFalse thus ?case by (auto intro: execn.intros) qed (simp_all) } with this [OF exec_c c noFault] c2 show ?case by auto next case Call thus ?case by (simp add: inter_guards_Call) next case (DynCom f1) have noFault: "\ isFault t" by fact have "(DynCom f1 \\<^sub>g c2) = Some c" by fact then obtain f2 f where c2: "c2=DynCom f2" and f_defined: "\s. ((f1 s) \\<^sub>g (f2 s)) \ None" and c: "c=DynCom (\s. the ((f1 s) \\<^sub>g (f2 s)))" by (auto simp add: inter_guards_DynCom) have "\\\c,Normal s\ =n\ t" by fact with c have "\\\DynCom (\s. the ((f1 s) \\<^sub>g (f2 s))),Normal s\ =n\ t" by simp then show ?case proof (cases) assume exec_f: "\\\the (f1 s \\<^sub>g f2 s),Normal s\ =n\ t" from f_defined obtain f where "(f1 s \\<^sub>g f2 s) = Some f" by auto with DynCom.hyps this exec_f c2 noFault show ?thesis using execn.DynCom by fastforce qed next case Guard thus ?case by (fastforce elim: execn_Normal_elim_cases intro: execn.intros simp add: inter_guards_Guard) next case Throw thus ?case by (fastforce elim: execn_Normal_elim_cases simp add: inter_guards_Throw) next case (Catch a1 a2) have noFault: "\ isFault t" by fact have "(Catch a1 a2 \\<^sub>g c2) = Some c" by fact then obtain b1 b2 d1 d2 where c2: "c2=Catch b1 b2" and d1: "(a1 \\<^sub>g b1) = Some d1" and d2: "(a2 \\<^sub>g b2) = Some d2" and c: "c=Catch d1 d2" by (auto simp add: inter_guards_Catch) have "\\\c,Normal s\ =n\ t" by fact with c have "\\\Catch d1 d2,Normal s\ =n\ t" by simp then show ?case proof (cases) fix s' assume "\\\d1,Normal s\ =n\ Abrupt s'" with d1 Catch.hyps obtain "\\\a1,Normal s\ =n\ Abrupt s'" and "\\\b1,Normal s\ =n\ Abrupt s'" by auto moreover assume "\\\d2,Normal s'\ =n\ t" with d2 Catch.hyps noFault obtain "\\\a2,Normal s'\ =n\ t" and "\\\b2,Normal s'\ =n\ t" by auto ultimately show ?thesis using c2 by (auto intro: execn.intros) next assume "\ isAbr t" moreover assume "\\\d1,Normal s\ =n\ t" with d1 Catch.hyps noFault obtain "\\\a1,Normal s\ =n\ t" and "\\\b1,Normal s\ =n\ t" by auto ultimately show ?thesis using c2 by (auto intro: execn.intros) qed qed lemma inter_guards_execn_noFault: assumes c: "(c1 \\<^sub>g c2) = Some c" assumes exec_c: "\\\c,s\ =n\ t" assumes noFault: "\ isFault t" shows "\\\c1,s\ =n\ t \ \\\c2,s\ =n\ t" proof (cases s) case (Fault f) with exec_c have "t = Fault f" by (auto intro: execn_Fault_end) with noFault show ?thesis by simp next case (Abrupt s') with exec_c have "t=Abrupt s'" by (simp add: execn_Abrupt_end) with Abrupt show ?thesis by auto next case Stuck with exec_c have "t=Stuck" by (simp add: execn_Stuck_end) with Stuck show ?thesis by auto next case (Normal s') with exec_c noFault inter_guards_execn_Normal_noFault [OF c] show ?thesis by blast qed lemma inter_guards_exec_noFault: assumes c: "(c1 \\<^sub>g c2) = Some c" assumes exec_c: "\\\c,s\ \ t" assumes noFault: "\ isFault t" shows "\\\c1,s\ \ t \ \\\c2,s\ \ t" proof - from exec_c obtain n where "\\\c,s\ =n\ t" by (auto simp add: exec_iff_execn) from c this noFault have "\\\c1,s\ =n\ t \ \\\c2,s\ =n\ t" by (rule inter_guards_execn_noFault) thus ?thesis by (auto intro: execn_to_exec) qed lemma inter_guards_execn_Normal_Fault: "\c c2 s n. \(c1 \\<^sub>g c2) = Some c; \\\c,Normal s\ =n\ Fault f\ \ (\\\c1,Normal s\ =n\ Fault f \ \\\c2,Normal s\ =n\ Fault f)" proof (induct c1) case Skip thus ?case by (fastforce simp add: inter_guards_Skip) next case (Basic f) thus ?case by (fastforce simp add: inter_guards_Basic) next case (Spec r) thus ?case by (fastforce simp add: inter_guards_Spec) next case (Seq a1 a2) have "(Seq a1 a2 \\<^sub>g c2) = Some c" by fact then obtain b1 b2 d1 d2 where c2: "c2=Seq b1 b2" and d1: "(a1 \\<^sub>g b1) = Some d1" and d2: "(a2 \\<^sub>g b2) = Some d2" and c: "c=Seq d1 d2" by (auto simp add: inter_guards_Seq) have "\\\c,Normal s\ =n\ Fault f" by fact with c obtain s' where exec_d1: "\\\d1,Normal s\ =n\ s'" and exec_d2: "\\\d2,s'\ =n\ Fault f" by (auto elim: execn_Normal_elim_cases) show ?case proof (cases s') case (Fault f') with exec_d2 have "f'=f" by (auto dest: execn_Fault_end) with Fault d1 exec_d1 have "\\\a1,Normal s\ =n\ Fault f \ \\\b1,Normal s\ =n\ Fault f" by (auto dest: Seq.hyps) thus ?thesis proof (cases rule: disjE [consumes 1]) assume "\\\a1,Normal s\ =n\ Fault f" hence "\\\Seq a1 a2,Normal s\ =n\ Fault f" by (auto intro: execn.intros) thus ?thesis by simp next assume "\\\b1,Normal s\ =n\ Fault f" hence "\\\Seq b1 b2,Normal s\ =n\ Fault f" by (auto intro: execn.intros) with c2 show ?thesis by simp qed next case Abrupt with exec_d2 show ?thesis by (auto dest: execn_Abrupt_end) next case Stuck with exec_d2 show ?thesis by (auto dest: execn_Stuck_end) next case (Normal s'') with inter_guards_execn_noFault [OF d1 exec_d1] obtain exec_a1: "\\\a1,Normal s\ =n\ Normal s''" and exec_b1: "\\\b1,Normal s\ =n\ Normal s''" by simp moreover from d2 exec_d2 Normal have "\\\a2,Normal s''\ =n\ Fault f \ \\\b2,Normal s''\ =n\ Fault f" by (auto dest: Seq.hyps) ultimately show ?thesis using c2 by (auto intro: execn.intros) qed next case (Cond b t1 e1) have "(Cond b t1 e1 \\<^sub>g c2) = Some c" by fact then obtain t2 e2 t e where c2: "c2=Cond b t2 e2" and t: "(t1 \\<^sub>g t2) = Some t" and e: "(e1 \\<^sub>g e2) = Some e" and c: "c=Cond b t e" by (auto simp add: inter_guards_Cond) have "\\\c,Normal s\ =n\ Fault f" by fact with c have "\\\Cond b t e,Normal s\ =n\ Fault f" by simp thus ?case proof (cases) assume "s \ b" moreover assume "\\\t,Normal s\ =n\ Fault f" with t have "\\\t1,Normal s\ =n\ Fault f \ \\\t2,Normal s\ =n\ Fault f" by (auto dest: Cond.hyps) ultimately show ?thesis using c2 c by (fastforce intro: execn.intros) next assume "s \ b" moreover assume "\\\e,Normal s\ =n\ Fault f" with e have "\\\e1,Normal s\ =n\ Fault f \ \\\e2,Normal s\ =n\ Fault f" by (auto dest: Cond.hyps) ultimately show ?thesis using c2 c by (fastforce intro: execn.intros) qed next case (While b bdy1) have "(While b bdy1 \\<^sub>g c2) = Some c" by fact then obtain bdy2 bdy where c2: "c2=While b bdy2" and bdy: "(bdy1 \\<^sub>g bdy2) = Some bdy" and c: "c=While b bdy" by (auto simp add: inter_guards_While) have exec_c: "\\\c,Normal s\ =n\ Fault f" by fact { fix s t n w w1 w2 assume exec_w: "\\\w,Normal s\ =n\ t" assume w: "w=While b bdy" assume Fault: "t=Fault f" from exec_w w Fault have "\\\While b bdy1,Normal s\ =n\ Fault f\ \\\While b bdy2,Normal s\ =n\ Fault f" proof (induct) case (WhileTrue s b' bdy' n s' s'') have eqs: "While b' bdy' = While b bdy" by fact from WhileTrue have s_in_b: "s \ b" by simp have Fault_s'': "s''=Fault f" by fact from WhileTrue have exec_bdy: "\\\bdy,Normal s\ =n\ s'" by simp from WhileTrue have exec_w: "\\\While b bdy,s'\ =n\ s''" by simp show ?case proof (cases s') case (Fault f') with exec_w Fault_s'' have "f'=f" by (auto dest: execn_Fault_end) with Fault exec_bdy bdy While.hyps have "\\\bdy1,Normal s\ =n\ Fault f \ \\\bdy2,Normal s\ =n\ Fault f" by auto with s_in_b show ?thesis by (fastforce intro: execn.intros) next case (Normal s''') with inter_guards_execn_noFault [OF bdy exec_bdy] obtain "\\\bdy1,Normal s\ =n\ Normal s'''" "\\\bdy2,Normal s\ =n\ Normal s'''" by auto moreover from Normal WhileTrue have "\\\While b bdy1,Normal s'''\ =n\ Fault f \ \\\While b bdy2,Normal s'''\ =n\ Fault f" by simp ultimately show ?thesis using s_in_b by (fastforce intro: execn.intros) next case (Abrupt s''') with exec_w Fault_s'' show ?thesis by (fastforce dest: execn_Abrupt_end) next case Stuck with exec_w Fault_s'' show ?thesis by (fastforce dest: execn_Stuck_end) qed next case WhileFalse thus ?case by (auto intro: execn.intros) qed (simp_all) } with this [OF exec_c c] c2 show ?case by auto next case Call thus ?case by (fastforce simp add: inter_guards_Call) next case (DynCom f1) have "(DynCom f1 \\<^sub>g c2) = Some c" by fact then obtain f2 where c2: "c2=DynCom f2" and F_defined: "\s. ((f1 s) \\<^sub>g (f2 s)) \ None" and c: "c=DynCom (\s. the ((f1 s) \\<^sub>g (f2 s)))" by (auto simp add: inter_guards_DynCom) have "\\\c,Normal s\ =n\ Fault f" by fact with c have "\\\DynCom (\s. the ((f1 s) \\<^sub>g (f2 s))),Normal s\ =n\ Fault f" by simp then show ?case proof (cases) assume exec_F: "\\\the (f1 s \\<^sub>g f2 s),Normal s\ =n\ Fault f" from F_defined obtain F where "(f1 s \\<^sub>g f2 s) = Some F" by auto with DynCom.hyps this exec_F c2 show ?thesis by (fastforce intro: execn.intros) qed next case (Guard m g1 bdy1) have "(Guard m g1 bdy1 \\<^sub>g c2) = Some c" by fact then obtain g2 bdy2 bdy where c2: "c2=Guard m g2 bdy2" and bdy: "(bdy1 \\<^sub>g bdy2) = Some bdy" and c: "c=Guard m (g1 \ g2) bdy" by (auto simp add: inter_guards_Guard) have "\\\c,Normal s\ =n\ Fault f" by fact with c have "\\\Guard m (g1 \ g2) bdy,Normal s\ =n\ Fault f" by simp thus ?case proof (cases) assume f_m: "Fault f = Fault m" assume "s \ g1 \ g2" hence "s\g1 \ s\g2" by blast with c2 f_m show ?thesis by (auto intro: execn.intros) next assume "s \ g1 \ g2" moreover assume "\\\bdy,Normal s\ =n\ Fault f" with bdy have "\\\bdy1,Normal s\ =n\ Fault f \ \\\bdy2,Normal s\ =n\ Fault f" by (rule Guard.hyps) ultimately show ?thesis using c2 by (auto intro: execn.intros) qed next case Throw thus ?case by (fastforce simp add: inter_guards_Throw) next case (Catch a1 a2) have "(Catch a1 a2 \\<^sub>g c2) = Some c" by fact then obtain b1 b2 d1 d2 where c2: "c2=Catch b1 b2" and d1: "(a1 \\<^sub>g b1) = Some d1" and d2: "(a2 \\<^sub>g b2) = Some d2" and c: "c=Catch d1 d2" by (auto simp add: inter_guards_Catch) have "\\\c,Normal s\ =n\ Fault f" by fact with c have "\\\Catch d1 d2,Normal s\ =n\ Fault f" by simp thus ?case proof (cases) fix s' assume "\\\d1,Normal s\ =n\ Abrupt s'" from inter_guards_execn_noFault [OF d1 this] obtain exec_a1: "\\\a1,Normal s\ =n\ Abrupt s'" and exec_b1: "\\\b1,Normal s\ =n\ Abrupt s'" by simp moreover assume "\\\d2,Normal s'\ =n\ Fault f" with d2 have "\\\a2,Normal s'\ =n\ Fault f \ \\\b2,Normal s'\ =n\ Fault f" by (auto dest: Catch.hyps) ultimately show ?thesis using c2 by (fastforce intro: execn.intros) next assume "\\\d1,Normal s\ =n\ Fault f" with d1 have "\\\a1,Normal s\ =n\ Fault f \ \\\b1,Normal s\ =n\ Fault f" by (auto dest: Catch.hyps) with c2 show ?thesis by (fastforce intro: execn.intros) qed qed lemma inter_guards_execn_Fault: assumes c: "(c1 \\<^sub>g c2) = Some c" assumes exec_c: "\\\c,s\ =n\ Fault f" shows "\\\c1,s\ =n\ Fault f \ \\\c2,s\ =n\ Fault f" proof (cases s) case (Fault f) with exec_c show ?thesis by (auto dest: execn_Fault_end) next case (Abrupt s') with exec_c show ?thesis by (fastforce dest: execn_Abrupt_end) next case Stuck with exec_c show ?thesis by (fastforce dest: execn_Stuck_end) next case (Normal s') with exec_c inter_guards_execn_Normal_Fault [OF c] show ?thesis by blast qed lemma inter_guards_exec_Fault: assumes c: "(c1 \\<^sub>g c2) = Some c" assumes exec_c: "\\\c,s\ \ Fault f" shows "\\\c1,s\ \ Fault f \ \\\c2,s\ \ Fault f" proof - from exec_c obtain n where "\\\c,s\ =n\ Fault f" by (auto simp add: exec_iff_execn) from c this have "\\\c1,s\ =n\ Fault f \ \\\c2,s\ =n\ Fault f" by (rule inter_guards_execn_Fault) thus ?thesis by (auto intro: execn_to_exec) qed (* ************************************************************************* *) subsection "Restriction of Procedure Environment" (* ************************************************************************* *) lemma restrict_SomeD: "(m|\<^bsub>A\<^esub>) x = Some y \ m x = Some y" by (auto simp add: restrict_map_def split: if_split_asm) -(* FIXME: To Map *) +(* fixme: To Map *) lemma restrict_dom_same [simp]: "m|\<^bsub>dom m\<^esub> = m" apply (rule ext) apply (clarsimp simp add: restrict_map_def) apply (simp only: not_None_eq [symmetric]) apply rule apply (drule sym) apply blast done lemma restrict_in_dom: "x \ A \ (m|\<^bsub>A\<^esub>) x = m x" by (auto simp add: restrict_map_def) lemma exec_restrict_to_exec: assumes exec_restrict: "\|\<^bsub>A\<^esub>\\c,s\ \ t" assumes notStuck: "t\Stuck" shows "\\\c,s\ \ t" using exec_restrict notStuck by (induct) (auto intro: exec.intros dest: restrict_SomeD Stuck_end) lemma execn_restrict_to_execn: assumes exec_restrict: "\|\<^bsub>A\<^esub>\\c,s\ =n\ t" assumes notStuck: "t\Stuck" shows "\\\c,s\ =n\ t" using exec_restrict notStuck by (induct) (auto intro: execn.intros dest: restrict_SomeD execn_Stuck_end) lemma restrict_NoneD: "m x = None \ (m|\<^bsub>A\<^esub>) x = None" by (auto simp add: restrict_map_def split: if_split_asm) lemma execn_to_execn_restrict: assumes execn: "\\\c,s\ =n\ t" shows "\t'. \|\<^bsub>P\<^esub>\\c,s\ =n\ t' \ (t=Stuck \ t'=Stuck) \ (\f. t=Fault f \ t'\{Fault f,Stuck}) \ (t'\Stuck \ t'=t)" using execn proof (induct) case Skip show ?case by (blast intro: execn.Skip) next case Guard thus ?case by (auto intro: execn.Guard) next case GuardFault thus ?case by (auto intro: execn.GuardFault) next case FaultProp thus ?case by (auto intro: execn.FaultProp) next case Basic thus ?case by (auto intro: execn.Basic) next case Spec thus ?case by (auto intro: execn.Spec) next case SpecStuck thus ?case by (auto intro: execn.SpecStuck) next case Seq thus ?case by (metis insertCI execn.Seq StuckProp) next case CondTrue thus ?case by (auto intro: execn.CondTrue) next case CondFalse thus ?case by (auto intro: execn.CondFalse) next case WhileTrue thus ?case by (metis insertCI execn.WhileTrue StuckProp) next case WhileFalse thus ?case by (auto intro: execn.WhileFalse) next case (Call p bdy n s s') have "\ p = Some bdy" by fact show ?case proof (cases "p \ P") case True with Call have "(\|\<^bsub>P\<^esub>) p = Some bdy" by (simp) with Call show ?thesis by (auto intro: execn.intros) next case False hence "(\|\<^bsub>P\<^esub>) p = None" by simp thus ?thesis by (auto intro: execn.CallUndefined) qed next case (CallUndefined p n s) have "\ p = None" by fact hence "(\|\<^bsub>P\<^esub>) p = None" by (rule restrict_NoneD) thus ?case by (auto intro: execn.CallUndefined) next case StuckProp thus ?case by (auto intro: execn.StuckProp) next case DynCom thus ?case by (auto intro: execn.DynCom) next case Throw thus ?case by (auto intro: execn.Throw) next case AbruptProp thus ?case by (auto intro: execn.AbruptProp) next case (CatchMatch c1 s n s' c2 s'') from CatchMatch.hyps obtain t' t'' where exec_res_c1: "\|\<^bsub>P\<^esub>\\c1,Normal s\ =n\ t'" and t'_notStuck: "t' \ Stuck \ t' = Abrupt s'" and exec_res_c2: "\|\<^bsub>P\<^esub>\\c2,Normal s'\ =n\ t''" and s''_Stuck: "s'' = Stuck \ t'' = Stuck" and s''_Fault: "\f. s'' = Fault f \ t'' \ {Fault f, Stuck}" and t''_notStuck: "t'' \ Stuck \ t'' = s''" by auto show ?case proof (cases "t'=Stuck") case True with exec_res_c1 have "\|\<^bsub>P\<^esub>\\Catch c1 c2,Normal s\ =n\ Stuck" by (auto intro: execn.CatchMiss) thus ?thesis by auto next case False with t'_notStuck have "t'= Abrupt s'" by simp with exec_res_c1 exec_res_c2 have "\|\<^bsub>P\<^esub>\\Catch c1 c2,Normal s\ =n\ t''" by (auto intro: execn.CatchMatch) with s''_Stuck s''_Fault t''_notStuck show ?thesis by blast qed next case (CatchMiss c1 s n w c2) have exec_c1: "\\\c1,Normal s\ =n\ w" by fact from CatchMiss.hyps obtain w' where exec_c1': "\|\<^bsub>P\<^esub>\\c1,Normal s\ =n\ w'" and w_Stuck: "w = Stuck \ w' = Stuck" and w_Fault: "\f. w = Fault f \ w' \ {Fault f, Stuck}" and w'_noStuck: "w' \ Stuck \ w' = w" by auto have noAbr_w: "\ isAbr w" by fact show ?case proof (cases w') case (Normal s') with w'_noStuck have "w'=w" by simp with exec_c1' Normal w_Stuck w_Fault w'_noStuck show ?thesis by (fastforce intro: execn.CatchMiss) next case (Abrupt s') with w'_noStuck have "w'=w" by simp with noAbr_w Abrupt show ?thesis by simp next case (Fault f) with w'_noStuck have "w'=w" by simp with exec_c1' Fault w_Stuck w_Fault w'_noStuck show ?thesis by (fastforce intro: execn.CatchMiss) next case Stuck with exec_c1' w_Stuck w_Fault w'_noStuck show ?thesis by (fastforce intro: execn.CatchMiss) qed qed lemma exec_to_exec_restrict: assumes exec: "\\\c,s\ \ t" shows "\t'. \|\<^bsub>P\<^esub>\\c,s\ \ t' \ (t=Stuck \ t'=Stuck) \ (\f. t=Fault f\ t'\{Fault f,Stuck}) \ (t'\Stuck \ t'=t)" proof - from exec obtain n where execn_strip: "\\\c,s\ =n\ t" by (auto simp add: exec_iff_execn) from execn_to_execn_restrict [where P=P,OF this] obtain t' where "\|\<^bsub>P\<^esub>\\c,s\ =n\ t'" "t=Stuck \ t'=Stuck" "\f. t=Fault f\ t'\{Fault f,Stuck}" "t'\Stuck \ t'=t" by blast thus ?thesis by (blast intro: execn_to_exec) qed lemma notStuck_GuardD: "\\\\Guard m g c,Normal s\ \\{Stuck}; s \ g\ \ \\\c,Normal s\ \\{Stuck}" by (auto simp add: final_notin_def dest: exec.Guard ) lemma notStuck_SeqD1: "\\\\Seq c1 c2,Normal s\ \\{Stuck}\ \ \\\c1,Normal s\ \\{Stuck}" by (auto simp add: final_notin_def dest: exec.Seq ) lemma notStuck_SeqD2: "\\\\Seq c1 c2,Normal s\ \\{Stuck}; \\\c1,Normal s\ \s'\ \ \\\c2,s'\ \\{Stuck}" by (auto simp add: final_notin_def dest: exec.Seq ) lemma notStuck_SeqD: "\\\\Seq c1 c2,Normal s\ \\{Stuck}\ \ \\\c1,Normal s\ \\{Stuck} \ (\s'. \\\c1,Normal s\ \s' \ \\\c2,s'\ \\{Stuck})" by (auto simp add: final_notin_def dest: exec.Seq ) lemma notStuck_CondTrueD: "\\\\Cond b c1 c2,Normal s\ \\{Stuck}; s \ b\ \ \\\c1,Normal s\ \\{Stuck}" by (auto simp add: final_notin_def dest: exec.CondTrue) lemma notStuck_CondFalseD: "\\\\Cond b c1 c2,Normal s\ \\{Stuck}; s \ b\ \ \\\c2,Normal s\ \\{Stuck}" by (auto simp add: final_notin_def dest: exec.CondFalse) lemma notStuck_WhileTrueD1: "\\\\While b c,Normal s\ \\{Stuck}; s \ b\ \ \\\c,Normal s\ \\{Stuck}" by (auto simp add: final_notin_def dest: exec.WhileTrue) lemma notStuck_WhileTrueD2: "\\\\While b c,Normal s\ \\{Stuck}; \\\c,Normal s\ \s'; s \ b\ \ \\\While b c,s'\ \\{Stuck}" by (auto simp add: final_notin_def dest: exec.WhileTrue) lemma notStuck_CallD: "\\\\Call p ,Normal s\ \\{Stuck}; \ p = Some bdy\ \ \\\bdy,Normal s\ \\{Stuck}" by (auto simp add: final_notin_def dest: exec.Call) lemma notStuck_CallDefinedD: "\\\\Call p,Normal s\ \\{Stuck}\ \ \ p \ None" by (cases "\ p") (auto simp add: final_notin_def dest: exec.CallUndefined) lemma notStuck_DynComD: "\\\\DynCom c,Normal s\ \\{Stuck}\ \ \\\(c s),Normal s\ \\{Stuck}" by (auto simp add: final_notin_def dest: exec.DynCom) lemma notStuck_CatchD1: "\\\\Catch c1 c2,Normal s\ \\{Stuck}\ \ \\\c1,Normal s\ \\{Stuck}" by (auto simp add: final_notin_def dest: exec.CatchMatch exec.CatchMiss ) lemma notStuck_CatchD2: "\\\\Catch c1 c2,Normal s\ \\{Stuck}; \\\c1,Normal s\ \Abrupt s'\ \ \\\c2,Normal s'\ \\{Stuck}" by (auto simp add: final_notin_def dest: exec.CatchMatch) (* ************************************************************************* *) subsection "Miscellaneous" (* ************************************************************************* *) lemma execn_noguards_no_Fault: assumes execn: "\\\c,s\ =n\ t" assumes noguards_c: "noguards c" assumes noguards_\: "\p \ dom \. noguards (the (\ p))" assumes s_no_Fault: "\ isFault s" shows "\ isFault t" using execn noguards_c s_no_Fault proof (induct) case (Call p bdy n s t) with noguards_\ show ?case apply - apply (drule bspec [where x=p]) apply auto done qed (auto) lemma exec_noguards_no_Fault: assumes exec: "\\\c,s\ \ t" assumes noguards_c: "noguards c" assumes noguards_\: "\p \ dom \. noguards (the (\ p))" assumes s_no_Fault: "\ isFault s" shows "\ isFault t" using exec noguards_c s_no_Fault proof (induct) case (Call p bdy s t) with noguards_\ show ?case apply - apply (drule bspec [where x=p]) apply auto done qed auto lemma execn_nothrows_no_Abrupt: assumes execn: "\\\c,s\ =n\ t" assumes nothrows_c: "nothrows c" assumes nothrows_\: "\p \ dom \. nothrows (the (\ p))" assumes s_no_Abrupt: "\(isAbr s)" shows "\(isAbr t)" using execn nothrows_c s_no_Abrupt proof (induct) case (Call p bdy n s t) with nothrows_\ show ?case apply - apply (drule bspec [where x=p]) apply auto done qed (auto) lemma exec_nothrows_no_Abrupt: assumes exec: "\\\c,s\ \ t" assumes nothrows_c: "nothrows c" assumes nothrows_\: "\p \ dom \. nothrows (the (\ p))" assumes s_no_Abrupt: "\(isAbr s)" shows "\(isAbr t)" using exec nothrows_c s_no_Abrupt proof (induct) case (Call p bdy s t) with nothrows_\ show ?case apply - apply (drule bspec [where x=p]) apply auto done qed (auto) end diff --git a/thys/Simpl/Simpl.thy b/thys/Simpl/Simpl.thy --- a/thys/Simpl/Simpl.thy +++ b/thys/Simpl/Simpl.thy @@ -1,45 +1,24 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: Simpl.thy - Author: Norbert Schirmer, TU Muenchen - -Copyright (C) 2008 Norbert Schirmer -Some rights reserved, TU Muenchen -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (C) 2006-2008 Norbert Schirmer *) (*<*) theory Simpl imports StateSpace AlternativeSmallStep SyntaxTest "ex/VcgEx" "ex/VcgExSP" "ex/VcgExTotal" "ex/Quicksort" "ex/XVcgEx" "ex/ProcParEx" "ex/ProcParExSP" "ex/ClosureEx" "ex/ComposeEx" UserGuide begin end (*>*) diff --git a/thys/Simpl/Simpl_Heap.thy b/thys/Simpl/Simpl_Heap.thy --- a/thys/Simpl/Simpl_Heap.thy +++ b/thys/Simpl/Simpl_Heap.thy @@ -1,119 +1,98 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: Heap.thy - Author: Norbert Schirmer, TU Muenchen - -Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (C) 2006-2008 Norbert Schirmer *) theory Simpl_Heap imports Main begin subsection "References" definition "ref = (UNIV::nat set)" typedef ref = ref by (simp add: ref_def) code_datatype Abs_ref lemma finite_nat_ex_max: assumes fin: "finite (N::nat set)" shows "\m. \n\N. n < m" using fin proof (induct) case empty show ?case by auto next case (insert k N) have "\m. \n\N. n < m" by fact then obtain m where m_max: "\n\N. n < m".. show "\m. \n\insert k N. n < m" proof (rule exI [where x="Suc (max k m)"]) qed (insert m_max, auto simp add: max_def) qed lemma infinite_nat: "\finite (UNIV::nat set)" proof assume fin: "finite (UNIV::nat set)" then obtain m::nat where "\n\UNIV. n < m" by (rule finite_nat_ex_max [elim_format] ) auto moreover have "m\UNIV".. ultimately show False by blast qed lemma infinite_ref [simp,intro]: "\finite (UNIV::ref set)" proof assume "finite (UNIV::ref set)" hence "finite (range Rep_ref)" by simp moreover have "range Rep_ref = ref" proof show "range Rep_ref \ ref" by (simp add: ref_def) next show "ref \ range Rep_ref" proof fix x assume x: "x \ ref" show "x \ range Rep_ref" by (rule Rep_ref_induct) (auto simp add: ref_def) qed qed ultimately have "finite ref" by simp thus False by (simp add: ref_def infinite_nat) qed consts Null :: ref definition new :: "ref set \ ref" where "new A = (SOME a. a \ {Null} \ A)" text \ Constant @{const "Null"} can be defined later on. Conceptually @{const "Null"} and @{const "new"} are \fixes\ of a locale with @{prop "finite A \ new A \ A \ {Null}"}. But since definitions relative to a locale do not yet work in Isabelle2005 we use this workaround to avoid lots of parameters in definitions. \ lemma new_notin [simp,intro]: "finite A \ new (A) \ A" apply (unfold new_def) apply (rule someI2_ex) apply (fastforce intro: ex_new_if_finite) apply simp done lemma new_not_Null [simp,intro]: "finite A \ new (A) \ Null" apply (unfold new_def) apply (rule someI2_ex) apply (fastforce intro: ex_new_if_finite) apply simp done end diff --git a/thys/Simpl/SmallStep.thy b/thys/Simpl/SmallStep.thy --- a/thys/Simpl/SmallStep.thy +++ b/thys/Simpl/SmallStep.thy @@ -1,3120 +1,3099 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: SmallStep.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2006-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA *) section \Small-Step Semantics and Infinite Computations\ theory SmallStep imports Termination begin text \The redex of a statement is the substatement, which is actually altered by the next step in the small-step semantics. \ primrec redex:: "('s,'p,'f)com \ ('s,'p,'f)com" where "redex Skip = Skip" | "redex (Basic f) = (Basic f)" | "redex (Spec r) = (Spec r)" | "redex (Seq c\<^sub>1 c\<^sub>2) = redex c\<^sub>1" | "redex (Cond b c\<^sub>1 c\<^sub>2) = (Cond b c\<^sub>1 c\<^sub>2)" | "redex (While b c) = (While b c)" | "redex (Call p) = (Call p)" | "redex (DynCom d) = (DynCom d)" | "redex (Guard f b c) = (Guard f b c)" | "redex (Throw) = Throw" | "redex (Catch c\<^sub>1 c\<^sub>2) = redex c\<^sub>1" subsection \Small-Step Computation: \\\(c, s) \ (c', s')\\ type_synonym ('s,'p,'f) config = "('s,'p,'f)com \ ('s,'f) xstate" inductive "step"::"[('s,'p,'f) body,('s,'p,'f) config,('s,'p,'f) config] \ bool" ("_\ (_ \/ _)" [81,81,81] 100) for \::"('s,'p,'f) body" where Basic: "\\(Basic f,Normal s) \ (Skip,Normal (f s))" | Spec: "(s,t) \ r \ \\(Spec r,Normal s) \ (Skip,Normal t)" | SpecStuck: "\t. (s,t) \ r \ \\(Spec r,Normal s) \ (Skip,Stuck)" | Guard: "s\g \ \\(Guard f g c,Normal s) \ (c,Normal s)" | GuardFault: "s\g \ \\(Guard f g c,Normal s) \ (Skip,Fault f)" | Seq: "\\(c\<^sub>1,s) \ (c\<^sub>1',s') \ \\(Seq c\<^sub>1 c\<^sub>2,s) \ (Seq c\<^sub>1' c\<^sub>2, s')" | SeqSkip: "\\(Seq Skip c\<^sub>2,s) \ (c\<^sub>2, s)" | SeqThrow: "\\(Seq Throw c\<^sub>2,Normal s) \ (Throw, Normal s)" | CondTrue: "s\b \ \\(Cond b c\<^sub>1 c\<^sub>2,Normal s) \ (c\<^sub>1,Normal s)" | CondFalse: "s\b \ \\(Cond b c\<^sub>1 c\<^sub>2,Normal s) \ (c\<^sub>2,Normal s)" | WhileTrue: "\s\b\ \ \\(While b c,Normal s) \ (Seq c (While b c),Normal s)" | WhileFalse: "\s\b\ \ \\(While b c,Normal s) \ (Skip,Normal s)" | Call: "\ p=Some bdy \ \\(Call p,Normal s) \ (bdy,Normal s)" | CallUndefined: "\ p=None \ \\(Call p,Normal s) \ (Skip,Stuck)" | DynCom: "\\(DynCom c,Normal s) \ (c s,Normal s)" | Catch: "\\\(c\<^sub>1,s) \ (c\<^sub>1',s')\ \ \\(Catch c\<^sub>1 c\<^sub>2,s) \ (Catch c\<^sub>1' c\<^sub>2,s')" | CatchThrow: "\\(Catch Throw c\<^sub>2,Normal s) \ (c\<^sub>2,Normal s)" | CatchSkip: "\\(Catch Skip c\<^sub>2,s) \ (Skip,s)" | FaultProp: "\c\Skip; redex c = c\ \ \\(c,Fault f) \ (Skip,Fault f)" | StuckProp: "\c\Skip; redex c = c\ \ \\(c,Stuck) \ (Skip,Stuck)" | AbruptProp: "\c\Skip; redex c = c\ \ \\(c,Abrupt f) \ (Skip,Abrupt f)" lemmas step_induct = step.induct [of _ "(c,s)" "(c',s')", split_format (complete), case_names Basic Spec SpecStuck Guard GuardFault Seq SeqSkip SeqThrow CondTrue CondFalse WhileTrue WhileFalse Call CallUndefined DynCom Catch CatchThrow CatchSkip FaultProp StuckProp AbruptProp, induct set] inductive_cases step_elim_cases [cases set]: "\\(Skip,s) \ u" "\\(Guard f g c,s) \ u" "\\(Basic f,s) \ u" "\\(Spec r,s) \ u" "\\(Seq c1 c2,s) \ u" "\\(Cond b c1 c2,s) \ u" "\\(While b c,s) \ u" "\\(Call p,s) \ u" "\\(DynCom c,s) \ u" "\\(Throw,s) \ u" "\\(Catch c1 c2,s) \ u" inductive_cases step_Normal_elim_cases [cases set]: "\\(Skip,Normal s) \ u" "\\(Guard f g c,Normal s) \ u" "\\(Basic f,Normal s) \ u" "\\(Spec r,Normal s) \ u" "\\(Seq c1 c2,Normal s) \ u" "\\(Cond b c1 c2,Normal s) \ u" "\\(While b c,Normal s) \ u" "\\(Call p,Normal s) \ u" "\\(DynCom c,Normal s) \ u" "\\(Throw,Normal s) \ u" "\\(Catch c1 c2,Normal s) \ u" text \The final configuration is either of the form \(Skip,_)\ for normal termination, or @{term "(Throw,Normal s)"} in case the program was started in a @{term "Normal"} state and terminated abruptly. The @{const "Abrupt"} state is not used to model abrupt termination, in contrast to the big-step semantics. Only if the program starts in an @{const "Abrupt"} states it ends in the same @{term "Abrupt"} state.\ definition final:: "('s,'p,'f) config \ bool" where "final cfg = (fst cfg=Skip \ (fst cfg=Throw \ (\s. snd cfg=Normal s)))" abbreviation "step_rtrancl" :: "[('s,'p,'f) body,('s,'p,'f) config,('s,'p,'f) config] \ bool" ("_\ (_ \\<^sup>*/ _)" [81,81,81] 100) where "\\cf0 \\<^sup>* cf1 \ (CONST step \)\<^sup>*\<^sup>* cf0 cf1" abbreviation "step_trancl" :: "[('s,'p,'f) body,('s,'p,'f) config,('s,'p,'f) config] \ bool" ("_\ (_ \\<^sup>+/ _)" [81,81,81] 100) where "\\cf0 \\<^sup>+ cf1 \ (CONST step \)\<^sup>+\<^sup>+ cf0 cf1" (* ************************************************************************ *) subsection \Structural Properties of Small Step Computations\ (* ************************************************************************ *) lemma redex_not_Seq: "redex c = Seq c1 c2 \ P" apply (induct c) apply auto done lemma no_step_final: assumes step: "\\(c,s) \ (c',s')" shows "final (c,s) \ P" using step by induct (auto simp add: final_def) lemma no_step_final': assumes step: "\\cfg \ cfg'" shows "final cfg \ P" using step by (cases cfg, cases cfg') (auto intro: no_step_final) lemma step_Abrupt: assumes step: "\\ (c, s) \ (c', s')" shows "\x. s=Abrupt x \ s'=Abrupt x" using step by (induct) auto lemma step_Fault: assumes step: "\\ (c, s) \ (c', s')" shows "\f. s=Fault f \ s'=Fault f" using step by (induct) auto lemma step_Stuck: assumes step: "\\ (c, s) \ (c', s')" shows "\f. s=Stuck \ s'=Stuck" using step by (induct) auto lemma SeqSteps: assumes steps: "\\cfg\<^sub>1\\<^sup>* cfg\<^sub>2" shows "\ c\<^sub>1 s c\<^sub>1' s'. \cfg\<^sub>1 = (c\<^sub>1,s);cfg\<^sub>2=(c\<^sub>1',s')\ \ \\(Seq c\<^sub>1 c\<^sub>2,s) \\<^sup>* (Seq c\<^sub>1' c\<^sub>2, s')" using steps proof (induct rule: converse_rtranclp_induct [case_names Refl Trans]) case Refl thus ?case by simp next case (Trans cfg\<^sub>1 cfg'') have step: "\\ cfg\<^sub>1 \ cfg''" by fact have steps: "\\ cfg'' \\<^sup>* cfg\<^sub>2" by fact have cfg\<^sub>1: "cfg\<^sub>1 = (c\<^sub>1, s)" and cfg\<^sub>2: "cfg\<^sub>2 = (c\<^sub>1', s')" by fact+ obtain c\<^sub>1'' s'' where cfg'': "cfg''=(c\<^sub>1'',s'')" by (cases cfg'') auto from step cfg\<^sub>1 cfg'' have "\\ (c\<^sub>1,s) \ (c\<^sub>1'',s'')" by simp hence "\\ (Seq c\<^sub>1 c\<^sub>2,s) \ (Seq c\<^sub>1'' c\<^sub>2,s'')" by (rule step.Seq) also from Trans.hyps (3) [OF cfg'' cfg\<^sub>2] have "\\ (Seq c\<^sub>1'' c\<^sub>2, s'') \\<^sup>* (Seq c\<^sub>1' c\<^sub>2, s')" . finally show ?case . qed lemma CatchSteps: assumes steps: "\\cfg\<^sub>1\\<^sup>* cfg\<^sub>2" shows "\ c\<^sub>1 s c\<^sub>1' s'. \cfg\<^sub>1 = (c\<^sub>1,s); cfg\<^sub>2=(c\<^sub>1',s')\ \ \\(Catch c\<^sub>1 c\<^sub>2,s) \\<^sup>* (Catch c\<^sub>1' c\<^sub>2, s')" using steps proof (induct rule: converse_rtranclp_induct [case_names Refl Trans]) case Refl thus ?case by simp next case (Trans cfg\<^sub>1 cfg'') have step: "\\ cfg\<^sub>1 \ cfg''" by fact have steps: "\\ cfg'' \\<^sup>* cfg\<^sub>2" by fact have cfg\<^sub>1: "cfg\<^sub>1 = (c\<^sub>1, s)" and cfg\<^sub>2: "cfg\<^sub>2 = (c\<^sub>1', s')" by fact+ obtain c\<^sub>1'' s'' where cfg'': "cfg''=(c\<^sub>1'',s'')" by (cases cfg'') auto from step cfg\<^sub>1 cfg'' have s: "\\ (c\<^sub>1,s) \ (c\<^sub>1'',s'')" by simp hence "\\ (Catch c\<^sub>1 c\<^sub>2,s) \ (Catch c\<^sub>1'' c\<^sub>2,s'')" by (rule step.Catch) also from Trans.hyps (3) [OF cfg'' cfg\<^sub>2] have "\\ (Catch c\<^sub>1'' c\<^sub>2, s'') \\<^sup>* (Catch c\<^sub>1' c\<^sub>2, s')" . finally show ?case . qed lemma steps_Fault: "\\ (c, Fault f) \\<^sup>* (Skip, Fault f)" proof (induct c) case (Seq c\<^sub>1 c\<^sub>2) have steps_c\<^sub>1: "\\ (c\<^sub>1, Fault f) \\<^sup>* (Skip, Fault f)" by fact have steps_c\<^sub>2: "\\ (c\<^sub>2, Fault f) \\<^sup>* (Skip, Fault f)" by fact from SeqSteps [OF steps_c\<^sub>1 refl refl] have "\\ (Seq c\<^sub>1 c\<^sub>2, Fault f) \\<^sup>* (Seq Skip c\<^sub>2, Fault f)". also have "\\ (Seq Skip c\<^sub>2, Fault f) \ (c\<^sub>2, Fault f)" by (rule SeqSkip) also note steps_c\<^sub>2 finally show ?case by simp next case (Catch c\<^sub>1 c\<^sub>2) have steps_c\<^sub>1: "\\ (c\<^sub>1, Fault f) \\<^sup>* (Skip, Fault f)" by fact from CatchSteps [OF steps_c\<^sub>1 refl refl] have "\\ (Catch c\<^sub>1 c\<^sub>2, Fault f) \\<^sup>* (Catch Skip c\<^sub>2, Fault f)". also have "\\ (Catch Skip c\<^sub>2, Fault f) \ (Skip, Fault f)" by (rule CatchSkip) finally show ?case by simp qed (fastforce intro: step.intros)+ lemma steps_Stuck: "\\ (c, Stuck) \\<^sup>* (Skip, Stuck)" proof (induct c) case (Seq c\<^sub>1 c\<^sub>2) have steps_c\<^sub>1: "\\ (c\<^sub>1, Stuck) \\<^sup>* (Skip, Stuck)" by fact have steps_c\<^sub>2: "\\ (c\<^sub>2, Stuck) \\<^sup>* (Skip, Stuck)" by fact from SeqSteps [OF steps_c\<^sub>1 refl refl] have "\\ (Seq c\<^sub>1 c\<^sub>2, Stuck) \\<^sup>* (Seq Skip c\<^sub>2, Stuck)". also have "\\ (Seq Skip c\<^sub>2, Stuck) \ (c\<^sub>2, Stuck)" by (rule SeqSkip) also note steps_c\<^sub>2 finally show ?case by simp next case (Catch c\<^sub>1 c\<^sub>2) have steps_c\<^sub>1: "\\ (c\<^sub>1, Stuck) \\<^sup>* (Skip, Stuck)" by fact from CatchSteps [OF steps_c\<^sub>1 refl refl] have "\\ (Catch c\<^sub>1 c\<^sub>2, Stuck) \\<^sup>* (Catch Skip c\<^sub>2, Stuck)" . also have "\\ (Catch Skip c\<^sub>2, Stuck) \ (Skip, Stuck)" by (rule CatchSkip) finally show ?case by simp qed (fastforce intro: step.intros)+ lemma steps_Abrupt: "\\ (c, Abrupt s) \\<^sup>* (Skip, Abrupt s)" proof (induct c) case (Seq c\<^sub>1 c\<^sub>2) have steps_c\<^sub>1: "\\ (c\<^sub>1, Abrupt s) \\<^sup>* (Skip, Abrupt s)" by fact have steps_c\<^sub>2: "\\ (c\<^sub>2, Abrupt s) \\<^sup>* (Skip, Abrupt s)" by fact from SeqSteps [OF steps_c\<^sub>1 refl refl] have "\\ (Seq c\<^sub>1 c\<^sub>2, Abrupt s) \\<^sup>* (Seq Skip c\<^sub>2, Abrupt s)". also have "\\ (Seq Skip c\<^sub>2, Abrupt s) \ (c\<^sub>2, Abrupt s)" by (rule SeqSkip) also note steps_c\<^sub>2 finally show ?case by simp next case (Catch c\<^sub>1 c\<^sub>2) have steps_c\<^sub>1: "\\ (c\<^sub>1, Abrupt s) \\<^sup>* (Skip, Abrupt s)" by fact from CatchSteps [OF steps_c\<^sub>1 refl refl] have "\\ (Catch c\<^sub>1 c\<^sub>2, Abrupt s) \\<^sup>* (Catch Skip c\<^sub>2, Abrupt s)". also have "\\ (Catch Skip c\<^sub>2, Abrupt s) \ (Skip, Abrupt s)" by (rule CatchSkip) finally show ?case by simp qed (fastforce intro: step.intros)+ lemma step_Fault_prop: assumes step: "\\ (c, s) \ (c', s')" shows "\f. s=Fault f \ s'=Fault f" using step by (induct) auto lemma step_Abrupt_prop: assumes step: "\\ (c, s) \ (c', s')" shows "\x. s=Abrupt x \ s'=Abrupt x" using step by (induct) auto lemma step_Stuck_prop: assumes step: "\\ (c, s) \ (c', s')" shows "s=Stuck \ s'=Stuck" using step by (induct) auto lemma steps_Fault_prop: assumes step: "\\ (c, s) \\<^sup>* (c', s')" shows "s=Fault f \ s'=Fault f" using step proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans]) case Refl thus ?case by simp next case (Trans c s c'' s'') thus ?case by (auto intro: step_Fault_prop) qed lemma steps_Abrupt_prop: assumes step: "\\ (c, s) \\<^sup>* (c', s')" shows "s=Abrupt t \ s'=Abrupt t" using step proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans]) case Refl thus ?case by simp next case (Trans c s c'' s'') thus ?case by (auto intro: step_Abrupt_prop) qed lemma steps_Stuck_prop: assumes step: "\\ (c, s) \\<^sup>* (c', s')" shows "s=Stuck \ s'=Stuck" using step proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans]) case Refl thus ?case by simp next case (Trans c s c'' s'') thus ?case by (auto intro: step_Stuck_prop) qed (* ************************************************************************ *) subsection \Equivalence between Small-Step and Big-Step Semantics\ (* ************************************************************************ *) theorem exec_impl_steps: assumes exec: "\\\c,s\ \ t" shows "\c' t'. \\(c,s) \\<^sup>* (c',t') \ (case t of Abrupt x \ if s=t then c'=Skip \ t'=t else c'=Throw \ t'=Normal x | _ \ c'=Skip \ t'=t)" using exec proof (induct) case Skip thus ?case by simp next case Guard thus ?case by (blast intro: step.Guard rtranclp_trans) next case GuardFault thus ?case by (fastforce intro: step.GuardFault rtranclp_trans) next case FaultProp show ?case by (fastforce intro: steps_Fault) next case Basic thus ?case by (fastforce intro: step.Basic rtranclp_trans) next case Spec thus ?case by (fastforce intro: step.Spec rtranclp_trans) next case SpecStuck thus ?case by (fastforce intro: step.SpecStuck rtranclp_trans) next case (Seq c\<^sub>1 s s' c\<^sub>2 t) have exec_c\<^sub>1: "\\ \c\<^sub>1,Normal s\ \ s'" by fact have exec_c\<^sub>2: "\\ \c\<^sub>2,s'\ \ t" by fact show ?case proof (cases "\x. s'=Abrupt x") case False from False Seq.hyps (2) have "\\ (c\<^sub>1, Normal s) \\<^sup>* (Skip, s')" by (cases s') auto hence seq_c\<^sub>1: "\\ (Seq c\<^sub>1 c\<^sub>2, Normal s) \\<^sup>* (Seq Skip c\<^sub>2, s')" by (rule SeqSteps) auto from Seq.hyps (4) obtain c' t' where steps_c\<^sub>2: "\\ (c\<^sub>2, s') \\<^sup>* (c', t')" and t: "(case t of Abrupt x \ if s' = t then c' = Skip \ t' = t else c' = Throw \ t' = Normal x | _ \ c' = Skip \ t' = t)" by auto note seq_c\<^sub>1 also have "\\ (Seq Skip c\<^sub>2, s') \ (c\<^sub>2, s')" by (rule step.SeqSkip) also note steps_c\<^sub>2 finally have "\\ (Seq c\<^sub>1 c\<^sub>2, Normal s) \\<^sup>* (c', t')". with t False show ?thesis by (cases t) auto next case True then obtain x where s': "s'=Abrupt x" by blast from s' Seq.hyps (2) have "\\ (c\<^sub>1, Normal s) \\<^sup>* (Throw, Normal x)" by auto hence seq_c\<^sub>1: "\\ (Seq c\<^sub>1 c\<^sub>2, Normal s) \\<^sup>* (Seq Throw c\<^sub>2, Normal x)" by (rule SeqSteps) auto also have "\\ (Seq Throw c\<^sub>2, Normal x) \ (Throw, Normal x)" by (rule SeqThrow) finally have "\\ (Seq c\<^sub>1 c\<^sub>2, Normal s) \\<^sup>* (Throw, Normal x)". moreover from exec_c\<^sub>2 s' have "t=Abrupt x" by (auto intro: Abrupt_end) ultimately show ?thesis by auto qed next case CondTrue thus ?case by (blast intro: step.CondTrue rtranclp_trans) next case CondFalse thus ?case by (blast intro: step.CondFalse rtranclp_trans) next case (WhileTrue s b c s' t) have exec_c: "\\ \c,Normal s\ \ s'" by fact have exec_w: "\\ \While b c,s'\ \ t" by fact have b: "s \ b" by fact hence step: "\\ (While b c,Normal s) \ (Seq c (While b c),Normal s)" by (rule step.WhileTrue) show ?case proof (cases "\x. s'=Abrupt x") case False from False WhileTrue.hyps (3) have "\\ (c, Normal s) \\<^sup>* (Skip, s')" by (cases s') auto hence seq_c: "\\ (Seq c (While b c), Normal s) \\<^sup>* (Seq Skip (While b c), s')" by (rule SeqSteps) auto from WhileTrue.hyps (5) obtain c' t' where steps_c\<^sub>2: "\\ (While b c, s') \\<^sup>* (c', t')" and t: "(case t of Abrupt x \ if s' = t then c' = Skip \ t' = t else c' = Throw \ t' = Normal x | _ \ c' = Skip \ t' = t)" by auto note step also note seq_c also have "\\ (Seq Skip (While b c), s') \ (While b c, s')" by (rule step.SeqSkip) also note steps_c\<^sub>2 finally have "\\ (While b c, Normal s) \\<^sup>* (c', t')". with t False show ?thesis by (cases t) auto next case True then obtain x where s': "s'=Abrupt x" by blast note step also from s' WhileTrue.hyps (3) have "\\ (c, Normal s) \\<^sup>* (Throw, Normal x)" by auto hence seq_c: "\\ (Seq c (While b c), Normal s) \\<^sup>* (Seq Throw (While b c), Normal x)" by (rule SeqSteps) auto also have "\\ (Seq Throw (While b c), Normal x) \ (Throw, Normal x)" by (rule SeqThrow) finally have "\\ (While b c, Normal s) \\<^sup>* (Throw, Normal x)". moreover from exec_w s' have "t=Abrupt x" by (auto intro: Abrupt_end) ultimately show ?thesis by auto qed next case WhileFalse thus ?case by (fastforce intro: step.WhileFalse rtrancl_trans) next case Call thus ?case by (blast intro: step.Call rtranclp_trans) next case CallUndefined thus ?case by (fastforce intro: step.CallUndefined rtranclp_trans) next case StuckProp thus ?case by (fastforce intro: steps_Stuck) next case DynCom thus ?case by (blast intro: step.DynCom rtranclp_trans) next case Throw thus ?case by simp next case AbruptProp thus ?case by (fastforce intro: steps_Abrupt) next case (CatchMatch c\<^sub>1 s s' c\<^sub>2 t) from CatchMatch.hyps (2) have "\\ (c\<^sub>1, Normal s) \\<^sup>* (Throw, Normal s')" by simp hence "\\ (Catch c\<^sub>1 c\<^sub>2, Normal s) \\<^sup>* (Catch Throw c\<^sub>2, Normal s')" by (rule CatchSteps) auto also have "\\ (Catch Throw c\<^sub>2, Normal s') \ (c\<^sub>2, Normal s')" by (rule step.CatchThrow) also from CatchMatch.hyps (4) obtain c' t' where steps_c\<^sub>2: "\\ (c\<^sub>2, Normal s') \\<^sup>* (c', t')" and t: "(case t of Abrupt x \ if Normal s' = t then c' = Skip \ t' = t else c' = Throw \ t' = Normal x | _ \ c' = Skip \ t' = t)" by auto note steps_c\<^sub>2 finally show ?case using t by (auto split: xstate.splits) next case (CatchMiss c\<^sub>1 s t c\<^sub>2) have t: "\ isAbr t" by fact with CatchMiss.hyps (2) have "\\ (c\<^sub>1, Normal s) \\<^sup>* (Skip, t)" by (cases t) auto hence "\\ (Catch c\<^sub>1 c\<^sub>2, Normal s) \\<^sup>* (Catch Skip c\<^sub>2, t)" by (rule CatchSteps) auto also have "\\ (Catch Skip c\<^sub>2, t) \ (Skip, t)" by (rule step.CatchSkip) finally show ?case using t by (fastforce split: xstate.splits) qed corollary exec_impl_steps_Normal: assumes exec: "\\\c,s\ \ Normal t" shows "\\(c,s) \\<^sup>* (Skip, Normal t)" using exec_impl_steps [OF exec] by auto corollary exec_impl_steps_Normal_Abrupt: assumes exec: "\\\c,Normal s\ \ Abrupt t" shows "\\(c,Normal s) \\<^sup>* (Throw, Normal t)" using exec_impl_steps [OF exec] by auto corollary exec_impl_steps_Abrupt_Abrupt: assumes exec: "\\\c,Abrupt t\ \ Abrupt t" shows "\\(c,Abrupt t) \\<^sup>* (Skip, Abrupt t)" using exec_impl_steps [OF exec] by auto corollary exec_impl_steps_Fault: assumes exec: "\\\c,s\ \ Fault f" shows "\\(c,s) \\<^sup>* (Skip, Fault f)" using exec_impl_steps [OF exec] by auto corollary exec_impl_steps_Stuck: assumes exec: "\\\c,s\ \ Stuck" shows "\\(c,s) \\<^sup>* (Skip, Stuck)" using exec_impl_steps [OF exec] by auto lemma step_Abrupt_end: assumes step: "\\ (c\<^sub>1, s) \ (c\<^sub>1', s')" shows "s'=Abrupt x \ s=Abrupt x" using step by induct auto lemma step_Stuck_end: assumes step: "\\ (c\<^sub>1, s) \ (c\<^sub>1', s')" shows "s'=Stuck \ s=Stuck \ (\r x. redex c\<^sub>1 = Spec r \ s=Normal x \ (\t. (x,t)\r)) \ (\p x. redex c\<^sub>1=Call p \ s=Normal x \ \ p = None)" using step by induct auto lemma step_Fault_end: assumes step: "\\ (c\<^sub>1, s) \ (c\<^sub>1', s')" shows "s'=Fault f \ s=Fault f \ (\g c x. redex c\<^sub>1 = Guard f g c \ s=Normal x \ x \ g)" using step by induct auto lemma exec_redex_Stuck: "\\\redex c,s\ \ Stuck \ \\\c,s\ \ Stuck" proof (induct c) case Seq thus ?case by (cases s) (auto intro: exec.intros elim:exec_elim_cases) next case Catch thus ?case by (cases s) (auto intro: exec.intros elim:exec_elim_cases) qed simp_all lemma exec_redex_Fault: "\\\redex c,s\ \ Fault f \ \\\c,s\ \ Fault f" proof (induct c) case Seq thus ?case by (cases s) (auto intro: exec.intros elim:exec_elim_cases) next case Catch thus ?case by (cases s) (auto intro: exec.intros elim:exec_elim_cases) qed simp_all lemma step_extend: assumes step: "\\(c,s) \ (c', s')" shows "\t. \\\c',s'\ \ t \ \\\c,s\ \ t" using step proof (induct) case Basic thus ?case by (fastforce intro: exec.intros elim: exec_Normal_elim_cases) next case Spec thus ?case by (fastforce intro: exec.intros elim: exec_Normal_elim_cases) next case SpecStuck thus ?case by (fastforce intro: exec.intros elim: exec_Normal_elim_cases) next case Guard thus ?case by (fastforce intro: exec.intros elim: exec_Normal_elim_cases) next case GuardFault thus ?case by (fastforce intro: exec.intros elim: exec_Normal_elim_cases) next case (Seq c\<^sub>1 s c\<^sub>1' s' c\<^sub>2) have step: "\\ (c\<^sub>1, s) \ (c\<^sub>1', s')" by fact have exec': "\\ \Seq c\<^sub>1' c\<^sub>2,s'\ \ t" by fact show ?case proof (cases s) case (Normal x) note s_Normal = this show ?thesis proof (cases s') case (Normal x') from exec' [simplified Normal] obtain s'' where exec_c\<^sub>1': "\\ \c\<^sub>1',Normal x'\ \ s''" and exec_c\<^sub>2: "\\ \c\<^sub>2,s''\ \ t" by cases from Seq.hyps (2) Normal exec_c\<^sub>1' s_Normal have "\\ \c\<^sub>1,Normal x\ \ s''" by simp from exec.Seq [OF this exec_c\<^sub>2] s_Normal show ?thesis by simp next case (Abrupt x') with exec' have "t=Abrupt x'" by (auto intro:Abrupt_end) moreover from step Abrupt have "s=Abrupt x'" by (auto intro: step_Abrupt_end) ultimately show ?thesis by (auto intro: exec.intros) next case (Fault f) from step_Fault_end [OF step this] s_Normal obtain g c where redex_c\<^sub>1: "redex c\<^sub>1 = Guard f g c" and fail: "x \ g" by auto hence "\\ \redex c\<^sub>1,Normal x\ \ Fault f" by (auto intro: exec.intros) from exec_redex_Fault [OF this] have "\\ \c\<^sub>1,Normal x\ \ Fault f". moreover from Fault exec' have "t=Fault f" by (auto intro: Fault_end) ultimately show ?thesis using s_Normal by (auto intro: exec.intros) next case Stuck from step_Stuck_end [OF step this] s_Normal have "(\r. redex c\<^sub>1 = Spec r \ (\t. (x, t) \ r)) \ (\p. redex c\<^sub>1 = Call p \ \ p = None)" by auto moreover { fix r assume "redex c\<^sub>1 = Spec r" and "(\t. (x, t) \ r)" hence "\\ \redex c\<^sub>1,Normal x\ \ Stuck" by (auto intro: exec.intros) from exec_redex_Stuck [OF this] have "\\ \c\<^sub>1,Normal x\ \ Stuck". moreover from Stuck exec' have "t=Stuck" by (auto intro: Stuck_end) ultimately have ?thesis using s_Normal by (auto intro: exec.intros) } moreover { fix p assume "redex c\<^sub>1 = Call p" and "\ p = None" hence "\\ \redex c\<^sub>1,Normal x\ \ Stuck" by (auto intro: exec.intros) from exec_redex_Stuck [OF this] have "\\ \c\<^sub>1,Normal x\ \ Stuck". moreover from Stuck exec' have "t=Stuck" by (auto intro: Stuck_end) ultimately have ?thesis using s_Normal by (auto intro: exec.intros) } ultimately show ?thesis by auto qed next case (Abrupt x) from step_Abrupt [OF step this] have "s'=Abrupt x". with exec' have "t=Abrupt x" by (auto intro: Abrupt_end) with Abrupt show ?thesis by (auto intro: exec.intros) next case (Fault f) from step_Fault [OF step this] have "s'=Fault f". with exec' have "t=Fault f" by (auto intro: Fault_end) with Fault show ?thesis by (auto intro: exec.intros) next case Stuck from step_Stuck [OF step this] have "s'=Stuck". with exec' have "t=Stuck" by (auto intro: Stuck_end) with Stuck show ?thesis by (auto intro: exec.intros) qed next case (SeqSkip c\<^sub>2 s t) thus ?case by (cases s) (fastforce intro: exec.intros elim: exec_elim_cases)+ next case (SeqThrow c\<^sub>2 s t) thus ?case by (fastforce intro: exec.intros elim: exec_elim_cases)+ next case CondTrue thus ?case by (fastforce intro: exec.intros elim: exec_Normal_elim_cases) next case CondFalse thus ?case by (fastforce intro: exec.intros elim: exec_Normal_elim_cases) next case WhileTrue thus ?case by (fastforce intro: exec.intros elim: exec_Normal_elim_cases) next case WhileFalse thus ?case by (fastforce intro: exec.intros elim: exec_Normal_elim_cases) next case Call thus ?case by (fastforce intro: exec.intros elim: exec_Normal_elim_cases) next case CallUndefined thus ?case by (fastforce intro: exec.intros elim: exec_Normal_elim_cases) next case DynCom thus ?case by (fastforce intro: exec.intros elim: exec_Normal_elim_cases) next case (Catch c\<^sub>1 s c\<^sub>1' s' c\<^sub>2 t) have step: "\\ (c\<^sub>1, s) \ (c\<^sub>1', s')" by fact have exec': "\\ \Catch c\<^sub>1' c\<^sub>2,s'\ \ t" by fact show ?case proof (cases s) case (Normal x) note s_Normal = this show ?thesis proof (cases s') case (Normal x') from exec' [simplified Normal] show ?thesis proof (cases) fix s'' assume exec_c\<^sub>1': "\\ \c\<^sub>1',Normal x'\ \ Abrupt s''" assume exec_c\<^sub>2: "\\ \c\<^sub>2,Normal s''\ \ t" from Catch.hyps (2) Normal exec_c\<^sub>1' s_Normal have "\\ \c\<^sub>1,Normal x\ \ Abrupt s''" by simp from exec.CatchMatch [OF this exec_c\<^sub>2] s_Normal show ?thesis by simp next assume exec_c\<^sub>1': "\\ \c\<^sub>1',Normal x'\ \ t" assume t: "\ isAbr t" from Catch.hyps (2) Normal exec_c\<^sub>1' s_Normal have "\\ \c\<^sub>1,Normal x\ \ t" by simp from exec.CatchMiss [OF this t] s_Normal show ?thesis by simp qed next case (Abrupt x') with exec' have "t=Abrupt x'" by (auto intro:Abrupt_end) moreover from step Abrupt have "s=Abrupt x'" by (auto intro: step_Abrupt_end) ultimately show ?thesis by (auto intro: exec.intros) next case (Fault f) from step_Fault_end [OF step this] s_Normal obtain g c where redex_c\<^sub>1: "redex c\<^sub>1 = Guard f g c" and fail: "x \ g" by auto hence "\\ \redex c\<^sub>1,Normal x\ \ Fault f" by (auto intro: exec.intros) from exec_redex_Fault [OF this] have "\\ \c\<^sub>1,Normal x\ \ Fault f". moreover from Fault exec' have "t=Fault f" by (auto intro: Fault_end) ultimately show ?thesis using s_Normal by (auto intro: exec.intros) next case Stuck from step_Stuck_end [OF step this] s_Normal have "(\r. redex c\<^sub>1 = Spec r \ (\t. (x, t) \ r)) \ (\p. redex c\<^sub>1 = Call p \ \ p = None)" by auto moreover { fix r assume "redex c\<^sub>1 = Spec r" and "(\t. (x, t) \ r)" hence "\\ \redex c\<^sub>1,Normal x\ \ Stuck" by (auto intro: exec.intros) from exec_redex_Stuck [OF this] have "\\ \c\<^sub>1,Normal x\ \ Stuck". moreover from Stuck exec' have "t=Stuck" by (auto intro: Stuck_end) ultimately have ?thesis using s_Normal by (auto intro: exec.intros) } moreover { fix p assume "redex c\<^sub>1 = Call p" and "\ p = None" hence "\\ \redex c\<^sub>1,Normal x\ \ Stuck" by (auto intro: exec.intros) from exec_redex_Stuck [OF this] have "\\ \c\<^sub>1,Normal x\ \ Stuck". moreover from Stuck exec' have "t=Stuck" by (auto intro: Stuck_end) ultimately have ?thesis using s_Normal by (auto intro: exec.intros) } ultimately show ?thesis by auto qed next case (Abrupt x) from step_Abrupt [OF step this] have "s'=Abrupt x". with exec' have "t=Abrupt x" by (auto intro: Abrupt_end) with Abrupt show ?thesis by (auto intro: exec.intros) next case (Fault f) from step_Fault [OF step this] have "s'=Fault f". with exec' have "t=Fault f" by (auto intro: Fault_end) with Fault show ?thesis by (auto intro: exec.intros) next case Stuck from step_Stuck [OF step this] have "s'=Stuck". with exec' have "t=Stuck" by (auto intro: Stuck_end) with Stuck show ?thesis by (auto intro: exec.intros) qed next case CatchThrow thus ?case by (fastforce intro: exec.intros elim: exec_Normal_elim_cases) next case CatchSkip thus ?case by (fastforce intro: exec.intros elim: exec_elim_cases) next case FaultProp thus ?case by (fastforce intro: exec.intros elim: exec_elim_cases) next case StuckProp thus ?case by (fastforce intro: exec.intros elim: exec_elim_cases) next case AbruptProp thus ?case by (fastforce intro: exec.intros elim: exec_elim_cases) qed theorem steps_Skip_impl_exec: assumes steps: "\\(c,s) \\<^sup>* (Skip,t)" shows "\\\c,s\ \ t" using steps proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans]) case Refl thus ?case by (cases t) (auto intro: exec.intros) next case (Trans c s c' s') have "\\ (c, s) \ (c', s')" and "\\ \c',s'\ \ t" by fact+ thus ?case by (rule step_extend) qed theorem steps_Throw_impl_exec: assumes steps: "\\(c,s) \\<^sup>* (Throw,Normal t)" shows "\\\c,s\ \ Abrupt t" using steps proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans]) case Refl thus ?case by (auto intro: exec.intros) next case (Trans c s c' s') have "\\ (c, s) \ (c', s')" and "\\ \c',s'\ \ Abrupt t" by fact+ thus ?case by (rule step_extend) qed (* ************************************************************************ *) subsection \Infinite Computations: \\\(c, s) \ \(\)\\ (* ************************************************************************ *) definition inf:: "('s,'p,'f) body \ ('s,'p,'f) config \ bool" ("_\ _ \ \'(\')" [60,80] 100) where "\\ cfg \ \(\) \ (\f. f (0::nat) = cfg \ (\i. \\f i \ f (i+1)))" lemma not_infI: "\\f. \f 0 = cfg; \i. \\f i \ f (Suc i)\ \ False\ \ \\\ cfg \ \(\)" by (auto simp add: inf_def) (* ************************************************************************ *) subsection \Equivalence between Termination and the Absence of Infinite Computations\ (* ************************************************************************ *) lemma step_preserves_termination: assumes step: "\\(c,s) \ (c',s')" shows "\\c\s \ \\c'\s'" using step proof (induct) case Basic thus ?case by (fastforce intro: terminates.intros) next case Spec thus ?case by (fastforce intro: terminates.intros) next case SpecStuck thus ?case by (fastforce intro: terminates.intros) next case Guard thus ?case by (fastforce intro: terminates.intros elim: terminates_Normal_elim_cases) next case GuardFault thus ?case by (fastforce intro: terminates.intros) next case (Seq c\<^sub>1 s c\<^sub>1' s' c\<^sub>2) thus ?case apply (cases s) apply (cases s') apply (fastforce intro: terminates.intros step_extend elim: terminates_Normal_elim_cases) apply (fastforce intro: terminates.intros dest: step_Abrupt_prop step_Fault_prop step_Stuck_prop)+ done next case (SeqSkip c\<^sub>2 s) thus ?case apply (cases s) apply (fastforce intro: terminates.intros exec.intros elim: terminates_Normal_elim_cases )+ done next case (SeqThrow c\<^sub>2 s) thus ?case by (fastforce intro: terminates.intros exec.intros elim: terminates_Normal_elim_cases ) next case CondTrue thus ?case by (fastforce intro: terminates.intros exec.intros elim: terminates_Normal_elim_cases ) next case CondFalse thus ?case by (fastforce intro: terminates.intros elim: terminates_Normal_elim_cases ) next case WhileTrue thus ?case by (fastforce intro: terminates.intros elim: terminates_Normal_elim_cases ) next case WhileFalse thus ?case by (fastforce intro: terminates.intros elim: terminates_Normal_elim_cases ) next case Call thus ?case by (fastforce intro: terminates.intros elim: terminates_Normal_elim_cases ) next case CallUndefined thus ?case by (fastforce intro: terminates.intros elim: terminates_Normal_elim_cases ) next case DynCom thus ?case by (fastforce intro: terminates.intros elim: terminates_Normal_elim_cases ) next case (Catch c\<^sub>1 s c\<^sub>1' s' c\<^sub>2) thus ?case apply (cases s) apply (cases s') apply (fastforce intro: terminates.intros step_extend elim: terminates_Normal_elim_cases) apply (fastforce intro: terminates.intros dest: step_Abrupt_prop step_Fault_prop step_Stuck_prop)+ done next case CatchThrow thus ?case by (fastforce intro: terminates.intros exec.intros elim: terminates_Normal_elim_cases ) next case (CatchSkip c\<^sub>2 s) thus ?case by (cases s) (fastforce intro: terminates.intros)+ next case FaultProp thus ?case by (fastforce intro: terminates.intros) next case StuckProp thus ?case by (fastforce intro: terminates.intros) next case AbruptProp thus ?case by (fastforce intro: terminates.intros) qed lemma steps_preserves_termination: assumes steps: "\\(c,s) \\<^sup>* (c',s')" shows "\\c\s \ \\c'\s'" using steps proof (induct rule: rtranclp_induct2 [consumes 1, case_names Refl Trans]) case Refl thus ?case . next case Trans thus ?case by (blast dest: step_preserves_termination) qed ML \ ML_Thms.bind_thm ("tranclp_induct2", Split_Rule.split_rule @{context} (Rule_Insts.read_instantiate @{context} [((("a", 0), Position.none), "(aa,ab)"), ((("b", 0), Position.none), "(ba,bb)")] [] @{thm tranclp_induct})); \ lemma steps_preserves_termination': assumes steps: "\\(c,s) \\<^sup>+ (c',s')" shows "\\c\s \ \\c'\s'" using steps proof (induct rule: tranclp_induct2 [consumes 1, case_names Step Trans]) case Step thus ?case by (blast intro: step_preserves_termination) next case Trans thus ?case by (blast dest: step_preserves_termination) qed definition head_com:: "('s,'p,'f) com \ ('s,'p,'f) com" where "head_com c = (case c of Seq c\<^sub>1 c\<^sub>2 \ c\<^sub>1 | Catch c\<^sub>1 c\<^sub>2 \ c\<^sub>1 | _ \ c)" definition head:: "('s,'p,'f) config \ ('s,'p,'f) config" where "head cfg = (head_com (fst cfg), snd cfg)" lemma le_Suc_cases: "\\i. \i < k\ \ P i; P k\ \ \i<(Suc k). P i" apply clarify apply (case_tac "i=k") apply auto done lemma redex_Seq_False: "\c' c''. (redex c = Seq c'' c') = False" by (induct c) auto lemma redex_Catch_False: "\c' c''. (redex c = Catch c'' c') = False" by (induct c) auto lemma infinite_computation_extract_head_Seq: assumes inf_comp: "\i::nat. \\f i \ f (i+1)" assumes f_0: "f 0 = (Seq c\<^sub>1 c\<^sub>2,s)" assumes not_fin: "\i final (head (f i))" shows "\ic' s'. f (i + 1) = (Seq c' c\<^sub>2, s')) \ \\head (f i) \ head (f (i+1))" (is "\ii final (head (f i))" by fact from this[rule_format] have not_fin_k: "\i final (head (f i))" apply clarify apply (subgoal_tac "i < Suc k") apply blast apply simp done from Suc.hyps [OF this] have hyp: "\ic' s'. f (i + 1) = (Seq c' c\<^sub>2, s')) \ \\ head (f i) \ head (f (i + 1))". show ?case proof (rule le_Suc_cases) fix i assume "i < k" then show "?P i" by (rule hyp [rule_format]) next show "?P k" proof - from hyp [rule_format, of "k - 1"] f_0 obtain c' fs' L' s' where f_k: "f k = (Seq c' c\<^sub>2, s')" by (cases k) auto from inf_comp [rule_format, of k] f_k have "\\(Seq c' c\<^sub>2, s') \ f (k + 1)" by simp moreover from not_fin_Suc [rule_format, of k] f_k have "\ final (c',s')" by (simp add: final_def head_def head_com_def) ultimately obtain c'' s'' where "\\(c', s') \ (c'', s'')" and "f (k + 1) = (Seq c'' c\<^sub>2, s'')" by cases (auto simp add: redex_Seq_False final_def) with f_k show ?thesis by (simp add: head_def head_com_def) qed qed qed lemma infinite_computation_extract_head_Catch: assumes inf_comp: "\i::nat. \\f i \ f (i+1)" assumes f_0: "f 0 = (Catch c\<^sub>1 c\<^sub>2,s)" assumes not_fin: "\i final (head (f i))" shows "\ic' s'. f (i + 1) = (Catch c' c\<^sub>2, s')) \ \\head (f i) \ head (f (i+1))" (is "\ii final (head (f i))" by fact from this[rule_format] have not_fin_k: "\i final (head (f i))" apply clarify apply (subgoal_tac "i < Suc k") apply blast apply simp done from Suc.hyps [OF this] have hyp: "\ic' s'. f (i + 1) = (Catch c' c\<^sub>2, s')) \ \\ head (f i) \ head (f (i + 1))". show ?case proof (rule le_Suc_cases) fix i assume "i < k" then show "?P i" by (rule hyp [rule_format]) next show "?P k" proof - from hyp [rule_format, of "k - 1"] f_0 obtain c' fs' L' s' where f_k: "f k = (Catch c' c\<^sub>2, s')" by (cases k) auto from inf_comp [rule_format, of k] f_k have "\\(Catch c' c\<^sub>2, s') \ f (k + 1)" by simp moreover from not_fin_Suc [rule_format, of k] f_k have "\ final (c',s')" by (simp add: final_def head_def head_com_def) ultimately obtain c'' s'' where "\\(c', s') \ (c'', s'')" and "f (k + 1) = (Catch c'' c\<^sub>2, s'')" by cases (auto simp add: redex_Catch_False final_def)+ with f_k show ?thesis by (simp add: head_def head_com_def) qed qed qed lemma no_inf_Throw: "\ \\(Throw,s) \ \(\)" proof assume "\\ (Throw, s) \ \(\)" then obtain f where step [rule_format]: "\i::nat. \\f i \ f (i+1)" and f_0: "f 0 = (Throw, s)" by (auto simp add: inf_def) from step [of 0, simplified f_0] step [of 1] show False by cases (auto elim: step_elim_cases) qed lemma split_inf_Seq: assumes inf_comp: "\\(Seq c\<^sub>1 c\<^sub>2,s) \ \(\)" shows "\\(c\<^sub>1,s) \ \(\) \ (\s'. \\(c\<^sub>1,s) \\<^sup>* (Skip,s') \ \\(c\<^sub>2,s') \ \(\))" proof - from inf_comp obtain f where step: "\i::nat. \\f i \ f (i+1)" and f_0: "f 0 = (Seq c\<^sub>1 c\<^sub>2, s)" by (auto simp add: inf_def) from f_0 have head_f_0: "head (f 0) = (c\<^sub>1,s)" by (simp add: head_def head_com_def) show ?thesis proof (cases "\i. final (head (f i))") case True define k where "k = (LEAST i. final (head (f i)))" have less_k: "\i final (head (f i))" apply (intro allI impI) apply (unfold k_def) apply (drule not_less_Least) apply auto done from infinite_computation_extract_head_Seq [OF step f_0 this] obtain step_head: "\i\ head (f i) \ head (f (i + 1))" and conf: "\ic' s'. f (i + 1) = (Seq c' c\<^sub>2, s'))" by blast from True have final_f_k: "final (head (f k))" apply - apply (erule exE) apply (drule LeastI) apply (simp add: k_def) done moreover from f_0 conf [rule_format, of "k - 1"] obtain c' s' where f_k: "f k = (Seq c' c\<^sub>2,s')" by (cases k) auto moreover from step_head have steps_head: "\\head (f 0) \\<^sup>* head (f k)" proof (induct k) case 0 thus ?case by simp next case (Suc m) have step: "\i\ head (f i) \ head (f (i + 1))" by fact hence "\i\ head (f i) \ head (f (i + 1))" by auto hence "\\ head (f 0) \\<^sup>* head (f m)" by (rule Suc.hyps) also from step [rule_format, of m] have "\\ head (f m) \ head (f (m + 1))" by simp finally show ?case by simp qed { assume f_k: "f k = (Seq Skip c\<^sub>2, s')" with steps_head have "\\(c\<^sub>1,s) \\<^sup>* (Skip,s')" using head_f_0 by (simp add: head_def head_com_def) moreover from step [rule_format, of k] f_k obtain "\\(Seq Skip c\<^sub>2,s') \ (c\<^sub>2,s')" and f_Suc_k: "f (k + 1) = (c\<^sub>2,s')" by (fastforce elim: step.cases intro: step.intros) define g where "g i = f (i + (k + 1))" for i from f_Suc_k have g_0: "g 0 = (c\<^sub>2,s')" by (simp add: g_def) from step have "\i. \\g i \ g (i + 1)" by (simp add: g_def) with g_0 have "\\(c\<^sub>2,s') \ \(\)" by (auto simp add: inf_def) ultimately have ?thesis by auto } moreover { fix x assume s': "s'=Normal x" and f_k: "f k = (Seq Throw c\<^sub>2, s')" from step [rule_format, of k] f_k s' obtain "\\(Seq Throw c\<^sub>2,s') \ (Throw,s')" and f_Suc_k: "f (k + 1) = (Throw,s')" by (fastforce elim: step_elim_cases intro: step.intros) define g where "g i = f (i + (k + 1))" for i from f_Suc_k have g_0: "g 0 = (Throw,s')" by (simp add: g_def) from step have "\i. \\g i \ g (i + 1)" by (simp add: g_def) with g_0 have "\\(Throw,s') \ \(\)" by (auto simp add: inf_def) with no_inf_Throw have ?thesis by auto } ultimately show ?thesis by (auto simp add: final_def head_def head_com_def) next case False then have not_fin: "\i. \ final (head (f i))" by blast have "\i. \\head (f i) \ head (f (i + 1))" proof fix k from not_fin have "\i<(Suc k). \ final (head (f i))" by simp from infinite_computation_extract_head_Seq [OF step f_0 this ] show "\\ head (f k) \ head (f (k + 1))" by simp qed with head_f_0 have "\\(c\<^sub>1,s) \ \(\)" by (auto simp add: inf_def) thus ?thesis by simp qed qed lemma split_inf_Catch: assumes inf_comp: "\\(Catch c\<^sub>1 c\<^sub>2,s) \ \(\)" shows "\\(c\<^sub>1,s) \ \(\) \ (\s'. \\(c\<^sub>1,s) \\<^sup>* (Throw,Normal s') \ \\(c\<^sub>2,Normal s') \ \(\))" proof - from inf_comp obtain f where step: "\i::nat. \\f i \ f (i+1)" and f_0: "f 0 = (Catch c\<^sub>1 c\<^sub>2, s)" by (auto simp add: inf_def) from f_0 have head_f_0: "head (f 0) = (c\<^sub>1,s)" by (simp add: head_def head_com_def) show ?thesis proof (cases "\i. final (head (f i))") case True define k where "k = (LEAST i. final (head (f i)))" have less_k: "\i final (head (f i))" apply (intro allI impI) apply (unfold k_def) apply (drule not_less_Least) apply auto done from infinite_computation_extract_head_Catch [OF step f_0 this] obtain step_head: "\i\ head (f i) \ head (f (i + 1))" and conf: "\ic' s'. f (i + 1) = (Catch c' c\<^sub>2, s'))" by blast from True have final_f_k: "final (head (f k))" apply - apply (erule exE) apply (drule LeastI) apply (simp add: k_def) done moreover from f_0 conf [rule_format, of "k - 1"] obtain c' s' where f_k: "f k = (Catch c' c\<^sub>2,s')" by (cases k) auto moreover from step_head have steps_head: "\\head (f 0) \\<^sup>* head (f k)" proof (induct k) case 0 thus ?case by simp next case (Suc m) have step: "\i\ head (f i) \ head (f (i + 1))" by fact hence "\i\ head (f i) \ head (f (i + 1))" by auto hence "\\ head (f 0) \\<^sup>* head (f m)" by (rule Suc.hyps) also from step [rule_format, of m] have "\\ head (f m) \ head (f (m + 1))" by simp finally show ?case by simp qed { assume f_k: "f k = (Catch Skip c\<^sub>2, s')" with steps_head have "\\(c\<^sub>1,s) \\<^sup>* (Skip,s')" using head_f_0 by (simp add: head_def head_com_def) moreover from step [rule_format, of k] f_k obtain "\\(Catch Skip c\<^sub>2,s') \ (Skip,s')" and f_Suc_k: "f (k + 1) = (Skip,s')" by (fastforce elim: step.cases intro: step.intros) from step [rule_format, of "k+1", simplified f_Suc_k] have ?thesis by (rule no_step_final') (auto simp add: final_def) } moreover { fix x assume s': "s'=Normal x" and f_k: "f k = (Catch Throw c\<^sub>2, s')" with steps_head have "\\(c\<^sub>1,s) \\<^sup>* (Throw,s')" using head_f_0 by (simp add: head_def head_com_def) moreover from step [rule_format, of k] f_k s' obtain "\\(Catch Throw c\<^sub>2,s') \ (c\<^sub>2,s')" and f_Suc_k: "f (k + 1) = (c\<^sub>2,s')" by (fastforce elim: step_elim_cases intro: step.intros) define g where "g i = f (i + (k + 1))" for i from f_Suc_k have g_0: "g 0 = (c\<^sub>2,s')" by (simp add: g_def) from step have "\i. \\g i \ g (i + 1)" by (simp add: g_def) with g_0 have "\\(c\<^sub>2,s') \ \(\)" by (auto simp add: inf_def) ultimately have ?thesis using s' by auto } ultimately show ?thesis by (auto simp add: final_def head_def head_com_def) next case False then have not_fin: "\i. \ final (head (f i))" by blast have "\i. \\head (f i) \ head (f (i + 1))" proof fix k from not_fin have "\i<(Suc k). \ final (head (f i))" by simp from infinite_computation_extract_head_Catch [OF step f_0 this ] show "\\ head (f k) \ head (f (k + 1))" by simp qed with head_f_0 have "\\(c\<^sub>1,s) \ \(\)" by (auto simp add: inf_def) thus ?thesis by simp qed qed lemma Skip_no_step: "\\(Skip,s) \ cfg \ P" apply (erule no_step_final') apply (simp add: final_def) done lemma not_inf_Stuck: "\ \\(c,Stuck) \ \(\)" proof (induct c) case Skip show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Skip, Stuck)" from f_step [of 0] f_0 show False by (auto elim: Skip_no_step) qed next case (Basic g) thus ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Basic g, Stuck)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Spec r) thus ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Spec r, Stuck)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Seq c\<^sub>1 c\<^sub>2) show ?case proof assume "\\ (Seq c\<^sub>1 c\<^sub>2, Stuck) \ \(\)" from split_inf_Seq [OF this] Seq.hyps show False by (auto dest: steps_Stuck_prop) qed next case (Cond b c\<^sub>1 c\<^sub>2) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Cond b c\<^sub>1 c\<^sub>2, Stuck)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (While b c) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (While b c, Stuck)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Call p) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Call p, Stuck)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (DynCom d) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (DynCom d, Stuck)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Guard m g c) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Guard m g c, Stuck)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case Throw show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Throw, Stuck)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Catch c\<^sub>1 c\<^sub>2) show ?case proof assume "\\ (Catch c\<^sub>1 c\<^sub>2, Stuck) \ \(\)" from split_inf_Catch [OF this] Catch.hyps show False by (auto dest: steps_Stuck_prop) qed qed lemma not_inf_Fault: "\ \\(c,Fault x) \ \(\)" proof (induct c) case Skip show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Skip, Fault x)" from f_step [of 0] f_0 show False by (auto elim: Skip_no_step) qed next case (Basic g) thus ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Basic g, Fault x)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Spec r) thus ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Spec r, Fault x)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Seq c\<^sub>1 c\<^sub>2) show ?case proof assume "\\ (Seq c\<^sub>1 c\<^sub>2, Fault x) \ \(\)" from split_inf_Seq [OF this] Seq.hyps show False by (auto dest: steps_Fault_prop) qed next case (Cond b c\<^sub>1 c\<^sub>2) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Cond b c\<^sub>1 c\<^sub>2, Fault x)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (While b c) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (While b c, Fault x)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Call p) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Call p, Fault x)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (DynCom d) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (DynCom d, Fault x)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Guard m g c) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Guard m g c, Fault x)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case Throw show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Throw, Fault x)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Catch c\<^sub>1 c\<^sub>2) show ?case proof assume "\\ (Catch c\<^sub>1 c\<^sub>2, Fault x) \ \(\)" from split_inf_Catch [OF this] Catch.hyps show False by (auto dest: steps_Fault_prop) qed qed lemma not_inf_Abrupt: "\ \\(c,Abrupt s) \ \(\)" proof (induct c) case Skip show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Skip, Abrupt s)" from f_step [of 0] f_0 show False by (auto elim: Skip_no_step) qed next case (Basic g) thus ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Basic g, Abrupt s)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Spec r) thus ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Spec r, Abrupt s)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Seq c\<^sub>1 c\<^sub>2) show ?case proof assume "\\ (Seq c\<^sub>1 c\<^sub>2, Abrupt s) \ \(\)" from split_inf_Seq [OF this] Seq.hyps show False by (auto dest: steps_Abrupt_prop) qed next case (Cond b c\<^sub>1 c\<^sub>2) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Cond b c\<^sub>1 c\<^sub>2, Abrupt s)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (While b c) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (While b c, Abrupt s)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Call p) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Call p, Abrupt s)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (DynCom d) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (DynCom d, Abrupt s)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Guard m g c) show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Guard m g c, Abrupt s)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case Throw show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Throw, Abrupt s)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Catch c\<^sub>1 c\<^sub>2) show ?case proof assume "\\ (Catch c\<^sub>1 c\<^sub>2, Abrupt s) \ \(\)" from split_inf_Catch [OF this] Catch.hyps show False by (auto dest: steps_Abrupt_prop) qed qed theorem terminates_impl_no_infinite_computation: assumes termi: "\\c \ s" shows "\ \\(c,s) \ \(\)" using termi proof (induct) case (Skip s) thus ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Skip, Normal s)" from f_step [of 0] f_0 show False by (auto elim: Skip_no_step) qed next case (Basic g s) thus ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Basic g, Normal s)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Spec r s) thus ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Spec r, Normal s)" from f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Guard s g c m) have g: "s \ g" by fact have hyp: "\ \\ (c, Normal s) \ \(\)" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Guard m g c, Normal s)" from f_step [of 0] f_0 g have "f 1 = (c,Normal s)" by (fastforce elim: step_elim_cases) with f_step have "\\ (c, Normal s) \ \(\)" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) by simp with hyp show False .. qed next case (GuardFault s g m c) have g: "s \ g" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Guard m g c, Normal s)" from g f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Fault c m) thus ?case by (rule not_inf_Fault) next case (Seq c\<^sub>1 s c\<^sub>2) show ?case proof assume "\\ (Seq c\<^sub>1 c\<^sub>2, Normal s) \ \(\)" from split_inf_Seq [OF this] Seq.hyps show False by (auto intro: steps_Skip_impl_exec) qed next case (CondTrue s b c1 c2) have b: "s \ b" by fact have hyp_c1: "\ \\ (c1, Normal s) \ \(\)" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Cond b c1 c2, Normal s)" from b f_step [of 0] f_0 have "f 1 = (c1,Normal s)" by (auto elim: step_Normal_elim_cases) with f_step have "\\ (c1, Normal s) \ \(\)" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) by simp with hyp_c1 show False by simp qed next case (CondFalse s b c2 c1) have b: "s \ b" by fact have hyp_c2: "\ \\ (c2, Normal s) \ \(\)" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Cond b c1 c2, Normal s)" from b f_step [of 0] f_0 have "f 1 = (c2,Normal s)" by (auto elim: step_Normal_elim_cases) with f_step have "\\ (c2, Normal s) \ \(\)" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) by simp with hyp_c2 show False by simp qed next case (WhileTrue s b c) have b: "s \ b" by fact have hyp_c: "\ \\ (c, Normal s) \ \(\)" by fact have hyp_w: "\s'. \\ \c,Normal s\ \ s' \ \\While b c \ s' \ \ \\ (While b c, s') \ \(\)" by fact have not_inf_Seq: "\ \\ (Seq c (While b c), Normal s) \ \(\)" proof assume "\\ (Seq c (While b c), Normal s) \ \(\)" from split_inf_Seq [OF this] hyp_c hyp_w show False by (auto intro: steps_Skip_impl_exec) qed show ?case proof assume "\\ (While b c, Normal s) \ \(\)" then obtain f where f_step: "\i. \\f i \ f (Suc i)" and f_0: "f 0 = (While b c, Normal s)" by (auto simp add: inf_def) from f_step [of 0] f_0 b have "f 1 = (Seq c (While b c),Normal s)" by (auto elim: step_Normal_elim_cases) with f_step have "\\ (Seq c (While b c), Normal s) \ \(\)" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) by simp with not_inf_Seq show False by simp qed next case (WhileFalse s b c) have b: "s \ b" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (While b c, Normal s)" from b f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Call p bdy s) have bdy: "\ p = Some bdy" by fact have hyp: "\ \\ (bdy, Normal s) \ \(\)" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Call p, Normal s)" from bdy f_step [of 0] f_0 have "f 1 = (bdy,Normal s)" by (auto elim: step_Normal_elim_cases) with f_step have "\\ (bdy, Normal s) \ \(\)" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) by simp with hyp show False by simp qed next case (CallUndefined p s) have no_bdy: "\ p = None" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Call p, Normal s)" from no_bdy f_step [of 0] f_0 f_step [of 1] show False by (fastforce elim: Skip_no_step step_elim_cases) qed next case (Stuck c) show ?case by (rule not_inf_Stuck) next case (DynCom c s) have hyp: "\ \\ (c s, Normal s) \ \(\)" by fact show ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (DynCom c, Normal s)" from f_step [of 0] f_0 have "f (Suc 0) = (c s, Normal s)" by (auto elim: step_elim_cases) with f_step have "\\ (c s, Normal s) \ \(\)" apply (simp add: inf_def) apply (rule_tac x="\i. f (Suc i)" in exI) by simp with hyp show False by simp qed next case (Throw s) thus ?case proof (rule not_infI) fix f assume f_step: "\i. \\f i \ f (Suc i)" assume f_0: "f 0 = (Throw, Normal s)" from f_step [of 0] f_0 show False by (auto elim: step_elim_cases) qed next case (Abrupt c) show ?case by (rule not_inf_Abrupt) next case (Catch c\<^sub>1 s c\<^sub>2) show ?case proof assume "\\ (Catch c\<^sub>1 c\<^sub>2, Normal s) \ \(\)" from split_inf_Catch [OF this] Catch.hyps show False by (auto intro: steps_Throw_impl_exec) qed qed definition termi_call_steps :: "('s,'p,'f) body \ (('s \ 'p) \ ('s \ 'p))set" where "termi_call_steps \ = {((t,q),(s,p)). \\Call p\Normal s \ (\c. \\(Call p,Normal s) \\<^sup>+ (c,Normal t) \ redex c = Call q)}" primrec subst_redex:: "('s,'p,'f)com \ ('s,'p,'f)com \ ('s,'p,'f)com" where "subst_redex Skip c = c" | "subst_redex (Basic f) c = c" | "subst_redex (Spec r) c = c" | "subst_redex (Seq c\<^sub>1 c\<^sub>2) c = Seq (subst_redex c\<^sub>1 c) c\<^sub>2" | "subst_redex (Cond b c\<^sub>1 c\<^sub>2) c = c" | "subst_redex (While b c') c = c" | "subst_redex (Call p) c = c" | "subst_redex (DynCom d) c = c" | "subst_redex (Guard f b c') c = c" | "subst_redex (Throw) c = c" | "subst_redex (Catch c\<^sub>1 c\<^sub>2) c = Catch (subst_redex c\<^sub>1 c) c\<^sub>2" lemma subst_redex_redex: "subst_redex c (redex c) = c" by (induct c) auto lemma redex_subst_redex: "redex (subst_redex c r) = redex r" by (induct c) auto lemma step_redex': shows "\\(redex c,s) \ (r',s') \ \\(c,s) \ (subst_redex c r',s')" by (induct c) (auto intro: step.Seq step.Catch) lemma step_redex: shows "\\(r,s) \ (r',s') \ \\(subst_redex c r,s) \ (subst_redex c r',s')" by (induct c) (auto intro: step.Seq step.Catch) lemma steps_redex: assumes steps: "\\ (r, s) \\<^sup>* (r', s')" shows "\c. \\(subst_redex c r,s) \\<^sup>* (subst_redex c r',s')" using steps proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans]) case Refl show "\\ (subst_redex c r', s') \\<^sup>* (subst_redex c r', s')" by simp next case (Trans r s r'' s'') have "\\ (r, s) \ (r'', s'')" by fact from step_redex [OF this] have "\\ (subst_redex c r, s) \ (subst_redex c r'', s'')". also have "\\ (subst_redex c r'', s'') \\<^sup>* (subst_redex c r', s')" by fact finally show ?case . qed ML \ ML_Thms.bind_thm ("trancl_induct2", Split_Rule.split_rule @{context} (Rule_Insts.read_instantiate @{context} [((("a", 0), Position.none), "(aa, ab)"), ((("b", 0), Position.none), "(ba, bb)")] [] @{thm trancl_induct})); \ lemma steps_redex': assumes steps: "\\ (r, s) \\<^sup>+ (r', s')" shows "\c. \\(subst_redex c r,s) \\<^sup>+ (subst_redex c r',s')" using steps proof (induct rule: tranclp_induct2 [consumes 1,case_names Step Trans]) case (Step r' s') have "\\ (r, s) \ (r', s')" by fact then have "\\ (subst_redex c r, s) \ (subst_redex c r', s')" by (rule step_redex) then show "\\ (subst_redex c r, s) \\<^sup>+ (subst_redex c r', s')".. next case (Trans r' s' r'' s'') have "\\ (subst_redex c r, s) \\<^sup>+ (subst_redex c r', s')" by fact also have "\\ (r', s') \ (r'', s'')" by fact hence "\\ (subst_redex c r', s') \ (subst_redex c r'', s'')" by (rule step_redex) finally show "\\ (subst_redex c r, s) \\<^sup>+ (subst_redex c r'', s'')" . qed primrec seq:: "(nat \ ('s,'p,'f)com) \ 'p \ nat \ ('s,'p,'f)com" where "seq c p 0 = Call p" | "seq c p (Suc i) = subst_redex (seq c p i) (c i)" lemma renumber': assumes f: "\i. (a,f i) \ r\<^sup>* \ (f i,f(Suc i)) \ r" assumes a_b: "(a,b) \ r\<^sup>*" shows "b = f 0 \ (\f. f 0 = a \ (\i. (f i, f(Suc i)) \ r))" using a_b proof (induct rule: converse_rtrancl_induct [consumes 1]) assume "b = f 0" with f show "\f. f 0 = b \ (\i. (f i, f (Suc i)) \ r)" by blast next fix a z assume a_z: "(a, z) \ r" and "(z, b) \ r\<^sup>*" assume "b = f 0 \ \f. f 0 = z \ (\i. (f i, f (Suc i)) \ r)" "b = f 0" then obtain f where f0: "f 0 = z" and seq: "\i. (f i, f (Suc i)) \ r" by iprover { fix i have "((\i. case i of 0 \ a | Suc i \ f i) i, f i) \ r" using seq a_z f0 by (cases i) auto } then show "\f. f 0 = a \ (\i. (f i, f (Suc i)) \ r)" by - (rule exI [where x="\i. case i of 0 \ a | Suc i \ f i"],simp) qed lemma renumber: "\i. (a,f i) \ r\<^sup>* \ (f i,f(Suc i)) \ r \ \f. f 0 = a \ (\i. (f i, f(Suc i)) \ r)" by (blast dest:renumber') lemma lem: "\y. r\<^sup>+\<^sup>+ a y \ P a \ P y \ ((b,a) \ {(y,x). P x \ r x y}\<^sup>+) = ((b,a) \ {(y,x). P x \ r\<^sup>+\<^sup>+ x y})" apply(rule iffI) apply clarify apply(erule trancl_induct) apply blast apply(blast intro:tranclp_trans) apply clarify apply(erule tranclp_induct) apply blast apply(blast intro:trancl_trans) done corollary terminates_impl_no_infinite_trans_computation: assumes terminates: "\\c\s" shows "\(\f. f 0 = (c,s) \ (\i. \\f i \\<^sup>+ f(Suc i)))" proof - have "wf({(y,x). \\(c,s) \\<^sup>* x \ \\x \ y}\<^sup>+)" proof (rule wf_trancl) show "wf {(y, x). \\(c,s) \\<^sup>* x \ \\x \ y}" proof (simp only: wf_iff_no_infinite_down_chain,clarify,simp) fix f assume "\i. \\(c,s) \\<^sup>* f i \ \\f i \ f (Suc i)" hence "\f. f (0::nat) = (c,s) \ (\i. \\f i \ f (Suc i))" by (rule renumber [to_pred]) moreover from terminates_impl_no_infinite_computation [OF terminates] have "\ (\f. f (0::nat) = (c, s) \ (\i. \\f i \ f (Suc i)))" by (simp add: inf_def) ultimately show False by simp qed qed hence "\ (\f. \i. (f (Suc i), f i) \ {(y, x). \\(c, s) \\<^sup>* x \ \\x \ y}\<^sup>+)" by (simp add: wf_iff_no_infinite_down_chain) thus ?thesis proof (rule contrapos_nn) assume "\f. f (0::nat) = (c, s) \ (\i. \\f i \\<^sup>+ f (Suc i))" then obtain f where f0: "f 0 = (c, s)" and seq: "\i. \\f i \\<^sup>+ f (Suc i)" by iprover show "\f. \i. (f (Suc i), f i) \ {(y, x). \\(c, s) \\<^sup>* x \ \\x \ y}\<^sup>+" proof (rule exI [where x=f],rule allI) fix i show "(f (Suc i), f i) \ {(y, x). \\(c, s) \\<^sup>* x \ \\x \ y}\<^sup>+" proof - { fix i have "\\(c,s) \\<^sup>* f i" proof (induct i) case 0 show "\\(c, s) \\<^sup>* f 0" by (simp add: f0) next case (Suc n) have "\\(c, s) \\<^sup>* f n" by fact with seq show "\\(c, s) \\<^sup>* f (Suc n)" by (blast intro: tranclp_into_rtranclp rtranclp_trans) qed } hence "\\(c,s) \\<^sup>* f i" by iprover with seq have "(f (Suc i), f i) \ {(y, x). \\(c, s) \\<^sup>* x \ \\x \\<^sup>+ y}" by clarsimp moreover have "\y. \\f i \\<^sup>+ y\\\(c, s) \\<^sup>* f i\\\(c, s) \\<^sup>* y" by (blast intro: tranclp_into_rtranclp rtranclp_trans) ultimately show ?thesis by (subst lem ) qed qed qed qed theorem wf_termi_call_steps: "wf (termi_call_steps \)" proof (simp only: termi_call_steps_def wf_iff_no_infinite_down_chain, clarify,simp) fix f assume inf: "\i. (\(t, q) (s, p). \\Call p \ Normal s \ (\c. \\ (Call p, Normal s) \\<^sup>+ (c, Normal t) \ redex c = Call q)) (f (Suc i)) (f i)" define s where "s i = fst (f i)" for i :: nat define p where "p i = (snd (f i)::'b)" for i :: nat from inf have inf': "\i. \\Call (p i) \ Normal (s i) \ (\c. \\ (Call (p i), Normal (s i)) \\<^sup>+ (c, Normal (s (i+1))) \ redex c = Call (p (i+1)))" apply - apply (rule allI) apply (erule_tac x=i in allE) apply (auto simp add: s_def p_def) done show False proof - from inf' have "\c. \i. \\Call (p i) \ Normal (s i) \ \\ (Call (p i), Normal (s i)) \\<^sup>+ (c i, Normal (s (i+1))) \ redex (c i) = Call (p (i+1))" apply - apply (rule choice) by blast then obtain c where termi_c: "\i. \\Call (p i) \ Normal (s i)" and steps_c: "\i. \\ (Call (p i), Normal (s i)) \\<^sup>+ (c i, Normal (s (i+1)))" and red_c: "\i. redex (c i) = Call (p (i+1))" by auto define g where "g i = (seq c (p 0) i,Normal (s i)::('a,'c) xstate)" for i from red_c [rule_format, of 0] have "g 0 = (Call (p 0), Normal (s 0))" by (simp add: g_def) moreover { fix i have "redex (seq c (p 0) i) = Call (p i)" by (induct i) (auto simp add: redex_subst_redex red_c) from this [symmetric] have "subst_redex (seq c (p 0) i) (Call (p i)) = (seq c (p 0) i)" by (simp add: subst_redex_redex) } note subst_redex_seq = this have "\i. \\ (g i) \\<^sup>+ (g (i+1))" proof fix i from steps_c [rule_format, of i] have "\\ (Call (p i), Normal (s i)) \\<^sup>+ (c i, Normal (s (i + 1)))". from steps_redex' [OF this, of "(seq c (p 0) i)"] have "\\ (subst_redex (seq c (p 0) i) (Call (p i)), Normal (s i)) \\<^sup>+ (subst_redex (seq c (p 0) i) (c i), Normal (s (i + 1)))" . hence "\\ (seq c (p 0) i, Normal (s i)) \\<^sup>+ (seq c (p 0) (i+1), Normal (s (i + 1)))" by (simp add: subst_redex_seq) thus "\\ (g i) \\<^sup>+ (g (i+1))" by (simp add: g_def) qed moreover from terminates_impl_no_infinite_trans_computation [OF termi_c [rule_format, of 0]] have "\ (\f. f 0 = (Call (p 0), Normal (s 0)) \ (\i. \\ f i \\<^sup>+ f (Suc i)))" . ultimately show False by auto qed qed lemma no_infinite_computation_implies_wf: assumes not_inf: "\ \\ (c, s) \ \(\)" shows "wf {(c2,c1). \ \ (c,s) \\<^sup>* c1 \ \ \ c1 \ c2}" proof (simp only: wf_iff_no_infinite_down_chain,clarify, simp) fix f assume "\i. \\(c, s) \\<^sup>* f i \ \\f i \ f (Suc i)" hence "\f. f 0 = (c, s) \ (\i. \\f i \ f (Suc i))" by (rule renumber [to_pred]) moreover from not_inf have "\ (\f. f 0 = (c, s) \ (\i. \\f i \ f (Suc i)))" by (simp add: inf_def) ultimately show False by simp qed lemma not_final_Stuck_step: "\ final (c,Stuck) \ \c' s'. \\ (c, Stuck) \ (c',s')" by (induct c) (fastforce intro: step.intros simp add: final_def)+ lemma not_final_Abrupt_step: "\ final (c,Abrupt s) \ \c' s'. \\ (c, Abrupt s) \ (c',s')" by (induct c) (fastforce intro: step.intros simp add: final_def)+ lemma not_final_Fault_step: "\ final (c,Fault f) \ \c' s'. \\ (c, Fault f) \ (c',s')" by (induct c) (fastforce intro: step.intros simp add: final_def)+ lemma not_final_Normal_step: "\ final (c,Normal s) \ \c' s'. \\ (c, Normal s) \ (c',s')" proof (induct c) case Skip thus ?case by (fastforce intro: step.intros simp add: final_def) next case Basic thus ?case by (fastforce intro: step.intros) next case (Spec r) thus ?case by (cases "\t. (s,t) \ r") (fastforce intro: step.intros)+ next case (Seq c\<^sub>1 c\<^sub>2) thus ?case by (cases "final (c\<^sub>1,Normal s)") (fastforce intro: step.intros simp add: final_def)+ next case (Cond b c1 c2) show ?case by (cases "s \ b") (fastforce intro: step.intros)+ next case (While b c) show ?case by (cases "s \ b") (fastforce intro: step.intros)+ next case (Call p) show ?case by (cases "\ p") (fastforce intro: step.intros)+ next case DynCom thus ?case by (fastforce intro: step.intros) next case (Guard f g c) show ?case by (cases "s \ g") (fastforce intro: step.intros)+ next case Throw thus ?case by (fastforce intro: step.intros simp add: final_def) next case (Catch c\<^sub>1 c\<^sub>2) thus ?case by (cases "final (c\<^sub>1,Normal s)") (fastforce intro: step.intros simp add: final_def)+ qed lemma final_termi: "final (c,s) \ \\c\s" by (cases s) (auto simp add: final_def terminates.intros) lemma split_computation: assumes steps: "\\ (c, s) \\<^sup>* (c\<^sub>f, s\<^sub>f)" assumes not_final: "\ final (c,s)" assumes final: "final (c\<^sub>f,s\<^sub>f)" shows "\c' s'. \\ (c, s) \ (c',s') \ \\ (c', s') \\<^sup>* (c\<^sub>f, s\<^sub>f)" using steps not_final final proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans]) case Refl thus ?case by simp next case (Trans c s c' s') thus ?case by auto qed lemma wf_implies_termi_reach_step_case: assumes hyp: "\c' s'. \\ (c, Normal s) \ (c', s') \ \\c' \ s'" shows "\\c \ Normal s" using hyp proof (induct c) case Skip show ?case by (fastforce intro: terminates.intros) next case Basic show ?case by (fastforce intro: terminates.intros) next case (Spec r) show ?case by (cases "\t. (s,t)\r") (fastforce intro: terminates.intros)+ next case (Seq c\<^sub>1 c\<^sub>2) have hyp: "\c' s'. \\ (Seq c\<^sub>1 c\<^sub>2, Normal s) \ (c', s') \ \\c' \ s'" by fact show ?case proof (rule terminates.Seq) { fix c' s' assume step_c\<^sub>1: "\\ (c\<^sub>1, Normal s) \ (c', s')" have "\\c' \ s'" proof - from step_c\<^sub>1 have "\\ (Seq c\<^sub>1 c\<^sub>2, Normal s) \ (Seq c' c\<^sub>2, s')" by (rule step.Seq) from hyp [OF this] have "\\Seq c' c\<^sub>2 \ s'". thus "\\c'\ s'" by cases auto qed } from Seq.hyps (1) [OF this] show "\\c\<^sub>1 \ Normal s". next show "\s'. \\ \c\<^sub>1,Normal s\ \ s' \ \\c\<^sub>2 \ s'" proof (intro allI impI) fix s' assume exec_c\<^sub>1: "\\ \c\<^sub>1,Normal s\ \ s'" show "\\c\<^sub>2 \ s'" proof (cases "final (c\<^sub>1,Normal s)") case True hence "c\<^sub>1=Skip \ c\<^sub>1=Throw" by (simp add: final_def) thus ?thesis proof assume Skip: "c\<^sub>1=Skip" have "\\(Seq Skip c\<^sub>2,Normal s) \ (c\<^sub>2,Normal s)" by (rule step.SeqSkip) from hyp [simplified Skip, OF this] have "\\c\<^sub>2 \ Normal s" . moreover from exec_c\<^sub>1 Skip have "s'=Normal s" by (auto elim: exec_Normal_elim_cases) ultimately show ?thesis by simp next assume Throw: "c\<^sub>1=Throw" with exec_c\<^sub>1 have "s'=Abrupt s" by (auto elim: exec_Normal_elim_cases) thus ?thesis by auto qed next case False from exec_impl_steps [OF exec_c\<^sub>1] obtain c\<^sub>f t where steps_c\<^sub>1: "\\ (c\<^sub>1, Normal s) \\<^sup>* (c\<^sub>f, t)" and fin:"(case s' of Abrupt x \ c\<^sub>f = Throw \ t = Normal x | _ \ c\<^sub>f = Skip \ t = s')" by (fastforce split: xstate.splits) with fin have final: "final (c\<^sub>f,t)" by (cases s') (auto simp add: final_def) from split_computation [OF steps_c\<^sub>1 False this] obtain c'' s'' where first: "\\ (c\<^sub>1, Normal s) \ (c'', s'')" and rest: "\\ (c'', s'') \\<^sup>* (c\<^sub>f, t)" by blast from step.Seq [OF first] have "\\ (Seq c\<^sub>1 c\<^sub>2, Normal s) \ (Seq c'' c\<^sub>2, s'')". from hyp [OF this] have termi_s'': "\\Seq c'' c\<^sub>2 \ s''". show ?thesis proof (cases s'') case (Normal x) from termi_s'' [simplified Normal] have termi_c\<^sub>2: "\t. \\ \c'',Normal x\ \ t \ \\c\<^sub>2 \ t" by cases show ?thesis proof (cases "\x'. s'=Abrupt x'") case False with fin obtain "c\<^sub>f=Skip" "t=s'" by (cases s') auto from steps_Skip_impl_exec [OF rest [simplified this]] Normal have "\\ \c'',Normal x\ \ s'" by simp from termi_c\<^sub>2 [rule_format, OF this] show "\\c\<^sub>2 \ s'" . next case True with fin obtain x' where s': "s'=Abrupt x'" and "c\<^sub>f=Throw" "t=Normal x'" by auto from steps_Throw_impl_exec [OF rest [simplified this]] Normal have "\\ \c'',Normal x\ \ Abrupt x'" by simp from termi_c\<^sub>2 [rule_format, OF this] s' show "\\c\<^sub>2 \ s'" by simp qed next case (Abrupt x) from steps_Abrupt_prop [OF rest this] have "t=Abrupt x" by simp with fin have "s'=Abrupt x" by (cases s') auto thus "\\c\<^sub>2 \ s'" by auto next case (Fault f) from steps_Fault_prop [OF rest this] have "t=Fault f" by simp with fin have "s'=Fault f" by (cases s') auto thus "\\c\<^sub>2 \ s'" by auto next case Stuck from steps_Stuck_prop [OF rest this] have "t=Stuck" by simp with fin have "s'=Stuck" by (cases s') auto thus "\\c\<^sub>2 \ s'" by auto qed qed qed qed next case (Cond b c\<^sub>1 c\<^sub>2) have hyp: "\c' s'. \\ (Cond b c\<^sub>1 c\<^sub>2, Normal s) \ (c', s') \ \\c' \ s'" by fact show ?case proof (cases "s\b") case True then have "\\ (Cond b c\<^sub>1 c\<^sub>2, Normal s) \ (c\<^sub>1, Normal s)" by (rule step.CondTrue) from hyp [OF this] have "\\c\<^sub>1 \ Normal s". with True show ?thesis by (auto intro: terminates.intros) next case False then have "\\ (Cond b c\<^sub>1 c\<^sub>2, Normal s) \ (c\<^sub>2, Normal s)" by (rule step.CondFalse) from hyp [OF this] have "\\c\<^sub>2 \ Normal s". with False show ?thesis by (auto intro: terminates.intros) qed next case (While b c) have hyp: "\c' s'. \\ (While b c, Normal s) \ (c', s') \ \\c' \ s'" by fact show ?case proof (cases "s\b") case True then have "\\ (While b c, Normal s) \ (Seq c (While b c), Normal s)" by (rule step.WhileTrue) from hyp [OF this] have "\\(Seq c (While b c)) \ Normal s". with True show ?thesis by (auto elim: terminates_Normal_elim_cases intro: terminates.intros) next case False thus ?thesis by (auto intro: terminates.intros) qed next case (Call p) have hyp: "\c' s'. \\ (Call p, Normal s) \ (c', s') \ \\c' \ s'" by fact show ?case proof (cases "\ p") case None thus ?thesis by (auto intro: terminates.intros) next case (Some bdy) then have "\\ (Call p, Normal s) \ (bdy, Normal s)" by (rule step.Call) from hyp [OF this] have "\\bdy \ Normal s". with Some show ?thesis by (auto intro: terminates.intros) qed next case (DynCom c) have hyp: "\c' s'. \\ (DynCom c, Normal s) \ (c', s') \ \\c' \ s'" by fact have "\\ (DynCom c, Normal s) \ (c s, Normal s)" by (rule step.DynCom) from hyp [OF this] have "\\c s \ Normal s". then show ?case by (auto intro: terminates.intros) next case (Guard f g c) have hyp: "\c' s'. \\ (Guard f g c, Normal s) \ (c', s') \ \\c' \ s'" by fact show ?case proof (cases "s\g") case True then have "\\ (Guard f g c, Normal s) \ (c, Normal s)" by (rule step.Guard) from hyp [OF this] have "\\c\ Normal s". with True show ?thesis by (auto intro: terminates.intros) next case False thus ?thesis by (auto intro: terminates.intros) qed next case Throw show ?case by (auto intro: terminates.intros) next case (Catch c\<^sub>1 c\<^sub>2) have hyp: "\c' s'. \\ (Catch c\<^sub>1 c\<^sub>2, Normal s) \ (c', s') \ \\c' \ s'" by fact show ?case proof (rule terminates.Catch) { fix c' s' assume step_c\<^sub>1: "\\ (c\<^sub>1, Normal s) \ (c', s')" have "\\c' \ s'" proof - from step_c\<^sub>1 have "\\ (Catch c\<^sub>1 c\<^sub>2, Normal s) \ (Catch c' c\<^sub>2, s')" by (rule step.Catch) from hyp [OF this] have "\\Catch c' c\<^sub>2 \ s'". thus "\\c'\ s'" by cases auto qed } from Catch.hyps (1) [OF this] show "\\c\<^sub>1 \ Normal s". next show "\s'. \\ \c\<^sub>1,Normal s\ \ Abrupt s' \ \\c\<^sub>2 \ Normal s'" proof (intro allI impI) fix s' assume exec_c\<^sub>1: "\\ \c\<^sub>1,Normal s\ \ Abrupt s'" show "\\c\<^sub>2 \ Normal s'" proof (cases "final (c\<^sub>1,Normal s)") case True with exec_c\<^sub>1 have Throw: "c\<^sub>1=Throw" by (auto simp add: final_def elim: exec_Normal_elim_cases) have "\\(Catch Throw c\<^sub>2,Normal s) \ (c\<^sub>2,Normal s)" by (rule step.CatchThrow) from hyp [simplified Throw, OF this] have "\\c\<^sub>2 \ Normal s". moreover from exec_c\<^sub>1 Throw have "s'=s" by (auto elim: exec_Normal_elim_cases) ultimately show ?thesis by simp next case False from exec_impl_steps [OF exec_c\<^sub>1] obtain c\<^sub>f t where steps_c\<^sub>1: "\\ (c\<^sub>1, Normal s) \\<^sup>* (Throw, Normal s')" by (fastforce split: xstate.splits) from split_computation [OF steps_c\<^sub>1 False] obtain c'' s'' where first: "\\ (c\<^sub>1, Normal s) \ (c'', s'')" and rest: "\\ (c'', s'') \\<^sup>* (Throw, Normal s')" by (auto simp add: final_def) from step.Catch [OF first] have "\\ (Catch c\<^sub>1 c\<^sub>2, Normal s) \ (Catch c'' c\<^sub>2, s'')". from hyp [OF this] have "\\Catch c'' c\<^sub>2 \ s''". moreover from steps_Throw_impl_exec [OF rest] have "\\ \c'',s''\ \ Abrupt s'". moreover from rest obtain x where "s''=Normal x" by (cases s'') (auto dest: steps_Fault_prop steps_Abrupt_prop steps_Stuck_prop) ultimately show ?thesis by (fastforce elim: terminates_elim_cases) qed qed qed qed lemma wf_implies_termi_reach: assumes wf: "wf {(cfg2,cfg1). \ \ (c,s) \\<^sup>* cfg1 \ \ \ cfg1 \ cfg2}" shows "\c1 s1. \\ \ (c,s) \\<^sup>* cfg1; cfg1=(c1,s1)\\ \\c1\s1" using wf proof (induct cfg1, simp) fix c1 s1 assume reach: "\\ (c, s) \\<^sup>* (c1, s1)" assume hyp_raw: "\y c2 s2. \\\ (c1, s1) \ (c2, s2); \\ (c, s) \\<^sup>* (c2, s2); y = (c2, s2)\ \ \\c2 \ s2" have hyp: "\c2 s2. \\ (c1, s1) \ (c2, s2) \ \\c2 \ s2" apply - apply (rule hyp_raw) apply assumption using reach apply simp apply (rule refl) done show "\\c1 \ s1" proof (cases s1) case (Normal s1') with wf_implies_termi_reach_step_case [OF hyp [simplified Normal]] show ?thesis by auto qed (auto intro: terminates.intros) qed theorem no_infinite_computation_impl_terminates: assumes not_inf: "\ \\ (c, s) \ \(\)" shows "\\c\s" proof - from no_infinite_computation_implies_wf [OF not_inf] have wf: "wf {(c2, c1). \\(c, s) \\<^sup>* c1 \ \\c1 \ c2}". show ?thesis by (rule wf_implies_termi_reach [OF wf]) auto qed corollary terminates_iff_no_infinite_computation: "\\c\s = (\ \\ (c, s) \ \(\))" apply (rule) apply (erule terminates_impl_no_infinite_computation) apply (erule no_infinite_computation_impl_terminates) done (* ************************************************************************* *) subsection \Generalised Redexes\ (* ************************************************************************* *) text \ For an important lemma for the completeness proof of the Hoare-logic for total correctness we need a generalisation of @{const "redex"} that not only yield the redex itself but all the enclosing statements as well. \ primrec redexes:: "('s,'p,'f)com \ ('s,'p,'f)com set" where "redexes Skip = {Skip}" | "redexes (Basic f) = {Basic f}" | "redexes (Spec r) = {Spec r}" | "redexes (Seq c\<^sub>1 c\<^sub>2) = {Seq c\<^sub>1 c\<^sub>2} \ redexes c\<^sub>1" | "redexes (Cond b c\<^sub>1 c\<^sub>2) = {Cond b c\<^sub>1 c\<^sub>2}" | "redexes (While b c) = {While b c}" | "redexes (Call p) = {Call p}" | "redexes (DynCom d) = {DynCom d}" | "redexes (Guard f b c) = {Guard f b c}" | "redexes (Throw) = {Throw}" | "redexes (Catch c\<^sub>1 c\<^sub>2) = {Catch c\<^sub>1 c\<^sub>2} \ redexes c\<^sub>1" lemma root_in_redexes: "c \ redexes c" apply (induct c) apply auto done lemma redex_in_redexes: "redex c \ redexes c" apply (induct c) apply auto done lemma redex_redexes: "\c'. \c' \ redexes c; redex c' = c'\ \ redex c = c'" apply (induct c) apply auto done lemma step_redexes: shows "\r r'. \\\(r,s) \ (r',s'); r \ redexes c\ \ \c'. \\(c,s) \ (c',s') \ r' \ redexes c'" proof (induct c) case Skip thus ?case by (fastforce intro: step.intros elim: step_elim_cases) next case Basic thus ?case by (fastforce intro: step.intros elim: step_elim_cases) next case Spec thus ?case by (fastforce intro: step.intros elim: step_elim_cases) next case (Seq c\<^sub>1 c\<^sub>2) have "r \ redexes (Seq c\<^sub>1 c\<^sub>2)" by fact hence r: "r = Seq c\<^sub>1 c\<^sub>2 \ r \ redexes c\<^sub>1" by simp have step_r: "\\ (r, s) \ (r', s')" by fact from r show ?case proof assume "r = Seq c\<^sub>1 c\<^sub>2" with step_r show ?case by (auto simp add: root_in_redexes) next assume r: "r \ redexes c\<^sub>1" from Seq.hyps (1) [OF step_r this] obtain c' where step_c\<^sub>1: "\\ (c\<^sub>1, s) \ (c', s')" and r': "r' \ redexes c'" by blast from step.Seq [OF step_c\<^sub>1] have "\\ (Seq c\<^sub>1 c\<^sub>2, s) \ (Seq c' c\<^sub>2, s')". with r' show ?case by auto qed next case Cond thus ?case by (fastforce intro: step.intros elim: step_elim_cases simp add: root_in_redexes) next case While thus ?case by (fastforce intro: step.intros elim: step_elim_cases simp add: root_in_redexes) next case Call thus ?case by (fastforce intro: step.intros elim: step_elim_cases simp add: root_in_redexes) next case DynCom thus ?case by (fastforce intro: step.intros elim: step_elim_cases simp add: root_in_redexes) next case Guard thus ?case by (fastforce intro: step.intros elim: step_elim_cases simp add: root_in_redexes) next case Throw thus ?case by (fastforce intro: step.intros elim: step_elim_cases simp add: root_in_redexes) next case (Catch c\<^sub>1 c\<^sub>2) have "r \ redexes (Catch c\<^sub>1 c\<^sub>2)" by fact hence r: "r = Catch c\<^sub>1 c\<^sub>2 \ r \ redexes c\<^sub>1" by simp have step_r: "\\ (r, s) \ (r', s')" by fact from r show ?case proof assume "r = Catch c\<^sub>1 c\<^sub>2" with step_r show ?case by (auto simp add: root_in_redexes) next assume r: "r \ redexes c\<^sub>1" from Catch.hyps (1) [OF step_r this] obtain c' where step_c\<^sub>1: "\\ (c\<^sub>1, s) \ (c', s')" and r': "r' \ redexes c'" by blast from step.Catch [OF step_c\<^sub>1] have "\\ (Catch c\<^sub>1 c\<^sub>2, s) \ (Catch c' c\<^sub>2, s')". with r' show ?case by auto qed qed lemma steps_redexes: assumes steps: "\\ (r, s) \\<^sup>* (r', s')" shows "\c. r \ redexes c \ \c'. \\(c,s) \\<^sup>* (c',s') \ r' \ redexes c'" using steps proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans]) case Refl then show "\c'. \\ (c, s') \\<^sup>* (c', s') \ r' \ redexes c'" by auto next case (Trans r s r'' s'') have "\\ (r, s) \ (r'', s'')" "r \ redexes c" by fact+ from step_redexes [OF this] obtain c' where step: "\\ (c, s) \ (c', s'')" and r'': "r'' \ redexes c'" by blast note step also from Trans.hyps (3) [OF r''] obtain c'' where steps: "\\ (c', s'') \\<^sup>* (c'', s')" and r': "r' \ redexes c''" by blast note steps finally show ?case using r' by blast qed lemma steps_redexes': assumes steps: "\\ (r, s) \\<^sup>+ (r', s')" shows "\c. r \ redexes c \ \c'. \\(c,s) \\<^sup>+ (c',s') \ r' \ redexes c'" using steps proof (induct rule: tranclp_induct2 [consumes 1, case_names Step Trans]) case (Step r' s' c') have "\\ (r, s) \ (r', s')" "r \ redexes c'" by fact+ from step_redexes [OF this] show ?case by (blast intro: r_into_trancl) next case (Trans r' s' r'' s'') from Trans obtain c' where steps: "\\ (c, s) \\<^sup>+ (c', s')" and r': "r' \ redexes c'" by blast note steps moreover have "\\ (r', s') \ (r'', s'')" by fact from step_redexes [OF this r'] obtain c'' where step: "\\ (c', s') \ (c'', s'')" and r'': "r'' \ redexes c''" by blast note step finally show ?case using r'' by blast qed lemma step_redexes_Seq: assumes step: "\\(r,s) \ (r',s')" assumes Seq: "Seq r c\<^sub>2 \ redexes c" shows "\c'. \\(c,s) \ (c',s') \ Seq r' c\<^sub>2 \ redexes c'" proof - from step.Seq [OF step] have "\\ (Seq r c\<^sub>2, s) \ (Seq r' c\<^sub>2, s')". from step_redexes [OF this Seq] show ?thesis . qed lemma steps_redexes_Seq: assumes steps: "\\ (r, s) \\<^sup>* (r', s')" shows "\c. Seq r c\<^sub>2 \ redexes c \ \c'. \\(c,s) \\<^sup>* (c',s') \ Seq r' c\<^sub>2 \ redexes c'" using steps proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans]) case Refl then show ?case by (auto) next case (Trans r s r'' s'') have "\\ (r, s) \ (r'', s'')" "Seq r c\<^sub>2 \ redexes c" by fact+ from step_redexes_Seq [OF this] obtain c' where step: "\\ (c, s) \ (c', s'')" and r'': "Seq r'' c\<^sub>2 \ redexes c'" by blast note step also from Trans.hyps (3) [OF r''] obtain c'' where steps: "\\ (c', s'') \\<^sup>* (c'', s')" and r': "Seq r' c\<^sub>2 \ redexes c''" by blast note steps finally show ?case using r' by blast qed lemma steps_redexes_Seq': assumes steps: "\\ (r, s) \\<^sup>+ (r', s')" shows "\c. Seq r c\<^sub>2 \ redexes c \ \c'. \\(c,s) \\<^sup>+ (c',s') \ Seq r' c\<^sub>2 \ redexes c'" using steps proof (induct rule: tranclp_induct2 [consumes 1, case_names Step Trans]) case (Step r' s' c') have "\\ (r, s) \ (r', s')" "Seq r c\<^sub>2 \ redexes c'" by fact+ from step_redexes_Seq [OF this] show ?case by (blast intro: r_into_trancl) next case (Trans r' s' r'' s'') from Trans obtain c' where steps: "\\ (c, s) \\<^sup>+ (c', s')" and r': "Seq r' c\<^sub>2 \ redexes c'" by blast note steps moreover have "\\ (r', s') \ (r'', s'')" by fact from step_redexes_Seq [OF this r'] obtain c'' where step: "\\ (c', s') \ (c'', s'')" and r'': "Seq r'' c\<^sub>2 \ redexes c''" by blast note step finally show ?case using r'' by blast qed lemma step_redexes_Catch: assumes step: "\\(r,s) \ (r',s')" assumes Catch: "Catch r c\<^sub>2 \ redexes c" shows "\c'. \\(c,s) \ (c',s') \ Catch r' c\<^sub>2 \ redexes c'" proof - from step.Catch [OF step] have "\\ (Catch r c\<^sub>2, s) \ (Catch r' c\<^sub>2, s')". from step_redexes [OF this Catch] show ?thesis . qed lemma steps_redexes_Catch: assumes steps: "\\ (r, s) \\<^sup>* (r', s')" shows "\c. Catch r c\<^sub>2 \ redexes c \ \c'. \\(c,s) \\<^sup>* (c',s') \ Catch r' c\<^sub>2 \ redexes c'" using steps proof (induct rule: converse_rtranclp_induct2 [case_names Refl Trans]) case Refl then show ?case by (auto) next case (Trans r s r'' s'') have "\\ (r, s) \ (r'', s'')" "Catch r c\<^sub>2 \ redexes c" by fact+ from step_redexes_Catch [OF this] obtain c' where step: "\\ (c, s) \ (c', s'')" and r'': "Catch r'' c\<^sub>2 \ redexes c'" by blast note step also from Trans.hyps (3) [OF r''] obtain c'' where steps: "\\ (c', s'') \\<^sup>* (c'', s')" and r': "Catch r' c\<^sub>2 \ redexes c''" by blast note steps finally show ?case using r' by blast qed lemma steps_redexes_Catch': assumes steps: "\\ (r, s) \\<^sup>+ (r', s')" shows "\c. Catch r c\<^sub>2 \ redexes c \ \c'. \\(c,s) \\<^sup>+ (c',s') \ Catch r' c\<^sub>2 \ redexes c'" using steps proof (induct rule: tranclp_induct2 [consumes 1, case_names Step Trans]) case (Step r' s' c') have "\\ (r, s) \ (r', s')" "Catch r c\<^sub>2 \ redexes c'" by fact+ from step_redexes_Catch [OF this] show ?case by (blast intro: r_into_trancl) next case (Trans r' s' r'' s'') from Trans obtain c' where steps: "\\ (c, s) \\<^sup>+ (c', s')" and r': "Catch r' c\<^sub>2 \ redexes c'" by blast note steps moreover have "\\ (r', s') \ (r'', s'')" by fact from step_redexes_Catch [OF this r'] obtain c'' where step: "\\ (c', s') \ (c'', s'')" and r'': "Catch r'' c\<^sub>2 \ redexes c''" by blast note step finally show ?case using r'' by blast qed lemma redexes_subset:"\c'. c' \ redexes c \ redexes c' \ redexes c" by (induct c) auto lemma redexes_preserves_termination: assumes termi: "\\c\s" shows "\c'. c' \ redexes c \ \\c'\s" using termi by induct (auto intro: terminates.intros) end diff --git a/thys/Simpl/StateSpace.thy b/thys/Simpl/StateSpace.thy --- a/thys/Simpl/StateSpace.thy +++ b/thys/Simpl/StateSpace.thy @@ -1,46 +1,38 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: StateSpace.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. +Copyright (c) 2022 Apple Inc. All rights reserved. -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA *) section \State Space Template\ theory StateSpace imports Hoare begin record 'g state = "globals"::'g definition upd_globals:: "('g \ 'g) \ ('g,'z) state_scheme \ ('g,'z) state_scheme" where "upd_globals upd s = s\globals := upd (globals s)\" +named_theorems state_simp + +lemma upd_globals_conv [state_simp]: "upd_globals f = (\s. s\globals := f (globals s)\)" + by (rule ext) (simp add: upd_globals_def) + +record ('g, 'l) state_locals = "'g state" + + locals :: 'l + +(* record ('g, 'n, 'val) stateSP = "'g state" + locals :: "'n \ 'val" +*) -lemma upd_globals_conv: "upd_globals f = (\s. s\globals := f (globals s)\)" - by (rule ext) (simp add: upd_globals_def) +type_synonym ('g, 'n, 'val) stateSP = "('g, 'n \ 'val) state_locals" +type_synonym ('g, 'n, 'val, 'x) stateSP_scheme = "('g, 'n \ 'val, 'x) state_locals_scheme" + end diff --git a/thys/Simpl/SyntaxTest.thy b/thys/Simpl/SyntaxTest.thy --- a/thys/Simpl/SyntaxTest.thy +++ b/thys/Simpl/SyntaxTest.thy @@ -1,97 +1,77 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: SyntaxTest.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2006-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. +*) -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA -*) (*<*) theory SyntaxTest imports HeapList Vcg begin record "globals" = alloc_' :: "ref list" free_':: nat GA_' :: "ref list list" next_' :: "ref \ ref" cont_' :: "ref \ nat" record 'g vars = "'g state" + A_' :: "nat list" AA_' :: "nat list list" I_' :: nat M_' :: nat N_' :: nat R_' :: int S_' :: int B_' :: bool Abr_':: string p_' :: ref q_' :: ref procedures Foo (p,I|p) = "\p :== \p" term "\I :==\<^sub>g 3 - 1" term "\R :==\<^sub>g 3 - 1" term "\I :==\<^sub>g \A!i" term " \A!i :== j" term " \AA :== \AA!![i,j]" term " \AA!![i,j] :== \AA" term "\A!i :==\<^sub>g j" term "\p :==\<^sub>g \GA!i!j" term "\GA!i!j :==\<^sub>g \p" term "\p :==\<^sub>g \p \ \next" term "\p \ \next :==\<^sub>g \p" term "\p \ \next \ \next :==\<^sub>g \p" term "\p :== NEW sz [\next :== Null,\cont :== 0]" term "\p\\next :==\<^sub>g NEW sz [\next :== Null,\cont :== 0]" term "\p :== NNEW sz [\next :== Null,\cont :== 0]" term "\p\\next :==\<^sub>g NNEW sz [\next :== Null,\cont :== 0]" term "\B :==\<^sub>g \N + 1 < 0 \ \M + \N < n" term "\B :==\<^sub>g \N + 1 < 0 \ \M + \N < n" term "\I :==\<^sub>g \N mod n" term "\I :==\<^sub>g \N div n" term "\R :==\<^sub>g \R div n" term "\R :==\<^sub>g \R mod n" term "\R :==\<^sub>g \R * n" term "\I :==\<^sub>g \I - \N" term "\R :==\<^sub>g \R - \S" term "\R :==\<^sub>g int \I" term "\I :==\<^sub>g nat \R" term "\R :==\<^sub>g -\R" term "IF\<^sub>g \A!i < \N THEN c1 ELSE c2 FI" term "WHILE\<^sub>g \A!i < \N DO c OD" term "WHILE\<^sub>g \A!i < \N INV \foo\ DO c OD" term "WHILE\<^sub>g \A!i < \N INV \foo\ VAR bar x DO c OD" term "WHILE\<^sub>g \A!i < \N INV \foo\ VAR bar x DO c OD;;c" term "c;;WHILE\<^sub>g \A!i < \N INV \foo\ VAR MEASURE \N + \M DO c;;c OD;;c" context Foo_impl begin term "\q :== CALL Foo(\p,\M)" term "\q :== CALL\<^sub>g Foo(\p,\M + 1)" term "\q :== CALL Foo(\p\\next,\M)" term "\q \ \next :== CALL Foo(\p,\M)" end end (*>*) diff --git a/thys/Simpl/Termination.thy b/thys/Simpl/Termination.thy --- a/thys/Simpl/Termination.thy +++ b/thys/Simpl/Termination.thy @@ -1,2334 +1,2407 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: Termination.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (c) 2022 Apple Inc. All rights reserved. *) section \Terminating Programs\ theory Termination imports Semantic begin subsection \Inductive Characterisation: \\\c\s\\ inductive "terminates"::"('s,'p,'f) body \ ('s,'p,'f) com \ ('s,'f) xstate \ bool" ("_\_ \ _" [60,20,60] 89) for \::"('s,'p,'f) body" where Skip: "\\Skip \(Normal s)" | Basic: "\\Basic f \(Normal s)" | Spec: "\\Spec r \(Normal s)" | Guard: "\s\g; \\c\(Normal s)\ \ \\Guard f g c\(Normal s)" | GuardFault: "s\g \ \\Guard f g c\(Normal s)" | Fault [intro,simp]: "\\c\Fault f" | Seq: "\\\c\<^sub>1\Normal s; \s'. \\\c\<^sub>1,Normal s\ \ s' \ \\c\<^sub>2\s'\ \ \\Seq c\<^sub>1 c\<^sub>2\(Normal s)" | CondTrue: "\s \ b; \\c\<^sub>1\(Normal s)\ \ \\Cond b c\<^sub>1 c\<^sub>2\(Normal s)" | CondFalse: "\s \ b; \\c\<^sub>2\(Normal s)\ \ \\Cond b c\<^sub>1 c\<^sub>2\(Normal s)" | WhileTrue: "\s \ b; \\c\(Normal s); \s'. \\\c,Normal s \ \ s' \ \\While b c\s'\ \ \\While b c\(Normal s)" | WhileFalse: "\s \ b\ \ \\While b c\(Normal s)" | Call: "\\ p=Some bdy;\\bdy\(Normal s)\ \ \\Call p\(Normal s)" | CallUndefined: "\\ p = None\ \ \\Call p\(Normal s)" | Stuck [intro,simp]: "\\c\Stuck" | DynCom: "\\\(c s)\(Normal s)\ \ \\DynCom c\(Normal s)" | Throw: "\\Throw\(Normal s)" | Abrupt [intro,simp]: "\\c\Abrupt s" | Catch: "\\\c\<^sub>1\Normal s; \s'. \\\c\<^sub>1,Normal s \ \ Abrupt s' \ \\c\<^sub>2\Normal s'\ \ \\Catch c\<^sub>1 c\<^sub>2\Normal s" inductive_cases terminates_elim_cases [cases set]: "\\Skip \ s" "\\Guard f g c \ s" "\\Basic f \ s" "\\Spec r \ s" "\\Seq c1 c2 \ s" "\\Cond b c1 c2 \ s" "\\While b c \ s" "\\Call p \ s" "\\DynCom c \ s" "\\Throw \ s" "\\Catch c1 c2 \ s" inductive_cases terminates_Normal_elim_cases [cases set]: "\\Skip \ Normal s" "\\Guard f g c \ Normal s" "\\Basic f \ Normal s" "\\Spec r \ Normal s" "\\Seq c1 c2 \ Normal s" "\\Cond b c1 c2 \ Normal s" "\\While b c \ Normal s" "\\Call p \ Normal s" "\\DynCom c \ Normal s" "\\Throw \ Normal s" "\\Catch c1 c2 \ Normal s" lemma terminates_Skip': "\\Skip \ s" by (cases s) (auto intro: terminates.intros) lemma terminates_Call_body: "\ p=Some bdy\\\Call p \s = \\(the (\ p))\s" by (cases s) (auto elim: terminates_Normal_elim_cases intro: terminates.intros) lemma terminates_Normal_Call_body: "p \ dom \ \ \\Call p \Normal s = \\(the (\ p))\Normal s" by (auto elim: terminates_Normal_elim_cases intro: terminates.intros) lemma terminates_implies_exec: assumes terminates: "\\c\s" shows "\t. \\\c,s\ \ t" using terminates proof (induct) case Skip thus ?case by (iprover intro: exec.intros) next case Basic thus ?case by (iprover intro: exec.intros) next case (Spec r s) thus ?case by (cases "\t. (s,t)\ r") (auto intro: exec.intros) next case Guard thus ?case by (iprover intro: exec.intros) next case GuardFault thus ?case by (iprover intro: exec.intros) next case Fault thus ?case by (iprover intro: exec.intros) next case Seq thus ?case by (iprover intro: exec_Seq') next case CondTrue thus ?case by (iprover intro: exec.intros) next case CondFalse thus ?case by (iprover intro: exec.intros) next case WhileTrue thus ?case by (iprover intro: exec.intros) next case WhileFalse thus ?case by (iprover intro: exec.intros) next case (Call p bdy s) then obtain s' where "\\\bdy,Normal s \ \ s'" by iprover moreover have "\ p = Some bdy" by fact ultimately show ?case by (cases s') (iprover intro: exec.intros)+ next case CallUndefined thus ?case by (iprover intro: exec.intros) next case Stuck thus ?case by (iprover intro: exec.intros) next case DynCom thus ?case by (iprover intro: exec.intros) next case Throw thus ?case by (iprover intro: exec.intros) next case Abrupt thus ?case by (iprover intro: exec.intros) next case (Catch c1 s c2) then obtain s' where exec_c1: "\\\c1,Normal s \ \ s'" by iprover thus ?case proof (cases s') case (Normal s'') with exec_c1 show ?thesis by (auto intro!: exec.intros) next case (Abrupt s'') with exec_c1 Catch.hyps obtain t where "\\\c2,Normal s'' \ \ t" by auto with exec_c1 Abrupt show ?thesis by (auto intro: exec.intros) next case Fault with exec_c1 show ?thesis by (auto intro!: exec.CatchMiss) next case Stuck with exec_c1 show ?thesis by (auto intro!: exec.CatchMiss) qed qed +lemma terminates_block_exn: +"\\\bdy \ Normal (init s); + \t. \\\bdy,Normal (init s)\ \ Normal t \ \\c s t \ Normal (return s t)\ + \ \\block_exn init bdy return result_exn c \ Normal s" +apply (unfold block_exn_def) +apply (fastforce intro: terminates.intros elim!: exec_Normal_elim_cases + dest!: not_isAbrD) + done + lemma terminates_block: "\\\bdy \ Normal (init s); \t. \\\bdy,Normal (init s)\ \ Normal t \ \\c s t \ Normal (return s t)\ \ \\block init bdy return c \ Normal s" -apply (unfold block_def) -apply (fastforce intro: terminates.intros elim!: exec_Normal_elim_cases - dest!: not_isAbrD) -done + unfolding block_def + by (rule terminates_block_exn) -lemma terminates_block_elim [cases set, consumes 1]: -assumes termi: "\\block init bdy return c \ Normal s" +lemma terminates_block_exn_elim [cases set, consumes 1]: +assumes termi: "\\block_exn init bdy return result_exn c \ Normal s" assumes e: "\\\bdy \ Normal (init s); \t. \\\bdy,Normal (init s)\ \ Normal t \ \\c s t \ Normal (return s t) \ \ P" shows P proof - have "\\\Basic init,Normal s\ \ Normal (init s)" by (auto intro: exec.intros) with termi have "\\bdy \ Normal (init s)" - apply (unfold block_def) + apply (unfold block_exn_def) apply (elim terminates_Normal_elim_cases) by simp moreover { fix t assume exec_bdy: "\\\bdy,Normal (init s)\ \ Normal t" have "\\c s t \ Normal (return s t)" proof - from exec_bdy have "\\\Catch (Seq (Basic init) bdy) - (Seq (Basic (return s)) Throw),Normal s\ \ Normal t" + (Seq (Basic (\t. result_exn (return s t) t)) Throw),Normal s\ \ Normal t" by (fastforce intro: exec.intros) with termi have "\\DynCom (\t. Seq (Basic (return s)) (c s t)) \ Normal t" - apply (unfold block_def) + apply (unfold block_exn_def) apply (elim terminates_Normal_elim_cases) by simp thus ?thesis apply (elim terminates_Normal_elim_cases) apply (auto intro: exec.intros) done qed } ultimately show P by (iprover intro: e) qed +lemma terminates_block_elim [cases set, consumes 1]: +assumes termi: "\\block init bdy return c \ Normal s" +assumes e: "\\\bdy \ Normal (init s); + \t. \\\bdy,Normal (init s)\ \ Normal t \ \\c s t \ Normal (return s t) + \ \ P" +shows P + using termi e unfolding block_def by (rule terminates_block_exn_elim) + lemma terminates_call: "\\ p = Some bdy; \\bdy \ Normal (init s); \t. \\\bdy,Normal (init s)\ \ Normal t \ \\c s t \ Normal (return s t)\ \ \\call init p return c \ Normal s" apply (unfold call_def) apply (rule terminates_block) apply (iprover intro: terminates.intros) apply (auto elim: exec_Normal_elim_cases) done +lemma terminates_call_exn: +"\\ p = Some bdy; \\bdy \ Normal (init s); + \t. \\\bdy,Normal (init s)\ \ Normal t \ \\c s t \ Normal (return s t)\ + \ \\call_exn init p return result_exn c \ Normal s" + apply (unfold call_exn_def) + apply (rule terminates_block_exn) + apply (iprover intro: terminates.intros) + apply (auto elim: exec_Normal_elim_cases) + done + lemma terminates_callUndefined: "\\ p = None\ \ \\call init p return result \ Normal s" apply (unfold call_def) apply (rule terminates_block) apply (iprover intro: terminates.intros) apply (auto elim: exec_Normal_elim_cases) done -lemma terminates_call_elim [cases set, consumes 1]: -assumes termi: "\\call init p return c \ Normal s" +lemma terminates_call_exnUndefined: +"\\ p = None\ + \ \\call_exn init p return result_exn result \ Normal s" + apply (unfold call_exn_def) + apply (rule terminates_block_exn) + apply (iprover intro: terminates.intros) + apply (auto elim: exec_Normal_elim_cases) + done + +lemma terminates_call_exn_elim [cases set, consumes 1]: +assumes termi: "\\call_exn init p return result_exn c \ Normal s" assumes bdy: "\bdy. \\ p = Some bdy; \\bdy \ Normal (init s); \t. \\\bdy,Normal (init s)\ \ Normal t \ \\c s t \ Normal (return s t)\ \ P" assumes undef: "\\ p = None\ \ P" shows P apply (cases "\ p") apply (erule undef) using termi -apply (unfold call_def) -apply (erule terminates_block_elim) +apply (unfold call_exn_def) +apply (erule terminates_block_exn_elim) apply (erule terminates_Normal_elim_cases) apply simp apply (frule (1) bdy) apply (fastforce intro: exec.intros) apply assumption apply simp done +lemma terminates_call_elim [cases set, consumes 1]: +assumes termi: "\\call init p return c \ Normal s" +assumes bdy: "\bdy. \\ p = Some bdy; \\bdy \ Normal (init s); + \t. \\\bdy,Normal (init s)\ \ Normal t \ \\c s t \ Normal (return s t)\ \ P" +assumes undef: "\\ p = None\ \ P" +shows P + using termi bdy undef unfolding call_call_exn by (rule terminates_call_exn_elim) + + lemma terminates_dynCall: "\\\call init (p s) return c \ Normal s\ \ \\dynCall init p return c \ Normal s" apply (unfold dynCall_def) apply (auto intro: terminates.intros terminates_call) done +lemma terminates_guards: "\\c \ Normal s \ \\guards gs c \ Normal s" + by (induct gs) (auto intro: terminates.intros) + +lemma terminates_guards_Fault: "find (\(f, g). s \ g) gs = Some (f, g) \ \\guards gs c \ Normal s" + by (induct gs) (auto intro: terminates.intros split: if_split_asm prod.splits) + +lemma terminates_maybe_guard_Fault: "s \ g \ \\maybe_guard f g c \ Normal s" + by (metis UNIV_I maybe_guard_def terminates.GuardFault) + +lemma terminates_guards_DynCom: "\\(c s) \ Normal s \ \\guards gs (DynCom c) \ Normal s" + by (induct gs) (auto intro: terminates.intros) + +lemma terminates_maybe_guard_DynCom: "\\(c s) \ Normal s \ \\maybe_guard f g (DynCom c) \ Normal s" + by (metis maybe_guard_def terminates.DynCom terminates.Guard terminates.GuardFault) + + +lemma terminates_dynCall_exn: +"\\\call_exn init (p s) return result_exn c \ Normal s\ + \ \\dynCall_exn f g init p return result_exn c \ Normal s" + apply (unfold dynCall_exn_def) + by (rule terminates_maybe_guard_DynCom) + lemma terminates_dynCall_elim [cases set, consumes 1]: assumes termi: "\\dynCall init p return c \ Normal s" assumes "\\\call init (p s) return c \ Normal s\ \ P" shows P using termi apply (unfold dynCall_def) apply (elim terminates_Normal_elim_cases) apply fact done +lemma terminates_guards_elim [cases set, consumes 1, case_names noFault someFault]: + assumes termi: "\\guards gs c \ Normal s" + assumes noFault: "\\f g. (f, g) \ set gs \ s \ g; \\c \ Normal s\ \ P" + assumes someFault: "\f g. find (\(f,g). s \ g) gs = Some (f, g) \ P" + shows P + using termi noFault someFault + by (induct gs) + (auto elim: terminates_Normal_elim_cases split: if_split_asm prod.splits) + +lemma terminates_maybe_guard_elim [cases set, consumes 1, case_names noFault someFault]: + assumes termi: "\\maybe_guard f g c \ Normal s" + assumes noFault: "\s \ g; \\c \ Normal s\ \ P" + assumes someFault: "s \ g \ P" + shows P + using termi noFault someFault + by (metis maybe_guard_def terminates_Normal_elim_cases(2)) + +lemma terminates_dynCall_exn_elim [cases set, consumes 1, case_names noFault someFault]: +assumes termi: "\\dynCall_exn f g init p return result_exn c \ Normal s" +assumes noFault: "\s \ g; + \\call_exn init (p s) return result_exn c \ Normal s\ \ P" +assumes someFault: "s \ g \ P" +shows P +using termi noFault someFault + apply (unfold dynCall_exn_def) + apply (erule terminates_maybe_guard_elim) + apply (auto elim: terminates_Normal_elim_cases) + done (* ************************************************************************* *) subsection \Lemmas about @{const "sequence"}, @{const "flatten"} and @{const "normalize"}\ (* ************************************************************************ *) lemma terminates_sequence_app: "\s. \\\sequence Seq xs \ Normal s; \s'. \\\sequence Seq xs,Normal s \ \ s' \ \\sequence Seq ys \ s'\ \ \\sequence Seq (xs @ ys) \ Normal s" proof (induct xs) case Nil thus ?case by (auto intro: exec.intros) next case (Cons x xs) have termi_x_xs: "\\sequence Seq (x # xs) \ Normal s" by fact have termi_ys: "\s'. \\\sequence Seq (x # xs),Normal s \ \ s' \ \\sequence Seq ys \ s'" by fact show ?case proof (cases xs) case Nil with termi_x_xs termi_ys show ?thesis by (cases ys) (auto intro: terminates.intros) next case Cons from termi_x_xs Cons have "\\x \ Normal s" by (auto elim: terminates_Normal_elim_cases) moreover { fix s' assume exec_x: "\\\x,Normal s \ \ s'" have "\\sequence Seq (xs @ ys) \ s'" proof - from exec_x termi_x_xs Cons have termi_xs: "\\sequence Seq xs \ s'" by (auto elim: terminates_Normal_elim_cases) show ?thesis proof (cases s') case (Normal s'') with exec_x termi_ys Cons have "\s'. \\\sequence Seq xs,Normal s'' \ \ s' \ \\sequence Seq ys \ s'" by (auto intro: exec.intros) from Cons.hyps [OF termi_xs [simplified Normal] this] have "\\sequence Seq (xs @ ys) \ Normal s''". with Normal show ?thesis by simp next case Abrupt thus ?thesis by (auto intro: terminates.intros) next case Fault thus ?thesis by (auto intro: terminates.intros) next case Stuck thus ?thesis by (auto intro: terminates.intros) qed qed } ultimately show ?thesis using Cons by (auto intro: terminates.intros) qed qed lemma terminates_sequence_appD: "\s. \\sequence Seq (xs @ ys) \ Normal s \ \\sequence Seq xs \ Normal s \ (\s'. \\\sequence Seq xs,Normal s \ \ s' \ \\sequence Seq ys \ s')" proof (induct xs) case Nil thus ?case by (auto elim: terminates_Normal_elim_cases exec_Normal_elim_cases intro: terminates.intros) next case (Cons x xs) have termi_x_xs_ys: "\\sequence Seq ((x # xs) @ ys) \ Normal s" by fact show ?case proof (cases xs) case Nil with termi_x_xs_ys show ?thesis by (cases ys) (auto elim: terminates_Normal_elim_cases exec_Normal_elim_cases intro: terminates_Skip') next case Cons with termi_x_xs_ys obtain termi_x: "\\x \ Normal s" and termi_xs_ys: "\s'. \\\x,Normal s \ \ s' \ \\sequence Seq (xs@ys) \ s'" by (auto elim: terminates_Normal_elim_cases) have "\\Seq x (sequence Seq xs) \ Normal s" proof (rule terminates.Seq [rule_format]) show "\\x \ Normal s" by (rule termi_x) next fix s' assume exec_x: "\\\x,Normal s \ \ s'" show "\\sequence Seq xs \ s'" proof - from termi_xs_ys [rule_format, OF exec_x] have termi_xs_ys': "\\sequence Seq (xs@ys) \ s'" . show ?thesis proof (cases s') case (Normal s'') from Cons.hyps [OF termi_xs_ys' [simplified Normal]] show ?thesis using Normal by auto next case Abrupt thus ?thesis by (auto intro: terminates.intros) next case Fault thus ?thesis by (auto intro: terminates.intros) next case Stuck thus ?thesis by (auto intro: terminates.intros) qed qed qed moreover { fix s' assume exec_x_xs: "\\\Seq x (sequence Seq xs),Normal s \ \ s'" have "\\sequence Seq ys \ s'" proof - from exec_x_xs obtain t where exec_x: "\\\x,Normal s \ \ t" and exec_xs: "\\\sequence Seq xs,t \ \ s'" by cases show ?thesis proof (cases t) case (Normal t') with exec_x termi_xs_ys have "\\sequence Seq (xs@ys) \ Normal t'" by auto from Cons.hyps [OF this] exec_xs Normal show ?thesis by auto next case (Abrupt t') with exec_xs have "s'=Abrupt t'" by (auto dest: Abrupt_end) thus ?thesis by (auto intro: terminates.intros) next case (Fault f) with exec_xs have "s'=Fault f" by (auto dest: Fault_end) thus ?thesis by (auto intro: terminates.intros) next case Stuck with exec_xs have "s'=Stuck" by (auto dest: Stuck_end) thus ?thesis by (auto intro: terminates.intros) qed qed } ultimately show ?thesis using Cons by auto qed qed lemma terminates_sequence_appE [consumes 1]: "\\\sequence Seq (xs @ ys) \ Normal s; \\\sequence Seq xs \ Normal s; \s'. \\\sequence Seq xs,Normal s \ \ s' \ \\sequence Seq ys \ s'\ \ P\ \ P" by (auto dest: terminates_sequence_appD) lemma terminates_to_terminates_sequence_flatten: assumes termi: "\\c\s" shows "\\sequence Seq (flatten c)\s" using termi by (induct) (auto intro: terminates.intros terminates_sequence_app exec_sequence_flatten_to_exec) lemma terminates_to_terminates_normalize: assumes termi: "\\c\s" shows "\\normalize c\s" using termi proof induct case Seq thus ?case by (fastforce intro: terminates.intros terminates_sequence_app terminates_to_terminates_sequence_flatten dest: exec_sequence_flatten_to_exec exec_normalize_to_exec) next case WhileTrue thus ?case by (fastforce intro: terminates.intros terminates_sequence_app terminates_to_terminates_sequence_flatten dest: exec_sequence_flatten_to_exec exec_normalize_to_exec) next case Catch thus ?case by (fastforce intro: terminates.intros terminates_sequence_app terminates_to_terminates_sequence_flatten dest: exec_sequence_flatten_to_exec exec_normalize_to_exec) qed (auto intro: terminates.intros) lemma terminates_sequence_flatten_to_terminates: shows "\s. \\sequence Seq (flatten c)\s \ \\c\s" proof (induct c) case (Seq c1 c2) have "\\sequence Seq (flatten (Seq c1 c2)) \ s" by fact hence termi_app: "\\sequence Seq (flatten c1 @ flatten c2) \ s" by simp show ?case proof (cases s) case (Normal s') have "\\Seq c1 c2 \ Normal s'" proof (rule terminates.Seq [rule_format]) from termi_app [simplified Normal] have "\\sequence Seq (flatten c1) \ Normal s'" by (cases rule: terminates_sequence_appE) with Seq.hyps show "\\c1 \ Normal s'" by simp next fix s'' assume "\\\c1,Normal s' \ \ s''" from termi_app [simplified Normal] exec_to_exec_sequence_flatten [OF this] have "\\sequence Seq (flatten c2) \ s''" by (cases rule: terminates_sequence_appE) auto with Seq.hyps show "\\c2 \ s''" by simp qed with Normal show ?thesis by simp qed (auto intro: terminates.intros) qed (auto intro: terminates.intros) lemma terminates_normalize_to_terminates: shows "\s. \\normalize c\s \ \\c\s" proof (induct c) case Skip thus ?case by (auto intro: terminates_Skip') next case Basic thus ?case by (cases s) (auto intro: terminates.intros) next case Spec thus ?case by (cases s) (auto intro: terminates.intros) next case (Seq c1 c2) have "\\normalize (Seq c1 c2) \ s" by fact hence termi_app: "\\sequence Seq (flatten (normalize c1) @ flatten (normalize c2)) \ s" by simp show ?case proof (cases s) case (Normal s') have "\\Seq c1 c2 \ Normal s'" proof (rule terminates.Seq [rule_format]) from termi_app [simplified Normal] have "\\sequence Seq (flatten (normalize c1)) \ Normal s'" by (cases rule: terminates_sequence_appE) from terminates_sequence_flatten_to_terminates [OF this] Seq.hyps show "\\c1 \ Normal s'" by simp next fix s'' assume "\\\c1,Normal s' \ \ s''" from exec_to_exec_normalize [OF this] have "\\\normalize c1,Normal s' \ \ s''" . from termi_app [simplified Normal] exec_to_exec_sequence_flatten [OF this] have "\\sequence Seq (flatten (normalize c2)) \ s''" by (cases rule: terminates_sequence_appE) auto from terminates_sequence_flatten_to_terminates [OF this] Seq.hyps show "\\c2 \ s''" by simp qed with Normal show ?thesis by simp qed (auto intro: terminates.intros) next case (Cond b c1 c2) thus ?case by (cases s) (auto intro: terminates.intros elim!: terminates_Normal_elim_cases) next case (While b c) have "\\normalize (While b c) \ s" by fact hence termi_norm_w: "\\While b (normalize c) \ s" by simp { fix t w assume termi_w: "\\ w \ t" have "w=While b (normalize c) \ \\While b c \ t" using termi_w proof (induct) case (WhileTrue t' b' c') from WhileTrue obtain t'_b: "t' \ b" and termi_norm_c: "\\normalize c \ Normal t'" and termi_norm_w': "\s'. \\\normalize c,Normal t' \ \ s' \ \\While b c \ s'" by auto from While.hyps [OF termi_norm_c] have "\\c \ Normal t'". moreover from termi_norm_w' have "\s'. \\\c,Normal t' \ \ s' \ \\While b c \ s'" by (auto intro: exec_to_exec_normalize) ultimately show ?case using t'_b by (auto intro: terminates.intros) qed (auto intro: terminates.intros) } from this [OF termi_norm_w] show ?case by auto next case Call thus ?case by simp next case DynCom thus ?case by (cases s) (auto intro: terminates.intros rangeI elim: terminates_Normal_elim_cases) next case Guard thus ?case by (cases s) (auto intro: terminates.intros elim: terminates_Normal_elim_cases) next case Throw thus ?case by (cases s) (auto intro: terminates.intros) next case Catch thus ?case by (cases s) (auto dest: exec_to_exec_normalize elim!: terminates_Normal_elim_cases intro!: terminates.Catch) qed lemma terminates_iff_terminates_normalize: "\\normalize c\s = \\c\s" by (auto intro: terminates_to_terminates_normalize terminates_normalize_to_terminates) (* ************************************************************************* *) subsection \Lemmas about @{const "strip_guards"}\ (* ************************************************************************* *) lemma terminates_strip_guards_to_terminates: "\s. \\strip_guards F c\s \ \\c\s" proof (induct c) case Skip thus ?case by simp next case Basic thus ?case by simp next case Spec thus ?case by simp next case (Seq c1 c2) hence "\\Seq (strip_guards F c1) (strip_guards F c2) \ s" by simp thus "\\Seq c1 c2 \ s" proof (cases) fix f assume "s=Fault f" thus ?thesis by simp next assume "s=Stuck" thus ?thesis by simp next fix s' assume "s=Abrupt s'" thus ?thesis by simp next fix s' assume s: "s=Normal s'" assume "\\strip_guards F c1 \ Normal s'" hence "\\c1 \ Normal s'" by (rule Seq.hyps) moreover assume c2: "\s''. \\\strip_guards F c1,Normal s'\ \ s'' \ \\strip_guards F c2\s''" { fix s'' assume exec_c1: "\\\c1,Normal s' \ \ s''" have " \\c2 \ s''" proof (cases s'') case (Normal s''') with exec_c1 have "\\\strip_guards F c1,Normal s' \ \ s''" by (auto intro: exec_to_exec_strip_guards) with c2 show ?thesis by (iprover intro: Seq.hyps) next case (Abrupt s''') with exec_c1 have "\\\strip_guards F c1,Normal s' \ \ s''" by (auto intro: exec_to_exec_strip_guards ) with c2 show ?thesis by (iprover intro: Seq.hyps) next case Fault thus ?thesis by simp next case Stuck thus ?thesis by simp qed } ultimately show ?thesis using s by (iprover intro: terminates.intros) qed next case (Cond b c1 c2) hence "\\Cond b (strip_guards F c1) (strip_guards F c2) \ s" by simp thus "\\Cond b c1 c2 \ s" proof (cases) fix f assume "s=Fault f" thus ?thesis by simp next assume "s=Stuck" thus ?thesis by simp next fix s' assume "s=Abrupt s'" thus ?thesis by simp next fix s' assume "s'\b" "\\strip_guards F c1 \ Normal s'" "s = Normal s'" thus ?thesis by (iprover intro: terminates.intros Cond.hyps) next fix s' assume "s'\b" "\\strip_guards F c2 \ Normal s'" "s = Normal s'" thus ?thesis by (iprover intro: terminates.intros Cond.hyps) qed next case (While b c) have hyp_c: "\s. \\strip_guards F c \ s \ \\c \ s" by fact have "\\While b (strip_guards F c) \ s" using While.prems by simp moreover { fix sw assume "\\sw\s" then have "sw=While b (strip_guards F c) \ \\While b c \ s" proof (induct) case (WhileTrue s b' c') have eqs: "While b' c' = While b (strip_guards F c)" by fact with \s\b'\ have b: "s\b" by simp from eqs \\\c' \ Normal s\ have "\\strip_guards F c \ Normal s" by simp hence term_c: "\\c \ Normal s" by (rule hyp_c) moreover { fix t assume exec_c: "\\\c,Normal s \ \ t" have "\\While b c \ t" proof (cases t) case Fault thus ?thesis by simp next case Stuck thus ?thesis by simp next case (Abrupt t') thus ?thesis by simp next case (Normal t') with exec_c have "\\\strip_guards F c,Normal s \ \ Normal t'" by (auto intro: exec_to_exec_strip_guards) with WhileTrue.hyps eqs Normal show ?thesis by fastforce qed } ultimately show ?case using b by (auto intro: terminates.intros) next case WhileFalse thus ?case by (auto intro: terminates.intros) qed simp_all } ultimately show "\\While b c \ s" by auto next case Call thus ?case by simp next case DynCom thus ?case by (cases s) (auto elim: terminates_Normal_elim_cases intro: terminates.intros rangeI) next case Guard thus ?case by (cases s) (auto elim: terminates_Normal_elim_cases intro: terminates.intros split: if_split_asm) next case Throw thus ?case by simp next case (Catch c1 c2) hence "\\Catch (strip_guards F c1) (strip_guards F c2) \ s" by simp thus "\\Catch c1 c2 \ s" proof (cases) fix f assume "s=Fault f" thus ?thesis by simp next assume "s=Stuck" thus ?thesis by simp next fix s' assume "s=Abrupt s'" thus ?thesis by simp next fix s' assume s: "s=Normal s'" assume "\\strip_guards F c1 \ Normal s'" hence "\\c1 \ Normal s'" by (rule Catch.hyps) moreover assume c2: "\s''. \\\strip_guards F c1,Normal s'\ \ Abrupt s'' \ \\strip_guards F c2\Normal s''" { fix s'' assume exec_c1: "\\\c1,Normal s' \ \ Abrupt s''" have " \\c2 \ Normal s''" proof - from exec_c1 have "\\\strip_guards F c1,Normal s' \ \ Abrupt s''" by (auto intro: exec_to_exec_strip_guards) with c2 show ?thesis by (auto intro: Catch.hyps) qed } ultimately show ?thesis using s by (iprover intro: terminates.intros) qed qed lemma terminates_strip_to_terminates: assumes termi_strip: "strip F \\c\s" shows "\\c\s" using termi_strip proof induct case (Seq c1 s c2) have "\\c1 \ Normal s" by fact moreover { fix s' assume exec: "\\ \c1,Normal s\ \ s'" have "\\c2 \ s'" proof (cases "isFault s'") case True thus ?thesis by (auto elim: isFaultE) next case False from exec_to_exec_strip [OF exec this] Seq.hyps show ?thesis by auto qed } ultimately show ?case by (auto intro: terminates.intros) next case (WhileTrue s b c) have "\\c \ Normal s" by fact moreover { fix s' assume exec: "\\ \c,Normal s\ \ s'" have "\\While b c \ s'" proof (cases "isFault s'") case True thus ?thesis by (auto elim: isFaultE) next case False from exec_to_exec_strip [OF exec this] WhileTrue.hyps show ?thesis by auto qed } ultimately show ?case by (auto intro: terminates.intros) next case (Catch c1 s c2) have "\\c1 \ Normal s" by fact moreover { fix s' assume exec: "\\ \c1,Normal s\ \ Abrupt s'" from exec_to_exec_strip [OF exec] Catch.hyps have "\\c2 \ Normal s'" by auto } ultimately show ?case by (auto intro: terminates.intros) next case Call thus ?case by (auto intro: terminates.intros terminates_strip_guards_to_terminates) qed (auto intro: terminates.intros) (* ************************************************************************* *) subsection \Lemmas about @{term "c\<^sub>1 \\<^sub>g c\<^sub>2"}\ (* ************************************************************************* *) lemma inter_guards_terminates: "\c c2 s. \(c1 \\<^sub>g c2) = Some c; \\c1\s \ \ \\c\s" proof (induct c1) case Skip thus ?case by (fastforce simp add: inter_guards_Skip) next case (Basic f) thus ?case by (fastforce simp add: inter_guards_Basic) next case (Spec r) thus ?case by (fastforce simp add: inter_guards_Spec) next case (Seq a1 a2) have "(Seq a1 a2 \\<^sub>g c2) = Some c" by fact then obtain b1 b2 d1 d2 where c2: "c2=Seq b1 b2" and d1: "(a1 \\<^sub>g b1) = Some d1" and d2: "(a2 \\<^sub>g b2) = Some d2" and c: "c=Seq d1 d2" by (auto simp add: inter_guards_Seq) have termi_c1: "\\Seq a1 a2 \ s" by fact have "\\Seq d1 d2 \ s" proof (cases s) case Fault thus ?thesis by simp next case Stuck thus ?thesis by simp next case Abrupt thus ?thesis by simp next case (Normal s') note Normal_s = this with d1 termi_c1 have "\\d1 \ Normal s'" by (auto elim: terminates_Normal_elim_cases intro: Seq.hyps) moreover { fix t assume exec_d1: "\\\d1,Normal s' \ \ t" have "\\d2 \ t" proof (cases t) case Fault thus ?thesis by simp next case Stuck thus ?thesis by simp next case Abrupt thus ?thesis by simp next case (Normal t') with inter_guards_exec_noFault [OF d1 exec_d1] have "\\\a1,Normal s' \ \ Normal t'" by simp with termi_c1 Normal_s have "\\a2 \ Normal t'" by (auto elim: terminates_Normal_elim_cases) with d2 have "\\d2 \ Normal t'" by (auto intro: Seq.hyps) with Normal show ?thesis by simp qed } ultimately have "\\Seq d1 d2 \ Normal s'" by (fastforce intro: terminates.intros) with Normal show ?thesis by simp qed with c show ?case by simp next case Cond thus ?case by - (cases s, auto intro: terminates.intros elim!: terminates_Normal_elim_cases simp add: inter_guards_Cond) next case (While b bdy1) have "(While b bdy1 \\<^sub>g c2) = Some c" by fact then obtain bdy2 bdy where c2: "c2=While b bdy2" and bdy: "(bdy1 \\<^sub>g bdy2) = Some bdy" and c: "c=While b bdy" by (auto simp add: inter_guards_While) have "\\While b bdy1 \ s" by fact moreover { fix s w w1 w2 assume termi_w: "\\w \ s" assume w: "w=While b bdy1" from termi_w w have "\\While b bdy \ s" proof (induct) case (WhileTrue s b' bdy1') have eqs: "While b' bdy1' = While b bdy1" by fact from WhileTrue have s_in_b: "s \ b" by simp from WhileTrue have termi_bdy1: "\\bdy1 \ Normal s" by simp show ?case proof - from bdy termi_bdy1 have "\\bdy\(Normal s)" by (rule While.hyps) moreover { fix t assume exec_bdy: "\\\bdy,Normal s \ \ t" have "\\While b bdy\t" proof (cases t) case Fault thus ?thesis by simp next case Stuck thus ?thesis by simp next case Abrupt thus ?thesis by simp next case (Normal t') with inter_guards_exec_noFault [OF bdy exec_bdy] have "\\\bdy1,Normal s \ \ Normal t'" by simp with WhileTrue have "\\While b bdy \ Normal t'" by simp with Normal show ?thesis by simp qed } ultimately show ?thesis using s_in_b by (blast intro: terminates.WhileTrue) qed next case WhileFalse thus ?case by (blast intro: terminates.WhileFalse) qed (simp_all) } ultimately show ?case using c by simp next case Call thus ?case by (simp add: inter_guards_Call) next case (DynCom f1) have "(DynCom f1 \\<^sub>g c2) = Some c" by fact then obtain f2 f where c2: "c2=DynCom f2" and f_defined: "\s. ((f1 s) \\<^sub>g (f2 s)) \ None" and c: "c=DynCom (\s. the ((f1 s) \\<^sub>g (f2 s)))" by (auto simp add: inter_guards_DynCom) have termi: "\\DynCom f1 \ s" by fact show ?case proof (cases s) case Fault thus ?thesis by simp next case Stuck thus ?thesis by simp next case Abrupt thus ?thesis by simp next case (Normal s') from f_defined obtain f where f: "((f1 s') \\<^sub>g (f2 s')) = Some f" by auto from Normal termi have "\\f1 s'\ (Normal s')" by (auto elim: terminates_Normal_elim_cases) from DynCom.hyps f this have "\\f\ (Normal s')" by blast with c f Normal show ?thesis by (auto intro: terminates.intros) qed next case (Guard f g1 bdy1) have "(Guard f g1 bdy1 \\<^sub>g c2) = Some c" by fact then obtain g2 bdy2 bdy where c2: "c2=Guard f g2 bdy2" and bdy: "(bdy1 \\<^sub>g bdy2) = Some bdy" and c: "c=Guard f (g1 \ g2) bdy" by (auto simp add: inter_guards_Guard) have termi_c1: "\\Guard f g1 bdy1 \ s" by fact show ?case proof (cases s) case Fault thus ?thesis by simp next case Stuck thus ?thesis by simp next case Abrupt thus ?thesis by simp next case (Normal s') show ?thesis proof (cases "s' \ g1") case False with Normal c show ?thesis by (auto intro: terminates.GuardFault) next case True note s_in_g1 = this show ?thesis proof (cases "s' \ g2") case False with Normal c show ?thesis by (auto intro: terminates.GuardFault) next case True with termi_c1 s_in_g1 Normal have "\\bdy1 \ Normal s'" by (auto elim: terminates_Normal_elim_cases) with c bdy Guard.hyps Normal True s_in_g1 show ?thesis by (auto intro: terminates.Guard) qed qed qed next case Throw thus ?case by (auto simp add: inter_guards_Throw) next case (Catch a1 a2) have "(Catch a1 a2 \\<^sub>g c2) = Some c" by fact then obtain b1 b2 d1 d2 where c2: "c2=Catch b1 b2" and d1: "(a1 \\<^sub>g b1) = Some d1" and d2: "(a2 \\<^sub>g b2) = Some d2" and c: "c=Catch d1 d2" by (auto simp add: inter_guards_Catch) have termi_c1: "\\Catch a1 a2 \ s" by fact have "\\Catch d1 d2 \ s" proof (cases s) case Fault thus ?thesis by simp next case Stuck thus ?thesis by simp next case Abrupt thus ?thesis by simp next case (Normal s') note Normal_s = this with d1 termi_c1 have "\\d1 \ Normal s'" by (auto elim: terminates_Normal_elim_cases intro: Catch.hyps) moreover { fix t assume exec_d1: "\\\d1,Normal s' \ \ Abrupt t" have "\\d2 \ Normal t" proof - from inter_guards_exec_noFault [OF d1 exec_d1] have "\\\a1,Normal s' \ \ Abrupt t" by simp with termi_c1 Normal_s have "\\a2 \ Normal t" by (auto elim: terminates_Normal_elim_cases) with d2 have "\\d2 \ Normal t" by (auto intro: Catch.hyps) with Normal show ?thesis by simp qed } ultimately have "\\Catch d1 d2 \ Normal s'" by (fastforce intro: terminates.intros) with Normal show ?thesis by simp qed with c show ?case by simp qed lemma inter_guards_terminates': assumes c: "(c1 \\<^sub>g c2) = Some c" assumes termi_c2: "\\c2\s" shows "\\c\s" proof - from c have "(c2 \\<^sub>g c1) = Some c" by (rule inter_guards_sym) from this termi_c2 show ?thesis by (rule inter_guards_terminates) qed (* ************************************************************************* *) subsection \Lemmas about @{const "mark_guards"}\ (* ************************************************************************ *) lemma terminates_to_terminates_mark_guards: assumes termi: "\\c\s" shows "\\mark_guards f c\s" using termi proof (induct) case Skip thus ?case by (fastforce intro: terminates.intros) next case Basic thus ?case by (fastforce intro: terminates.intros) next case Spec thus ?case by (fastforce intro: terminates.intros) next case Guard thus ?case by (fastforce intro: terminates.intros) next case GuardFault thus ?case by (fastforce intro: terminates.intros) next case Fault thus ?case by (fastforce intro: terminates.intros) next case (Seq c1 s c2) have "\\mark_guards f c1 \ Normal s" by fact moreover { fix t assume exec_mark: "\\\mark_guards f c1,Normal s \ \ t" have "\\mark_guards f c2 \ t" proof - from exec_mark_guards_to_exec [OF exec_mark] obtain t' where exec_c1: "\\\c1,Normal s \ \ t'" and t_Fault: "isFault t \ isFault t'" and t'_Fault_f: "t' = Fault f \ t' = t" and t'_Fault: "isFault t' \ isFault t" and t'_noFault: "\ isFault t' \ t' = t" by blast show ?thesis proof (cases "isFault t'") case True with t'_Fault have "isFault t" by simp thus ?thesis by (auto elim: isFaultE) next case False with t'_noFault have "t'=t" by simp with exec_c1 Seq.hyps show ?thesis by auto qed qed } ultimately show ?case by (auto intro: terminates.intros) next case CondTrue thus ?case by (fastforce intro: terminates.intros) next case CondFalse thus ?case by (fastforce intro: terminates.intros) next case (WhileTrue s b c) have s_in_b: "s \ b" by fact have "\\mark_guards f c \ Normal s" by fact moreover { fix t assume exec_mark: "\\\mark_guards f c,Normal s \ \ t" have "\\mark_guards f (While b c) \ t" proof - from exec_mark_guards_to_exec [OF exec_mark] obtain t' where exec_c1: "\\\c,Normal s \ \ t'" and t_Fault: "isFault t \ isFault t'" and t'_Fault_f: "t' = Fault f \ t' = t" and t'_Fault: "isFault t' \ isFault t" and t'_noFault: "\ isFault t' \ t' = t" by blast show ?thesis proof (cases "isFault t'") case True with t'_Fault have "isFault t" by simp thus ?thesis by (auto elim: isFaultE) next case False with t'_noFault have "t'=t" by simp with exec_c1 WhileTrue.hyps show ?thesis by auto qed qed } ultimately show ?case by (auto intro: terminates.intros) next case WhileFalse thus ?case by (fastforce intro: terminates.intros) next case Call thus ?case by (fastforce intro: terminates.intros) next case CallUndefined thus ?case by (fastforce intro: terminates.intros) next case Stuck thus ?case by (fastforce intro: terminates.intros) next case DynCom thus ?case by (fastforce intro: terminates.intros) next case Throw thus ?case by (fastforce intro: terminates.intros) next case Abrupt thus ?case by (fastforce intro: terminates.intros) next case (Catch c1 s c2) have "\\mark_guards f c1 \ Normal s" by fact moreover { fix t assume exec_mark: "\\\mark_guards f c1,Normal s \ \ Abrupt t" have "\\mark_guards f c2 \ Normal t" proof - from exec_mark_guards_to_exec [OF exec_mark] obtain t' where exec_c1: "\\\c1,Normal s \ \ t'" and t'_Fault_f: "t' = Fault f \ t' = Abrupt t" and t'_Fault: "isFault t' \ isFault (Abrupt t)" and t'_noFault: "\ isFault t' \ t' = Abrupt t" by fastforce show ?thesis proof (cases "isFault t'") case True with t'_Fault have "isFault (Abrupt t)" by simp thus ?thesis by simp next case False with t'_noFault have "t'=Abrupt t" by simp with exec_c1 Catch.hyps show ?thesis by auto qed qed } ultimately show ?case by (auto intro: terminates.intros) qed lemma terminates_mark_guards_to_terminates_Normal: "\s. \\mark_guards f c\Normal s \ \\c\Normal s" proof (induct c) case Skip thus ?case by (fastforce intro: terminates.intros) next case Basic thus ?case by (fastforce intro: terminates.intros) next case Spec thus ?case by (fastforce intro: terminates.intros) next case (Seq c1 c2) have "\\mark_guards f (Seq c1 c2) \ Normal s" by fact then obtain termi_merge_c1: "\\mark_guards f c1 \ Normal s" and termi_merge_c2: "\s'. \\\mark_guards f c1,Normal s \ \ s' \ \\mark_guards f c2 \ s'" by (auto elim: terminates_Normal_elim_cases) from termi_merge_c1 Seq.hyps have "\\c1 \ Normal s" by iprover moreover { fix s' assume exec_c1: "\\\c1,Normal s \ \ s'" have "\\ c2 \ s'" proof (cases "isFault s'") case True thus ?thesis by (auto elim: isFaultE) next case False from exec_to_exec_mark_guards [OF exec_c1 False] have "\\\mark_guards f c1,Normal s \ \ s'" . from termi_merge_c2 [rule_format, OF this] Seq.hyps show ?thesis by (cases s') (auto) qed } ultimately show ?case by (auto intro: terminates.intros) next case Cond thus ?case by (fastforce intro: terminates.intros elim: terminates_Normal_elim_cases) next case (While b c) { fix u c' assume termi_c': "\\c' \ Normal u" assume c': "c' = mark_guards f (While b c)" have "\\While b c \ Normal u" using termi_c' c' proof (induct) case (WhileTrue s b' c') have s_in_b: "s \ b" using WhileTrue by simp have "\\mark_guards f c \ Normal s" using WhileTrue by (auto elim: terminates_Normal_elim_cases) with While.hyps have "\\c \ Normal s" by auto moreover have hyp_w: "\w. \\\mark_guards f c,Normal s \ \ w \ \\While b c \ w" using WhileTrue by simp hence "\w. \\\c,Normal s \ \ w \ \\While b c \ w" apply - apply (rule allI) apply (case_tac "w") apply (auto dest: exec_to_exec_mark_guards) done ultimately show ?case using s_in_b by (auto intro: terminates.intros) next case WhileFalse thus ?case by (auto intro: terminates.intros) qed auto } with While show ?case by simp next case Call thus ?case by (fastforce intro: terminates.intros ) next case DynCom thus ?case by (fastforce intro: terminates.intros elim: terminates_Normal_elim_cases) next case (Guard f g c) thus ?case by (fastforce intro: terminates.intros elim: terminates_Normal_elim_cases) next case Throw thus ?case by (fastforce intro: terminates.intros ) next case (Catch c1 c2) have "\\mark_guards f (Catch c1 c2) \ Normal s" by fact then obtain termi_merge_c1: "\\mark_guards f c1 \ Normal s" and termi_merge_c2: "\s'. \\\mark_guards f c1,Normal s \ \ Abrupt s' \ \\mark_guards f c2 \ Normal s'" by (auto elim: terminates_Normal_elim_cases) from termi_merge_c1 Catch.hyps have "\\c1 \ Normal s" by iprover moreover { fix s' assume exec_c1: "\\\c1,Normal s \ \ Abrupt s'" have "\\ c2 \ Normal s'" proof - from exec_to_exec_mark_guards [OF exec_c1] have "\\\mark_guards f c1,Normal s \ \ Abrupt s'" by simp from termi_merge_c2 [rule_format, OF this] Catch.hyps show ?thesis by iprover qed } ultimately show ?case by (auto intro: terminates.intros) qed lemma terminates_mark_guards_to_terminates: "\\mark_guards f c\s \ \\c\ s" by (cases s) (auto intro: terminates_mark_guards_to_terminates_Normal) (* ************************************************************************* *) subsection \Lemmas about @{const "merge_guards"}\ (* ************************************************************************ *) lemma terminates_to_terminates_merge_guards: assumes termi: "\\c\s" shows "\\merge_guards c\s" using termi proof (induct) case (Guard s g c f) have s_in_g: "s \ g" by fact have termi_merge_c: "\\merge_guards c \ Normal s" by fact show ?case proof (cases "\f' g' c'. merge_guards c = Guard f' g' c'") case False hence "merge_guards (Guard f g c) = Guard f g (merge_guards c)" by (cases "merge_guards c") (auto simp add: Let_def) with s_in_g termi_merge_c show ?thesis by (auto intro: terminates.intros) next case True then obtain f' g' c' where mc: "merge_guards c = Guard f' g' c'" by blast show ?thesis proof (cases "f=f'") case False with mc have "merge_guards (Guard f g c) = Guard f g (merge_guards c)" by (simp add: Let_def) with s_in_g termi_merge_c show ?thesis by (auto intro: terminates.intros) next case True with mc have "merge_guards (Guard f g c) = Guard f (g \ g') c'" by simp with s_in_g mc True termi_merge_c show ?thesis by (cases "s \ g'") (auto intro: terminates.intros elim: terminates_Normal_elim_cases) qed qed next case (GuardFault s g f c) have "s \ g" by fact thus ?case by (cases "merge_guards c") (auto intro: terminates.intros split: if_split_asm simp add: Let_def) qed (fastforce intro: terminates.intros dest: exec_merge_guards_to_exec)+ lemma terminates_merge_guards_to_terminates_Normal: shows "\s. \\merge_guards c\Normal s \ \\c\Normal s" proof (induct c) case Skip thus ?case by (fastforce intro: terminates.intros) next case Basic thus ?case by (fastforce intro: terminates.intros) next case Spec thus ?case by (fastforce intro: terminates.intros) next case (Seq c1 c2) have "\\merge_guards (Seq c1 c2) \ Normal s" by fact then obtain termi_merge_c1: "\\merge_guards c1 \ Normal s" and termi_merge_c2: "\s'. \\\merge_guards c1,Normal s \ \ s' \ \\merge_guards c2 \ s'" by (auto elim: terminates_Normal_elim_cases) from termi_merge_c1 Seq.hyps have "\\c1 \ Normal s" by iprover moreover { fix s' assume exec_c1: "\\\c1,Normal s \ \ s'" have "\\ c2 \ s'" proof - from exec_to_exec_merge_guards [OF exec_c1] have "\\\merge_guards c1,Normal s \ \ s'" . from termi_merge_c2 [rule_format, OF this] Seq.hyps show ?thesis by (cases s') (auto) qed } ultimately show ?case by (auto intro: terminates.intros) next case Cond thus ?case by (fastforce intro: terminates.intros elim: terminates_Normal_elim_cases) next case (While b c) { fix u c' assume termi_c': "\\c' \ Normal u" assume c': "c' = merge_guards (While b c)" have "\\While b c \ Normal u" using termi_c' c' proof (induct) case (WhileTrue s b' c') have s_in_b: "s \ b" using WhileTrue by simp have "\\merge_guards c \ Normal s" using WhileTrue by (auto elim: terminates_Normal_elim_cases) with While.hyps have "\\c \ Normal s" by auto moreover have hyp_w: "\w. \\\merge_guards c,Normal s \ \ w \ \\While b c \ w" using WhileTrue by simp hence "\w. \\\c,Normal s \ \ w \ \\While b c \ w" by (simp add: exec_iff_exec_merge_guards [symmetric]) ultimately show ?case using s_in_b by (auto intro: terminates.intros) next case WhileFalse thus ?case by (auto intro: terminates.intros) qed auto } with While show ?case by simp next case Call thus ?case by (fastforce intro: terminates.intros ) next case DynCom thus ?case by (fastforce intro: terminates.intros elim: terminates_Normal_elim_cases) next case (Guard f g c) have termi_merge: "\\merge_guards (Guard f g c) \ Normal s" by fact show ?case proof (cases "\f' g' c'. merge_guards c = Guard f' g' c'") case False hence m: "merge_guards (Guard f g c) = Guard f g (merge_guards c)" by (cases "merge_guards c") (auto simp add: Let_def) from termi_merge Guard.hyps show ?thesis by (simp only: m) (fastforce intro: terminates.intros elim: terminates_Normal_elim_cases) next case True then obtain f' g' c' where mc: "merge_guards c = Guard f' g' c'" by blast show ?thesis proof (cases "f=f'") case False with mc have m: "merge_guards (Guard f g c) = Guard f g (merge_guards c)" by (simp add: Let_def) from termi_merge Guard.hyps show ?thesis by (simp only: m) (fastforce intro: terminates.intros elim: terminates_Normal_elim_cases) next case True with mc have m: "merge_guards (Guard f g c) = Guard f (g \ g') c'" by simp from termi_merge Guard.hyps show ?thesis by (simp only: m mc) (auto intro: terminates.intros elim: terminates_Normal_elim_cases) qed qed next case Throw thus ?case by (fastforce intro: terminates.intros ) next case (Catch c1 c2) have "\\merge_guards (Catch c1 c2) \ Normal s" by fact then obtain termi_merge_c1: "\\merge_guards c1 \ Normal s" and termi_merge_c2: "\s'. \\\merge_guards c1,Normal s \ \ Abrupt s' \ \\merge_guards c2 \ Normal s'" by (auto elim: terminates_Normal_elim_cases) from termi_merge_c1 Catch.hyps have "\\c1 \ Normal s" by iprover moreover { fix s' assume exec_c1: "\\\c1,Normal s \ \ Abrupt s'" have "\\ c2 \ Normal s'" proof - from exec_to_exec_merge_guards [OF exec_c1] have "\\\merge_guards c1,Normal s \ \ Abrupt s'" . from termi_merge_c2 [rule_format, OF this] Catch.hyps show ?thesis by iprover qed } ultimately show ?case by (auto intro: terminates.intros) qed lemma terminates_merge_guards_to_terminates: "\\merge_guards c\ s \ \\c\ s" by (cases s) (auto intro: terminates_merge_guards_to_terminates_Normal) theorem terminates_iff_terminates_merge_guards: "\\c\ s = \\merge_guards c\ s" by (iprover intro: terminates_to_terminates_merge_guards terminates_merge_guards_to_terminates) (* ************************************************************************* *) subsection \Lemmas about @{term "c\<^sub>1 \\<^sub>g c\<^sub>2"}\ (* ************************************************************************ *) lemma terminates_fewer_guards_Normal: shows "\c s. \\\c'\Normal s; c \\<^sub>g c'; \\\c',Normal s \ \\Fault ` UNIV\ \ \\c\Normal s" proof (induct c') case Skip thus ?case by (auto intro: terminates.intros dest: subseteq_guardsD) next case Basic thus ?case by (auto intro: terminates.intros dest: subseteq_guardsD) next case Spec thus ?case by (auto intro: terminates.intros dest: subseteq_guardsD) next case (Seq c1' c2') have termi: "\\Seq c1' c2' \ Normal s" by fact then obtain termi_c1': "\\c1'\ Normal s" and termi_c2': "\s'. \\\c1',Normal s \ \ s' \ \\c2'\ s'" by (auto elim: terminates_Normal_elim_cases) have noFault: "\\\Seq c1' c2',Normal s \ \\Fault ` UNIV" by fact hence noFault_c1': "\\\c1',Normal s \ \\Fault ` UNIV" by (auto intro: exec.intros simp add: final_notin_def) have "c \\<^sub>g Seq c1' c2'" by fact from subseteq_guards_Seq [OF this] obtain c1 c2 where c: "c = Seq c1 c2" and c1_c1': "c1 \\<^sub>g c1'" and c2_c2': "c2 \\<^sub>g c2'" by blast from termi_c1' c1_c1' noFault_c1' have "\\c1\ Normal s" by (rule Seq.hyps) moreover { fix t assume exec_c1: "\\\c1,Normal s \ \ t" have "\\c2\ t" proof - from exec_to_exec_subseteq_guards [OF c1_c1' exec_c1] obtain t' where exec_c1': "\\\c1',Normal s \ \ t'" and t_Fault: "isFault t \ isFault t'" and t'_noFault: "\ isFault t' \ t' = t" by blast show ?thesis proof (cases "isFault t'") case True with exec_c1' noFault_c1' have False by (fastforce elim: isFaultE dest: Fault_end simp add: final_notin_def) thus ?thesis .. next case False with t'_noFault have t': "t'=t" by simp with termi_c2' exec_c1' have termi_c2': "\\c2'\ t" by auto show ?thesis proof (cases t) case Fault thus ?thesis by auto next case Abrupt thus ?thesis by auto next case Stuck thus ?thesis by auto next case (Normal u) with noFault exec_c1' t' have "\\\c2',Normal u \ \\Fault ` UNIV" by (auto intro: exec.intros simp add: final_notin_def) from termi_c2' [simplified Normal] c2_c2' this have "\\c2 \ Normal u" by (rule Seq.hyps) with Normal exec_c1 show ?thesis by simp qed qed qed } ultimately show ?case using c by (auto intro: terminates.intros) next case (Cond b c1' c2') have noFault: "\\\Cond b c1' c2',Normal s \ \\Fault ` UNIV" by fact have termi: "\\Cond b c1' c2' \ Normal s" by fact have "c \\<^sub>g Cond b c1' c2'" by fact from subseteq_guards_Cond [OF this] obtain c1 c2 where c: "c = Cond b c1 c2" and c1_c1': "c1 \\<^sub>g c1'" and c2_c2': "c2 \\<^sub>g c2'" by blast thus ?case proof (cases "s \ b") case True with termi have termi_c1': "\\c1' \ Normal s" by (auto elim: terminates_Normal_elim_cases) from True noFault have "\\\c1',Normal s \ \\Fault ` UNIV" by (auto intro: exec.intros simp add: final_notin_def) from termi_c1' c1_c1' this have "\\c1 \ Normal s" by (rule Cond.hyps) with True c show ?thesis by (auto intro: terminates.intros) next case False with termi have termi_c2': "\\c2' \ Normal s" by (auto elim: terminates_Normal_elim_cases) from False noFault have "\\\c2',Normal s \ \\Fault ` UNIV" by (auto intro: exec.intros simp add: final_notin_def) from termi_c2' c2_c2' this have "\\c2 \ Normal s" by (rule Cond.hyps) with False c show ?thesis by (auto intro: terminates.intros) qed next case (While b c') have noFault: "\\\While b c',Normal s \ \\Fault ` UNIV" by fact have termi: "\\While b c' \ Normal s" by fact have "c \\<^sub>g While b c'" by fact from subseteq_guards_While [OF this] obtain c'' where c: "c = While b c''" and c''_c': "c'' \\<^sub>g c'" by blast { fix d u assume termi: "\\d \ u" assume d: "d = While b c'" assume noFault: "\\\While b c',u \ \\Fault ` UNIV" have "\\While b c'' \ u" using termi d noFault proof (induct) case (WhileTrue u b' c''') have u_in_b: "u \ b" using WhileTrue by simp have termi_c': "\\c' \ Normal u" using WhileTrue by simp have noFault: "\\\While b c',Normal u \ \\Fault ` UNIV" using WhileTrue by simp hence noFault_c': "\\\c',Normal u \ \\Fault ` UNIV" using u_in_b by (auto intro: exec.intros simp add: final_notin_def) from While.hyps [OF termi_c' c''_c' this] have "\\c'' \ Normal u". moreover from WhileTrue have hyp_w: "\s'. \\\c',Normal u \ \ s' \ \\\While b c',s' \ \\Fault ` UNIV \ \\While b c'' \ s'" by simp { fix v assume exec_c'': "\\\c'',Normal u \ \ v" have "\\While b c'' \ v" proof - from exec_to_exec_subseteq_guards [OF c''_c' exec_c''] obtain v' where exec_c': "\\\c',Normal u \ \ v'" and v_Fault: "isFault v \ isFault v'" and v'_noFault: "\ isFault v' \ v' = v" by auto show ?thesis proof (cases "isFault v'") case True with exec_c' noFault u_in_b have False by (fastforce simp add: final_notin_def intro: exec.intros elim: isFaultE) thus ?thesis .. next case False with v'_noFault have v': "v'=v" by simp with noFault exec_c' u_in_b have "\\\While b c',v \ \\Fault ` UNIV" by (fastforce simp add: final_notin_def intro: exec.intros) from hyp_w [rule_format, OF exec_c' [simplified v'] this] show "\\While b c'' \ v" . qed qed } ultimately show ?case using u_in_b by (auto intro: terminates.intros) next case WhileFalse thus ?case by (auto intro: terminates.intros) qed auto } with c noFault termi show ?case by auto next case Call thus ?case by (auto intro: terminates.intros dest: subseteq_guardsD) next case (DynCom C') have termi: "\\DynCom C' \ Normal s" by fact hence termi_C': "\\C' s \ Normal s" by cases have noFault: "\\\DynCom C',Normal s \ \\Fault ` UNIV" by fact hence noFault_C': "\\\C' s,Normal s \ \\Fault ` UNIV" by (auto intro: exec.intros simp add: final_notin_def) have "c \\<^sub>g DynCom C'" by fact from subseteq_guards_DynCom [OF this] obtain C where c: "c = DynCom C" and C_C': "\s. C s \\<^sub>g C' s" by blast from DynCom.hyps termi_C' C_C' [rule_format] noFault_C' have "\\C s \ Normal s" by fast with c show ?case by (auto intro: terminates.intros) next case (Guard f' g' c') have noFault: "\\\Guard f' g' c',Normal s \ \\Fault ` UNIV" by fact have termi: "\\Guard f' g' c' \ Normal s" by fact have "c \\<^sub>g Guard f' g' c'" by fact hence c_cases: "(c \\<^sub>g c') \ (\c''. c = Guard f' g' c'' \ (c'' \\<^sub>g c'))" by (rule subseteq_guards_Guard) thus ?case proof (cases "s \ g'") case True note s_in_g' = this with noFault have noFault_c': "\\\c',Normal s \ \\Fault ` UNIV" by (auto simp add: final_notin_def intro: exec.intros) from termi s_in_g' have termi_c': "\\c' \ Normal s" by cases auto from c_cases show ?thesis proof assume "c \\<^sub>g c'" from termi_c' this noFault_c' show "\\c \ Normal s" by (rule Guard.hyps) next assume "\c''. c = Guard f' g' c'' \ (c'' \\<^sub>g c')" then obtain c'' where c: "c = Guard f' g' c''" and c''_c': "c'' \\<^sub>g c'" by blast from termi_c' c''_c' noFault_c' have "\\c'' \ Normal s" by (rule Guard.hyps) with s_in_g' c show ?thesis by (auto intro: terminates.intros) qed next case False with noFault have False by (auto intro: exec.intros simp add: final_notin_def) thus ?thesis .. qed next case Throw thus ?case by (auto intro: terminates.intros dest: subseteq_guardsD) next case (Catch c1' c2') have termi: "\\Catch c1' c2' \ Normal s" by fact then obtain termi_c1': "\\c1'\ Normal s" and termi_c2': "\s'. \\\c1',Normal s \ \ Abrupt s' \ \\c2'\ Normal s'" by (auto elim: terminates_Normal_elim_cases) have noFault: "\\\Catch c1' c2',Normal s \ \\Fault ` UNIV" by fact hence noFault_c1': "\\\c1',Normal s \ \\Fault ` UNIV" by (fastforce intro: exec.intros simp add: final_notin_def) have "c \\<^sub>g Catch c1' c2'" by fact from subseteq_guards_Catch [OF this] obtain c1 c2 where c: "c = Catch c1 c2" and c1_c1': "c1 \\<^sub>g c1'" and c2_c2': "c2 \\<^sub>g c2'" by blast from termi_c1' c1_c1' noFault_c1' have "\\c1\ Normal s" by (rule Catch.hyps) moreover { fix t assume exec_c1: "\\\c1,Normal s \ \ Abrupt t" have "\\c2\ Normal t" proof - from exec_to_exec_subseteq_guards [OF c1_c1' exec_c1] obtain t' where exec_c1': "\\\c1',Normal s \ \ t'" and t'_noFault: "\ isFault t' \ t' = Abrupt t" by blast show ?thesis proof (cases "isFault t'") case True with exec_c1' noFault_c1' have False by (fastforce elim: isFaultE dest: Fault_end simp add: final_notin_def) thus ?thesis .. next case False with t'_noFault have t': "t'=Abrupt t" by simp with termi_c2' exec_c1' have termi_c2': "\\c2'\ Normal t" by auto with noFault exec_c1' t' have "\\\c2',Normal t \ \\Fault ` UNIV" by (auto intro: exec.intros simp add: final_notin_def) from termi_c2' c2_c2' this show "\\c2 \ Normal t" by (rule Catch.hyps) qed qed } ultimately show ?case using c by (auto intro: terminates.intros) qed theorem terminates_fewer_guards: shows "\\\c'\s; c \\<^sub>g c'; \\\c',s \ \\Fault ` UNIV\ \ \\c\s" by (cases s) (auto intro: terminates_fewer_guards_Normal) lemma terminates_noFault_strip_guards: assumes termi: "\\c\Normal s" shows "\\\\c,Normal s \ \\Fault ` F\ \ \\strip_guards F c\Normal s" using termi proof (induct) case Skip thus ?case by (auto intro: terminates.intros) next case Basic thus ?case by (auto intro: terminates.intros) next case Spec thus ?case by (auto intro: terminates.intros) next case (Guard s g c f) have s_in_g: "s \ g" by fact have "\\c \ Normal s" by fact have "\\\Guard f g c,Normal s \ \\Fault ` F" by fact with s_in_g have "\\\c,Normal s \ \\Fault ` F" by (fastforce simp add: final_notin_def intro: exec.intros) with Guard.hyps have "\\strip_guards F c \ Normal s" by simp with s_in_g show ?case by (auto intro: terminates.intros) next case GuardFault thus ?case by (auto intro: terminates.intros exec.intros simp add: final_notin_def ) next case Fault thus ?case by (auto intro: terminates.intros) next case (Seq c1 s c2) have noFault_Seq: "\\\Seq c1 c2,Normal s \ \\Fault ` F" by fact hence noFault_c1: "\\\c1,Normal s \ \\Fault ` F" by (auto simp add: final_notin_def intro: exec.intros) with Seq.hyps have "\\strip_guards F c1 \ Normal s" by simp moreover { fix s' assume exec_strip_guards_c1: "\\\strip_guards F c1,Normal s \ \ s'" have "\\strip_guards F c2 \ s'" proof (cases "isFault s'") case True thus ?thesis by (auto elim: isFaultE intro: terminates.intros) next case False with exec_strip_guards_to_exec [OF exec_strip_guards_c1] noFault_c1 have *: "\\\c1,Normal s \ \ s'" by (auto simp add: final_notin_def elim!: isFaultE) with noFault_Seq have "\\\c2,s' \ \\Fault ` F" by (auto simp add: final_notin_def intro: exec.intros) with * show ?thesis using Seq.hyps by simp qed } ultimately show ?case by (auto intro: terminates.intros) next case CondTrue thus ?case by (fastforce intro: terminates.intros exec.intros simp add: final_notin_def ) next case CondFalse thus ?case by (fastforce intro: terminates.intros exec.intros simp add: final_notin_def ) next case (WhileTrue s b c) have s_in_b: "s \ b" by fact have noFault_while: "\\\While b c,Normal s \ \\Fault ` F" by fact with s_in_b have noFault_c: "\\\c,Normal s \ \\Fault ` F" by (auto simp add: final_notin_def intro: exec.intros) with WhileTrue.hyps have "\\strip_guards F c \ Normal s" by simp moreover { fix s' assume exec_strip_guards_c: "\\\strip_guards F c,Normal s \ \ s'" have "\\strip_guards F (While b c) \ s'" proof (cases "isFault s'") case True thus ?thesis by (auto elim: isFaultE intro: terminates.intros) next case False with exec_strip_guards_to_exec [OF exec_strip_guards_c] noFault_c have *: "\\\c,Normal s \ \ s'" by (auto simp add: final_notin_def elim!: isFaultE) with s_in_b noFault_while have "\\\While b c,s' \ \\Fault ` F" by (auto simp add: final_notin_def intro: exec.intros) with * show ?thesis using WhileTrue.hyps by simp qed } ultimately show ?case using WhileTrue.hyps by (auto intro: terminates.intros) next case WhileFalse thus ?case by (auto intro: terminates.intros) next case Call thus ?case by (auto intro: terminates.intros) next case CallUndefined thus ?case by (auto intro: terminates.intros) next case Stuck thus ?case by (auto intro: terminates.intros) next case DynCom thus ?case by (auto intro: terminates.intros exec.intros simp add: final_notin_def ) next case Throw thus ?case by (auto intro: terminates.intros) next case Abrupt thus ?case by (auto intro: terminates.intros) next case (Catch c1 s c2) have noFault_Catch: "\\\Catch c1 c2,Normal s \ \\Fault ` F" by fact hence noFault_c1: "\\\c1,Normal s \ \\Fault ` F" by (fastforce simp add: final_notin_def intro: exec.intros) with Catch.hyps have "\\strip_guards F c1 \ Normal s" by simp moreover { fix s' assume exec_strip_guards_c1: "\\\strip_guards F c1,Normal s \ \ Abrupt s'" have "\\strip_guards F c2 \ Normal s'" proof - from exec_strip_guards_to_exec [OF exec_strip_guards_c1] noFault_c1 have *: "\\\c1,Normal s \ \ Abrupt s'" by (auto simp add: final_notin_def elim!: isFaultE) with noFault_Catch have "\\\c2,Normal s' \ \\Fault ` F" by (auto simp add: final_notin_def intro: exec.intros) with * show ?thesis using Catch.hyps by simp qed } ultimately show ?case using Catch.hyps by (auto intro: terminates.intros) qed (* ************************************************************************* *) subsection \Lemmas about @{const "strip_guards"}\ (* ************************************************************************* *) lemma terminates_noFault_strip: assumes termi: "\\c\Normal s" shows "\\\\c,Normal s \ \\Fault ` F\ \ strip F \\c\Normal s" using termi proof (induct) case Skip thus ?case by (auto intro: terminates.intros) next case Basic thus ?case by (auto intro: terminates.intros) next case Spec thus ?case by (auto intro: terminates.intros) next case (Guard s g c f) have s_in_g: "s \ g" by fact have "\\\Guard f g c,Normal s \ \\Fault ` F" by fact with s_in_g have "\\\c,Normal s \ \\Fault ` F" by (fastforce simp add: final_notin_def intro: exec.intros) then have "strip F \\c \ Normal s" by (simp add: Guard.hyps) with s_in_g show ?case by (auto intro: terminates.intros simp del: strip_simp) next case GuardFault thus ?case by (auto intro: terminates.intros exec.intros simp add: final_notin_def ) next case Fault thus ?case by (auto intro: terminates.intros) next case (Seq c1 s c2) have noFault_Seq: "\\\Seq c1 c2,Normal s \ \\Fault ` F" by fact hence noFault_c1: "\\\c1,Normal s \ \\Fault ` F" by (auto simp add: final_notin_def intro: exec.intros) then have "strip F \\c1 \ Normal s" by (simp add: Seq.hyps) moreover { fix s' assume exec_strip_c1: "strip F \\\c1,Normal s \ \ s'" have "strip F \\c2 \ s'" proof (cases "isFault s'") case True thus ?thesis by (auto elim: isFaultE intro: terminates.intros) next case False with exec_strip_to_exec [OF exec_strip_c1] noFault_c1 have *: "\\\c1,Normal s \ \ s'" by (auto simp add: final_notin_def elim!: isFaultE) with noFault_Seq have "\\\c2,s' \ \\Fault ` F" by (auto simp add: final_notin_def intro: exec.intros) with * show ?thesis using Seq.hyps by (simp del: strip_simp) qed } ultimately show ?case by (fastforce intro: terminates.intros) next case CondTrue thus ?case by (fastforce intro: terminates.intros exec.intros simp add: final_notin_def ) next case CondFalse thus ?case by (fastforce intro: terminates.intros exec.intros simp add: final_notin_def ) next case (WhileTrue s b c) have s_in_b: "s \ b" by fact have noFault_while: "\\\While b c,Normal s \ \\Fault ` F" by fact with s_in_b have noFault_c: "\\\c,Normal s \ \\Fault ` F" by (auto simp add: final_notin_def intro: exec.intros) then have "strip F \\c \ Normal s" by (simp add: WhileTrue.hyps) moreover { fix s' assume exec_strip_c: "strip F \\\c,Normal s \ \ s'" have "strip F \\While b c \ s'" proof (cases "isFault s'") case True thus ?thesis by (auto elim: isFaultE intro: terminates.intros) next case False with exec_strip_to_exec [OF exec_strip_c] noFault_c have *: "\\\c,Normal s \ \ s'" by (auto simp add: final_notin_def elim!: isFaultE) with s_in_b noFault_while have "\\\While b c,s' \ \\Fault ` F" by (auto simp add: final_notin_def intro: exec.intros) with * show ?thesis using WhileTrue.hyps by (simp del: strip_simp) qed } ultimately show ?case using WhileTrue.hyps by (auto intro: terminates.intros simp del: strip_simp) next case WhileFalse thus ?case by (auto intro: terminates.intros) next case (Call p bdy s) have bdy: "\ p = Some bdy" by fact have "\\\Call p,Normal s \ \\Fault ` F" by fact with bdy have bdy_noFault: "\\\bdy,Normal s \ \\Fault ` F" by (auto intro: exec.intros simp add: final_notin_def) then have strip_bdy_noFault: "strip F \\\bdy,Normal s \ \\Fault ` F" by (auto simp add: final_notin_def dest!: exec_strip_to_exec elim!: isFaultE) from bdy_noFault have "strip F \\bdy \ Normal s" by (simp add: Call.hyps) from terminates_noFault_strip_guards [OF this strip_bdy_noFault] have "strip F \\strip_guards F bdy \ Normal s". with bdy show ?case by (fastforce intro: terminates.Call) next case CallUndefined thus ?case by (auto intro: terminates.intros) next case Stuck thus ?case by (auto intro: terminates.intros) next case DynCom thus ?case by (auto intro: terminates.intros exec.intros simp add: final_notin_def ) next case Throw thus ?case by (auto intro: terminates.intros) next case Abrupt thus ?case by (auto intro: terminates.intros) next case (Catch c1 s c2) have noFault_Catch: "\\\Catch c1 c2,Normal s \ \\Fault ` F" by fact hence noFault_c1: "\\\c1,Normal s \ \\Fault ` F" by (fastforce simp add: final_notin_def intro: exec.intros) then have "strip F \\c1 \ Normal s" by (simp add: Catch.hyps) moreover { fix s' assume exec_strip_c1: "strip F \\\c1,Normal s \ \ Abrupt s'" have "strip F \\c2 \ Normal s'" proof - from exec_strip_to_exec [OF exec_strip_c1] noFault_c1 have *: "\\\c1,Normal s \ \ Abrupt s'" by (auto simp add: final_notin_def elim!: isFaultE) with * noFault_Catch have "\\\c2,Normal s' \ \\Fault ` F" by (auto simp add: final_notin_def intro: exec.intros) with * show ?thesis using Catch.hyps by (simp del: strip_simp) qed } ultimately show ?case using Catch.hyps by (auto intro: terminates.intros simp del: strip_simp) qed (* ************************************************************************* *) subsection \Miscellaneous\ (* ************************************************************************* *) lemma terminates_while_lemma: assumes termi: "\\w\fk" shows "\k b c. \fk = Normal (f k); w=While b c; \i. \\\c,Normal (f i) \ \ Normal (f (Suc i))\ \ \i. f i \ b" using termi proof (induct) case WhileTrue thus ?case by blast next case WhileFalse thus ?case by blast qed simp_all lemma terminates_while: "\\\(While b c)\Normal (f k); \i. \\\c,Normal (f i) \ \ Normal (f (Suc i))\ \ \i. f i \ b" by (blast intro: terminates_while_lemma) lemma wf_terminates_while: "wf {(t,s). \\(While b c)\Normal s \ s\b \ \\\c,Normal s \ \ Normal t}" apply(subst wf_iff_no_infinite_down_chain) apply(rule notI) apply clarsimp apply(insert terminates_while) apply blast done lemma terminates_restrict_to_terminates: assumes terminates_res: "\|\<^bsub>M\<^esub>\ c \ s" assumes not_Stuck: "\|\<^bsub>M\<^esub>\\c,s \ \\{Stuck}" shows "\\ c \ s" using terminates_res not_Stuck proof (induct) case Skip show ?case by (rule terminates.Skip) next case Basic show ?case by (rule terminates.Basic) next case Spec show ?case by (rule terminates.Spec) next case Guard thus ?case by (auto intro: terminates.Guard dest: notStuck_GuardD) next case GuardFault thus ?case by (auto intro: terminates.GuardFault) next case Fault show ?case by (rule terminates.Fault) next case (Seq c1 s c2) have not_Stuck: "\|\<^bsub>M\<^esub>\\Seq c1 c2,Normal s \ \\{Stuck}" by fact hence c1_notStuck: "\|\<^bsub>M\<^esub>\\c1,Normal s \ \\{Stuck}" by (rule notStuck_SeqD1) show "\\Seq c1 c2 \ Normal s" proof (rule terminates.Seq,safe) from c1_notStuck show "\\c1 \ Normal s" by (rule Seq.hyps) next fix s' assume exec: "\\\c1,Normal s \ \ s'" show "\\c2 \ s'" proof - from exec_to_exec_restrict [OF exec] obtain t' where exec_res: "\|\<^bsub>M\<^esub>\\c1,Normal s \ \ t'" and t'_notStuck: "t' \ Stuck \ t' = s'" by blast show ?thesis proof (cases "t'=Stuck") case True with c1_notStuck exec_res have "False" by (auto simp add: final_notin_def) thus ?thesis .. next case False with t'_notStuck have t': "t'=s'" by simp with not_Stuck exec_res have "\|\<^bsub>M\<^esub>\\c2,s' \ \\{Stuck}" by (auto dest: notStuck_SeqD2) with exec_res t' Seq.hyps show ?thesis by auto qed qed qed next case CondTrue thus ?case by (auto intro: terminates.CondTrue dest: notStuck_CondTrueD) next case CondFalse thus ?case by (auto intro: terminates.CondFalse dest: notStuck_CondFalseD) next case (WhileTrue s b c) have s: "s \ b" by fact have not_Stuck: "\|\<^bsub>M\<^esub>\\While b c,Normal s \ \\{Stuck}" by fact with WhileTrue have c_notStuck: "\|\<^bsub>M\<^esub>\\c,Normal s \ \\{Stuck}" by (iprover intro: notStuck_WhileTrueD1) show ?case proof (rule terminates.WhileTrue [OF s],safe) from c_notStuck show "\\c \ Normal s" by (rule WhileTrue.hyps) next fix s' assume exec: "\\\c,Normal s \ \ s'" show "\\While b c \ s'" proof - from exec_to_exec_restrict [OF exec] obtain t' where exec_res: "\|\<^bsub>M\<^esub>\\c,Normal s \ \ t'" and t'_notStuck: "t' \ Stuck \ t' = s'" by blast show ?thesis proof (cases "t'=Stuck") case True with c_notStuck exec_res have "False" by (auto simp add: final_notin_def) thus ?thesis .. next case False with t'_notStuck have t': "t'=s'" by simp with not_Stuck exec_res s have "\|\<^bsub>M\<^esub>\\While b c,s' \ \\{Stuck}" by (auto dest: notStuck_WhileTrueD2) with exec_res t' WhileTrue.hyps show ?thesis by auto qed qed qed next case WhileFalse then show ?case by (iprover intro: terminates.WhileFalse) next case Call thus ?case by (auto intro: terminates.Call dest: notStuck_CallD restrict_SomeD) next case CallUndefined thus ?case by (auto dest: notStuck_CallDefinedD) next case Stuck show ?case by (rule terminates.Stuck) next case DynCom thus ?case by (auto intro: terminates.DynCom dest: notStuck_DynComD) next case Throw show ?case by (rule terminates.Throw) next case Abrupt show ?case by (rule terminates.Abrupt) next case (Catch c1 s c2) have not_Stuck: "\|\<^bsub>M\<^esub>\\Catch c1 c2,Normal s \ \\{Stuck}" by fact hence c1_notStuck: "\|\<^bsub>M\<^esub>\\c1,Normal s \ \\{Stuck}" by (rule notStuck_CatchD1) show "\\Catch c1 c2 \ Normal s" proof (rule terminates.Catch,safe) from c1_notStuck show "\\c1 \ Normal s" by (rule Catch.hyps) next fix s' assume exec: "\\\c1,Normal s \ \ Abrupt s'" show "\\c2 \ Normal s'" proof - from exec_to_exec_restrict [OF exec] obtain t' where exec_res: "\|\<^bsub>M\<^esub>\\c1,Normal s \ \ t'" and t'_notStuck: "t' \ Stuck \ t' = Abrupt s'" by blast show ?thesis proof (cases "t'=Stuck") case True with c1_notStuck exec_res have "False" by (auto simp add: final_notin_def) thus ?thesis .. next case False with t'_notStuck have t': "t'=Abrupt s'" by simp with not_Stuck exec_res have "\|\<^bsub>M\<^esub>\\c2,Normal s' \ \\{Stuck}" by (auto dest: notStuck_CatchD2) with exec_res t' Catch.hyps show ?thesis by auto qed qed qed qed end diff --git a/thys/Simpl/UserGuide.thy b/thys/Simpl/UserGuide.thy --- a/thys/Simpl/UserGuide.thy +++ b/thys/Simpl/UserGuide.thy @@ -1,1583 +1,1564 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: UserGuide.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (c) 2022 Apple Inc. All rights reserved. *) section \User Guide \label{sec:UserGuide}\ (*<*) theory UserGuide imports HeapList Vcg "HOL-Statespace.StateSpaceSyntax" "HOL-Library.LaTeXsugar" begin (*>*) (*<*) syntax "_statespace_updates" :: "('a \ 'b) \ updbinds \ ('a \ 'b)" ("_\_\" [900,0] 900) (*>*) text \ We introduce the verification environment with a couple of examples that illustrate how to use the different bits and pieces to verify programs. \ subsection \Basics\ text \ First of all we have to decide how to represent the state space. There are currently two implementations. One is based on records the other one on the concept called `statespace' that was introduced with Isabelle 2007 (see \texttt{HOL/Statespace}) . In contrast to records a 'satespace' does not define a new type, but provides a notion of state, based on locales. Logically the state is modelled as a function from (abstract) names to (abstract) values and the statespace infrastructure organises distinctness of names an projection/injection of concrete values into the abstract one. Towards the user the interface of records and statespaces is quite similar. However, statespaces offer more flexibility, inherited from the locale infrastructure, in particular multiple inheritance and renaming of components. In this user guide we prefer statespaces, but give some comments on the usage of records in Section \ref{sec:records}. \ hoarestate vars = A :: nat I :: nat M :: nat N :: nat R :: nat S :: nat text (in vars) \The command \isacommand{hoarestate} is a simple preprocessor for the command \isacommand{statespaces} which decorates the state components with the suffix \_'\, to avoid cluttering the namespace. Also note that underscores are printed as hyphens in this documentation. So what you see as @{term "A_'"} in this document is actually \texttt{A\_'}. Every component name becomes a fixed variable in the locale \vars\ and can no longer be used for logical variables. Lookup of a component @{term "A_'"} in a state @{term "s"} is written as @{term "s\A_'"}, and update with a value @{term "term v"} as @{term "s\A_' := v\"}. To deal with local and global variables in the context of procedures the program state is organised as a record containing the two componets @{const "locals"} and @{const "globals"}. The variables defined in hoarestate \vars\ reside in the @{const "locals"} part. \ text \ Here is a first example. \ lemma (in vars) "\\ \\N = 5\ \N :== 2 * \N \\N = 10\" apply vcg txt \@{subgoals}\ apply simp txt \@{subgoals}\ done text \We enable the locale of statespace \vars\ by the \texttt{in vars} directive. The verification condition generator is invoked via the \vcg\ method and leaves us with the expected subgoal that can be proved by simplification.\ text (in vars) \ If we refer to components (variables) of the state-space of the program we always mark these with \\\ (in assertions and also in the program itself). It is the acute-symbol and is present on most keyboards. The assertions of the Hoare tuple are ordinary Isabelle sets. As we usually want to refer to the state space in the assertions, we provide special brackets for them. They can be written as {\verb+{| |}+} in ASCII or \\ \\ with symbols. Internally, marking variables has two effects. First of all we refer to the implicit state and secondary we get rid of the suffix \_'\. So the assertion @{term "{|\N = 5|}"} internally gets expanded to \{s. locals s \N_' = 5}\ written in ordinary set comprehension notation of Isabelle. It describes the set of states where the \N_'\ component is equal to \5\. An empty context and an empty postcondition for abrupt termination can be omitted. The lemma above is a shorthand for \\,{}\ \\N = 5\ \N :== 2 * \N \\N = 10\,{}\. \ text \We can step through verification condition generation by the method \vcg_step\. \ lemma (in vars) "\,{}\ \\N = 5\ \N :== 2 * \N \\N = 10\" apply vcg_step txt \@{subgoals}\ txt \The last step of verification condition generation, transforms the inclusion of state sets to the corresponding predicate on components of the state space. \ apply vcg_step txt \@{subgoals}\ by simp text \ Although our assertions work semantically on the state space, stepping through verification condition generation ``feels'' like the expected syntactic substitutions of traditional Hoare logic. This is achieved by light simplification on the assertions calculated by the Hoare rules. \ lemma (in vars) "\\ \\N = 5\ \N :== 2 * \N \\N = 10\" apply (rule HoarePartial.Basic) txt \@{subgoals}\ apply (simp only: mem_Collect_eq) txt \@{subgoals}\ apply (tactic \Hoare.BasicSimpTac @{context} Hoare.Function false [] (K all_tac) 1\) txt \@{subgoals}\ by simp text \The next example shows how we deal with the while loop. Note the invariant annotation. \ lemma (in vars) "\,{}\ \\M = 0 \ \S = 0\ WHILE \M \ a INV \\S = \M * b\ DO \S :== \S + b;; \M :== \M + 1 OD \\S = a * b\" apply vcg txt \@{subgoals [display]}\ txt \The verification condition generator gives us three proof obligations, stemming from the path from the precondition to the invariant, from the invariant together with loop condition through the loop body to the invariant, and finally from the invariant together with the negated loop condition to the postcondition.\ apply auto done subsection \Procedures\ subsubsection \Declaration\ text \ Our first procedure is a simple square procedure. We provide the command \isacommand{procedures}, to declare and define a procedure. \ procedures Square (N::nat|R::nat) where I::nat in "\R :== \N * \N" text \A procedure is given by the signature of the procedure followed by the procedure body. The signature consists of the name of the procedure and a list of parameters together with their types. The parameters in front of the pipe \|\ are value parameters and behind the pipe are the result parameters. Value parameters model call by value semantics. The value of a result parameter at the end of the procedure is passed back to the caller. Local variables follow the \where\. If there are no local variables the \where \ in\ can be omitted. The variable @{term "I"} is actually unused in the body, but is used in the examples below.\ text \ The procedures command provides convenient syntax for procedure calls (that creates the proper @{term init}, @{term return} and @{term result} functions on the fly) and creates locales and statespaces to reason about the procedure. The purpose of locales is to set up logical contexts to support modular reasoning. Locales can be seen as freeze-dried proof contexts that -get alive as you setup a new lemma or theorem (\<^cite>\"Ballarin-04-locales"\). +get alive as you setup a new lemma or theorem (\cite{Ballarin-04-locales}). The locale the user deals with is named \Square_impl\. It defines the procedure name (internally @{term "Square_'proc"}), the procedure body (named \Square_body\) and the statespaces for parameters and local and global variables. Moreover it contains the assumption @{term "\ Square_'proc = Some Square_body"}, which states that the procedure is properly defined in the procedure context. The purpose of the locale is to give us easy means to setup the context in which we prove programs correct. In this locale the procedure context @{term "\"} is fixed. So we always use this letter for the procedure specification. This is crucial, if we prove programs under the assumption of some procedure specifications. \ (*<*) context Square_impl begin (*>*) text \The procedures command generates syntax, so that we can either write \CALL Square(\I,\R)\ or @{term "\I :== CALL Square(\R)"} for the procedure call. The internal term is the following: \ (*<*) declare [[hoare_use_call_tr' = false]] (*>*) text \\small @{term [display] "CALL Square(\I,\R)"}\ (*<*) declare [[hoare_use_call_tr' = true]] (*>*) text \Note the additional decoration (with the procedure name) of the parameter and local variable names.\ (*<*) end (*>*) text \The abstract syntax for the procedure call is @{term "call init p return result"}. The @{term "init"} function copies the values of the actual parameters to the formal parameters, the @{term return} function copies the global variables back (in our case there are no global variables), and the @{term "result"} function additionally copies the values of the formal result parameters to the actual locations. Actual value parameters can be all kind of expressions, since we only need their value. But result parameters must be proper ``lvalues'': variables (including dereferenced pointers) or array locations, since we have to assign values to them. \ subsubsection \Verification\ text (in Square_impl) \ A procedure specification is an ordinary Hoare tuple. We use the parameterless call for the specification; \\R :== PROC Square(\N)\ is syntactic sugar for \Call Square_'proc\. This emphasises that the specification describes the internal behaviour of the procedure, whereas parameter passing corresponds to the procedure call. The following precondition fixes the current value \\N\ to the logical variable @{term n}. Universal quantification of @{term "n"} enables us to adapt the specification to an actual parameter. The specification is used in the rule for procedure call when we come upon a call to @{term Square}. Thus @{term "n"} plays the role of the auxiliary variable @{term "Z"}. \ text \To verify the procedure we need to verify the body. We use a derived variant of the general recursion rule, tailored for non recursive procedures: @{thm [source] HoarePartial.ProcNoRec1}: \begin{center} @{thm [mode=Rule,mode=ParenStmt] HoarePartial.ProcNoRec1 [no_vars]} \end{center} The naming convention for the rule is the following: The \1\ expresses that we look at one procedure, and \NoRec\ that the procedure is non recursive. \ lemma (in Square_impl) shows "\n. \\\\N = n\ \R :== PROC Square(\N) \\R = n * n\" txt \The directive \in\ has the effect that the context of the locale @{term "Square_impl"} is included to the current lemma, and that the lemma is added as a fact to the locale, after it is proven. The next time locale @{term "Square_impl"} is invoked this lemma is immediately available as fact, which the verification condition generator can use. \ apply (hoare_rule HoarePartial.ProcNoRec1) txt "@{subgoals[display]}" txt \The method \hoare_rule\, like \rule\ applies a single rule, but additionally does some ``obvious'' steps: It solves the canonical side-conditions of various Hoare-rules and it automatically expands the procedure body: With @{thm [source] Square_impl}: @{thm [names_short] Square_impl [no_vars]} we get the procedure body out of the procedure context @{term "\"}; with @{thm [source] Square_body_def}: @{thm [names_short] Square_body_def [no_vars]} we can unfold the definition of the body. The proof is finished by the vcg and simp. \ txt "@{subgoals[display]}" by vcg simp text \If the procedure is non recursive and there is no specification given, the verification condition generator automatically expands the body.\ lemma (in Square_impl) Square_spec: shows "\n. \\\\N = n\ \R :== PROC Square(\N) \\R = n * n\" by vcg simp text \An important naming convention is to name the specification as \_spec\. The verification condition generator refers to this name in order to search for a specification in the theorem database. \ subsubsection \Usage\ text\Let us see how we can use procedure specifications.\ -(* FIXME: maybe don't show this at all *) +(* fixme: maybe don't show this at all *) lemma (in Square_impl) shows "\\\\I = 2\ \R :== CALL Square(\I) \\R = 4\" txt \Remember that we have already proven @{thm [source] "Square_spec"} in the locale \Square_impl\. This is crucial for verification condition generation. When reaching a procedure call, it looks for the specification (by its name) and applies the rule @{thm [source,mode=ParenStmt] HoarePartial.ProcSpec} instantiated with the specification (as last premise). Before we apply the verification condition generator, let us take some time to think of what we can expect. Let's look at the specification @{thm [source] Square_spec} again: @{thm [display] Square_spec [no_vars]} The specification talks about the formal parameters @{term "N"} and @{term R}. The precondition @{term "\\N = n\"} just fixes the initial value of \N\. The actual parameters are @{term "I"} and @{term "R"}. We have to adapt the specification to this calling context. @{term "\n. \\ \\I = n\ \R :== CALL Square(\I) \\R = n * n\"}. From the postcondition @{term "\\R = n * n\"} we have to derive the actual postcondition @{term "\\R = 4\"}. So we gain something like: @{term "\n * n = (4::nat)\"}. The precondition is @{term "\\I = 2\"} and the specification tells us @{term "\\I = n\"} for the pre-state. So the value of @{term n} is the value of @{term I} in the pre-state. So we arrive at @{term "\\I = 2\ \ \\I * \I = 4\"}. \ apply vcg_step txt "@{subgoals[display]}" txt \ The second set looks slightly more involved: @{term "\\t. \<^bsup>t\<^esup>R = \I * \I \ \I * \I = 4\"}, this is an artefact from the procedure call rule. Originally \\I * \I = 4\ was \\<^bsup>t\<^esup>R = 4\. Where @{term "t"} denotes the final state of the procedure and the superscript notation allows to select a component from a particular state. \ apply vcg_step txt "@{subgoals[display]}" by simp text \ The adaption of the procedure specification to the actual calling context is done due to the @{term init}, @{term return} and @{term result} functions in the rule @{thm [source] HoarePartial.ProcSpec} (or in the variant @{thm [source] HoarePartial.ProcSpecNoAbrupt} which already incorporates the fact that the postcondition for abrupt termination is the empty set). For the readers interested in the internals, here a version without vcg. \ lemma (in Square_impl) shows "\\\\I = 2\ \R :== CALL Square(\I) \\R = 4\" apply (rule HoarePartial.ProcSpecNoAbrupt [OF _ _ Square_spec]) txt "@{subgoals[display]}" txt \This is the raw verification condition, It is interesting to see how the auxiliary variable @{term "Z"} is actually used. It is unified with @{term n} of the specification and fixes the state after parameter passing. \ apply simp txt "@{subgoals[display]}" prefer 2 apply vcg_step txt "@{subgoals[display]}" apply (auto intro: ext) done subsubsection \Recursion\ text \We want to define a procedure for the factorial. We first define a HOL function that calculates it, to specify the procedure later on. \ primrec fac:: "nat \ nat" where "fac 0 = 1" | "fac (Suc n) = (Suc n) * fac n" (*<*) lemma fac_simp [simp]: "0 < i \ fac i = i * fac (i - 1)" by (cases i) simp_all (*>*) text \Now we define the procedure.\ procedures Fac (N::nat | R::nat) "IF \N = 0 THEN \R :== 1 ELSE \R :== CALL Fac(\N - 1);; \R :== \N * \R FI" text \ Now let us prove that our implementation of @{term "Fac"} meets its specification. \ lemma (in Fac_impl) shows "\n. \\ \\N = n\ \R :== PROC Fac(\N) \\R = fac n\" apply (hoare_rule HoarePartial.ProcRec1) txt "@{subgoals[display]}" apply vcg txt "@{subgoals[display]}" apply simp done text \ Since the factorial is implemented recursively, the main ingredient of this proof is, to assume that the specification holds for the recursive call of @{term Fac} and prove the body correct. The assumption for recursive calls is added to the context by the rule @{thm [source] HoarePartial.ProcRec1} (also derived from the general rule for mutually recursive procedures): \begin{center} @{thm [mode=Rule,mode=ParenStmt] HoarePartial.ProcRec1 [no_vars]} \end{center} The verification condition generator infers the specification out of the context @{term "\"} when it encounters a recursive call of the factorial. \ subsection \Global Variables and Heap \label{sec:VcgHeap}\ text \ Now we define and verify some procedures on heap-lists. We consider list structures consisting of two fields, a content element @{term "cont"} and a reference to the next list element @{term "next"}. We model this by the following state space where every field has its own heap. \ hoarestate globals_heap = "next" :: "ref \ ref" cont :: "ref \ nat" text \It is mandatory to start the state name with `globals'. This is exploited by the syntax translations to store the components in the @{const globals} part of the state. \ text \Updates to global components inside a procedure are always propagated to the caller. This is implicitly done by the parameter passing syntax translations. \ text \We first define an append function on lists. It takes two references as parameters. It appends the list referred to by the first parameter with the list referred to by the second parameter. The statespace of the global variables has to be imported. \ procedures (imports globals_heap) append(p :: ref, q::ref | p::ref) "IF \p=Null THEN \p :== \q ELSE \p\\next :== CALL append(\p\\next,\q) FI" (*<*) context append_impl begin (*>*) text \ The difference of a global and a local variable is that global variables are automatically copied back to the procedure caller. We can study this effect on the translation of @{term "\p :== CALL append(\p,\q)"}: \ (*<*) declare [[hoare_use_call_tr' = false]] (*>*) text \ @{term [display] "\p :== CALL append(\p,\q)"} \ (*<*) declare [[hoare_use_call_tr' = true]] end (*>*) text \Below we give two specifications this time. One captures the functional behaviour and focuses on the entities that are potentially modified by the procedure, the second one is a pure frame condition. \ text \ The functional specification below introduces two logical variables besides the state space variable @{term "\"}, namely @{term "Ps"} and @{term "Qs"}. They are universally quantified and range over both the pre-and the postcondition, so that we are able to properly instantiate the specification during the proofs. The syntax \\\. \\\ is a shorthand to fix the current state: \{s. \ = s \}\. Moreover \\<^bsup>\\<^esup>x\ abbreviates the lookup of variable \x\ in the state \\\. The approach to specify procedures on lists -basically follows \<^cite>\"MehtaN-CADE03"\. From the pointer structure +basically follows \cite{MehtaN-CADE03}. From the pointer structure in the heap we (relationally) abstract to HOL lists of references. Then we can specify further properties on the level of HOL lists, rather then on the heap. The basic abstractions are: @{thm [display] Path.simps [no_vars]} @{term [show_types] "Path x h y ps"}: @{term ps} is a list of references that we can obtain out of the heap @{term h} by starting with the reference @{term x}, following the references in @{term h} up to the reference @{term y}. @{thm [display] List_def [no_vars]} A list @{term "List p h ps"} is a path starting in @{term p} and ending up in @{term Null}. \ lemma (in append_impl) append_spec1: shows "\\ Ps Qs. \\ \\. List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {}\ \p :== PROC append(\p,\q) \List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoarePartial.ProcRec1) txt \@{subgoals [margin=80,display]} Note that @{term "hoare_rule"} takes care of multiple auxiliary variables! @{thm [source] HoarePartial.ProcRec1} has only one auxiliary variable, namely @{term Z}. But the type of @{term Z} can be instantiated arbitrarily. So \hoare_rule\ instantiates @{term Z} with the tuple @{term "(\,Ps,Qs)"} and derives a proper variant of the rule. Therefore \hoare_rule\ depends on the proper quantification of auxiliary variables! \ apply vcg txt \@{subgoals [display]} For each branch of the \IF\ statement we have one conjunct to prove. The \THEN\ branch starts with \p = Null \ \\ and the \ELSE\ branch with \p \ Null \ \\. Let us focus on the \ELSE\ branch, were the recursive call to append occurs. First of all we have to prove that the precondition for the recursive call is fulfilled. That means we have to provide some witnesses for the lists @{term Psa} and @{term Qsa} which are referenced by \p\next\ (now written as @{term "next p"}) and @{term q}. Then we have to show that we can derive the overall postcondition from the postcondition of the recursive call. The state components that have changed by the recursive call are the ones with the suffix \a\, like \nexta\ and \pa\. \ apply fastforce done text \If the verification condition generator works on a procedure call it checks whether it can find a modifies clause in the context. If one is present the procedure call is simplified before the Hoare rule @{thm [source] HoarePartial.ProcSpec} is applied. Simplification of the procedure call means that the ``copy back'' of the global components is simplified. Only those components that occur in the modifies clause are actually copied back. This simplification is justified by the rule @{thm [source] HoarePartial.ProcModifyReturn}. So after this simplification all global components that do not appear in the modifies clause are treated as local variables.\ text \We study the effect of the modifies clause on the following examples, where we want to prove that @{term "append"} does not change the @{term "cont"} part of the heap. \ lemma (in append_impl) shows "\\ \\cont=c\ \p :== CALL append(Null,Null) \\cont=c\" proof - note append_spec = append_spec1 show ?thesis apply vcg txt \@{subgoals [display]}\ txt \Only focus on the very last line: @{term conta} is the heap component after the procedure call, and @{term cont} the heap component before the procedure call. Since we have not added the modified clause we do not know that they have to be equal. \ oops text \ We now add the frame condition. The list in the modifies clause names all global state components that may be changed by the procedure. Note that we know from the modifies clause that the @{term cont} parts are not changed. Also a small side note on the syntax. We use ordinary brackets in the postcondition of the modifies clause, and also the state components do not carry the acute, because we explicitly note the state @{term t} here. \ lemma (in append_impl) append_modifies: shows "\\. \\\<^bsub>/UNIV\<^esub> {\} \p :== PROC append(\p,\q) {t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcRec1) apply (vcg spec=modifies) done text \We tell the verification condition generator to use only the modifies clauses and not to search for functional specifications by the parameter \spec=modifies\. It also tries to solve the verification conditions automatically. Again it is crucial to name the lemma with this naming scheme, since the verfication condition generator searches for these names. \ text \The modifies clause is equal to a state update specification of the following form. \ lemma (in append_impl) shows "{t. t may_only_modify_globals Z in [next]} = {t. \next. globals t=update id id next_' (K_statefun next) (globals Z)}" apply (unfold mex_def meq_def) apply simp done text \Now that we have proven the frame-condition, it is available within the locale \append_impl\ and the \vcg\ exploits it.\ lemma (in append_impl) shows "\\ \\cont=c\ \p :== CALL append(Null,Null) \\cont=c\" proof - note append_spec = append_spec1 show ?thesis apply vcg txt \@{subgoals [display]}\ txt \With a modifies clause present we know that no change to @{term cont} has occurred. \ by simp qed text \ Of course we could add the modifies clause to the functional specification as well. But separating both has the advantage that we split up the verification work. We can make use of the modifies clause before we apply the functional specification in a fully automatic fashion. \ text \ To prove that a procedure respects the modifies clause, we only need the modifies clauses of the procedures called in the body. We do not need the functional specifications. So we can always prove the modifies clause without functional specifications, but we may need the modifies clause to prove the functional specifications. So usually the modifies clause is proved before the proof of the functional specification, so that it can already be used by the verification condition generator. \ subsection \Total Correctness\ text \When proving total correctness the additional proof burden to the user is to come up with a well-founded relation and to prove that certain states get smaller according to this relation. Proving that a relation is well-founded can be quite hard. But fortunately there are ways to construct and stick together relations so that they are well-founded by construction. This infrastructure is already present in Isabelle/HOL. For example, @{term "measure f"} is always well-founded; the lexicographic product of two well-founded relations is again well-founded and the inverse image construction @{term "inv_image"} of a well-founded relation is again well-founded. The constructions are best explained by some equations: @{thm in_measure_iff [no_vars]}\\ @{thm in_lex_iff [no_vars]}\\ @{thm in_inv_image_iff [no_vars]} Another useful construction is \<*mlex*>\ which is a combination of a measure and a lexicographic product: @{thm in_mlex_iff [no_vars]}\\ In contrast to the lexicographic product it does not construct a product type. The state may either decrease according to the measure function @{term f} or the measure stays the same and the state decreases because of the relation @{term r}. Lets look at a loop: \ lemma (in vars) "\\\<^sub>t \\M = 0 \ \S = 0\ WHILE \M \ a INV \\S = \M * b \ \M \ a\ VAR MEASURE a - \M DO \S :== \S + b;; \M :== \M + 1 OD \\S = a * b\" apply vcg txt \@{subgoals [display]} The first conjunct of the second subgoal is the proof obligation that the variant decreases in the loop body. \ by auto text \The variant annotation is preceded by \VAR\. The capital \MEASURE\ is a shorthand for \measure (\s. a - \<^bsup>s\<^esup>M)\. Analogous there is a capital \<*MLEX*>\. \ lemma (in Fac_impl) Fac_spec': shows "\\. \\\<^sub>t {\} \R :== PROC Fac(\N) \\R = fac \<^bsup>\\<^esup>N\" apply (hoare_rule HoareTotal.ProcRec1 [where r="measure (\(s,p). \<^bsup>s\<^esup>N)"]) txt \In case of the factorial the parameter @{term N} decreases in every call. This is easily expressed by the measure function. Note that the well-founded relation for recursive procedures is formally defined on tuples containing the state space and the procedure name. \ txt \@{subgoals [display]} The initial call to the factorial is in state @{term "\"}. Note that in the precondition @{term "{\} \ {\'}"}, @{term "\'"} stems from the lemma we want to prove and @{term "\"} stems from the recursion rule for total correctness. Both are synonym for the initial state. To use the assumption in the Hoare context we have to show that the call to the factorial is invoked on a smaller @{term N} compared to the initial \\<^bsup>\\<^esup>N\. \ apply vcg txt \@{subgoals [display]} The tribute to termination is that we have to show \N - 1 < N\ in case of the recursive call. \ by simp lemma (in append_impl) append_spec2: shows "\\ Ps Qs. \\\<^sub>t \\. List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {}\ \p :== PROC append(\p,\q) \List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoareTotal.ProcRec1 [where r="measure (\(s,p). length (list \<^bsup>s\<^esup>p \<^bsup>s\<^esup>next))"]) txt \In case of the append function the length of the list referenced by @{term p} decreases in every recursive call. \ txt \@{subgoals [margin=80,display]}\ apply vcg apply (fastforce simp add: List_list) done text \ In case of the lists above, we have used a relational list abstraction @{term List} to construct the HOL lists @{term Ps} and @{term Qs} for the pre- and postcondition. To supply a proper measure function we use a functional abstraction @{term list}. The functional abstraction can be defined by means of the relational list abstraction, since the lists are already uniquely determined by the relational abstraction: @{thm islist_def [no_vars]}\\ @{thm list_def [no_vars]} \isacommand{lemma} @{thm List_conv_islist_list [no_vars]} \ text \ -The next contrived example is taken from \<^cite>\"Homeier-95-vcg"\, to illustrate +The next contrived example is taken from \cite{Homeier-95-vcg}, to illustrate a more complex termination criterion for mutually recursive procedures. The procedures do not calculate anything useful. \ procedures pedal(N::nat,M::nat) "IF 0 < \N THEN IF 0 < \M THEN CALL coast(\N- 1,\M- 1) FI;; CALL pedal(\N- 1,\M) FI" and coast(N::nat,M::nat) "CALL pedal(\N,\M);; IF 0 < \M THEN CALL coast(\N,\M- 1) FI" text \ In the recursive calls in procedure \pedal\ the first argument always decreases. In the body of \coast\ in the recursive call of \coast\ the second argument decreases, but in the call to \pedal\ no argument decreases. Therefore an relation only on the state space is insufficient. We have to take the procedure names into account, too. We consider the procedure \coast\ to be ``bigger'' than \pedal\ when we construct a well-founded relation on the product of state space and procedure names. \ ML \ML_Thms.bind_thm ("HoareTotal_ProcRec2", Hoare.gen_proc_rec @{context} Hoare.Total 2)\ text \ We provide the ML function {\tt gen\_proc\_rec} to automatically derive a convenient rule for recursion for a given number of mutually recursive procedures. \ lemma (in pedal_coast_clique) shows "(\\. \\\<^sub>t {\} PROC pedal(\N,\M) UNIV) \ (\\. \\\<^sub>t {\} PROC coast(\N,\M) UNIV)" apply (hoare_rule HoareTotal_ProcRec2 [where r= "((\(s,p). \<^bsup>s\<^esup>N) <*mlex*> (\(s,p). \<^bsup>s\<^esup>M) <*mlex*> measure (\(s,p). if p = coast_'proc then 1 else 0))"]) txt \We can directly express the termination condition described above with the \<*mlex*>\ construction. Either state component \N\ decreases, or it stays the same and \M\ decreases or this also stays the same, but then the procedure name has to decrease.\ txt \@{subgoals [margin=80,display]}\ apply simp_all txt \@{subgoals [margin=75,display]}\ by (vcg,simp)+ text \We can achieve the same effect without \<*mlex*>\ by using the ordinary lexicographic product \<*lex*>\, \inv_image\ and \measure\ \ lemma (in pedal_coast_clique) shows "(\\. \\\<^sub>t {\} PROC pedal(\N,\M) UNIV) \ (\\. \\\<^sub>t {\} PROC coast(\N,\M) UNIV)" apply (hoare_rule HoareTotal_ProcRec2 [where r= "inv_image (measure (\m. m) <*lex*> measure (\m. m) <*lex*> measure (\p. if p = coast_'proc then 1 else 0)) (\(s,p). (\<^bsup>s\<^esup>N,\<^bsup>s\<^esup>M,p))"]) txt \With the lexicographic product we construct a well-founded relation on triples of type @{typ "(nat\nat\string)"}. With @{term inv_image} we project the components out of the state-space and the procedure names to this triple. \ txt \@{subgoals [margin=75,display]}\ apply simp_all - by (vcg,force)+ +by (vcg,simp)+ + text \By doing some arithmetic we can express the termination condition with a single measure function. \ lemma (in pedal_coast_clique) shows "(\\. \\\<^sub>t {\} PROC pedal(\N,\M) UNIV) \ (\\. \\\<^sub>t {\} PROC coast(\N,\M) UNIV)" apply(hoare_rule HoareTotal_ProcRec2 [where r= "measure (\(s,p). \<^bsup>s\<^esup>N + \<^bsup>s\<^esup>M + (if p = coast_'proc then 1 else 0))"]) apply simp_all txt \@{subgoals [margin=75,display]}\ by (vcg,simp,arith?)+ subsection \Guards\ text (in vars) \The purpose of a guard is to guard the {\bf (sub-) expressions} of a statement against runtime faults. Typical runtime faults are array bound violations, dereferencing null pointers or arithmetical overflow. Guards make the potential runtime faults explicit, since the expressions themselves never ``fail'' because they are ordinary HOL expressions. To relieve the user from typing in lots of standard guards for every subexpression, we supply some input syntax for the common language constructs that automatically generate the guards. For example the guarded assignment \\M :==\<^sub>g (\M + 1) div \N\ gets expanded to guarded command @{term "\M :==\<^sub>g (\M + 1) div \N"}. Here @{term "in_range"} is uninterpreted by now. \ lemma (in vars) "\\\True\ \M :==\<^sub>g (\M + 1) div \N \True\" apply vcg txt \@{subgoals}\ oops text \ The user can supply on (overloaded) definition of \in_range\ to fit to his needs. Currently guards are generated for: \begin{itemize} \item overflow and underflow of numbers (\in_range\). For subtraction of natural numbers \a - b\ the guard \b \ a\ is generated instead of \in_range\ to guard against underflows. \item division by \0\ \item dereferencing of @{term Null} pointers \item array bound violations \end{itemize} Following (input) variants of guarded statements are available: \begin{itemize} \item Assignment: \\ :==\<^sub>g \\ \item If: \IF\<^sub>g \\ \item While: \WHILE\<^sub>g \\ \item Call: \CALL\<^sub>g \\ or \\ :== CALL\<^sub>g \\ \end{itemize} \ subsection \Miscellaneous Techniques\ subsubsection \Modifies Clause\ text \We look at some issues regarding the modifies clause with the example of insertion sort for heap lists. \ primrec sorted:: "('a \ 'a \ bool) \ 'a list \ bool" where "sorted le [] = True" | "sorted le (x#xs) = ((\y\set xs. le x y) \ sorted le xs)" procedures (imports globals_heap) insert(r::ref,p::ref | p::ref) "IF \r=Null THEN SKIP ELSE IF \p=Null THEN \p :== \r;; \p\\next :== Null ELSE IF \r\\cont \ \p\\cont THEN \r\\next :== \p;; \p:==\r ELSE \p\\next :== CALL insert(\r,\p\\next) FI FI FI" lemma (in insert_impl) insert_modifies: "\\. \\\<^bsub>/UNIV\<^esub> {\} \p :== PROC insert(\r,\p) {t. t may_only_modify_globals \ in [next]}" by (hoare_rule HoarePartial.ProcRec1) (vcg spec=modifies) lemma (in insert_impl) insert_spec: "\\ Ps . \\ \\. List \p \next Ps \ sorted (\) (map \cont Ps) \ \r \ Null \ \r \ set Ps\ \p :== PROC insert(\r,\p) \\Qs. List \p \next Qs \ sorted (\) (map \<^bsup>\\<^esup>cont Qs) \ set Qs = insert \<^bsup>\\<^esup>r (set Ps) \ (\x. x \ set Qs \ \next x = \<^bsup>\\<^esup>next x)\" (*<*) apply (hoare_rule HoarePartial.ProcRec1) apply vcg apply (intro conjI impI) apply fastforce apply fastforce apply fastforce apply (clarsimp) apply force done (*>*) text \ In the postcondition of the functional specification there is a small but important subtlety. Whenever we talk about the @{term "cont"} part we refer to the one of the pre-state. The reason is that we have separated out the information that @{term "cont"} is not modified by the procedure, to the modifies clause. So whenever we talk about unmodified parts in the postcondition we have to use the pre-state part, or explicitly state an equality in the postcondition. The reason is simple. If the postcondition would talk about \\cont\ instead of \mbox{\\<^bsup>\\<^esup>cont\}, we get a new instance of \cont\ during verification and the postcondition would only state something about this new instance. But as the verification condition generator uses the modifies clause the caller of @{term "insert"} instead still has the old \cont\ after the call. Thats the sense of the modifies clause. So the caller and the specification simply talk about two different things, without being able to relate them (unless an explicit equality is added to the specification). \ subsubsection \Annotations\ text \ Annotations (like loop invariants) are mere syntactic sugar of statements that are used by the \vcg\. Logically a statement with an annotation is equal to the statement without it. Hence annotations can be introduced by the user while building a proof: @{thm [source] HoarePartial.annotateI}: @{thm [mode=Rule] HoarePartial.annotateI [no_vars]} When introducing annotations it can easily happen that these mess around with the nesting of sequential composition. Then after stripping the annotations the resulting statement is no longer syntactically identical to original one, only equivalent modulo associativity of sequential composition. The following rule also deals with this case: @{thm [source] HoarePartial.annotate_normI}: @{thm [mode=Rule] HoarePartial.annotate_normI [no_vars]} \ text_raw \\paragraph{Loop Annotations} \mbox{} \medskip \mbox{} \ procedures (imports globals_heap) insertSort(p::ref| p::ref) where r::ref q::ref in "\r:==Null;; WHILE (\p \ Null) DO \q :== \p;; \p :== \p\\next;; \r :== CALL insert(\q,\r) OD;; \p:==\r" lemma (in insertSort_impl) insertSort_modifies: shows "\\. \\\<^bsub>/UNIV\<^esub> {\} \p :== PROC insertSort(\p) {t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcRec1) apply (vcg spec=modifies) done text \Insertion sort is not implemented recursively here, but with a loop. Note that the while loop is not annotated with an invariant in the procedure definition. The invariant only comes into play during verification. Therefore we annotate the loop first, before we run the \vcg\. \ lemma (in insertSort_impl) insertSort_spec: shows "\\ Ps. \\ \\. List \p \next Ps \ \p :== PROC insertSort(\p) \\Qs. List \p \next Qs \ sorted (\) (map \<^bsup>\\<^esup>cont Qs) \ set Qs = set Ps\" apply (hoare_rule HoarePartial.ProcRec1) apply (hoare_rule anno= "\r :== Null;; WHILE \p \ Null INV \\Qs Rs. List \p \next Qs \ List \r \next Rs \ set Qs \ set Rs = {} \ sorted (\) (map \cont Rs) \ set Qs \ set Rs = set Ps \ \cont = \<^bsup>\\<^esup>cont \ DO \q :== \p;; \p :== \p\\next;; \r :== CALL insert(\q,\r) OD;; \p :== \r" in HoarePartial.annotateI) apply vcg txt \\\\\ (*<*) apply fastforce prefer 2 apply fastforce apply (clarsimp) apply (rule_tac x=ps in exI) apply (intro conjI) apply (rule heap_eq_ListI1) apply assumption apply clarsimp apply (subgoal_tac "x\p \ x \ set Rs") apply auto done (*>*) text \The method \hoare_rule\ automatically solves the side-condition that the annotated program is the same as the original one after stripping the annotations.\ text_raw \\paragraph{Specification Annotations} \mbox{} \medskip \mbox{} \ text \ When verifying a larger block of program text, it might be useful to split up the block and to prove the parts in isolation. This is especially useful to isolate loops. On the level of the Hoare calculus the parts can then be combined with the consequence rule. To automate this process we introduce the derived command @{term specAnno}, which allows to introduce a Hoare tuple (inclusive auxiliary variables) in the program text: @{thm specAnno_def [no_vars]} The whole annotation reduces to the body @{term "c undefined"}. The type of the assertions @{term "P"}, @{term "Q"} and @{term "A"} is @{typ "'a \ 's set"} and the type of command @{term c} is @{typ "'a \ ('s,'p,'f) com"}. All entities formally depend on an auxiliary (logical) variable of type @{typ "'a"}. The body @{term "c"} formally also depends on this variable, since a nested annotation or loop invariant may also depend on this logical variable. But the raw body without annotations does not depend on the logical variable. The logical variable is only used by the verification condition generator. We express this by defining the whole @{term specAnno} to be equivalent with the body applied to an arbitrary variable. The Hoare rule for \specAnno\ is mainly an instance of the consequence rule: @{thm [mode=Rule,mode=ParenStmt] HoarePartial.SpecAnno [no_vars]} The side-condition @{term "\Z. c Z = c undefined"} expresses the intention of body @{term c} explained above: The raw body is independent of the auxiliary variable. This side-condition is solved automatically by the \vcg\. The concrete syntax for this specification annotation is shown in the following example: \ lemma (in vars) "\\ {\} \I :== \M;; ANNO \. \\. \I = \<^bsup>\\<^esup>M\ \M :== \N;; \N :== \I \\M = \<^bsup>\\<^esup>N \ \N = \<^bsup>\\<^esup>I\ \\M = \<^bsup>\\<^esup>N \ \N = \<^bsup>\\<^esup>M\" txt \With the annotation we can name an intermediate state @{term \}. Since the postcondition refers to @{term "\"} we have to link the information about the equivalence of \\<^bsup>\\<^esup>I\ and \\<^bsup>\\<^esup>M\ in the specification in order to be able to derive the postcondition. \ apply vcg_step apply vcg_step txt \@{subgoals [display]}\ txt \The first subgoal is the isolated Hoare tuple. The second one is the side-condition of the consequence rule that allows us to derive the outermost pre/post condition from our inserted specification. \\I = \<^bsup>\\<^esup>M\ is the precondition of the specification, The second conjunct is a simplified version of \\t. \<^bsup>t\<^esup>M = \N \ \<^bsup>t\<^esup>N = \I \ \<^bsup>t\<^esup>M = \<^bsup>\\<^esup>N \ \<^bsup>t\<^esup>N = \<^bsup>\\<^esup>M\ expressing that the postcondition of the specification implies the outermost postcondition. \ apply vcg txt \@{subgoals [display]}\ apply simp apply vcg txt \@{subgoals [display]}\ by simp lemma (in vars) "\\ {\} \I :== \M;; ANNO \. \\. \I = \<^bsup>\\<^esup>M\ \M :== \N;; \N :== \I \\M = \<^bsup>\\<^esup>N \ \N = \<^bsup>\\<^esup>I\ \\M = \<^bsup>\\<^esup>N \ \N = \<^bsup>\\<^esup>M\" apply vcg txt \@{subgoals [display]}\ by simp_all text \Note that \vcg_step\ changes the order of sequential composition, to allow the user to decompose sequences by repeated calls to \vcg_step\, whereas \vcg\ preserves the order. The above example illustrates how we can introduce a new logical state variable @{term "\"}. You can introduce multiple variables by using a tuple: \ lemma (in vars) "\\ {\} \I :== \M;; ANNO (n,i,m). \\I = \<^bsup>\\<^esup>M \ \N=n \ \I=i \ \M=m\ \M :== \N;; \N :== \I \\M = n \ \N = i\ \\M = \<^bsup>\\<^esup>N \ \N = \<^bsup>\\<^esup>M\" apply vcg txt \@{subgoals [display]}\ by simp_all text_raw \\paragraph{Lemma Annotations} \mbox{} \medskip \mbox{} \ text \ The specification annotations described before split the verification into several Hoare triples which result in several subgoals. If we instead want to proof the Hoare triples independently as separate lemmas we can use the \LEMMA\ annotation to plug together the lemmas. It inserts the lemma in the same fashion as the specification annotation. \ lemma (in vars) foo_lemma: "\n m. \\ \\N = n \ \M = m\ \N :== \N + 1;; \M :== \M + 1 \\N = n + 1 \ \M = m + 1\" apply vcg apply simp done lemma (in vars) "\\ \\N = n \ \M = m\ LEMMA foo_lemma \N :== \N + 1;; \M :== \M + 1 END;; \N :== \N + 1 \\N = n + 2 \ \M = m + 1\" apply vcg apply simp done lemma (in vars) "\\ \\N = n \ \M = m\ LEMMA foo_lemma \N :== \N + 1;; \M :== \M + 1 END;; LEMMA foo_lemma \N :== \N + 1;; \M :== \M + 1 END \\N = n + 2 \ \M = m + 2\" apply vcg apply simp done lemma (in vars) "\\ \\N = n \ \M = m\ \N :== \N + 1;; \M :== \M + 1;; \N :== \N + 1;; \M :== \M + 1 \\N = n + 2 \ \M = m + 2\" apply (hoare_rule anno= "LEMMA foo_lemma \N :== \N + 1;; \M :== \M + 1 END;; LEMMA foo_lemma \N :== \N + 1;; \M :== \M + 1 END" in HoarePartial.annotate_normI) apply vcg apply simp done subsubsection \Total Correctness of Nested Loops\ text \ When proving termination of nested loops it is sometimes necessary to express that the loop variable of the outer loop is not modified in the inner loop. To express this one has to fix the value of the outer loop variable before the inner loop and use this value in the invariant of the inner loop. This can be achieved by surrounding the inner while loop with an \ANNO\ specification as explained previously. However, this leads to repeating the invariant of the inner loop three times: in the invariant itself and in the the pre- and postcondition of the \ANNO\ specification. Moreover one has to deal with the additional subgoal introduced by \ANNO\ that expresses how the pre- and postcondition is connected to the invariant. To avoid this extra specification and verification work, we introduce an variant of the annotated while-loop, where one can introduce logical variables by \FIX\. As for the \ANNO\ specification multiple logical variables can be introduced via a tuple (\FIX (a,b,c).\). The Hoare logic rule for the augmented while-loop is a mixture of the invariant rule for loops and the consequence rule for \ANNO\: \begin{center} @{thm [mode=Rule,mode=ParenStmt] HoareTotal.WhileAnnoFix' [no_vars]} \end{center} The first premise expresses that the precondition implies the invariant and that the invariant together with the negated loop condition implies the postcondition. Since both implications may depend on the choice of the auxiliary variable @{term "Z"} these two implications are expressed in a single premise and not in two of them as for the usual while rule. The second premise is the preservation of the invariant by the loop body. And the third premise is the side-condition that the computational part of the body does not depend on the auxiliary variable. Finally the last premise is the well-foundedness of the variant. The last two premises are usually discharged automatically by the verification condition generator. Hence usually two subgoals remain for the user, stemming from the first two premises. The following example illustrates the usage of this rule. The outer loop increments the loop variable @{term "M"} while the inner loop increments @{term "N"}. To discharge the proof obligation for the termination of the outer loop, we need to know that the inner loop does not mess around with @{term "M"}. This is expressed by introducing the logical variable @{term "m"} and fixing the value of @{term "M"} to it. \ lemma (in vars) "\\\<^sub>t \\M=0 \ \N=0\ WHILE (\M < i) INV \\M \ i \ (\M \ 0 \ \N = j) \ \N \ j\ VAR MEASURE (i - \M) DO \N :== 0;; WHILE (\N < j) FIX m. INV \\M=m \ \N \ j\ VAR MEASURE (j - \N) DO \N :== \N + 1 OD;; \M :== \M + 1 OD \\M=i \ (\M\0 \ \N=j)\" apply vcg txt \@{subgoals [display]} The first subgoal is from the precondition to the invariant of the outer loop. The fourth subgoal is from the invariant together with the negated loop condition of the outer loop to the postcondition. The subgoals two and three are from the body of the outer while loop which is mainly the inner while loop. Because we introduce the logical variable @{term "m"} here, the while Rule described above is used instead of the ordinary while Rule. That is why we end up with two subgoals for the inner loop. Subgoal two is from the invariant and the loop condition of the outer loop to the invariant of the inner loop. And at the same time from the invariant of the inner loop to the invariant of the outer loop (together with the proof obligation that the measure of the outer loop decreases). The universal quantified variables @{term "Ma"} and @{term "N"} are the ``fresh'' state variables introduced for the final state of the inner loop. The equality @{term "Ma=M"} is the result of the equality \\M=m\ in the inner invariant. Subgoal three is the preservation of the invariant by the inner loop body (together with the proof obligation that the measure of the inner loop decreases). \ (*<*) apply (simp) apply (simp,arith) apply (simp,arith) done (*>*) subsection \Functional Correctness, Termination and Runtime Faults\ text \ Total correctness of a program with guards conceptually leads to three verification tasks. \begin{itemize} \item functional (partial) correctness \item absence of runtime faults \item termination \end{itemize} In case of a modifies specification the functional correctness part can be solved automatically. But the absence of runtime faults and termination may be non trivial. Fortunately the modifies clause is usually just a helpful companion of another specification that expresses the ``real'' functional behaviour. Therefor the task to prove the absence of runtime faults and termination can be dealt with during the proof of this functional specification. In most cases the absence of runtime faults and termination heavily build on the functional specification parts. So after all there is no reason why we should again prove the absence of runtime faults and termination for the modifies clause. Therefor it suffices to have partial correctness of the modifies clause for a program were all guards are ignored. This leads to the following pattern:\ procedures foo (N::nat|M::nat) "\M :== \M \ \think of body with guards instead\" foo_spec: "\\. \\\<^sub>t (P \) \M :== PROC foo(\N) (Q \)" foo_modifies: "\\. \\\<^bsub>/UNIV\<^esub> {\} \M :== PROC foo(\N) {t. t may_only_modify_globals \ in []}" text \ The verification condition generator can solve those modifies clauses automatically and can use them to simplify calls to \foo\ even in the context of total correctness. \ subsection \Procedures and Locales \label{sec:Locales}\ text \ Verification of a larger program is organised on the granularity of procedures. We proof the procedures in a bottom up fashion. Of course you can also always use Isabelle's dummy proof \sorry\ to prototype your formalisation. So you can write the theory in a bottom up fashion but actually prove the lemmas in any other order. Here are some explanations of handling of locales. In the examples below, consider \proc\<^sub>1\ and \proc\<^sub>2\ to be ``leaf'' procedures, which do not call any other procedure. Procedure \proc\ directly calls \proc\<^sub>1\ and \proc\<^sub>2\. \isacommand{lemma} (\isacommand{in} \proc\<^sub>1_impl\) \proc\<^sub>1_modifies\:\\ \isacommand{shows} \\\ After the proof of \proc\<^sub>1_modifies\, the \isacommand{in} directive stores the lemma in the locale \proc\<^sub>1_impl\. When we later on include \proc\<^sub>1_impl\ or prove another theorem in locale \proc\<^sub>1_impl\ the lemma \proc\<^sub>1_modifies\ will already be available as fact. \isacommand{lemma} (\isacommand{in} \proc\<^sub>1_impl\) \proc\<^sub>1_spec\:\\ \isacommand{shows} \\\ \isacommand{lemma} (\isacommand{in} \proc\<^sub>2_impl\) \proc\<^sub>2_modifies\:\\ \isacommand{shows} \\\ \isacommand{lemma} (\isacommand{in} \proc\<^sub>2_impl\) \proc\<^sub>2_spec\:\\ \isacommand{shows} \\\ \isacommand{lemma} (\isacommand{in} \proc_impl\) \proc_modifies\:\\ \isacommand{shows} \\\ Note that we do not explicitly include anything about \proc\<^sub>1\ or \proc\<^sub>2\ here. This is handled automatically. When defining an \impl\-locale it imports all \impl\-locales of procedures that are called in the body. In case of \proc_impl\ this means, that \proc\<^sub>1_impl\ and \proc\<^sub>2_impl\ are imported. This has the neat effect that all theorems that are proven in \proc\<^sub>1_impl\ and \proc\<^sub>2_impl\ are also present in \proc_impl\. \isacommand{lemma} (\isacommand{in} \proc_impl\) \proc_spec\:\\ \isacommand{shows} \\\ As we have seen in this example you only have to prove a procedure in its own \impl\ locale. You do not have to include any other locale. \ subsection \Records \label{sec:records}\ text \ Before @{term "statespaces"} where introduced the state was represented as a @{term "record"}. This is still supported. Compared to the flexibility of statespaces there are some drawbacks in particular with respect to modularity. Even names of local variables and parameters are globally visible and records can only be extended in a linear fashion, whereas statespaces also allow multiple inheritance. The usage of records is quite similar to the usage of statespaces. We repeat the example of an append function for heap lists. First we define the global components. Again the appearance of the prefix `globals' is mandatory. This is the way the syntax layer distinguishes local and global variables. \ record globals_list = next_' :: "ref \ ref" cont_' :: "ref \ nat" text \The local variables also have to be defined as a record before the actual definition of the procedure. The parent record \state\ defines a generic @{term "globals"} field as a place-holder for the record of global components. In contrast to the statespace approach there is no single @{term "locals"} slot. The local components are just added to the record. \ record 'g list_vars = "'g state" + p_' :: "ref" q_' :: "ref" r_' :: "ref" root_' :: "ref" tmp_' :: "ref" text \Since the parameters and local variables are determined by the record, there are no type annotations or definitions of local variables while defining a procedure. \ procedures append'(p,q|p) = "IF \p=Null THEN \p :== \q ELSE \p \\next:== CALL append'(\p\\next,\q) FI" text \As in the statespace approach, a locale called \append'_impl\ is created. Note that we do not give any explicit information which global or local state-record to use. Since the records are already defined we rely on Isabelle's type inference. Dealing with the locale is analogous to the case with statespaces. \ lemma (in append'_impl) append'_modifies: shows "\\. \\ {\} \p :== PROC append'(\p,\q) {t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcRec1) apply (vcg spec=modifies) done lemma (in append'_impl) append'_spec: shows "\\ Ps Qs. \\ \\. List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {}\ \p :== PROC append'(\p,\q) \List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoarePartial.ProcRec1) apply vcg apply fastforce done text \ However, in some corner cases the inferred state type in a procedure definition can be too general which raises problems when attempting to proof a suitable specifications in the locale. Consider for example the simple procedure body @{term "\p :== NULL"} for a procedure \init\. \ procedures init (|p) = "\p:== Null" text \ Here Isabelle can only infer the local variable record. Since no reference to any global variable is made the type fixed for the global variables (in the locale \init'_impl\) is a type variable say @{typ "'g"} and not a @{term "globals_list"} record. Any specification mentioning @{term "next"} or @{term "cont"} restricts the state type and cannot be added to the locale \init_impl\. Hence we have to restrict the body @{term "\p :== NULL"} in the first place by adding a typing annotation: \ procedures init' (|p) = "\p:== Null::(('a globals_list_scheme, 'b) list_vars_scheme, char list, 'c) com" subsubsection \Extending State Spaces\ text \ The records in Isabelle are -extensible \<^cite>\"Nipkow-02-hol" and "NaraschewskiW-TPHOLs98"\. In principle this can be exploited +extensible \cite{Nipkow-02-hol,NaraschewskiW-TPHOLs98}. In principle this can be exploited during verification. The state space can be extended while we we add procedures. But there is one major drawback: \begin{itemize} \item records can only be extended in a linear fashion (there is no multiple inheritance) \end{itemize} You can extend both the main state record as well as the record for the global variables. \ subsubsection \Mapping Variables to Record Fields\ text \ Generally the state space (global and local variables) is flat and all components are accessible from everywhere. Locality or globality of variables is achieved by the proper \init\ and \return\/\result\ functions in procedure calls. What is the best way to map programming language variables to the state records? One way is to disambiguate all names, by using the procedure names as prefix or the structure names for heap components. This leads to long names and lots of record components. But for local variables this is not necessary, since variable @{term i} of procedure @{term A} and variable @{term "i"} of procedure @{term B} can be mapped to the same record component, without any harm, provided they have the same logical type. Therefor for local variables it is preferable to map them per type. You only have to distinguish a variable with the same name if they have a different type. Note that all pointers just have logical type \ref\. So you even do not have to distinguish between a pointer \p\ to a integer and a pointer \p\ to a list. For global components (global variables and heap structures) you have to disambiguate the name. But hopefully the field names of structures have different names anyway. Also note that there is no notion of hiding of a global component by a local one in the logic. You have to disambiguate global and local names! As the names of the components show up in the specifications and the proof obligations, names are even more important as for programming. Try to find meaningful and short names, to avoid cluttering up your reasoning. \ (*<*) text \ in locales, includes, spec or impl? Names: per type not per procedure\ downgrading total to partial\ \ (*>*) text \\ (*<*) end (*>*) diff --git a/thys/Simpl/Vcg.thy b/thys/Simpl/Vcg.thy --- a/thys/Simpl/Vcg.thy +++ b/thys/Simpl/Vcg.thy @@ -1,663 +1,656 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: Vcg.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. +Copyright (c) 2022 Apple Inc. All rights reserved. -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA *) section \Facilitating the Hoare Logic\ theory Vcg imports StateSpace "HOL-Statespace.StateSpaceLocale" Generalise keywords "procedures" "hoarestate" :: thy_defn begin axiomatization NoBody::"('s,'p,'f) com" ML_file \hoare.ML\ method_setup hoare = "Hoare.hoare" "raw verification condition generator for Hoare Logic" method_setup hoare_raw = "Hoare.hoare_raw" "even more raw verification condition generator for Hoare Logic" method_setup vcg = "Hoare.vcg" "verification condition generator for Hoare Logic" method_setup vcg_step = "Hoare.vcg_step" "single verification condition generation step with light simplification" method_setup hoare_rule = "Hoare.hoare_rule" "apply single hoare rule and solve certain sideconditions" text \Variables of the programming language are represented as components of a record. To avoid cluttering up the namespace of Isabelle with lots of typical variable names, we append a unusual suffix at the end of each name by parsing \ definition list_multsel:: "'a list \ nat list \ 'a list" (infixl "!!" 100) where "xs !! ns = map (nth xs) ns" definition list_multupd:: "'a list \ nat list \ 'a list \ 'a list" where "list_multupd xs ns ys = foldl (\xs (n,v). xs[n:=v]) xs (zip ns ys)" nonterminal lmupdbinds and lmupdbind syntax \ \multiple list update\ "_lmupdbind":: "['a, 'a] => lmupdbind" ("(2_ [:=]/ _)") "" :: "lmupdbind => lmupdbinds" ("_") "_lmupdbinds" :: "[lmupdbind, lmupdbinds] => lmupdbinds" ("_,/ _") "_LMUpdate" :: "['a, lmupdbinds] => 'a" ("_/[(_)]" [900,0] 900) translations "_LMUpdate xs (_lmupdbinds b bs)" == "_LMUpdate (_LMUpdate xs b) bs" "xs[is[:=]ys]" == "CONST list_multupd xs is ys" subsection \Some Fancy Syntax\ (* priority guidline: * If command should be atomic for the guard it must have at least priority 21. *) text \reverse application\ definition rapp:: "'a \ ('a \ 'b) \ 'b" (infixr "|>" 60) where "rapp x f = f x" nonterminal newinit and newinits and locinit and locinits and switchcase and switchcases and grds and grd and bdy and basics and basic and basicblock notation Skip ("SKIP") and Throw ("THROW") syntax "_raise":: "'c \ 'c \ ('a,'b,'f) com" ("(RAISE _ :==/ _)" [30, 30] 23) "_seq"::"('s,'p,'f) com \ ('s,'p,'f) com \ ('s,'p,'f) com" ("_;;/ _" [20, 21] 20) "_guarantee" :: "'s set \ grd" ("_\" [1000] 1000) "_guaranteeStrip":: "'s set \ grd" ("_#" [1000] 1000) "_grd" :: "'s set \ grd" ("_" [1000] 1000) "_last_grd" :: "grd \ grds" ("_" 1000) "_grds" :: "[grd, grds] \ grds" ("_,/ _" [999,1000] 1000) "_guards" :: "grds \ ('s,'p,'f) com \ ('s,'p,'f) com" ("(_/\ _)" [60, 21] 23) "_quote" :: "'b => ('a => 'b)" "_antiquoteCur0" :: "('a => 'b) => 'b" ("\_" [1000] 1000) "_antiquoteCur" :: "('a => 'b) => 'b" "_antiquoteOld0" :: "('a => 'b) => 'a => 'b" ("\<^bsup>_\<^esup>_" [1000,1000] 1000) "_antiquoteOld" :: "('a => 'b) => 'a => 'b" "_Assert" :: "'a => 'a set" ("(\_\)" [0] 1000) "_AssertState" :: "idt \ 'a => 'a set" ("(\_. _\)" [1000,0] 1000) "_Assign" :: "'b => 'b => ('a,'p,'f) com" ("(_ :==/ _)" [30, 30] 23) "_Init" :: "ident \ 'c \ 'b \ ('a,'p,'f) com" ("(\_ :==\<^bsub>_\<^esub>/ _)" [30,1000, 30] 23) "_GuardedAssign":: "'b => 'b => ('a,'p,'f) com" ("(_ :==\<^sub>g/ _)" [30, 30] 23) "_newinit" :: "[ident,'a] \ newinit" ("(2\_ :==/ _)") "" :: "newinit \ newinits" ("_") "_newinits" :: "[newinit, newinits] \ newinits" ("_,/ _") "_New" :: "['a, 'b, newinits] \ ('a,'b,'f) com" ("(_ :==/(2 NEW _/ [_]))" [30, 65, 0] 23) "_GuardedNew" :: "['a, 'b, newinits] \ ('a,'b,'f) com" ("(_ :==\<^sub>g/(2 NEW _/ [_]))" [30, 65, 0] 23) "_NNew" :: "['a, 'b, newinits] \ ('a,'b,'f) com" ("(_ :==/(2 NNEW _/ [_]))" [30, 65, 0] 23) "_GuardedNNew" :: "['a, 'b, newinits] \ ('a,'b,'f) com" ("(_ :==\<^sub>g/(2 NNEW _/ [_]))" [30, 65, 0] 23) "_Cond" :: "'a bexp => ('a,'p,'f) com => ('a,'p,'f) com => ('a,'p,'f) com" ("(0IF (_)/ (2THEN/ _)/ (2ELSE _)/ FI)" [0, 0, 0] 71) "_Cond_no_else":: "'a bexp => ('a,'p,'f) com => ('a,'p,'f) com" ("(0IF (_)/ (2THEN/ _)/ FI)" [0, 0] 71) "_GuardedCond" :: "'a bexp => ('a,'p,'f) com => ('a,'p,'f) com => ('a,'p,'f) com" ("(0IF\<^sub>g (_)/ (2THEN _)/ (2ELSE _)/ FI)" [0, 0, 0] 71) "_GuardedCond_no_else":: "'a bexp => ('a,'p,'f) com => ('a,'p,'f) com" ("(0IF\<^sub>g (_)/ (2THEN _)/ FI)" [0, 0] 71) "_While_inv_var" :: "'a bexp => 'a assn \ ('a \ 'a) set \ bdy \ ('a,'p,'f) com" ("(0WHILE (_)/ INV (_)/ VAR (_) /_)" [25, 0, 0, 81] 71) "_WhileFix_inv_var" :: "'a bexp => pttrn \ ('z \ 'a assn) \ ('z \ ('a \ 'a) set) \ bdy \ ('a,'p,'f) com" ("(0WHILE (_)/ FIX _./ INV (_)/ VAR (_) /_)" [25, 0, 0, 0, 81] 71) "_WhileFix_inv" :: "'a bexp => pttrn \ ('z \ 'a assn) \ bdy \ ('a,'p,'f) com" ("(0WHILE (_)/ FIX _./ INV (_) /_)" [25, 0, 0, 81] 71) "_GuardedWhileFix_inv_var" :: "'a bexp => pttrn \ ('z \ 'a assn) \ ('z \ ('a \ 'a) set) \ bdy \ ('a,'p,'f) com" ("(0WHILE\<^sub>g (_)/ FIX _./ INV (_)/ VAR (_) /_)" [25, 0, 0, 0, 81] 71) "_GuardedWhileFix_inv_var_hook" :: "'a bexp \ ('z \ 'a assn) \ ('z \ ('a \ 'a) set) \ bdy \ ('a,'p,'f) com" "_GuardedWhileFix_inv" :: "'a bexp => pttrn \ ('z \ 'a assn) \ bdy \ ('a,'p,'f) com" ("(0WHILE\<^sub>g (_)/ FIX _./ INV (_)/_)" [25, 0, 0, 81] 71) "_GuardedWhile_inv_var":: "'a bexp => 'a assn \ ('a \ 'a) set \ bdy \ ('a,'p,'f) com" ("(0WHILE\<^sub>g (_)/ INV (_)/ VAR (_) /_)" [25, 0, 0, 81] 71) "_While_inv" :: "'a bexp => 'a assn => bdy => ('a,'p,'f) com" ("(0WHILE (_)/ INV (_) /_)" [25, 0, 81] 71) "_GuardedWhile_inv" :: "'a bexp => 'a assn => ('a,'p,'f) com => ('a,'p,'f) com" ("(0WHILE\<^sub>g (_)/ INV (_) /_)" [25, 0, 81] 71) "_While" :: "'a bexp => bdy => ('a,'p,'f) com" ("(0WHILE (_) /_)" [25, 81] 71) "_GuardedWhile" :: "'a bexp => bdy => ('a,'p,'f) com" ("(0WHILE\<^sub>g (_) /_)" [25, 81] 71) "_While_guard" :: "grds => 'a bexp => bdy => ('a,'p,'f) com" ("(0WHILE (_/\ (1_)) /_)" [1000,25,81] 71) "_While_guard_inv":: "grds \'a bexp\'a assn\bdy \ ('a,'p,'f) com" ("(0WHILE (_/\ (1_)) INV (_) /_)" [1000,25,0,81] 71) "_While_guard_inv_var":: "grds \'a bexp\'a assn\('a\'a) set \bdy \ ('a,'p,'f) com" ("(0WHILE (_/\ (1_)) INV (_)/ VAR (_) /_)" [1000,25,0,0,81] 71) "_WhileFix_guard_inv_var":: "grds \'a bexp\pttrn\('z\'a assn)\('z\('a\'a) set) \bdy \ ('a,'p,'f) com" ("(0WHILE (_/\ (1_)) FIX _./ INV (_)/ VAR (_) /_)" [1000,25,0,0,0,81] 71) "_WhileFix_guard_inv":: "grds \'a bexp\pttrn\('z\'a assn) \bdy \ ('a,'p,'f) com" ("(0WHILE (_/\ (1_)) FIX _./ INV (_)/_)" [1000,25,0,0,81] 71) "_Try_Catch":: "('a,'p,'f) com \('a,'p,'f) com \ ('a,'p,'f) com" ("(0TRY (_)/ (2CATCH _)/ END)" [0,0] 71) "_DoPre" :: "('a,'p,'f) com \ ('a,'p,'f) com" "_Do" :: "('a,'p,'f) com \ bdy" ("(2DO/ (_)) /OD" [0] 1000) "_Lab":: "'a bexp \ ('a,'p,'f) com \ bdy" ("_\/_" [1000,71] 81) "":: "bdy \ ('a,'p,'f) com" ("_") "_Spec":: "pttrn \ 's set \ ('s,'p,'f) com \ 's set \ 's set \ ('s,'p,'f) com" ("(ANNO _. _/ (_)/ _,/_)" [0,1000,20,1000,1000] 60) "_SpecNoAbrupt":: "pttrn \ 's set \ ('s,'p,'f) com \ 's set \ ('s,'p,'f) com" ("(ANNO _. _/ (_)/ _)" [0,1000,20,1000] 60) "_LemAnno":: "'n \ ('s,'p,'f) com \ ('s,'p,'f) com" ("(0 LEMMA (_)/ _ END)" [1000,0] 71) "_locnoinit" :: "ident \ locinit" ("\_") "_locinit" :: "[ident,'a] \ locinit" ("(2\_ :==/ _)") "" :: "locinit \ locinits" ("_") "_locinits" :: "[locinit, locinits] \ locinits" ("_,/ _") "_Loc":: "[locinits,('s,'p,'f) com] \ ('s,'p,'f) com" ("(2 LOC _;;/ (_) COL)" [0,0] 71) "_Switch":: "('s \ 'v) \ switchcases \ ('s,'p,'f) com" ("(0 SWITCH (_)/ _ END)" [22,0] 71) "_switchcase":: "'v set \ ('s,'p,'f) com \ switchcase" ("_\/ _" ) "_switchcasesSingle" :: "switchcase \ switchcases" ("_") "_switchcasesCons":: "switchcase \ switchcases \ switchcases" ("_/ | _") "_Basic":: "basicblock \ ('s,'p,'f) com" ("(0BASIC/ (_)/ END)" [22] 71) "_BasicBlock":: "basics \ basicblock" ("_") "_BAssign" :: "'b => 'b => basic" ("(_ :==/ _)" [30, 30] 23) "" :: "basic \ basics" ("_") "_basics" :: "[basic, basics] \ basics" ("_,/ _") syntax (ASCII) "_Assert" :: "'a => 'a set" ("({|_|})" [0] 1000) "_AssertState" :: "idt \ 'a \ 'a set" ("({|_. _|})" [1000,0] 1000) "_While_guard" :: "grds => 'a bexp => bdy \ ('a,'p,'f) com" ("(0WHILE (_|-> /_) /_)" [0,0,1000] 71) "_While_guard_inv":: "grds\'a bexp\'a assn\bdy \ ('a,'p,'f) com" ("(0WHILE (_|-> /_) INV (_) /_)" [0,0,0,1000] 71) "_guards" :: "grds \ ('s,'p,'f) com \ ('s,'p,'f) com" ("(_|->_ )" [60, 21] 23) syntax (output) "_hidden_grds" :: "grds" ("\") translations "_Do c" => "c" "b\ c" => "CONST condCatch c b SKIP" "b\ (_DoPre c)" <= "CONST condCatch c b SKIP" "l\ (CONST whileAnnoG gs b I V c)" <= "l\ (_DoPre (CONST whileAnnoG gs b I V c))" "l\ (CONST whileAnno b I V c)" <= "l\ (_DoPre (CONST whileAnno b I V c))" "CONST condCatch c b SKIP" <= "(_DoPre (CONST condCatch c b SKIP))" "_Do c" <= "_DoPre c" "c;; d" == "CONST Seq c d" "_guarantee g" => "(CONST True, g)" "_guaranteeStrip g" == "CONST guaranteeStripPair (CONST True) g" "_grd g" => "(CONST False, g)" "_grds g gs" => "g#gs" "_last_grd g" => "[g]" "_guards gs c" == "CONST guards gs c" "{|s. P|}" == "{|_antiquoteCur((=) s) \ P |}" "{|b|}" => "CONST Collect (_quote b)" "IF b THEN c1 ELSE c2 FI" => "CONST Cond {|b|} c1 c2" "IF b THEN c1 FI" == "IF b THEN c1 ELSE SKIP FI" "IF\<^sub>g b THEN c1 FI" == "IF\<^sub>g b THEN c1 ELSE SKIP FI" "_While_inv_var b I V c" => "CONST whileAnno {|b|} I V c" "_While_inv_var b I V (_DoPre c)" <= "CONST whileAnno {|b|} I V c" "_While_inv b I c" == "_While_inv_var b I (CONST undefined) c" "_While b c" == "_While_inv b {|CONST undefined|} c" "_While_guard_inv_var gs b I V c" => "CONST whileAnnoG gs {|b|} I V c" (* "_While_guard_inv_var gs b I V (_DoPre c)" <= "CONST whileAnnoG gs {|b|} I V c"*) "_While_guard_inv gs b I c" == "_While_guard_inv_var gs b I (CONST undefined) c" "_While_guard gs b c" == "_While_guard_inv gs b {|CONST undefined|} c" "_GuardedWhile_inv b I c" == "_GuardedWhile_inv_var b I (CONST undefined) c" "_GuardedWhile b c" == "_GuardedWhile_inv b {|CONST undefined|} c" (* "\<^bsup>s\<^esup>A" => "A s"*) "TRY c1 CATCH c2 END" == "CONST Catch c1 c2" "ANNO s. P c Q,A" => "CONST specAnno (\s. P) (\s. c) (\s. Q) (\s. A)" "ANNO s. P c Q" == "ANNO s. P c Q,{}" "_WhileFix_inv_var b z I V c" => "CONST whileAnnoFix {|b|} (\z. I) (\z. V) (\z. c)" "_WhileFix_inv_var b z I V (_DoPre c)" <= "_WhileFix_inv_var {|b|} z I V c" "_WhileFix_inv b z I c" == "_WhileFix_inv_var b z I (CONST undefined) c" "_GuardedWhileFix_inv b z I c" == "_GuardedWhileFix_inv_var b z I (CONST undefined) c" "_GuardedWhileFix_inv_var b z I V c" => "_GuardedWhileFix_inv_var_hook {|b|} (\z. I) (\z. V) (\z. c)" "_WhileFix_guard_inv_var gs b z I V c" => "CONST whileAnnoGFix gs {|b|} (\z. I) (\z. V) (\z. c)" "_WhileFix_guard_inv_var gs b z I V (_DoPre c)" <= "_WhileFix_guard_inv_var gs {|b|} z I V c" "_WhileFix_guard_inv gs b z I c" == "_WhileFix_guard_inv_var gs b z I (CONST undefined) c" "LEMMA x c END" == "CONST lem x c" translations "(_switchcase V c)" => "(V,c)" "(_switchcasesSingle b)" => "[b]" "(_switchcasesCons b bs)" => "CONST Cons b bs" "(_Switch v vs)" => "CONST switch (_quote v) vs" parse_ast_translation \ let fun tr c asts = Ast.mk_appl (Ast.Constant c) (map Ast.strip_positions asts) in [(@{syntax_const "_antiquoteCur0"}, K (tr @{syntax_const "_antiquoteCur"})), (@{syntax_const "_antiquoteOld0"}, K (tr @{syntax_const "_antiquoteOld"}))] end \ print_ast_translation \ let fun tr c asts = Ast.mk_appl (Ast.Constant c) asts in [(@{syntax_const "_antiquoteCur"}, K (tr @{syntax_const "_antiquoteCur0"})), (@{syntax_const "_antiquoteOld"}, K (tr @{syntax_const "_antiquoteOld0"}))] end \ print_ast_translation \ let fun dest_abs (Ast.Appl [Ast.Constant @{syntax_const "_abs"}, x, t]) = (x, t) | dest_abs _ = raise Match; fun spec_tr' [P, c, Q, A] = let val (x',P') = dest_abs P; val (_ ,c') = dest_abs c; val (_ ,Q') = dest_abs Q; val (_ ,A') = dest_abs A; in if (A' = Ast.Constant @{const_syntax bot}) then Ast.mk_appl (Ast.Constant @{syntax_const "_SpecNoAbrupt"}) [x', P', c', Q'] else Ast.mk_appl (Ast.Constant @{syntax_const "_Spec"}) [x', P', c', Q', A'] end; fun whileAnnoFix_tr' [b, I, V, c] = let val (x',I') = dest_abs I; val (_ ,V') = dest_abs V; val (_ ,c') = dest_abs c; in Ast.mk_appl (Ast.Constant @{syntax_const "_WhileFix_inv_var"}) [b, x', I', V', c'] end; in [(@{const_syntax specAnno}, K spec_tr'), (@{const_syntax whileAnnoFix}, K whileAnnoFix_tr')] end \ syntax "_faccess" :: "'ref \ ('ref \ 'v) \ 'v" ("_\_" [65,1000] 100) syntax (ASCII) "_faccess" :: "'ref \ ('ref \ 'v) \ 'v" ("_->_" [65,1000] 100) translations "p\f" => "f p" "g\(_antiquoteCur f)" <= "_antiquoteCur f g" nonterminal par and pars and actuals syntax "_par" :: "'a \ par" ("_") "" :: "par \ pars" ("_") "_pars" :: "[par,pars] \ pars" ("_,/_") "_actuals" :: "pars \ actuals" ("'(_')") "_actuals_empty" :: "actuals" ("'(')") syntax "_Call" :: "'p \ actuals \ (('a,string,'f) com)" ("CALL __" [1000,1000] 21) "_GuardedCall" :: "'p \ actuals \ (('a,string,'f) com)" ("CALL\<^sub>g __" [1000,1000] 21) "_CallAss":: "'a \ 'p \ actuals \ (('a,string,'f) com)" ("_ :== CALL __" [30,1000,1000] 21) + "_Call_exn" :: "'p \ actuals \ (('a,string,'f) com)" ("CALL\<^sub>e __" [1000,1000] 21) + "_CallAss_exn":: "'a \ 'p \ actuals \ (('a,string,'f) com)" + ("_ :== CALL\<^sub>e __" [30,1000,1000] 21) "_Proc" :: "'p \ actuals \ (('a,string,'f) com)" ("PROC __" 21) "_ProcAss":: "'a \ 'p \ actuals \ (('a,string,'f) com)" ("_ :== PROC __" [30,1000,1000] 21) "_GuardedCallAss":: "'a \ 'p \ actuals \ (('a,string,'f) com)" ("_ :== CALL\<^sub>g __" [30,1000,1000] 21) "_DynCall" :: "'p \ actuals \ (('a,string,'f) com)" ("DYNCALL __" [1000,1000] 21) "_GuardedDynCall" :: "'p \ actuals \ (('a,string,'f) com)" ("DYNCALL\<^sub>g __" [1000,1000] 21) "_DynCallAss":: "'a \ 'p \ actuals \ (('a,string,'f) com)" ("_ :== DYNCALL __" [30,1000,1000] 21) + "_DynCall_exn" :: "'p \ actuals \ (('a,string,'f) com)" ("DYNCALL\<^sub>e __" [1000,1000] 21) + "_DynCallAss_exn":: "'a \ 'p \ actuals \ (('a,string,'f) com)" + ("_ :== DYNCALL\<^sub>e __" [30,1000,1000] 21) "_GuardedDynCallAss":: "'a \ 'p \ actuals \ (('a,string,'f) com)" ("_ :== DYNCALL\<^sub>g __" [30,1000,1000] 21) "_Bind":: "['s \ 'v, idt, 'v \ ('s,'p,'f) com] \ ('s,'p,'f) com" ("_ \ _./ _" [22,1000,21] 21) "_bseq"::"('s,'p,'f) com \ ('s,'p,'f) com \ ('s,'p,'f) com" ("_\/ _" [22, 21] 21) "_FCall" :: "['p,actuals,idt,(('a,string,'f) com)]\ (('a,string,'f) com)" ("CALL __ \ _./ _" [1000,1000,1000,21] 21) translations "_Bind e i c" == "CONST bind (_quote e) (\i. c)" "_FCall p acts i c" == "_FCall p acts (\i. c)" "_bseq c d" == "CONST bseq c d" nonterminal modifyargs syntax "_may_modify" :: "['a,'a,modifyargs] \ bool" ("_ may'_only'_modify'_globals _ in [_]" [100,100,0] 100) "_may_not_modify" :: "['a,'a] \ bool" ("_ may'_not'_modify'_globals _" [100,100] 100) "_may_modify_empty" :: "['a,'a] \ bool" ("_ may'_only'_modify'_globals _ in []" [100,100] 100) "_modifyargs" :: "[id,modifyargs] \ modifyargs" ("_,/ _") "" :: "id => modifyargs" ("_") translations "s may_only_modify_globals Z in []" => "s may_not_modify_globals Z" definition Let':: "['a, 'a => 'b] => 'b" where "Let' = Let" ML_file \hoare_syntax.ML\ parse_translation \ let val argsC = @{syntax_const "_modifyargs"}; val globalsN = "globals"; val ex = @{const_syntax mex}; val eq = @{const_syntax meq}; val varn = Hoare.varname; fun extract_args (Const (argsC,_)$Free (n,_)$t) = varn n::extract_args t | extract_args (Free (n,_)) = [varn n] | extract_args t = raise TERM ("extract_args", [t]) fun idx [] y = error "idx: element not in list" | idx (x::xs) y = if x=y then 0 else (idx xs y)+1 fun gen_update ctxt names (name,t) = - Hoare_Syntax.update_comp ctxt [] false true name (Bound (idx names name)) t + Hoare_Syntax.update_comp ctxt NONE [] false true name (Bound (idx names name)) t fun gen_updates ctxt names t = Library.foldr (gen_update ctxt names) (names,t) fun gen_ex (name,t) = Syntax.const ex $ Abs (name,dummyT,t) fun gen_exs names t = Library.foldr gen_ex (names,t) fun tr ctxt s Z names = let val upds = gen_updates ctxt (rev names) (Syntax.free globalsN$Z); val eq = Syntax.const eq $ (Syntax.free globalsN$s) $ upds; in gen_exs names eq end; fun may_modify_tr ctxt [s,Z,names] = tr ctxt s Z (sort_strings (extract_args names)) fun may_not_modify_tr ctxt [s,Z] = tr ctxt s Z [] in [(@{syntax_const "_may_modify"}, may_modify_tr), (@{syntax_const "_may_not_modify"}, may_not_modify_tr)] end \ print_translation \ let val argsC = @{syntax_const "_modifyargs"}; val chop = Hoare.chopsfx Hoare.deco; fun get_state ( _ $ _ $ t) = get_state t (* for record-updates*) | get_state ( _ $ _ $ _ $ _ $ _ $ t) = get_state t (* for statespace-updates *) | get_state (globals$(s as Const (@{syntax_const "_free"},_) $ Free _)) = s | get_state (globals$(s as Const (@{syntax_const "_bound"},_) $ Free _)) = s | get_state (globals$(s as Const (@{syntax_const "_var"},_) $ Var _)) = s | get_state (globals$(s as Const _)) = s | get_state (globals$(s as Free _)) = s | get_state (globals$(s as Bound _)) = s | get_state t = raise Match; fun mk_args [n] = Syntax.free (chop n) | mk_args (n::ns) = Syntax.const argsC $ Syntax.free (chop n) $ mk_args ns | mk_args _ = raise Match; fun tr' names (Abs (n,_,t)) = tr' (n::names) t | tr' names (Const (@{const_syntax mex},_) $ t) = tr' names t | tr' names (Const (@{const_syntax meq},_) $ (globals$s) $ upd) = let val Z = get_state upd; in (case names of [] => Syntax.const @{syntax_const "_may_not_modify"} $ s $ Z | xs => Syntax.const @{syntax_const "_may_modify"} $ s $ Z $ mk_args (rev names)) end; fun may_modify_tr' [t] = tr' [] t fun may_not_modify_tr' [_$s,_$Z] = Syntax.const @{syntax_const "_may_not_modify"} $ s $ Z in [(@{const_syntax mex}, K may_modify_tr'), (@{const_syntax meq}, K may_not_modify_tr')] end \ (* decorate state components with suffix *) (* parse_ast_translation {* [(@{syntax_const "_Free"}, K Hoare_Syntax.free_arg_ast_tr), (@{syntax_const "_In"}, K Hoare_Syntax.in_arg_ast_tr), (@{syntax_const "_Where"}, K Hoare_Syntax.where_arg_ast_tr @{syntax_const "_Where"}), (@{syntax_const "_WhereElse"}, K Hoare_Syntax.where_arg_ast_tr @{syntax_const "_WhereElse"}) ] *} *) (* parse_ast_translation {* [(@{syntax_const "_antiquoteOld"}, Hoare_Syntax.antiquoteOld_varname_tr @{syntax_const "_antiquoteOld"})] *} *) parse_translation \ [(@{syntax_const "_antiquoteCur"}, K (Hoare_Syntax.antiquote_varname_tr @{syntax_const "_antiquoteCur"}))] \ parse_translation \ [(@{syntax_const "_antiquoteOld"}, Hoare_Syntax.antiquoteOld_tr), - (@{syntax_const "_Call"}, Hoare_Syntax.call_tr false false), + (@{syntax_const "_Call"}, Hoare_Syntax.call_tr false false []), + (@{syntax_const "_Call_exn"}, Hoare_Syntax.call_tr false true []), (@{syntax_const "_FCall"}, Hoare_Syntax.fcall_tr), - (@{syntax_const "_CallAss"}, Hoare_Syntax.call_ass_tr false false), - (@{syntax_const "_GuardedCall"}, Hoare_Syntax.call_tr false true), - (@{syntax_const "_GuardedCallAss"}, Hoare_Syntax.call_ass_tr false true), + (@{syntax_const "_CallAss"}, Hoare_Syntax.call_ass_tr false false []), + (@{syntax_const "_CallAss_exn"}, Hoare_Syntax.call_ass_tr false true []), + (@{syntax_const "_GuardedCall"}, Hoare_Syntax.call_tr false true []), + (@{syntax_const "_GuardedCallAss"}, Hoare_Syntax.call_ass_tr false true []), (@{syntax_const "_Proc"}, Hoare_Syntax.proc_tr), (@{syntax_const "_ProcAss"}, Hoare_Syntax.proc_ass_tr), - (@{syntax_const "_DynCall"}, Hoare_Syntax.call_tr true false), - (@{syntax_const "_DynCallAss"}, Hoare_Syntax.call_ass_tr true false), - (@{syntax_const "_GuardedDynCall"}, Hoare_Syntax.call_tr true true), - (@{syntax_const "_GuardedDynCallAss"}, Hoare_Syntax.call_ass_tr true true), + (@{syntax_const "_DynCall"}, Hoare_Syntax.call_tr true false []), + (@{syntax_const "_DynCall_exn"}, Hoare_Syntax.call_tr true true []), + (@{syntax_const "_DynCallAss"}, Hoare_Syntax.call_ass_tr true false []), + (@{syntax_const "_DynCallAss_exn"}, Hoare_Syntax.call_ass_tr true true []), + (@{syntax_const "_GuardedDynCall"}, Hoare_Syntax.call_tr true true []), + (@{syntax_const "_GuardedDynCallAss"}, Hoare_Syntax.call_ass_tr true true []), (@{syntax_const "_BasicBlock"}, Hoare_Syntax.basic_assigns_tr)] \ (* (@{syntax_const "_Call"}, Hoare_Syntax.call_ast_tr), (@{syntax_const "_CallAss"}, Hoare_Syntax.call_ass_ast_tr), (@{syntax_const "_GuardedCall"}, Hoare_Syntax.guarded_call_ast_tr), (@{syntax_const "_GuardedCallAss"}, Hoare_Syntax.guarded_call_ass_ast_tr) *) parse_translation \ let fun quote_tr ctxt [t] = Hoare_Syntax.quote_tr ctxt @{syntax_const "_antiquoteCur"} t | quote_tr ctxt ts = raise TERM ("quote_tr", ts); in [(@{syntax_const "_quote"}, quote_tr)] end \ parse_translation \ [(@{syntax_const "_Assign"}, Hoare_Syntax.assign_tr), (@{syntax_const "_raise"}, Hoare_Syntax.raise_tr), (@{syntax_const "_New"}, Hoare_Syntax.new_tr), (@{syntax_const "_NNew"}, Hoare_Syntax.nnew_tr), (@{syntax_const "_GuardedAssign"}, Hoare_Syntax.guarded_Assign_tr), (@{syntax_const "_GuardedNew"}, Hoare_Syntax.guarded_New_tr), (@{syntax_const "_GuardedNNew"}, Hoare_Syntax.guarded_NNew_tr), (@{syntax_const "_GuardedWhile_inv_var"}, Hoare_Syntax.guarded_While_tr), (@{syntax_const "_GuardedWhileFix_inv_var_hook"}, Hoare_Syntax.guarded_WhileFix_tr), (@{syntax_const "_GuardedCond"}, Hoare_Syntax.guarded_Cond_tr), (@{syntax_const "_Basic"}, Hoare_Syntax.basic_tr)] \ parse_translation \ [(@{syntax_const "_Init"}, Hoare_Syntax.init_tr), (@{syntax_const "_Loc"}, Hoare_Syntax.loc_tr)] \ print_translation \ [(@{const_syntax Basic}, Hoare_Syntax.assign_tr'), (@{const_syntax raise}, Hoare_Syntax.raise_tr'), (@{const_syntax Basic}, Hoare_Syntax.new_tr'), (@{const_syntax Basic}, Hoare_Syntax.init_tr'), (@{const_syntax Spec}, Hoare_Syntax.nnew_tr'), (@{const_syntax block}, Hoare_Syntax.loc_tr'), (@{const_syntax Collect}, Hoare_Syntax.assert_tr'), (@{const_syntax Cond}, Hoare_Syntax.bexp_tr' "_Cond"), (@{const_syntax switch}, Hoare_Syntax.switch_tr'), (@{const_syntax Basic}, Hoare_Syntax.basic_tr'), (@{const_syntax guards}, Hoare_Syntax.guards_tr'), (@{const_syntax whileAnnoG}, Hoare_Syntax.whileAnnoG_tr'), (@{const_syntax whileAnnoGFix}, Hoare_Syntax.whileAnnoGFix_tr'), (@{const_syntax bind}, Hoare_Syntax.bind_tr')] \ print_translation \ let fun spec_tr' ctxt ((coll as Const _)$ ((splt as Const _) $ (t as (Abs (s,T,p))))::ts) = let fun selector (Const (c, T)) = Hoare.is_state_var c | selector (Const (@{syntax_const "_free"}, _) $ (Free (c, T))) = Hoare.is_state_var c | selector _ = false; in if Hoare_Syntax.antiquote_applied_only_to selector p then Syntax.const @{const_syntax Spec} $ coll $ (splt $ Hoare_Syntax.quote_mult_tr' ctxt selector Hoare_Syntax.antiquoteCur Hoare_Syntax.antiquoteOld (Abs (s,T,t))) else raise Match end | spec_tr' _ ts = raise Match in [(@{const_syntax Spec}, spec_tr')] end \ syntax "_Measure":: "('a \ nat) \ ('a \ 'a) set" ("MEASURE _" [22] 1) "_Mlex":: "('a \ nat) \ ('a \ 'a) set \ ('a \ 'a) set" (infixr "<*MLEX*>" 30) translations "MEASURE f" => "(CONST measure) (_quote f)" "f <*MLEX*> r" => "(_quote f) <*mlex*> r" print_translation \ let fun selector (Const (c,T)) = Hoare.is_state_var c | selector _ = false; fun measure_tr' ctxt ((t as (Abs (_,_,p)))::ts) = if Hoare_Syntax.antiquote_applied_only_to selector p then Hoare_Syntax.app_quote_tr' ctxt (Syntax.const @{syntax_const "_Measure"}) (t::ts) else raise Match | measure_tr' _ _ = raise Match fun mlex_tr' ctxt ((t as (Abs (_,_,p)))::r::ts) = if Hoare_Syntax.antiquote_applied_only_to selector p then Hoare_Syntax.app_quote_tr' ctxt (Syntax.const @{syntax_const "_Mlex"}) (t::r::ts) else raise Match | mlex_tr' _ _ = raise Match in [(@{const_syntax measure}, measure_tr'), (@{const_syntax mlex_prod}, mlex_tr')] end \ print_translation \ [(@{const_syntax call}, Hoare_Syntax.call_tr'), (@{const_syntax dynCall}, Hoare_Syntax.dyn_call_tr'), + (@{const_syntax call_exn}, Hoare_Syntax.call_exn_tr'), + (@{const_syntax dynCall_exn}, Hoare_Syntax.dyn_call_exn_tr'), (@{const_syntax fcall}, Hoare_Syntax.fcall_tr'), (@{const_syntax Call}, Hoare_Syntax.proc_tr')] \ end diff --git a/thys/Simpl/XVcg.thy b/thys/Simpl/XVcg.thy --- a/thys/Simpl/XVcg.thy +++ b/thys/Simpl/XVcg.thy @@ -1,57 +1,36 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: XVcg.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2006-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA *) theory XVcg imports Vcg begin text \We introduce a syntactic variant of the let-expression so that we can safely unfold it during verification condition generation. With the new theorem attribute \vcg_simp\ we can declare equalities to be used by the verification condition generator, while simplifying assertions. \ syntax "_Let'" :: "[letbinds, basicblock] => basicblock" ("(LET (_)/ IN (_))" 23) translations "_Let' (_binds b bs) e" == "_Let' b (_Let' bs e)" "_Let' (_bind x a) e" == "CONST Let' a (%x. e)" lemma Let'_unfold [vcg_simp]: "Let' x f = f x" by (simp add: Let'_def Let_def) lemma Let'_split_conv [vcg_simp]: "(Let' x (\p. (case_prod (f p) (g p)))) = (Let' x (\p. (f p) (fst (g p)) (snd (g p))))" by (simp add: split_def) end diff --git a/thys/Simpl/ex/Closure.thy b/thys/Simpl/ex/Closure.thy --- a/thys/Simpl/ex/Closure.thy +++ b/thys/Simpl/ex/Closure.thy @@ -1,352 +1,331 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: Closure.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2006-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA *) section "Experiments with Closures" theory Closure imports "../Hoare" begin definition "callClosure upd cl = Seq (Basic (upd (fst cl))) (Call (snd cl))" definition "dynCallClosure init upd cl return c = DynCom (\s. call (upd (fst (cl s)) \ init) (snd (cl s)) return c)" lemma dynCallClosure_sound: assumes adapt: "P \ {s. \P' Q' A'. \n. \,\\n:\<^bsub>/F \<^esub> P' (callClosure upd (cl s)) Q',A' \ init s \ P' \ (\t \ Q'. return s t \ R s t) \ (\t \ A'. return s t \ A)}" assumes res: "\s t n. \,\\n:\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" shows "\,\\n:\<^bsub>/F \<^esub>P (dynCallClosure init upd cl return c) Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\n:\<^bsub>/F\<^esub> P Call p Q,A" assume exec: "\\ \dynCallClosure init upd cl return c,Normal s\ =n\ t" from execn.Basic [where f="(upd (fst (cl s)))" and s="(init s)"] have exec_upd: "\\\Basic (upd (fst (cl s))),Normal (init s)\ =n\ Normal (((upd (fst (cl s))) \ init) s)" by auto assume P: "s \ P" from P adapt obtain P' Q' A' where valid: "\n. \,\\n:\<^bsub>/F \<^esub> P' (callClosure upd (cl s)) Q',A'" and init_P': "init s \ P'" and R: "(\t \ Q'. return s t \ R s t)" and A: "(\t \ A'. return s t \ A)" by auto assume t_notin_F: "t \ Fault ` F" from exec [simplified dynCallClosure_def] have exec_call: "\\\call (upd (fst (cl s)) \ init) (snd (cl s)) return c,Normal s\ =n\ t" by cases then show "t \ Normal ` Q \ Abrupt ` A" proof (cases rule: execn_call_Normal_elim) fix bdy m t' assume bdy: "\ (snd (cl s)) = Some bdy" assume exec_body: "\\\bdy,Normal ((upd (fst (cl s)) \ init) s)\ =m\ Normal t'" assume exec_c: "\\\c s t',Normal (return s t')\ =Suc m\ t" assume n: "n = Suc m" have "\\\Basic init,Normal s\ =m\ Normal (init s)" by (rule execn.Basic) from bdy exec_body have exec_callC: "\\\Call (snd (cl s)),Normal ((upd (fst (cl s)) \ init) s)\ =Suc m\ Normal t'" by (rule execn.Call) from execn.Seq [OF exec_upd [simplified n]exec_callC] have exec_closure: "\\ \callClosure upd (cl s),Normal (init s)\ =n\ Normal t'" by (simp add: callClosure_def n) from cnvalidD [OF valid [rule_format] ctxt exec_closure init_P'] have "t' \ Q'" by auto with R have "return s t' \ R s t'" by auto from cnvalidD [OF res [rule_format] ctxt exec_c [simplified n[symmetric]] this t_notin_F] show ?thesis by auto next fix bdy m t' assume bdy: "\ (snd (cl s)) = Some bdy" assume exec_body: "\\\bdy,Normal ((upd (fst (cl s)) \ init) s)\ =m\ Abrupt t'" assume t: "t=Abrupt (return s t')" assume n: "n = Suc m" from bdy exec_body have exec_callC: "\\\Call (snd (cl s)),Normal ((upd (fst (cl s)) \ init) s)\ =Suc m\ Abrupt t'" by (rule execn.Call) from execn.Seq [OF exec_upd [simplified n] exec_callC] have exec_closure: "\\ \callClosure upd (cl s),Normal (init s)\ =n\ Abrupt t'" by (simp add: callClosure_def n) from cnvalidD [OF valid [rule_format] ctxt exec_closure init_P'] have "t' \ A'" by auto with A have "return s t' \ A" by auto with t show ?thesis by auto next fix bdy m f assume bdy: "\ (snd (cl s)) = Some bdy" assume exec_body: "\\\bdy,Normal ((upd (fst (cl s)) \ init) s)\ =m\ Fault f" assume t: "t=Fault f" assume n: "n = Suc m" from bdy exec_body have exec_callC: "\\\Call (snd (cl s)),Normal ((upd (fst (cl s)) \ init) s)\ =Suc m\ Fault f" by (rule execn.Call) from execn.Seq [OF exec_upd [simplified n] exec_callC] have exec_closure: "\\ \callClosure upd (cl s),Normal (init s)\ =n\ Fault f" by (simp add: callClosure_def n) from cnvalidD [OF valid [rule_format] ctxt exec_closure init_P'] t_notin_F t have False by auto thus ?thesis .. next fix bdy m assume bdy: "\ (snd (cl s)) = Some bdy" assume exec_body: "\\\bdy,Normal ((upd (fst (cl s)) \ init) s)\ =m\ Stuck" assume t: "t=Stuck" assume n: "n = Suc m" from execn.Basic [where f="(upd (fst (cl s)))" and s="(init s)"] have exec_upd: "\\\Basic (upd (fst (cl s))),Normal (init s)\ =Suc m\ Normal (((upd (fst (cl s))) \ init) s)" by auto from bdy exec_body have exec_callC: "\\\Call (snd (cl s)),Normal ((upd (fst (cl s)) \ init) s)\ =Suc m\ Stuck" by (rule execn.Call) from execn.Seq [OF exec_upd [simplified n] exec_callC] have exec_closure: "\\ \callClosure upd (cl s),Normal (init s)\ =n\ Stuck" by (simp add: callClosure_def n) from cnvalidD [OF valid [rule_format] ctxt exec_closure init_P'] t_notin_F t have False by auto thus ?thesis .. next fix m assume no_bdy: "\ (snd (cl s)) = None" assume t: "t=Stuck" assume n: "n = Suc m" from no_bdy have exec_callC: "\\\Call (snd (cl s)),Normal ((upd (fst (cl s)) \ init) s)\ =Suc m\ Stuck" by (rule execn.CallUndefined) from execn.Seq [OF exec_upd [simplified n]exec_callC] have exec_closure: "\\ \callClosure upd (cl s),Normal (init s)\ =n\ Stuck" by (simp add: callClosure_def n) from cnvalidD [OF valid [rule_format] ctxt exec_closure init_P'] t_notin_F t have False by auto thus ?thesis .. qed qed lemma dynCallClosure: assumes adapt: "P \ {s. \P' Q' A'. \,\\\<^bsub>/F \<^esub> P' (callClosure upd (cl s)) Q',A' \ init s \ P' \ (\t \ Q'. return s t \ R s t) \ (\t \ A'. return s t \ A)}" assumes res: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" shows "\,\\\<^bsub>/F \<^esub>P (dynCallClosure init upd cl return c) Q,A" apply (rule hoare_complete') apply (rule allI) apply (rule dynCallClosure_sound [where R=R]) using adapt apply (blast intro: hoare_cnvalid) using res apply (blast intro: hoare_cnvalid) done lemma in_subsetD: "\P \ P'; x \ P\ \ x \ P'" by blast lemma dynCallClosureFix: assumes adapt: "P \ {s. \Z. cl'=cl s \ init s \ P' Z \ (\t \ Q' Z. return s t \ R s t) \ (\t \ A' Z. return s t \ A)}" assumes res: "\s t. \,\\\<^bsub>/F\<^esub> (R s t) (c s t) Q,A" assumes spec: "\Z. \,\\\<^bsub>/F \<^esub> (P' Z) (callClosure upd cl') (Q' Z),(A' Z)" shows "\,\\\<^bsub>/F \<^esub>P (dynCallClosure init upd cl return c) Q,A" apply (rule dynCallClosure [OF _ res]) using adapt spec apply clarsimp apply (drule (1) in_subsetD) apply clarsimp apply (rule_tac x="P' Z" in exI) apply (rule_tac x="Q' Z" in exI) apply (rule_tac x="A' Z" in exI) apply blast done lemma conseq_extract_pre: "\\s \ P. \,\\\<^bsub>/F\<^esub> ({s}) c Q,A\ \ \,\\\<^bsub>/F\<^esub> P c Q,A" apply (rule hoarep.Conseq) apply clarify apply (rule_tac x="{s}" in exI) apply (rule_tac x="Q" in exI) apply (rule_tac x="A" in exI) by simp lemma app_closure_sound: assumes adapt: "P \ {s. \P' Q' A'. \n. \,\\n:\<^bsub>/F\<^esub> P' (callClosure upd (e',p)) Q',A' \ upd x s \ P' \ Q' \ Q \ A' \ A}" assumes ap: "upd e = upd e' \ upd x" shows "\,\\n:\<^bsub>/F\<^esub> P (callClosure upd (e,p)) Q,A" proof (rule cnvalidI) fix s t assume ctxt: "\(P, p, Q, A)\\. \\n:\<^bsub>/F\<^esub> P Call p Q,A" assume exec_e: "\\ \callClosure upd (e, p),Normal s\ =n\ t" assume P: "s \ P" assume t: "t \ Fault ` F" from P adapt obtain P' Q' A' where valid: "\n. \,\\n:\<^bsub>/F \<^esub> P' (callClosure upd (e',p)) Q',A'" and init_P': "upd x s \ P'" and Q: "Q' \ Q" and A: "A' \ A" by auto from exec_e [simplified callClosure_def] obtain s' where exec_e: "\\ \Basic (upd (fst (e, p))),Normal s\ =n\ s'"and exec_p: "\\ \Call (snd (e, p)),s'\ =n\ t" by cases from exec_e [simplified] have s': "s'=Normal (upd e s)" by cases simp from ap obtain s'' where s'': "upd x s = s''" and upd_e': "upd e' s''=upd e s" by auto from ap s' execn.Basic [where f="(upd (fst (e', p)))" and s="upd x s" and \=\] have exec_e': "\\ \Basic (upd (fst (e', p))),Normal (upd x s)\ =n\ s'" by simp with exec_p have "\\ \callClosure upd (e', p),Normal (upd x s)\ =n\ t" by (auto simp add: callClosure_def intro: execn.Seq) from cnvalidD [OF valid [rule_format] ctxt this init_P'] t Q A show "t \ Normal ` Q \ Abrupt ` A" by auto qed lemma app_closure: assumes adapt: "P \ {s. \P' Q' A'. \,\\\<^bsub>/F\<^esub> P' (callClosure upd (e',p)) Q',A' \ upd x s \ P' \ Q' \ Q \ A' \ A}" assumes ap: "upd e = upd e' \ upd x" shows "\,\\\<^bsub>/F\<^esub> P (callClosure upd (e,p)) Q,A" apply (rule hoare_complete') apply (rule allI) apply (rule app_closure_sound [where x=x and e'=e', OF _ ap]) using adapt apply (blast intro: hoare_cnvalid) done lemma app_closure_spec: assumes adapt: "P \ {s. \Z. upd x s \ P' Z \ Q' Z \ Q \ A' Z \ A}" assumes ap: "upd e = upd e' \ upd x" assumes spec: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) (callClosure upd (e',p)) (Q' Z),(A' Z)" shows "\,\\\<^bsub>/F\<^esub> P (callClosure upd (e,p)) Q,A" apply (rule app_closure [OF _ ap]) apply clarsimp using adapt spec apply - apply (drule (1) in_subsetD) apply clarsimp apply (rule_tac x="P' Z" in exI) apply (rule_tac x="Q' Z" in exI) apply (rule_tac x="A' Z" in exI) apply blast done text \Implementation of closures as association lists.\ definition "gen_upd var es s = foldl (\s (x,i). the (var x) i s) s es" definition "ap es c \ (es@fst c,snd c)" lemma gen_upd_app: "\es'. gen_upd var (es@es') = gen_upd var es' \ gen_upd var es" apply (induct es) apply (rule ext) apply (simp add: gen_upd_def) apply (rule ext) apply (simp add: gen_upd_def) done lemma gen_upd_ap: "gen_upd var (fst (ap es (es',p))) = gen_upd var es' \ gen_upd var es" by (simp add: gen_upd_app ap_def) lemma ap_closure: assumes adapt: "P \ {s. \P' Q' A'. \,\\\<^bsub>/F\<^esub> P' (callClosure (gen_upd var) c) Q',A' \ gen_upd var es s \ P' \ Q' \ Q \ A' \ A}" shows "\,\\\<^bsub>/F\<^esub> P (callClosure (gen_upd var) (ap es c)) Q,A" proof - obtain es' p where c: "c=(es',p)" by (cases c) have "gen_upd var (fst (ap es (es',p))) = gen_upd var es' \ gen_upd var es" by (simp add: gen_upd_ap) from app_closure [OF adapt [simplified c] this] show ?thesis by (simp add: c ap_def) qed lemma ap_closure_spec: assumes adapt: "P \ {s. \Z. gen_upd var es s \ P' Z \ Q' Z \ Q \ A' Z \ A}" assumes spec: "\Z. \,\\\<^bsub>/F\<^esub> (P' Z) (callClosure (gen_upd var) c) (Q' Z),(A' Z)" shows "\,\\\<^bsub>/F\<^esub> P (callClosure (gen_upd var) (ap es c)) Q,A" proof - obtain es' p where c: "c=(es',p)" by (cases c) have "gen_upd var (fst (ap es (es',p))) = gen_upd var es' \ gen_upd var es" by (simp add: gen_upd_ap) from app_closure_spec [OF adapt [simplified c] this spec [simplified c]] show ?thesis by (simp add: c ap_def) qed end diff --git a/thys/Simpl/ex/ClosureEx.thy b/thys/Simpl/ex/ClosureEx.thy --- a/thys/Simpl/ex/ClosureEx.thy +++ b/thys/Simpl/ex/ClosureEx.thy @@ -1,265 +1,249 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: ClosureEx.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2006-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA *) theory ClosureEx imports "../Vcg" "../Simpl_Heap" Closure begin record globals = cnt_' :: "ref \ nat" alloc_' :: "ref list" free_' :: "nat" record 'g vars = "'g state" + p_':: ref r_':: nat n_':: nat m_':: nat c_':: "(string \ ref) list \ string" d_':: "(string \ ref) list \ string" e_':: "(string \ nat) list \ string" definition "var\<^sub>n = [''n''\ (\x. n_'_update (\_. x)), ''m''\ (\x. m_'_update (\_. x))]" definition "upd\<^sub>n = gen_upd var\<^sub>n" lemma upd\<^sub>n_ap: "upd\<^sub>n (fst (ap es (es',p))) = upd\<^sub>n es' \ upd\<^sub>n es" by (simp add: upd\<^sub>n_def gen_upd_ap) lemma "\\\\n=n\<^sub>0 \ (\i j. \\ \\n=i \ \m=j\ callClosure upd\<^sub>n \e \\r=i + j\)\ \e :== (ap [(''n'',\n)] \e) \\j. \\ \\m=j\ callClosure upd\<^sub>n \e \\r=n\<^sub>0 + j\\" apply vcg_step apply clarify apply (rule ap_closure [where var=var\<^sub>n, folded upd\<^sub>n_def]) apply clarsimp apply (rename_tac s s') apply (erule_tac x="n_' s" in allE) apply (erule_tac x="m_' s'" in allE) apply (rule exI) apply (rule exI) apply (rule conjI) apply (assumption) apply (simp add: upd\<^sub>n_def gen_upd_def var\<^sub>n_def) done definition "var = [''p''\ (\x. p_'_update (\_. x))]" definition "upd = gen_upd var" procedures Inc(p|r) = "\p\\cnt :== \p\\cnt + 1;; \r :== \p\\cnt" lemma (in Inc_impl) "\i p. \\ \\p\\cnt = i\ \r :== PROC Inc(\p) \\r=i+1 \ \p\\cnt = i+1\" apply vcg apply simp done procedures (imports Inc_signature) NewCounter(|c) = "\p :== NEW 1 [\cnt :== 0];; \c :== ([(''p'',\p)],Inc_'proc)" locale NewCounter_impl' = NewCounter_impl + Inc_impl lemma (in NewCounter_impl') shows "\alloc. \\ \1 \ \free\ \c :== PROC NewCounter() \\p. p\\cnt = 0 \ (\i. \\ \p\\cnt = i\ callClosure upd \c \\r=i+1 \ p\\cnt = i+1\)\" apply vcg apply simp apply (rule_tac x="new (set alloc)" in exI) apply simp apply (simp add: callClosure_def) apply vcg_step apply vcg_step apply vcg_step apply vcg_step apply (simp add: upd_def var_def gen_upd_def) done lemma (in NewCounter_impl') shows "\alloc. \\ \1 \ \free\ \c :== PROC NewCounter() \\p. p\\cnt = 0 \ (\i. \\ \p\\cnt = i\ callClosure upd \c \\r=i+1 \ p\\cnt = i+1\)\" apply vcg apply simp apply (rule_tac x="new (set alloc)" in exI) apply simp apply (simp add: callClosure_def) apply vcg_step apply vcg_step apply vcg_step apply vcg_step apply (simp add: upd_def var_def gen_upd_def) done lemma (in NewCounter_impl') shows NewCounter_spec: "\alloc. \\ \1 \ \free \ \alloc=alloc\ \c :== PROC NewCounter() \\p. p \ set alloc \ p \ set \alloc \ p \ Null \ p\\cnt = 0 \ (\i. \\ \p\\cnt = i\ callClosure upd \c \\r=i+1 \ p\\cnt = i+1\)\" apply vcg apply clarsimp apply (rule_tac x="new (set alloc)" in exI) apply simp apply (simp add: callClosure_def) apply vcg_step apply vcg_step apply vcg_step apply vcg_step apply (simp add: upd_def var_def gen_upd_def) done lemma "\\\\p. p \ Null \ p\\cnt = i \ (\i. \\ \p\\cnt = i\ callClosure upd \c \\r=i+1 \ p\\cnt = i+1\)\ dynCallClosure (\s. s) upd c_' (\s t. s\globals := globals t\) (\s t. Basic (\u. u\r_' := r_' t\)) \\r=i+1\" apply (rule conseq_extract_pre) apply clarify apply (rule dynCallClosureFix) apply (simp only: Ball_def) prefer 3 apply (assumption) prefer 2 apply vcg_step apply vcg_step apply (simp only: simp_thms) apply clarsimp done +declare [[hoare_trace = 1]] + +ML \ +val hoare_tacs = #hoare_tacs (Hoare.get_data @{context}); +\ lemma (in NewCounter_impl') shows "\\ \1 \ \free\ \c :== CALL NewCounter ();; dynCallClosure (\s. s) upd c_' (\s t. s\globals := globals t\) (\s t. Basic (\u. u\r_' := r_' t\)) \\r=1\" -apply vcg_step + apply vcg_step apply (rule dynCallClosure) prefer 2 apply vcg_step apply vcg_step apply vcg_step apply clarsimp apply (erule_tac x=0 in allE) apply (rule exI) apply (rule exI) apply (rule conjI) apply (assumption) apply simp done lemma (in NewCounter_impl') shows "\\ \1 \ \free\ \c :== CALL NewCounter ();; dynCallClosure (\s. s) upd c_' (\s t. s\globals := globals t\) (\s t. Basic (\u. u\r_' := r_' t\));; dynCallClosure (\s. s) upd c_' (\s t. s\globals := globals t\) (\s t. Basic (\u. u\r_' := r_' t\)) \\r=2\" apply vcg_step apply (rule dynCallClosure) prefer 2 apply vcg_step apply vcg_step apply vcg_step apply (rule dynCallClosure) apply vcg_step apply vcg_step apply vcg_step apply clarsimp apply (subgoal_tac "\\ \p\\cnt = 0\ callClosure upd (c_' t) \\r = Suc 0 \ p\\cnt = Suc 0\") apply (rule exI) apply (rule exI) apply (rule conjI) apply assumption apply clarsimp apply (erule_tac x=1 in allE) apply (rule exI) apply (rule exI) apply (rule conjI) apply assumption apply clarsimp apply (erule allE) apply assumption done lemma (in NewCounter_impl') shows "\\ \1 \ \free\ \c :== CALL NewCounter ();; \d :== \c;; dynCallClosure (\s. s) upd c_' (\s t. s\globals := globals t\) (\s t. Basic (\u. u\n_' := r_' t\));; dynCallClosure (\s. s) upd d_' (\s t. s\globals := globals t\) (\s t. Basic (\u. u\m_' := r_' t\));; \r :== \n + \m \\r=3\" apply vcg_step apply vcg_step apply (rule dynCallClosure) prefer 2 apply vcg_step apply vcg_step apply vcg_step apply (rule dynCallClosure) apply vcg_step apply vcg_step apply vcg_step apply vcg_step apply clarsimp apply (subgoal_tac "\\ \p\\cnt = 0\ callClosure upd (c_' t) \\r = Suc 0 \ p\\cnt = Suc 0\") apply (rule exI) apply (rule exI) apply (rule conjI) apply assumption apply clarsimp apply (erule_tac x=1 in allE) apply (rule exI) apply (rule exI) apply (rule conjI) apply assumption apply clarsimp apply (erule allE) apply assumption done end diff --git a/thys/Simpl/ex/Compose.thy b/thys/Simpl/ex/Compose.thy --- a/thys/Simpl/ex/Compose.thy +++ b/thys/Simpl/ex/Compose.thy @@ -1,1306 +1,1287 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: Compose.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2006-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. +Copyright (c) 2022 Apple Inc. All rights reserved. -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA *) section "Experiments on State Composition" theory Compose imports "../HoareTotalProps" begin text \ We develop some theory to support state-space modular development of programs. These experiments aim at the representation of state-spaces with records. If we use \statespaces\ instead we get this kind of compositionality for free. \ subsection \Changing the State-Space\ (* Lift a command on statespace 'b to work on statespace 'a *) definition lift\<^sub>f:: "('S \ 's) \ ('S \ 's \ 'S) \ ('s \ 's) \ ('S \ 'S)" where "lift\<^sub>f prj inject f = (\S. inject S (f (prj S)))" definition lift\<^sub>s:: "('S \ 's) \ 's set \ 'S set" where "lift\<^sub>s prj A = {S. prj S \ A}" definition lift\<^sub>r:: "('S \ 's) \ ('S \ 's \ 'S) \ ('s \ 's) set \ ('S \ 'S) set" where "lift\<^sub>r prj inject R = {(S,T). (prj S,prj T) \ R \ T=inject S (prj T)}" primrec lift\<^sub>c:: "('S \ 's) \ ('S \ 's \ 'S) \ ('s,'p,'f) com \ ('S,'p,'f) com" where "lift\<^sub>c prj inject Skip = Skip" | "lift\<^sub>c prj inject (Basic f) = Basic (lift\<^sub>f prj inject f)" | "lift\<^sub>c prj inject (Spec r) = Spec (lift\<^sub>r prj inject r)" | "lift\<^sub>c prj inject (Seq c\<^sub>1 c\<^sub>2) = (Seq (lift\<^sub>c prj inject c\<^sub>1) (lift\<^sub>c prj inject c\<^sub>2))" | "lift\<^sub>c prj inject (Cond b c\<^sub>1 c\<^sub>2) = Cond (lift\<^sub>s prj b) (lift\<^sub>c prj inject c\<^sub>1) (lift\<^sub>c prj inject c\<^sub>2)" | "lift\<^sub>c prj inject (While b c) = While (lift\<^sub>s prj b) (lift\<^sub>c prj inject c)" | "lift\<^sub>c prj inject (Call p) = Call p" | "lift\<^sub>c prj inject (DynCom c) = DynCom (\s. lift\<^sub>c prj inject (c (prj s)))" | "lift\<^sub>c prj inject (Guard f g c) = Guard f (lift\<^sub>s prj g) (lift\<^sub>c prj inject c)" | "lift\<^sub>c prj inject Throw = Throw" | "lift\<^sub>c prj inject (Catch c\<^sub>1 c\<^sub>2) = Catch (lift\<^sub>c prj inject c\<^sub>1) (lift\<^sub>c prj inject c\<^sub>2)" lemma lift\<^sub>c_Skip: "(lift\<^sub>c prj inject c = Skip) = (c = Skip)" by (cases c) auto lemma lift\<^sub>c_Basic: "(lift\<^sub>c prj inject c = Basic lf) = (\f. c = Basic f \ lf = lift\<^sub>f prj inject f)" by (cases c) auto lemma lift\<^sub>c_Spec: "(lift\<^sub>c prj inject c = Spec lr) = (\r. c = Spec r \ lr = lift\<^sub>r prj inject r)" by (cases c) auto lemma lift\<^sub>c_Seq: "(lift\<^sub>c prj inject c = Seq lc\<^sub>1 lc\<^sub>2) = (\ c\<^sub>1 c\<^sub>2. c = Seq c\<^sub>1 c\<^sub>2 \ lc\<^sub>1 = lift\<^sub>c prj inject c\<^sub>1 \ lc\<^sub>2 = lift\<^sub>c prj inject c\<^sub>2 )" by (cases c) auto lemma lift\<^sub>c_Cond: "(lift\<^sub>c prj inject c = Cond lb lc\<^sub>1 lc\<^sub>2) = (\b c\<^sub>1 c\<^sub>2. c = Cond b c\<^sub>1 c\<^sub>2 \ lb = lift\<^sub>s prj b \ lc\<^sub>1 = lift\<^sub>c prj inject c\<^sub>1 \ lc\<^sub>2 = lift\<^sub>c prj inject c\<^sub>2 )" by (cases c) auto lemma lift\<^sub>c_While: "(lift\<^sub>c prj inject c = While lb lc') = (\b c'. c = While b c' \ lb = lift\<^sub>s prj b \ lc' = lift\<^sub>c prj inject c')" by (cases c) auto lemma lift\<^sub>c_Call: "(lift\<^sub>c prj inject c = Call p) = (c = Call p)" by (cases c) auto lemma lift\<^sub>c_DynCom: "(lift\<^sub>c prj inject c = DynCom lc) = (\C. c=DynCom C \ lc = (\s. lift\<^sub>c prj inject (C (prj s))))" by (cases c) auto lemma lift\<^sub>c_Guard: "(lift\<^sub>c prj inject c = Guard f lg lc') = (\g c'. c = Guard f g c' \ lg = lift\<^sub>s prj g \ lc' = lift\<^sub>c prj inject c')" by (cases c) auto lemma lift\<^sub>c_Throw: "(lift\<^sub>c prj inject c = Throw) = (c = Throw)" by (cases c) auto lemma lift\<^sub>c_Catch: "(lift\<^sub>c prj inject c = Catch lc\<^sub>1 lc\<^sub>2) = (\ c\<^sub>1 c\<^sub>2. c = Catch c\<^sub>1 c\<^sub>2 \ lc\<^sub>1 = lift\<^sub>c prj inject c\<^sub>1 \ lc\<^sub>2 = lift\<^sub>c prj inject c\<^sub>2 )" by (cases c) auto definition xstate_map:: "('S \ 's) \ ('S,'f) xstate \ ('s,'f) xstate" where "xstate_map g x = (case x of Normal s \ Normal (g s) | Abrupt s \ Abrupt (g s) | Fault f \ Fault f | Stuck \ Stuck)" lemma xstate_map_simps [simp]: "xstate_map g (Normal s) = Normal (g s)" "xstate_map g (Abrupt s) = Abrupt (g s)" "xstate_map g (Fault f) = (Fault f)" "xstate_map g Stuck = Stuck" by (auto simp add: xstate_map_def) lemma xstate_map_Normal_conv: "xstate_map g S = Normal s = (\s'. S=Normal s' \ s = g s')" by (cases S) auto lemma xstate_map_Abrupt_conv: "xstate_map g S = Abrupt s = (\s'. S=Abrupt s' \ s = g s')" by (cases S) auto lemma xstate_map_Fault_conv: "xstate_map g S = Fault f = (S=Fault f)" by (cases S) auto lemma xstate_map_Stuck_conv: "xstate_map g S = Stuck = (S=Stuck)" by (cases S) auto lemmas xstate_map_convs = xstate_map_Normal_conv xstate_map_Abrupt_conv xstate_map_Fault_conv xstate_map_Stuck_conv definition state:: "('s,'f) xstate \ 's" where "state x = (case x of Normal s \ s | Abrupt s \ s | Fault g \ undefined | Stuck \ undefined)" lemma state_simps [simp]: "state (Normal s) = s" "state (Abrupt s) = s" by (auto simp add: state_def ) locale lift_state_space = fixes project::"'S \ 's" fixes "inject"::"'S \ 's \ 'S" fixes "project\<^sub>x"::"('S,'f) xstate \ ('s,'f) xstate" fixes "lift\<^sub>e"::"('s,'p,'f) body \ ('S,'p,'f) body" fixes lift\<^sub>c:: "('s,'p,'f) com \ ('S,'p,'f) com" fixes lift\<^sub>f:: "('s \ 's) \ ('S \ 'S)" fixes lift\<^sub>s:: "'s set \ 'S set" fixes lift\<^sub>r:: "('s \ 's) set \ ('S \ 'S) set" assumes proj_inj_commute: "\S s. project (inject S s) = s" defines "lift\<^sub>c \ Compose.lift\<^sub>c project inject" defines "project\<^sub>x \ xstate_map project" defines "lift\<^sub>e \ (\\ p. map_option lift\<^sub>c (\ p))" defines "lift\<^sub>f \ Compose.lift\<^sub>f project inject" defines "lift\<^sub>s \ Compose.lift\<^sub>s project" defines "lift\<^sub>r \ Compose.lift\<^sub>r project inject" lemma (in lift_state_space) lift\<^sub>f_simp: "lift\<^sub>f f \ \S. inject S (f (project S))" by (simp add: lift\<^sub>f_def Compose.lift\<^sub>f_def) lemma (in lift_state_space) lift\<^sub>s_simp: "lift\<^sub>s A \ {S. project S \ A}" by (simp add: lift\<^sub>s_def Compose.lift\<^sub>s_def) lemma (in lift_state_space) lift\<^sub>r_simp: "lift\<^sub>r R \ {(S,T). (project S,project T) \ R \ T=inject S (project T)}" by (simp add: lift\<^sub>r_def Compose.lift\<^sub>r_def) (* Causes loop when instantiating locale lemmas (in lift_state_space) lift\<^sub>f_simp = Compose.lift\<^sub>f_def [of project "inject", folded lift\<^sub>f_def] lemmas (in lift_state_space) lift\<^sub>s_simp = Compose.lift\<^sub>s_def [of project, folded lift\<^sub>s_def] lemmas (in lift_state_space) lift\<^sub>r_simp = Compose.lift\<^sub>r_def [of project "inject", folded lift\<^sub>r_def] *) lemma (in lift_state_space) lift\<^sub>c_Skip_simp [simp]: "lift\<^sub>c Skip = Skip" by (simp add: lift\<^sub>c_def) lemma (in lift_state_space) lift\<^sub>c_Basic_simp [simp]: "lift\<^sub>c (Basic f) = Basic (lift\<^sub>f f)" by (simp add: lift\<^sub>c_def lift\<^sub>f_def) lemma (in lift_state_space) lift\<^sub>c_Spec_simp [simp]: "lift\<^sub>c (Spec r) = Spec (lift\<^sub>r r)" by (simp add: lift\<^sub>c_def lift\<^sub>r_def) lemma (in lift_state_space) lift\<^sub>c_Seq_simp [simp]: "lift\<^sub>c (Seq c\<^sub>1 c\<^sub>2) = (Seq (lift\<^sub>c c\<^sub>1) (lift\<^sub>c c\<^sub>2))" by (simp add: lift\<^sub>c_def) lemma (in lift_state_space) lift\<^sub>c_Cond_simp [simp]: "lift\<^sub>c (Cond b c\<^sub>1 c\<^sub>2) = Cond (lift\<^sub>s b) (lift\<^sub>c c\<^sub>1) (lift\<^sub>c c\<^sub>2)" by (simp add: lift\<^sub>c_def lift\<^sub>s_def) lemma (in lift_state_space) lift\<^sub>c_While_simp [simp]: "lift\<^sub>c (While b c) = While (lift\<^sub>s b) (lift\<^sub>c c)" by (simp add: lift\<^sub>c_def lift\<^sub>s_def) lemma (in lift_state_space) lift\<^sub>c_Call_simp [simp]: "lift\<^sub>c (Call p) = Call p" by (simp add: lift\<^sub>c_def) lemma (in lift_state_space) lift\<^sub>c_DynCom_simp [simp]: "lift\<^sub>c (DynCom c) = DynCom (\s. lift\<^sub>c (c (project s)))" by (simp add: lift\<^sub>c_def) lemma (in lift_state_space) lift\<^sub>c_Guard_simp [simp]: "lift\<^sub>c (Guard f g c) = Guard f (lift\<^sub>s g) (lift\<^sub>c c)" by (simp add: lift\<^sub>c_def lift\<^sub>s_def) lemma (in lift_state_space) lift\<^sub>c_Throw_simp [simp]: "lift\<^sub>c Throw = Throw" by (simp add: lift\<^sub>c_def) lemma (in lift_state_space) lift\<^sub>c_Catch_simp [simp]: "lift\<^sub>c (Catch c\<^sub>1 c\<^sub>2) = Catch (lift\<^sub>c c\<^sub>1) (lift\<^sub>c c\<^sub>2)" by (simp add: lift\<^sub>c_def) lemma (in lift_state_space) project\<^sub>x_def': "project\<^sub>x s \ (case s of Normal s \ Normal (project s) | Abrupt s \ Abrupt (project s) | Fault f \ Fault f | Stuck \ Stuck)" by (simp add: xstate_map_def project\<^sub>x_def) lemma (in lift_state_space) lift\<^sub>e_def': "lift\<^sub>e \ p \ (case \ p of Some bdy \ Some (lift\<^sub>c bdy) | None \ None)" by (simp add: lift\<^sub>e_def map_option_case) text \ The problem is that @{term "(lift\<^sub>c project inject \ \)"} is quite a strong premise. The problem is that @{term "\"} is a function here. A map would be better. We only have to lift those procedures in the domain of @{term "\"}: \\ p = Some bdy \ \' p = Some lift\<^sub>c project inject bdy\. We then can com up with theorems that allow us to extend the domains of @{term \} and preserve validity. \ lemma (in lift_state_space) "{(S,T). \t. (project S,t) \ r \ T=inject S t} \ {(S,T). (project S,project T) \ r \ T=inject S (project T)}" apply clarsimp apply (rename_tac S t) apply (simp add: proj_inj_commute) done lemma (in lift_state_space) "{(S,T). (project S,project T) \ r \ T=inject S (project T)} \ {(S,T). \t. (project S,t) \ r \ T=inject S t}" apply clarsimp apply (rename_tac S T) apply (rule_tac x="project T" in exI) apply simp done lemma (in lift_state_space) lift_exec: assumes exec_lc: "(lift\<^sub>e \)\\lc,s\ \ t" shows "\c. \ lift\<^sub>c c = lc\ \ \\\c,project\<^sub>x s\ \ project\<^sub>x t" using exec_lc proof (induct) case Skip thus ?case by (auto simp add: project\<^sub>x_def lift\<^sub>c_Skip lift\<^sub>c_def intro: exec.Skip) next case Guard thus ?case by (auto simp add: project\<^sub>x_def lift\<^sub>s_def Compose.lift\<^sub>s_def lift\<^sub>c_Guard lift\<^sub>c_def intro: exec.Guard) next case GuardFault thus ?case by (auto simp add: project\<^sub>x_def lift\<^sub>s_def Compose.lift\<^sub>s_def lift\<^sub>c_Guard lift\<^sub>c_def intro: exec.GuardFault) next case FaultProp thus ?case by (fastforce simp add: project\<^sub>x_def) next case Basic thus ?case by (fastforce simp add: project\<^sub>x_def lift\<^sub>c_Basic lift\<^sub>f_def Compose.lift\<^sub>f_def lift\<^sub>c_def proj_inj_commute intro: exec.Basic) next case Spec thus ?case by (fastforce simp add: project\<^sub>x_def lift\<^sub>c_Spec lift\<^sub>f_def Compose.lift\<^sub>f_def lift\<^sub>r_def Compose.lift\<^sub>r_def lift\<^sub>c_def proj_inj_commute intro: exec.Spec) next case (SpecStuck s r) thus ?case apply (simp add: project\<^sub>x_def) apply (clarsimp simp add: lift\<^sub>c_Spec lift\<^sub>c_def) apply (unfold lift\<^sub>r_def Compose.lift\<^sub>r_def) apply (rule exec.SpecStuck) apply (rule allI) apply (erule_tac x="inject s t" in allE) apply clarsimp apply (simp add: proj_inj_commute) done next case Seq thus ?case by (fastforce simp add: project\<^sub>x_def lift\<^sub>c_Seq lift\<^sub>c_def intro: exec.intros) next case CondTrue thus ?case by (auto simp add: project\<^sub>x_def lift\<^sub>s_def Compose.lift\<^sub>s_def lift\<^sub>c_Cond lift\<^sub>c_def intro: exec.CondTrue) next case CondFalse thus ?case by (auto simp add: project\<^sub>x_def lift\<^sub>s_def Compose.lift\<^sub>s_def lift\<^sub>c_Cond lift\<^sub>c_def intro: exec.CondFalse) next case WhileTrue thus ?case by (fastforce simp add: project\<^sub>x_def lift\<^sub>s_def Compose.lift\<^sub>s_def lift\<^sub>c_While lift\<^sub>c_def intro: exec.WhileTrue) next case WhileFalse thus ?case by (fastforce simp add: project\<^sub>x_def lift\<^sub>s_def Compose.lift\<^sub>s_def lift\<^sub>c_While lift\<^sub>c_def intro: exec.WhileFalse) next case Call thus ?case by (fastforce simp add: project\<^sub>x_def lift\<^sub>c_Call lift\<^sub>f_def Compose.lift\<^sub>f_def lift\<^sub>c_def lift\<^sub>e_def intro: exec.Call) next case CallUndefined thus ?case by (fastforce simp add: project\<^sub>x_def lift\<^sub>c_Call lift\<^sub>f_def Compose.lift\<^sub>f_def lift\<^sub>c_def lift\<^sub>e_def intro: exec.CallUndefined) next case StuckProp thus ?case by (fastforce simp add: project\<^sub>x_def) next case DynCom thus ?case by (fastforce simp add: project\<^sub>x_def lift\<^sub>c_DynCom lift\<^sub>f_def Compose.lift\<^sub>f_def lift\<^sub>c_def intro: exec.DynCom) next case Throw thus ?case by (fastforce simp add: project\<^sub>x_def lift\<^sub>c_Throw lift\<^sub>c_def intro: exec.Throw) next case AbruptProp thus ?case by (fastforce simp add: project\<^sub>x_def) next case CatchMatch thus ?case by (fastforce simp add: project\<^sub>x_def lift\<^sub>c_Catch lift\<^sub>c_def intro: exec.CatchMatch) next case (CatchMiss c\<^sub>1 s t c\<^sub>2 c) thus ?case by (cases t) (fastforce simp add: project\<^sub>x_def lift\<^sub>c_Catch lift\<^sub>c_def intro: exec.CatchMiss)+ qed lemma (in lift_state_space) lift_exec': assumes exec_lc: "(lift\<^sub>e \)\\lift\<^sub>c c,s\ \ t" shows "\\\c,project\<^sub>x s\ \ project\<^sub>x t" using lift_exec [OF exec_lc] by simp lemma (in lift_state_space) lift_valid: assumes valid: "\\\<^bsub>/F\<^esub> P c Q,A" shows "(lift\<^sub>e \)\\<^bsub>/F\<^esub> (lift\<^sub>s P) (lift\<^sub>c c) (lift\<^sub>s Q),(lift\<^sub>s A)" proof (rule validI) fix s t assume lexec: "(lift\<^sub>e \)\\lift\<^sub>c c,Normal s\ \ t" assume lP: "s \ lift\<^sub>s P" assume noFault: "t \ Fault ` F" show "t \ Normal ` lift\<^sub>s Q \ Abrupt ` lift\<^sub>s A" proof - from lexec have "\\ \c,project\<^sub>x (Normal s)\ \ (project\<^sub>x t)" by (rule lift_exec) (simp_all) moreover from lP have "project s \ P" by (simp add: lift\<^sub>s_def Compose.lift\<^sub>s_def project\<^sub>x_def) ultimately have "project\<^sub>x t \ Normal ` Q \ Abrupt ` A" using valid noFault apply (clarsimp simp add: valid_def project\<^sub>x_def) apply (cases t) apply auto done thus ?thesis apply (simp add: lift\<^sub>s_def Compose.lift\<^sub>s_def) apply (cases t) apply (auto simp add: project\<^sub>x_def) done qed qed lemma (in lift_state_space) lift_hoarep: assumes deriv: "\,{}\\<^bsub>/F\<^esub> P c Q,A" shows "(lift\<^sub>e \),{}\\<^bsub>/F\<^esub> (lift\<^sub>s P) (lift\<^sub>c c) (lift\<^sub>s Q),(lift\<^sub>s A)" apply (rule hoare_complete) apply (insert hoare_sound [OF deriv]) apply (rule lift_valid) apply (simp add: cvalid_def) done lemma (in lift_state_space) lift_hoarep': "\Z. \,{}\\<^bsub>/F\<^esub> (P Z) c (Q Z),(A Z) \ \Z. (lift\<^sub>e \),{}\\<^bsub>/F\<^esub> (lift\<^sub>s (P Z)) (lift\<^sub>c c) (lift\<^sub>s (Q Z)),(lift\<^sub>s (A Z))" apply (iprover intro: lift_hoarep) done lemma (in lift_state_space) lift_termination: assumes termi: "\\c\s" shows "\S. project\<^sub>x S = s \ lift\<^sub>e \ \(lift\<^sub>c c)\S" using termi proof (induct) case Skip thus ?case by (clarsimp simp add: terminates.Skip project\<^sub>x_def xstate_map_convs) next case Basic thus ?case by (fastforce simp add: project\<^sub>x_def xstate_map_convs intro: terminates.intros) next case Spec thus ?case by (fastforce simp add: project\<^sub>x_def xstate_map_convs intro: terminates.intros) next case Guard thus ?case by (auto simp add: project\<^sub>x_def xstate_map_convs intro: terminates.intros) next case GuardFault thus ?case by (auto simp add: project\<^sub>x_def xstate_map_convs lift\<^sub>s_def Compose.lift\<^sub>s_def intro: terminates.intros) next case Fault thus ?case by (clarsimp simp add: project\<^sub>x_def xstate_map_convs) next case (Seq c1 s c2) have "project\<^sub>x S = Normal s" by fact then obtain s' where S: "S=Normal s'" and s: "s = project s'" by (auto simp add: project\<^sub>x_def xstate_map_convs) from Seq have "lift\<^sub>e \\lift\<^sub>c c1 \ S" by simp moreover { fix w assume exec_lc1: "lift\<^sub>e \\\lift\<^sub>c c1,Normal s'\ \ w" have "lift\<^sub>e \\lift\<^sub>c c2 \ w" proof (cases w) case (Normal w') with lift_exec [where c=c1, OF exec_lc1] s have "\\\c1,Normal s\ \ Normal (project w')" by (simp add: project\<^sub>x_def) from Seq.hyps (3) [rule_format, OF this] Normal show "lift\<^sub>e \\lift\<^sub>c c2 \ w" by (auto simp add: project\<^sub>x_def xstate_map_convs) qed (auto) } ultimately show ?case using S s by (auto intro: terminates.intros) next case CondTrue thus ?case by (fastforce simp add: project\<^sub>x_def lift\<^sub>s_def Compose.lift\<^sub>s_def xstate_map_convs intro: terminates.intros) next case CondFalse thus ?case by (fastforce simp add: project\<^sub>x_def lift\<^sub>s_def Compose.lift\<^sub>s_def xstate_map_convs intro: terminates.intros) next case (WhileTrue s b c) have "project\<^sub>x S = Normal s" by fact then obtain s' where S: "S=Normal s'" and s: "s = project s'" by (auto simp add: project\<^sub>x_def xstate_map_convs) from WhileTrue have "lift\<^sub>e \\lift\<^sub>c c \ S" by simp moreover { fix w assume exec_lc: "lift\<^sub>e \\\lift\<^sub>c c,Normal s'\ \ w" have "lift\<^sub>e \\lift\<^sub>c (While b c) \ w" proof (cases w) case (Normal w') with lift_exec [where c=c, OF exec_lc] s have "\\\c,Normal s\ \ Normal (project w')" by (simp add: project\<^sub>x_def) from WhileTrue.hyps (4) [rule_format, OF this] Normal show "lift\<^sub>e \\lift\<^sub>c (While b c) \ w" by (auto simp add: project\<^sub>x_def xstate_map_convs) qed (auto) } ultimately show ?case using S s by (auto intro: terminates.intros) next case WhileFalse thus ?case by (fastforce simp add: project\<^sub>x_def lift\<^sub>s_def Compose.lift\<^sub>s_def xstate_map_convs intro: terminates.intros) next case Call thus ?case by (fastforce simp add: project\<^sub>x_def xstate_map_convs lift\<^sub>e_def intro: terminates.intros) next case CallUndefined thus ?case by (fastforce simp add: project\<^sub>x_def xstate_map_convs lift\<^sub>e_def intro: terminates.intros) next case Stuck thus ?case by (fastforce simp add: project\<^sub>x_def xstate_map_convs) next case DynCom thus ?case by (fastforce simp add: project\<^sub>x_def xstate_map_convs intro: terminates.intros) next case Throw thus ?case by (fastforce simp add: project\<^sub>x_def xstate_map_convs intro: terminates.intros) next case Abrupt thus ?case by (fastforce simp add: project\<^sub>x_def xstate_map_convs intro: terminates.intros) next case (Catch c1 s c2) have "project\<^sub>x S = Normal s" by fact then obtain s' where S: "S=Normal s'" and s: "s = project s'" by (auto simp add: project\<^sub>x_def xstate_map_convs) from Catch have "lift\<^sub>e \\lift\<^sub>c c1 \ S" by simp moreover { fix w assume exec_lc1: "lift\<^sub>e \\\lift\<^sub>c c1,Normal s'\ \ Abrupt w" have "lift\<^sub>e \\lift\<^sub>c c2 \ Normal w" proof - from lift_exec [where c=c1, OF exec_lc1] s have "\\\c1,Normal s\ \ Abrupt (project w)" by (simp add: project\<^sub>x_def) from Catch.hyps (3) [rule_format, OF this] show "lift\<^sub>e \\lift\<^sub>c c2 \ Normal w" by (auto simp add: project\<^sub>x_def xstate_map_convs) qed } ultimately show ?case using S s by (auto intro: terminates.intros) qed lemma (in lift_state_space) lift_termination': assumes termi: "\\c\project\<^sub>x S" shows "lift\<^sub>e \ \(lift\<^sub>c c)\S" using lift_termination [OF termi] by iprover lemma (in lift_state_space) lift_validt: assumes valid: "\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" shows "(lift\<^sub>e \)\\<^sub>t\<^bsub>/F\<^esub> (lift\<^sub>s P) (lift\<^sub>c c) (lift\<^sub>s Q),(lift\<^sub>s A)" proof - from valid have "(lift\<^sub>e \)\\<^bsub>/F\<^esub> (lift\<^sub>s P) (lift\<^sub>c c) (lift\<^sub>s Q),(lift\<^sub>s A)" by (auto intro: lift_valid simp add: validt_def) moreover { fix S assume "S \ lift\<^sub>s P" hence "project S \ P" by (simp add: lift\<^sub>s_def Compose.lift\<^sub>s_def) with valid have "\\c \ project\<^sub>x (Normal S)" by (simp add: validt_def project\<^sub>x_def) hence "lift\<^sub>e \\lift\<^sub>c c \ Normal S" by (rule lift_termination') } ultimately show ?thesis by (simp add: validt_def) qed lemma (in lift_state_space) lift_hoaret: assumes deriv: "\,{}\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" shows "(lift\<^sub>e \),{}\\<^sub>t\<^bsub>/F\<^esub> (lift\<^sub>s P) (lift\<^sub>c c) (lift\<^sub>s Q),(lift\<^sub>s A)" apply (rule hoaret_complete) apply (insert hoaret_sound [OF deriv]) apply (rule lift_validt) apply (simp add: cvalidt_def) done locale lift_state_space_ext = lift_state_space + assumes inj_proj_commute: "\S. inject S (project S) = S" assumes inject_last: "\S s t. inject (inject S s) t = inject S t" (* \x. state t = inject (state s) x *) lemma (in lift_state_space_ext) lift_exec_inject_same: assumes exec_lc: "(lift\<^sub>e \)\\lc,s\ \ t" shows "\c. \lift\<^sub>c c = lc; t \ (Fault ` UNIV) \ {Stuck}\ \ state t = inject (state s) (project (state t))" using exec_lc proof (induct) case Skip thus ?case by (clarsimp simp add: inj_proj_commute) next case Guard thus ?case by (clarsimp simp add: lift\<^sub>c_Guard lift\<^sub>c_def) next case GuardFault thus ?case by simp next case FaultProp thus ?case by simp next case Basic thus ?case by (clarsimp simp add: lift\<^sub>f_def Compose.lift\<^sub>f_def proj_inj_commute lift\<^sub>c_Basic lift\<^sub>c_def) next case (Spec r) thus ?case by (clarsimp simp add: Compose.lift\<^sub>r_def lift\<^sub>c_Spec lift\<^sub>c_def) next case SpecStuck thus ?case by simp next case (Seq lc1 s s' lc2 t c) have t: "t \ Fault ` UNIV \ {Stuck}" by fact have "lift\<^sub>c c = Seq lc1 lc2" by fact then obtain c1 c2 where c: "c = Seq c1 c2" and lc1: "lc1 = lift\<^sub>c c1" and lc2: "lc2 = lift\<^sub>c c2" by (auto simp add: lift\<^sub>c_Seq lift\<^sub>c_def) show ?case proof (cases s') case (Normal s'') from Seq.hyps (2) [OF lc1 [symmetric]] this have "s'' = inject s (project s'')" by auto moreover from Seq.hyps (4) [OF lc2 [symmetric]] Normal t have "state t = inject s'' (project (state t))" by auto ultimately have "state t = inject (inject s (project s'')) (project (state t))" by simp then show ?thesis by (simp add: inject_last) next case (Abrupt s'') from Seq.hyps (2) [OF lc1 [symmetric]] this have "s'' = inject s (project s'')" by auto moreover from Seq.hyps (4) [OF lc2 [symmetric]] Abrupt t have "state t = inject s'' (project (state t))" by auto ultimately have "state t = inject (inject s (project s'')) (project (state t))" by simp then show ?thesis by (simp add: inject_last) next case (Fault f) with Seq have "t = Fault f" by (auto dest: Fault_end) with t have False by simp thus ?thesis .. next case Stuck with Seq have "t = Stuck" by (auto dest: Stuck_end) with t have False by simp thus ?thesis .. qed next case CondTrue thus ?case by (clarsimp simp add: lift\<^sub>c_Cond lift\<^sub>c_def) next case CondFalse thus ?case by (clarsimp simp add: lift\<^sub>c_Cond lift\<^sub>c_def) next case (WhileTrue s lb lc' s' t c) have t: "t \ Fault ` UNIV \ {Stuck}" by fact have lw: "lift\<^sub>c c = While lb lc'" by fact then obtain b c' where c: "c = While b c'" and lb: "lb = lift\<^sub>s b" and lc: "lc' = lift\<^sub>c c'" by (auto simp add: lift\<^sub>c_While lift\<^sub>s_def lift\<^sub>c_def) show ?case proof (cases s') case (Normal s'') from WhileTrue.hyps (3) [OF lc [symmetric]] this have "s'' = inject s (project s'')" by auto moreover from WhileTrue.hyps (5) [OF lw] Normal t have "state t = inject s'' (project (state t))" by auto ultimately have "state t = inject (inject s (project s'')) (project (state t))" by simp then show ?thesis by (simp add: inject_last) next case (Abrupt s'') from WhileTrue.hyps (3) [OF lc [symmetric]] this have "s'' = inject s (project s'')" by auto moreover from WhileTrue.hyps (5) [OF lw] Abrupt t have "state t = inject s'' (project (state t))" by auto ultimately have "state t = inject (inject s (project s'')) (project (state t))" by simp then show ?thesis by (simp add: inject_last) next case (Fault f) with WhileTrue have "t = Fault f" by (auto dest: Fault_end) with t have False by simp thus ?thesis .. next case Stuck with WhileTrue have "t = Stuck" by (auto dest: Stuck_end) with t have False by simp thus ?thesis .. qed next case WhileFalse thus ?case by (clarsimp simp add: lift\<^sub>c_While inj_proj_commute) next case Call thus ?case by (clarsimp simp add: inject_last lift\<^sub>c_Call lift\<^sub>e_def lift\<^sub>c_def) next case CallUndefined thus ?case by simp next case StuckProp thus ?case by simp next case DynCom thus ?case by (clarsimp simp add: lift\<^sub>c_DynCom lift\<^sub>c_def) next case Throw thus ?case by (simp add: inj_proj_commute) next case AbruptProp thus ?case by (simp add: inj_proj_commute) next case (CatchMatch lc1 s s' lc2 t c) have t: "t \ Fault ` UNIV \ {Stuck}" by fact have "lift\<^sub>c c = Catch lc1 lc2" by fact then obtain c1 c2 where c: "c = Catch c1 c2" and lc1: "lc1 = lift\<^sub>c c1" and lc2: "lc2 = lift\<^sub>c c2" by (auto simp add: lift\<^sub>c_Catch lift\<^sub>c_def) from CatchMatch.hyps (2) [OF lc1 [symmetric]] this have "s' = inject s (project s')" by auto moreover from CatchMatch.hyps (4) [OF lc2 [symmetric]] t have "state t = inject s' (project (state t))" by auto ultimately have "state t = inject (inject s (project s')) (project (state t))" by simp then show ?case by (simp add: inject_last) next case CatchMiss thus ?case by (clarsimp simp add: lift\<^sub>c_Catch lift\<^sub>c_def) qed lemma (in lift_state_space_ext) valid_inject_project: assumes noFaultStuck: "\\\c,Normal (project \)\ \\(Fault ` UNIV \ {Stuck})" shows "lift\<^sub>e \\\<^bsub>/F\<^esub> {\} lift\<^sub>c c {t. t=inject \ (project t)}, {t. t=inject \ (project t)}" proof (rule validI) fix s t assume exec: "lift\<^sub>e \\\lift\<^sub>c c,Normal s\ \ t" assume P: "s \ {\}" assume noFault: "t \ Fault ` F" show "t \ Normal ` {t. t = inject \ (project t)} \ Abrupt ` {t. t = inject \ (project t)}" proof - from lift_exec [OF exec] have "\\\c,project\<^sub>x (Normal s)\ \ project\<^sub>x t" by simp with noFaultStuck P have t: "t \ Fault ` UNIV \ {Stuck}" by (auto simp add: final_notin_def project\<^sub>x_def) from lift_exec_inject_same [OF exec refl this] P have "state t = inject \ (project (state t))" by simp with t show ?thesis by (cases t) auto qed qed lemma (in lift_state_space_ext) lift_exec_inject_same': assumes exec_lc: "(lift\<^sub>e \)\\lift\<^sub>c c,S\ \ T" shows "\c. \T \ (Fault ` UNIV) \ {Stuck}\ \ state T = inject (state S) (project (state T))" using lift_exec_inject_same [OF exec_lc] by simp lemma (in lift_state_space_ext) valid_lift_modifies: assumes valid: "\s. \\\<^bsub>/F\<^esub> {s} c (Modif s),(ModifAbr s)" shows "(lift\<^sub>e \)\\<^bsub>/F\<^esub> {S} (lift\<^sub>c c) {T. T \ lift\<^sub>s (Modif (project S)) \ T=inject S (project T)}, {T. T \ lift\<^sub>s (ModifAbr (project S)) \ T=inject S (project T)}" proof (rule validI) fix s t assume exec: "lift\<^sub>e \\\lift\<^sub>c c,Normal s\ \ t" assume P: "s \ {S}" assume noFault: "t \ Fault ` F" show "t \ Normal ` {t \ lift\<^sub>s (Modif (project S)). t = inject S (project t)} \ Abrupt ` {t \ lift\<^sub>s (ModifAbr (project S)). t = inject S (project t)}" proof - from lift_exec [OF exec] have "\\ \c,project\<^sub>x (Normal s)\ \ project\<^sub>x t" by auto moreover from noFault have "project\<^sub>x t \ Fault ` F" by (cases "t") (auto simp add: project\<^sub>x_def) ultimately have "project\<^sub>x t \ Normal ` (Modif (project s)) \ Abrupt ` (ModifAbr (project s))" using valid [rule_format, of "(project s)"] by (auto simp add: valid_def project\<^sub>x_def) hence t: "t \ Normal ` lift\<^sub>s (Modif (project s)) \ Abrupt ` lift\<^sub>s (ModifAbr (project s))" by (cases t) (auto simp add: project\<^sub>x_def lift\<^sub>s_def Compose.lift\<^sub>s_def) then have "t \ Fault ` UNIV \ {Stuck}" by (cases t) auto from lift_exec_inject_same [OF exec _ this] have "state t = inject (state (Normal s)) (project (state t))" by simp with t show ?thesis using P by auto qed qed lemma (in lift_state_space_ext) hoare_lift_modifies: assumes deriv: "\\. \,{}\\<^bsub>/F\<^esub> {\} c (Modif \),(ModifAbr \)" shows "\\. (lift\<^sub>e \),{}\\<^bsub>/F\<^esub> {\} (lift\<^sub>c c) {T. T \ lift\<^sub>s (Modif (project \)) \ T=inject \ (project T)}, {T. T \ lift\<^sub>s (ModifAbr (project \)) \ T=inject \ (project T)}" apply (rule allI) apply (rule hoare_complete) apply (rule valid_lift_modifies) apply (rule allI) apply (insert hoare_sound [OF deriv [rule_format]]) apply (simp add: cvalid_def) done lemma (in lift_state_space_ext) hoare_lift_modifies': assumes deriv: "\\. \,{}\\<^bsub>/F\<^esub> {\} c (Modif \),(ModifAbr \)" shows "\\. (lift\<^sub>e \),{}\\<^bsub>/F\<^esub> {\} (lift\<^sub>c c) {T. T \ lift\<^sub>s (Modif (project \)) \ (\T'. T=inject \ T')}, {T. T \ lift\<^sub>s (ModifAbr (project \)) \ (\T'. T=inject \ T')}" apply (rule allI) apply (rule HoarePartialDef.conseq [OF hoare_lift_modifies [OF deriv]]) apply blast done subsection \Renaming Procedures\ primrec rename:: "('p \ 'q) \ ('s,'p,'f) com \ ('s,'q,'f) com" where "rename N Skip = Skip" | "rename N (Basic f) = Basic f" | "rename N (Spec r) = Spec r" | "rename N (Seq c\<^sub>1 c\<^sub>2) = (Seq (rename N c\<^sub>1) (rename N c\<^sub>2))" | "rename N (Cond b c\<^sub>1 c\<^sub>2) = Cond b (rename N c\<^sub>1) (rename N c\<^sub>2)" | "rename N (While b c) = While b (rename N c)" | "rename N (Call p) = Call (N p)" | "rename N (DynCom c) = DynCom (\s. rename N (c s))" | "rename N (Guard f g c) = Guard f g (rename N c)" | "rename N Throw = Throw" | "rename N (Catch c\<^sub>1 c\<^sub>2) = Catch (rename N c\<^sub>1) (rename N c\<^sub>2)" lemma rename_Skip: "rename h c = Skip = (c=Skip)" by (cases c) auto lemma rename_Basic: "(rename h c = Basic f) = (c=Basic f)" by (cases c) auto lemma rename_Spec: "(rename h c = Spec r) = (c=Spec r)" by (cases c) auto lemma rename_Seq: "(rename h c = Seq rc\<^sub>1 rc\<^sub>2) = (\ c\<^sub>1 c\<^sub>2. c = Seq c\<^sub>1 c\<^sub>2 \ rc\<^sub>1 = rename h c\<^sub>1 \ rc\<^sub>2 = rename h c\<^sub>2 )" by (cases c) auto lemma rename_Cond: "(rename h c = Cond b rc\<^sub>1 rc\<^sub>2) = (\c\<^sub>1 c\<^sub>2. c = Cond b c\<^sub>1 c\<^sub>2 \ rc\<^sub>1 = rename h c\<^sub>1 \ rc\<^sub>2 = rename h c\<^sub>2 )" by (cases c) auto lemma rename_While: "(rename h c = While b rc') = (\c'. c = While b c' \ rc' = rename h c')" by (cases c) auto lemma rename_Call: "(rename h c = Call q) = (\p. c = Call p \ q=h p)" by (cases c) auto lemma rename_DynCom: "(rename h c = DynCom rc) = (\C. c=DynCom C \ rc = (\s. rename h (C s)))" by (cases c) auto lemma rename_Guard: "(rename h c = Guard f g rc') = (\c'. c = Guard f g c' \ rc' = rename h c')" by (cases c) auto lemma rename_Throw: "(rename h c = Throw) = (c = Throw)" by (cases c) auto lemma rename_Catch: "(rename h c = Catch rc\<^sub>1 rc\<^sub>2) = (\c\<^sub>1 c\<^sub>2. c = Catch c\<^sub>1 c\<^sub>2 \ rc\<^sub>1 = rename h c\<^sub>1 \ rc\<^sub>2 = rename h c\<^sub>2 )" by (cases c) auto lemma exec_rename_to_exec: assumes \: "\p bdy. \ p = Some bdy \ \' (h p) = Some (rename h bdy)" assumes exec: "\'\\rc,s\ \ t" shows "\c. rename h c = rc\ \t'. \\\c,s\ \ t' \ (t'=Stuck \ t'=t)" using exec proof (induct) case Skip thus ?case by (fastforce intro: exec.intros simp add: rename_Skip) next case Guard thus ?case by (fastforce intro: exec.intros simp add: rename_Guard) next case GuardFault thus ?case by (fastforce intro: exec.intros simp add: rename_Guard) next case FaultProp thus ?case by (fastforce intro: exec.intros) next case Basic thus ?case by (fastforce intro: exec.intros simp add: rename_Basic) next case Spec thus ?case by (fastforce intro: exec.intros simp add: rename_Spec) next case SpecStuck thus ?case by (fastforce intro: exec.intros simp add: rename_Spec) next case Seq thus ?case by (fastforce intro: exec.intros simp add: rename_Seq) next case CondTrue thus ?case by (fastforce intro: exec.intros simp add: rename_Cond) next case CondFalse thus ?case by (fastforce intro: exec.intros simp add: rename_Cond) next case WhileTrue thus ?case by (fastforce intro: exec.intros simp add: rename_While) next case WhileFalse thus ?case by (fastforce intro: exec.intros simp add: rename_While) next case (Call p rbdy s t) have rbdy: "\' p = Some rbdy" by fact have "rename h c = Call p" by fact then obtain q where c: "c=Call q" and p: "p=h q" by (auto simp add: rename_Call) show ?case proof (cases "\ q") case None with c show ?thesis by (auto intro: exec.CallUndefined) next case (Some bdy) from \ [rule_format, OF this] p rbdy have "rename h bdy = rbdy" by simp with Call.hyps c Some show ?thesis by (fastforce intro: exec.intros) qed next case (CallUndefined p s) have undef: "\' p = None" by fact have "rename h c = Call p" by fact then obtain q where c: "c=Call q" and p: "p=h q" by (auto simp add: rename_Call) from undef p \ have "\ q = None" by (cases "\ q") auto with p c show ?case by (auto intro: exec.intros) next case StuckProp thus ?case by (fastforce intro: exec.intros) next case DynCom thus ?case by (fastforce intro: exec.intros simp add: rename_DynCom) next case Throw thus ?case by (fastforce intro: exec.intros simp add: rename_Throw) next case AbruptProp thus ?case by (fastforce intro: exec.intros) next case CatchMatch thus ?case by (fastforce intro: exec.intros simp add: rename_Catch) next case CatchMiss thus ?case by (fastforce intro: exec.intros simp add: rename_Catch) qed lemma exec_rename_to_exec': assumes \: "\p bdy. \ p = Some bdy \ \' (N p) = Some (rename N bdy)" assumes exec: "\'\\rename N c,s\ \ t" shows "\t'. \\\c,s\ \ t' \ (t'=Stuck \ t'=t)" using exec_rename_to_exec [OF \ exec] by auto lemma valid_to_valid_rename: assumes \: "\p bdy. \ p = Some bdy \ \' (N p) = Some (rename N bdy)" assumes valid: "\\\<^bsub>/F\<^esub> P c Q,A" shows "\'\\<^bsub>/F\<^esub> P (rename N c) Q,A" proof (rule validI) fix s t assume execr: "\'\ \rename N c,Normal s\ \ t" assume P: "s \ P" assume noFault: "t \ Fault ` F" show "t \ Normal ` Q \ Abrupt ` A" proof - from exec_rename_to_exec [OF \ execr] obtain t' where exec: "\\ \c,Normal s\ \ t'" and t': "(t' = Stuck \ t' = t)" by auto with valid noFault P show ?thesis by (auto simp add: valid_def) qed qed lemma hoare_to_hoare_rename: assumes \: "\p bdy. \ p = Some bdy \ \' (N p) = Some (rename N bdy)" assumes deriv: "\,{}\\<^bsub>/F\<^esub> P c Q,A" shows "\',{}\\<^bsub>/F\<^esub> P (rename N c) Q,A" apply (rule hoare_complete) apply (insert hoare_sound [OF deriv]) apply (rule valid_to_valid_rename) apply (rule \) apply (simp add: cvalid_def) done lemma hoare_to_hoare_rename': assumes \: "\p bdy. \ p = Some bdy \ \' (N p) = Some (rename N bdy)" assumes deriv: "\Z. \,{}\\<^bsub>/F\<^esub> (P Z) c (Q Z),(A Z)" shows "\Z. \',{}\\<^bsub>/F\<^esub> (P Z) (rename N c) (Q Z),(A Z)" apply rule apply (rule hoare_to_hoare_rename [OF \]) apply (rule deriv[rule_format]) done lemma terminates_to_terminates_rename: assumes \: "\p bdy. \ p = Some bdy \ \' (N p) = Some (rename N bdy)" assumes termi: "\\ c \ s" assumes noStuck: "\\ \c,s\ \\{Stuck}" shows "\'\ rename N c \ s" using termi noStuck proof (induct) case Skip thus ?case by (fastforce intro: terminates.intros) next case Basic thus ?case by (fastforce intro: terminates.intros) next case Spec thus ?case by (fastforce intro: terminates.intros) next case Guard thus ?case by (fastforce intro: terminates.intros simp add: final_notin_def exec.intros) next case GuardFault thus ?case by (fastforce intro: terminates.intros) next case Fault thus ?case by (fastforce intro: terminates.intros) next case Seq thus ?case by (force intro!: terminates.intros exec.intros dest: exec_rename_to_exec [OF \] simp add: final_notin_def) next case CondTrue thus ?case by (fastforce intro: terminates.intros simp add: final_notin_def exec.intros) next case CondFalse thus ?case by (fastforce intro: terminates.intros simp add: final_notin_def exec.intros) next case (WhileTrue s b c) have s_in_b: "s \ b" by fact have noStuck: "\\ \While b c,Normal s\ \\{Stuck}" by fact with s_in_b have "\\ \c,Normal s\ \\{Stuck}" by (auto simp add: final_notin_def intro: exec.intros) with WhileTrue.hyps have "\'\rename N c \ Normal s" by simp moreover { fix t assume exec_rc: "\'\ \rename N c,Normal s\ \ t" have "\'\ While b (rename N c) \ t" proof - from exec_rename_to_exec [OF \ exec_rc] obtain t' where exec_c: "\\ \c,Normal s\ \ t'" and t': "(t' = Stuck \ t' = t)" by auto with s_in_b noStuck obtain "t'=t" and "\\ \While b c,t\ \\{Stuck}" by (auto simp add: final_notin_def intro: exec.intros) with exec_c WhileTrue.hyps show ?thesis by auto qed } ultimately show ?case using s_in_b by (auto intro: terminates.intros) next case WhileFalse thus ?case by (fastforce intro: terminates.intros) next case (Call p bdy s) have "\ p = Some bdy" by fact from \ [rule_format, OF this] have bdy': "\' (N p) = Some (rename N bdy)". from Call have "\'\rename N bdy \ Normal s" by (auto simp add: final_notin_def intro: exec.intros) with bdy' have "\'\Call (N p) \ Normal s" by (auto intro: terminates.intros) thus ?case by simp next case (CallUndefined p s) have "\ p = None" "\\ \Call p,Normal s\ \\{Stuck}" by fact+ hence False by (auto simp add: final_notin_def intro: exec.intros) thus ?case .. next case Stuck thus ?case by (fastforce intro: terminates.intros) next case DynCom thus ?case by (fastforce intro: terminates.intros simp add: final_notin_def exec.intros) next case Throw thus ?case by (fastforce intro: terminates.intros) next case Abrupt thus ?case by (fastforce intro: terminates.intros) next case (Catch c1 s c2) have noStuck: "\\ \Catch c1 c2,Normal s\ \\{Stuck}" by fact hence "\\ \c1,Normal s\ \\{Stuck}" by (fastforce simp add: final_notin_def intro: exec.intros) with Catch.hyps have "\'\rename N c1 \ Normal s" by auto moreover { fix t assume exec_rc1:"\'\ \rename N c1,Normal s\ \ Abrupt t" have "\'\rename N c2 \ Normal t" proof - from exec_rename_to_exec [OF \ exec_rc1] obtain t' where exec_c: "\\ \c1,Normal s\ \ t'" and "(t' = Stuck \ t' = Abrupt t)" by auto with noStuck have t': "t'=Abrupt t" by (fastforce simp add: final_notin_def intro: exec.intros) with exec_c noStuck have "\\ \c2,Normal t\ \\{Stuck}" by (auto simp add: final_notin_def intro: exec.intros) with exec_c t' Catch.hyps show ?thesis by auto qed } ultimately show ?case by (auto intro: terminates.intros) qed lemma validt_to_validt_rename: assumes \: "\p bdy. \ p = Some bdy \ \' (N p) = Some (rename N bdy)" assumes valid: "\\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" shows "\'\\<^sub>t\<^bsub>/F\<^esub> P (rename N c) Q,A" proof - from valid have "\'\\<^bsub>/F\<^esub> P (rename N c) Q,A" by (auto intro: valid_to_valid_rename [OF \] simp add: validt_def) moreover { fix s assume "s \ P" with valid obtain "\\c \ (Normal s)" "\\ \c,Normal s\ \\{Stuck}" by (auto simp add: validt_def valid_def final_notin_def) from terminates_to_terminates_rename [OF \ this] have "\'\rename N c \ Normal s" . } ultimately show ?thesis by (simp add: validt_def) qed lemma hoaret_to_hoaret_rename: assumes \: "\p bdy. \ p = Some bdy \ \' (N p) = Some (rename N bdy)" assumes deriv: "\,{}\\<^sub>t\<^bsub>/F\<^esub> P c Q,A" shows "\',{}\\<^sub>t\<^bsub>/F\<^esub> P (rename N c) Q,A" apply (rule hoaret_complete) apply (insert hoaret_sound [OF deriv]) apply (rule validt_to_validt_rename) apply (rule \) apply (simp add: cvalidt_def) done lemma hoaret_to_hoaret_rename': assumes \: "\p bdy. \ p = Some bdy \ \' (N p) = Some (rename N bdy)" assumes deriv: "\Z. \,{}\\<^sub>t\<^bsub>/F\<^esub> (P Z) c (Q Z),(A Z)" shows "\Z. \',{}\\<^sub>t\<^bsub>/F\<^esub> (P Z) (rename N c) (Q Z),(A Z)" apply rule apply (rule hoaret_to_hoaret_rename [OF \]) apply (rule deriv[rule_format]) done lemma lift\<^sub>c_whileAnno [simp]: "lift\<^sub>c prj inject (whileAnno b I V c) = whileAnno (lift\<^sub>s prj b) (lift\<^sub>s prj I) (lift\<^sub>r prj inject V) (lift\<^sub>c prj inject c)" by (simp add: whileAnno_def) lemma lift\<^sub>c_block [simp]: "lift\<^sub>c prj inject (block init bdy return c) = block (lift\<^sub>f prj inject init) (lift\<^sub>c prj inject bdy) (\s. (lift\<^sub>f prj inject (return (prj s)))) (\s t. lift\<^sub>c prj inject (c (prj s) (prj t)))" - by (simp add: block_def) + by (simp add: block_def block_exn_def) (* lemma lift\<^sub>c_block [simp]: "lift\<^sub>c prj inject (block init bdy return c) = block (lift\<^sub>f prj inject init) (lift\<^sub>c prj inject bdy) (\s t. inject s (return (prj s) (prj t))) (\s t. lift\<^sub>c prj inject (c (prj s) (prj t)))" apply (simp add: block_def) apply (simp add: lift\<^sub>f_def) *) lemma lift\<^sub>c_call [simp]: "lift\<^sub>c prj inject (call init p return c) = call (lift\<^sub>f prj inject init) p (\s. (lift\<^sub>f prj inject (return (prj s)))) (\s t. lift\<^sub>c prj inject (c (prj s) (prj t)))" by (simp add: call_def lift\<^sub>c_block) lemma rename_whileAnno [simp]: "rename h (whileAnno b I V c) = whileAnno b I V (rename h c)" by (simp add: whileAnno_def) lemma rename_block [simp]: "rename h (block init bdy return c) = block init (rename h bdy) return (\s t. rename h (c s t))" - by (simp add: block_def) + by (simp add: block_def block_exn_def) lemma rename_call [simp]: "rename h (call init p return c) = call init (h p) return (\s t. rename h (c s t))" by (simp add: call_def) end diff --git a/thys/Simpl/ex/ComposeEx.thy b/thys/Simpl/ex/ComposeEx.thy --- a/thys/Simpl/ex/ComposeEx.thy +++ b/thys/Simpl/ex/ComposeEx.thy @@ -1,248 +1,227 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: ComposeEx.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2006-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA *) theory ComposeEx imports Compose "../Vcg" "../HeapList" begin record globals_list = next_' :: "ref \ ref" record state_list = "globals_list state" + p_' :: "ref" sl_q_' :: "ref" r_' :: "ref" procedures Rev(p|sl_q) = "\sl_q :== Null;; WHILE \p \ Null DO \r :== \p;; \\p \ Null\\ \p :== \p\\next;; \\r \ Null\\ \r\\next :== \sl_q;; \sl_q :== \r OD" print_theorems lemma (in Rev_impl) Rev_modifies: "\\. \\\<^bsub>/UNIV \<^esub>{\} \sl_q :== PROC Rev(\p) {t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcNoRec1) apply (vcg spec=modifies) done lemma (in Rev_impl) shows Rev_spec: "\Ps. \\ \List \p \next Ps\ \sl_q :== PROC Rev(\p) \List \sl_q \next (rev Ps)\" apply (hoare_rule HoarePartial.ProcNoRec1) apply (hoare_rule anno = "\sl_q :== Null;; WHILE \p \ Null INV \\Ps' Qs'. List \p \next Ps' \ List \sl_q \next Qs' \ set Ps' \ set Qs' = {} \ rev Ps' @ Qs' = rev Ps\ DO \r :== \p;; \\p \ Null\\\p :== \p\\next;; \\r \ Null\\ \r\\next :== \sl_q;; \sl_q :== \r OD" in HoarePartial.annotateI) apply vcg apply clarsimp apply fastforce apply clarsimp done declare [[names_unique = false]] record globals = strnext_' :: "ref \ ref" chr_' :: "ref \ char" qnext_' :: "ref \ ref" cont_' :: "ref \ int" record state = "globals state" + str_' :: "ref" queue_':: "ref" q_' :: "ref" r_' :: "ref" definition project_globals_str:: "globals \ globals_list" where "project_globals_str g = \next_' = strnext_' g\" definition project_str:: "state \ state_list" where "project_str s = \globals = project_globals_str (globals s), state_list.p_' = str_' s, sl_q_' = q_' s, state_list.r_' = r_' s\" definition inject_globals_str:: "globals \ globals_list \ globals" where "inject_globals_str G g = G\strnext_' := next_' g\" definition "inject_str"::"state \ state_list \ state" where "inject_str S s = S\globals := inject_globals_str (globals S) (globals s), str_' := state_list.p_' s, q_' := sl_q_' s, r_' := state_list.r_' s\" lemma globals_inject_project_str_commutes: "inject_globals_str G (project_globals_str G) = G" by (simp add: inject_globals_str_def project_globals_str_def) lemma inject_project_str_commutes: "inject_str S (project_str S) = S" by (simp add: inject_str_def project_str_def globals_inject_project_str_commutes) lemma globals_project_inject_str_commutes: "project_globals_str (inject_globals_str G g) = g" by (simp add: inject_globals_str_def project_globals_str_def) lemma project_inject_str_commutes: "project_str (inject_str S s) = s" by (simp add: inject_str_def project_str_def globals_project_inject_str_commutes) lemma globals_inject_str_last: "inject_globals_str (inject_globals_str G g) g' = inject_globals_str G g'" by (simp add: inject_globals_str_def) lemma inject_str_last: "inject_str (inject_str S s) s' = inject_str S s'" by (simp add: inject_str_def globals_inject_str_last) definition "lift\<^sub>e = (\\ p. map_option (lift\<^sub>c project_str inject_str) (\ p))" print_locale lift_state_space interpretation ex: lift_state_space project_str inject_str "xstate_map project_str" lift\<^sub>e "lift\<^sub>c project_str inject_str" "lift\<^sub>f project_str inject_str" "lift\<^sub>s project_str" "lift\<^sub>r project_str inject_str" apply - apply (rule lift_state_space.intro) apply (rule project_inject_str_commutes) apply simp apply simp apply (simp add: lift\<^sub>e_def) apply simp apply simp apply simp done interpretation ex: lift_state_space_ext project_str inject_str "xstate_map project_str" lift\<^sub>e "lift\<^sub>c project_str inject_str" "lift\<^sub>f project_str inject_str" "lift\<^sub>s project_str" "lift\<^sub>r project_str inject_str" (* project_str "inject_str" _ lift\<^sub>e *) apply - apply intro_locales [1] apply (rule lift_state_space_ext_axioms.intro) apply (rule inject_project_str_commutes) apply (rule inject_str_last) apply (simp_all add: lift\<^sub>e_def) done (* apply (intro_locales) apply (rule lift_state_space_ext_axioms.intro) apply (rule inject_project_str_commutes) apply (rule inject_str_last) done *) (* declare lift_set_def [simp] project_def [simp] project_globals_def [simp] *) lemmas Rev_lift_spec = ex.lift_hoarep' [OF Rev_impl.Rev_spec,simplified lift\<^sub>s_def project_str_def project_globals_str_def,simplified, of _ "''Rev''"] print_theorems definition "\ p' p = (if p=''Rev'' then p' else '''')" procedures RevStr(str|q) = "rename (\ RevStr_'proc) (lift\<^sub>c project_str inject_str (Rev_body.Rev_body))" lemmas Rev_lift_spec' = Rev_lift_spec [of "[''Rev''\Rev_body.Rev_body]" , simplified Rev_impl_def Rev_clique_def,simplified] thm Rev_lift_spec' lemma Rev_lift_spec'': "\Ps. lift\<^sub>e [''Rev'' \ Rev_body.Rev_body] \ \List \str \strnext Ps\ Call ''Rev'' \List \q \strnext (rev Ps)\" by (rule Rev_lift_spec') lemma (in RevStr_impl) \_ok: "\p bdy. (lift\<^sub>e [''Rev'' \ Rev_body.Rev_body]) p = Some bdy \ \ (\ RevStr_'proc p) = Some (rename (\ RevStr_'proc) bdy)" apply (insert RevStr_impl) apply (auto simp add: RevStr_body_def lift\<^sub>e_def \_def) done context RevStr_impl begin thm hoare_to_hoare_rename'[OF _ Rev_lift_spec'', OF \_ok, simplified \_def, simplified ] end lemmas (in RevStr_impl) RevStr_spec = hoare_to_hoare_rename' [OF _ Rev_lift_spec'', OF \_ok, simplified \_def, simplified ] lemma (in RevStr_impl) RevStr_spec': "\Ps. \\ \List \str \strnext Ps\ \q :== PROC RevStr(\str) \List \q \strnext (rev Ps)\" by (rule RevStr_spec) lemmas Rev_modifies' = Rev_impl.Rev_modifies [of "[''Rev''\Rev_body.Rev_body]", simplified Rev_impl_def, simplified] thm Rev_modifies' context RevStr_impl begin lemmas RevStr_modifies' = hoare_to_hoare_rename' [OF _ ex.hoare_lift_modifies' [OF Rev_modifies'], OF \_ok, of "''Rev''", simplified \_def Rev_clique_def,simplified] end lemma (in RevStr_impl) RevStr_modifies: "\\. \\\<^bsub>/UNIV \<^esub>{\} \str :== PROC RevStr(\str) {t. t may_only_modify_globals \ in [strnext]}" apply (rule allI) apply (rule HoarePartialProps.ConseqMGT [OF RevStr_modifies']) apply (clarsimp simp add: lift\<^sub>s_def mex_def meq_def project_str_def inject_str_def project_globals_str_def inject_globals_str_def) apply blast done end diff --git a/thys/Simpl/ex/ProcParEx.thy b/thys/Simpl/ex/ProcParEx.thy --- a/thys/Simpl/ex/ProcParEx.thy +++ b/thys/Simpl/ex/ProcParEx.thy @@ -1,255 +1,234 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: ProcParEx.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2006-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA *) section "Examples for Procedures as Parameters" theory ProcParEx imports "../Vcg" begin lemma DynProcProcPar': assumes adapt: "P \ {s. p s = q \ (\Z. init s \ P' Z \ (\t \ Q' Z. return s t \ R s t) \ (\t \ A' Z. return s t \ A))}" assumes result: "\s t. \,\\\<^bsub>/F \<^esub>(R s t) result s t Q,A" assumes q: "\Z. \,\\\<^bsub>/F \<^esub>(P' Z) Call q (Q' Z),(A' Z)" shows "\,\\\<^bsub>/F \<^esub>P dynCall init p return result Q,A" apply (rule HoarePartial.DynProcProcPar [OF _ result q]) apply (insert adapt) apply fast done lemma conseq_exploit_pre': "\\s \ S. \,\ \ ({s} \ P) c Q,A\ \ \,\\ (P \ S)c Q,A" apply (rule HoarePartialDef.Conseq) apply clarify by (metis IntI insertI1 subset_refl) lemma conseq_exploit_pre'': "\\Z. \s \ S Z. \,\ \ ({s} \ P Z) c (Q Z),(A Z)\ \ \Z. \,\\ (P Z \ S Z)c (Q Z),(A Z)" apply (rule allI) apply (rule conseq_exploit_pre') apply blast done lemma conseq_exploit_pre''': "\\s \ S. \Z. \,\ \ ({s} \ P Z) c (Q Z),(A Z)\ \ \Z. \,\\ (P Z \ S)c (Q Z),(A Z)" apply (rule allI) apply (rule conseq_exploit_pre') apply blast done record 'g vars = "'g state" + compare_' :: string n_' :: nat m_' :: nat b_' :: bool k_' :: nat procedures compare(n,m|b) = "NoBody" print_locale! compare_signature context compare_signature begin declare [[hoare_use_call_tr' = false]] term "\b :== CALL compare(\n,\m)" term "\b :== DYNCALL \compare(\n,\m)" declare [[hoare_use_call_tr' = true]] term "\b :== DYNCALL \compare(\n,\m)" end procedures LEQ (n,m | b) = "\b :== \n \ \m" LEQ_spec: "\\. \\ {\} PROC LEQ(\n,\m,\b) \\b = (\<^bsup>\\<^esup>n \ \<^bsup>\\<^esup>m)\" LEQ_modifies: "\\. \\ {\} PROC LEQ(\n,\m,\b) {t. t may_only_modify_globals \ in []}" definition mx:: "('a \ 'a \ bool) \ 'a \ 'a \ 'a" where "mx leq a b = (if leq a b then a else b)" procedures Max (compare, n, m | k) = "\b :== DYNCALL \compare(\n,\m);; IF \b THEN \k :== \n ELSE \k :== \m FI" Max_spec: "\leq. \\. \\ ({\} \ {s. (\\. \\ {\} \b :== PROC \<^bsup>s\<^esup>compare(\n,\m) \\b = (leq \<^bsup>\\<^esup>n \<^bsup>\\<^esup>m)\) \ (\\. \\ {\} \b :== PROC \<^bsup>s\<^esup>compare(\n,\m) {t. t may_only_modify_globals \ in []})}) PROC Max(\compare,\n,\m,\k) \\k = mx leq \<^bsup>\\<^esup>n \<^bsup>\\<^esup>m\" lemma (in Max_impl ) Max_spec1: shows "\\ leq. \\ ({\} \ \ (\\. \\{\} \b :== PROC \compare(\n,\m) \\b = (leq \<^bsup>\\<^esup>n \<^bsup>\\<^esup>m)\) \ (\\. \\ {\} \b :== PROC \compare(\n,\m) {t. t may_only_modify_globals \ in []})\) \k :== PROC Max(\compare,\n,\m) \\k = mx leq \<^bsup>\\<^esup>n \<^bsup>\\<^esup>m\" apply (hoare_rule HoarePartial.ProcNoRec1) apply (intro allI) apply (rule conseq_exploit_pre') apply (rule) apply clarify proof - fix \:: "('a,'b) vars_scheme" and s::"('a,'b) vars_scheme" and leq assume compare_spec: "\\. \\{\} \b :== PROC \<^bsup>s\<^esup>compare(\n,\m) \\b = leq \<^bsup>\\<^esup>n \<^bsup>\\<^esup>m\" assume compare_modifies: "\\. \\{\} \b :== PROC \<^bsup>s\<^esup>compare(\n,\m) {t. t may_only_modify_globals \ in []}" show "\\({s} \ {\}) \b :== DYNCALL \compare (\n,\m);; IF \b THEN \k :== \n ELSE \k :== \m FI \\k = mx leq \<^bsup>\\<^esup>n \<^bsup>\\<^esup>m\" apply vcg apply (clarsimp simp add: mx_def) done qed lemma (in Max_impl) Max_spec2: shows "\\ leq. \\ ({\} \ \(\\. \\ {\} \b :== PROC \compare(\n,\m) \\b = (leq \<^bsup>\\<^esup>n \<^bsup>\\<^esup>m)\) \ (\\. \\ {\} \b :== PROC \compare(\n,\m) {t. t may_only_modify_globals \ in []})\) \k :== PROC Max(\compare,\n,\m) \\k = mx leq \<^bsup>\\<^esup>n \<^bsup>\\<^esup>m\" apply (hoare_rule HoarePartial.ProcNoRec1) apply (intro allI) apply (rule conseq_exploit_pre') apply (rule) apply clarify apply vcg apply (clarsimp simp add: mx_def) done lemma (in Max_impl) Max_spec3: shows "\n m leq. \\ (\\n=n \ \m=m\ \ \(\\. \\ {\} \b :== PROC \compare(\n,\m) \\b = (leq \<^bsup>\\<^esup>n \<^bsup>\\<^esup>m)\) \ (\\. \\ {\} \b :== PROC \compare(\n,\m) {t. t may_only_modify_globals \ in []})\) \k :== PROC Max(\compare,\n,\m) \\k = mx leq n m\" apply (hoare_rule HoarePartial.ProcNoRec1) apply (intro allI) apply (rule conseq_exploit_pre') apply (rule) apply clarify apply vcg apply (clarsimp simp add: mx_def) done lemma (in Max_impl) Max_spec4: shows "\n m leq. \\ (\\n=n \ \m=m\ \ \\\. \\ {\} \b :== PROC \compare(\n,\m) \\b = (leq \<^bsup>\\<^esup>n \<^bsup>\\<^esup>m)\\) \k :== PROC Max(\compare,\n,\m) \\k = mx leq n m\" apply (hoare_rule HoarePartial.ProcNoRec1) apply (intro allI) apply (rule conseq_exploit_pre') apply (rule) apply clarify apply vcg apply (clarsimp simp add: mx_def) done locale Max_test = Max_spec + LEQ_spec + LEQ_modifies lemma (in Max_test) shows "\\ {\} \k :== CALL Max(LEQ_'proc,\n,\m) \\k = mx (\) \<^bsup>\\<^esup>n \<^bsup>\\<^esup>m\" proof - note Max_spec = Max_spec [where leq="(\)"] show ?thesis apply vcg apply (clarsimp) apply (rule conjI) apply (rule LEQ_spec [simplified]) apply (rule LEQ_modifies [simplified]) done qed lemma (in Max_impl) Max_spec5: shows "\n m leq. \\ (\\n=n \ \m=m\ \ \\n' m'. \\ \\n=n' \ \m=m'\ \b :== PROC \compare(\n,\m) \\b = (leq n' m')\\) \k :== PROC Max(\compare,\n,\m) \\k = mx leq n m\" term "\{s. \<^bsup>s\<^esup>n = n' \ \<^bsup>s\<^esup>m = m'} = X\" apply (hoare_rule HoarePartial.ProcNoRec1) apply (intro allI) apply (rule conseq_exploit_pre') apply (rule) apply clarify apply vcg apply clarsimp apply (clarsimp simp add: mx_def) done lemma (in LEQ_impl) LEQ_spec: "\n m. \\ \\n=n \ \m=m\ PROC LEQ(\n,\m,\b) \\b = (n \ m)\" apply vcg done locale Max_test' = Max_impl + LEQ_impl lemma (in Max_test') shows "\n m. \\ \\n=n \ \m=m\ \k :== CALL Max(LEQ_'proc,\n,\m) \\k = mx (\) n m\" proof - note Max_spec = Max_spec5 show ?thesis apply vcg apply (rule_tac x="(\)" in exI) apply clarsimp apply (rule LEQ_spec [rule_format]) done qed end diff --git a/thys/Simpl/ex/ProcParExSP.thy b/thys/Simpl/ex/ProcParExSP.thy --- a/thys/Simpl/ex/ProcParExSP.thy +++ b/thys/Simpl/ex/ProcParExSP.thy @@ -1,278 +1,259 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: ProcParEx.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2007-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. +Copyright (c) 2022 Apple Inc. All rights reserved. -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA *) section "Examples for Procedures as Parameters using Statespaces" theory ProcParExSP imports "../Vcg" begin lemma DynProcProcPar': assumes adapt: "P \ {s. p s = q \ (\Z. init s \ P' Z \ (\t \ Q' Z. return s t \ R s t) \ (\t \ A' Z. return s t \ A))}" assumes result: "\s t. \,\\\<^bsub>/F \<^esub>(R s t) result s t Q,A" assumes q: "\Z. \,\\\<^bsub>/F \<^esub>(P' Z) Call q (Q' Z),(A' Z)" shows "\,\\\<^bsub>/F \<^esub>P dynCall init p return result Q,A" apply (rule HoarePartial.DynProcProcPar [OF _ result q]) apply (insert adapt) apply fast done lemma conseq_exploit_pre': "\\s \ S. \,\ \ ({s} \ P) c Q,A\ \ \,\\ (P \ S)c Q,A" apply (rule HoarePartialDef.Conseq) apply clarify by (metis IntI insertI1 subset_refl) lemma conseq_exploit_pre'': "\\Z. \s \ S Z. \,\ \ ({s} \ P Z) c (Q Z),(A Z)\ \ \Z. \,\\ (P Z \ S Z)c (Q Z),(A Z)" apply (rule allI) apply (rule conseq_exploit_pre') apply blast done lemma conseq_exploit_pre''': "\\s \ S. \Z. \,\ \ ({s} \ P Z) c (Q Z),(A Z)\ \ \Z. \,\\ (P Z \ S)c (Q Z),(A Z)" apply (rule allI) apply (rule conseq_exploit_pre') apply blast done procedures compare(i::nat,j::nat|r::bool) "NoBody" print_locale! compare_signature context compare_impl begin declare [[hoare_use_call_tr' = false]] term "\r :== CALL compare(\i,\j)" declare [[hoare_use_call_tr' = true]] end -(* FIXME: typing issue with modifies locale*) +(* fixme: typing issue with modifies locale*) procedures LEQ (i::nat,j::nat | r::bool) "\r :== \i \ \j" LEQ_spec: "\\. \\ {\} PROC LEQ(\i,\j,\r) \\r = (\<^bsup>\\<^esup>i \ \<^bsup>\\<^esup>j)\" LEQ_modifies: "\\. \\ {\} PROC LEQ(\i,\j,\r) {t. t may_only_modify_globals \ in []}" definition mx:: "('a \ 'a \ bool) \ 'a \ 'a \ 'a" where "mx leq a b = (if leq a b then a else b)" procedures (imports compare_signature) Max (compare::string, n::nat, m::nat | k::nat) where b::bool in "\b :== DYNCALL \compare(\n,\m);; IF \b THEN \k :== \n ELSE \k :== \m FI" Max_spec: "\leq. \\. \\ ({\} \ {s. (\\. \\ {\} \r :== PROC \<^bsup>s\<^esup>compare(\i,\j) \\r = (leq \<^bsup>\\<^esup>i \<^bsup>\\<^esup>j)\) \ (\\. \\ {\} \r :== PROC \<^bsup>s\<^esup>compare(\i,\j) {t. t may_only_modify_globals \ in []})}) PROC Max(\compare,\n,\m,\k) \\k = mx leq \<^bsup>\\<^esup>n \<^bsup>\\<^esup>m\" context Max_spec begin thm Max_spec end context Max_impl begin term "\b :== DYNCALL \compare(\n,\m)" declare [[hoare_use_call_tr' = false]] term "\b :== DYNCALL \compare(\n,\m)" declare [[hoare_use_call_tr' = true]] end lemma (in Max_impl ) Max_spec1: shows "\\ leq. \\ ({\} \ \ (\\. \\{\} \r :== PROC \compare(\i,\j) \\r = (leq \<^bsup>\\<^esup>i \<^bsup>\\<^esup>j)\) \ (\\. \\ {\} \r :== PROC \compare(\i,\j) {t. t may_only_modify_globals \ in []})\) \k :== PROC Max(\compare,\n,\m) \\k = mx leq \<^bsup>\\<^esup>n \<^bsup>\\<^esup>m\" apply (hoare_rule HoarePartial.ProcNoRec1) apply (intro allI) apply (rule conseq_exploit_pre') apply (rule) apply clarify proof - fix \:: "('a, 'b, 'c, 'd) stateSP_scheme" and s::"('a, 'b, 'c, 'd) stateSP_scheme" and leq assume compare_spec: "\\. \\{\} \r :== PROC \<^bsup>s\<^esup>compare(\i,\j) \\r = leq \<^bsup>\\<^esup>i \<^bsup>\\<^esup>j\" assume compare_modifies: "\\. \\{\} \r :== PROC \<^bsup>s\<^esup>compare(\i,\j) {t. t may_only_modify_globals \ in []}" show "\\({s} \ {\}) \b :== DYNCALL \compare (\n,\m);; IF \b THEN \k :== \n ELSE \k :== \m FI \\k = mx leq \<^bsup>\\<^esup>n \<^bsup>\\<^esup>m\" apply vcg apply (clarsimp simp add: mx_def) done qed lemma (in Max_impl) Max_spec2: shows "\\ leq. \\ ({\} \ \(\\. \\ {\} \r :== PROC \compare(\i,\j) \\r = (leq \<^bsup>\\<^esup>i \<^bsup>\\<^esup>j)\) \ (\\. \\ {\} \r :== PROC \compare(\i,\j) {t. t may_only_modify_globals \ in []})\) \k :== PROC Max(\compare,\n,\m) \\k = mx leq \<^bsup>\\<^esup>n \<^bsup>\\<^esup>m\" apply (hoare_rule HoarePartial.ProcNoRec1) apply (intro allI) apply (rule conseq_exploit_pre') apply (rule) apply clarify apply vcg apply (clarsimp simp add: mx_def) done lemma (in Max_impl) Max_spec3: shows "\n m leq. \\ (\\n=n \ \m=m\ \ \(\\. \\ {\} \r :== PROC \compare(\i,\j) \\r = (leq \<^bsup>\\<^esup>i \<^bsup>\\<^esup>j)\) \ (\\. \\ {\} \r :== PROC \compare(\i,\j) {t. t may_only_modify_globals \ in []})\) \k :== PROC Max(\compare,\n,\m) \\k = mx leq n m\" apply (hoare_rule HoarePartial.ProcNoRec1) apply (intro allI) apply (rule conseq_exploit_pre') apply (rule) apply clarify apply vcg apply (clarsimp simp add: mx_def) done lemma (in Max_impl) Max_spec4: shows "\n m leq. \\ (\\n=n \ \m=m\ \ \\\. \\ {\} \r :== PROC \compare(\i,\j) \\r = (leq \<^bsup>\\<^esup>i \<^bsup>\\<^esup>j)\\) \k :== PROC Max(\compare,\n,\m) \\k = mx leq n m\" apply (hoare_rule HoarePartial.ProcNoRec1) apply (intro allI) apply (rule conseq_exploit_pre') apply (rule) apply clarify apply vcg apply (clarsimp simp add: mx_def) done print_locale Max_spec (* We have to rename the parameters of the compare procedure to match the LEQ procedure *) locale Max_test = Max_spec where i_'compare_' = i_'LEQ_' and j_'compare_' = j_'LEQ_' and r_'compare_' = r_'LEQ_' + LEQ_spec + LEQ_modifies lemma (in Max_test) shows "\\ {\} \k :== CALL Max(LEQ_'proc,\n,\m) \\k = mx (\) \<^bsup>\\<^esup>n \<^bsup>\\<^esup>m\" proof - note Max_spec = Max_spec [where leq="(\)"] show ?thesis apply vcg apply (clarsimp) apply (rule conjI) apply (rule LEQ_spec) apply (rule LEQ_modifies) done qed lemma (in Max_impl) Max_spec5: shows "\n m leq. \\ (\\n=n \ \m=m\ \ \\n' m'. \\ \\i=n' \ \j=m'\ \r :== PROC \compare(\i,\j) \\r = (leq n' m')\\) \k :== PROC Max(\compare,\n,\m) \\k = mx leq n m\" apply (hoare_rule HoarePartial.ProcNoRec1) apply (intro allI) apply (rule conseq_exploit_pre') apply (rule) apply clarify apply vcg apply clarsimp apply (clarsimp simp add: mx_def) done lemma (in LEQ_impl) LEQ_spec: "\n m. \\ \\i=n \ \j=m\ PROC LEQ(\i,\j,\r) \\r = (n \ m)\" apply vcg apply simp done print_locale Max_impl locale Max_test' = Max_impl where i_'compare_' = i_'LEQ_' and j_'compare_' = j_'LEQ_' and r_'compare_' = r_'LEQ_' + LEQ_impl lemma (in Max_test') shows "\n m. \\ \\n=n \ \m=m\ \k :== CALL Max(LEQ_'proc,\n,\m) \\k = mx (\) n m\" proof - note Max_spec = Max_spec5 show ?thesis apply vcg apply (rule_tac x="(\)" in exI) apply clarsimp apply (rule LEQ_spec [rule_format]) done qed end diff --git a/thys/Simpl/ex/Quicksort.thy b/thys/Simpl/ex/Quicksort.thy --- a/thys/Simpl/ex/Quicksort.thy +++ b/thys/Simpl/ex/Quicksort.thy @@ -1,221 +1,200 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: Quicksort.thy - Author: Norbert Schirmer, TU Muenchen - -Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (C) 2006-2008 Norbert Schirmer *) section "Example: Quicksort on Heap Lists" theory Quicksort imports "../Vcg" "../HeapList" "HOL-Library.Multiset" begin record globals_heap = next_' :: "ref \ ref" cont_' :: "ref \ nat" record 'g vars = "'g state" + p_' :: "ref" q_' :: "ref" le_' :: "ref" gt_' :: "ref" hd_' :: "ref" tl_' :: "ref" procedures append(p,q|p) = "IF \p=Null THEN \p :== \q ELSE \p\\next :== CALL append(\p\\next,\q) FI" append_spec: "\\ Ps Qs. \\ \\. List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {}\ \p :== PROC append(\p,\q) \List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" append_modifies: "\\. \\ {\} \p :== PROC append(\p,\q){t. t may_only_modify_globals \ in [next]}" lemma (in append_impl) append_modifies: shows "\\. \\ {\} \p :== PROC append(\p,\q){t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcRec1) apply (vcg spec=modifies) done lemma (in append_impl) append_spec: shows "\\ Ps Qs. \\ \\. List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {}\ \p :== PROC append(\p,\q) \List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoarePartial.ProcRec1) apply vcg apply fastforce done primrec sorted:: "('a \ 'a \ bool) \ 'a list \ bool" where "sorted le [] = True" | "sorted le (x#xs) = ((\y\set xs. le x y) \ sorted le xs)" lemma sorted_append[simp]: "sorted le (xs@ys) = (sorted le xs \ sorted le ys \ (\x \ set xs. \y \ set ys. le x y))" by (induct xs, auto) procedures quickSort(p|p) = "IF \p=Null THEN SKIP ELSE \tl :== \p\\next;; \le :== Null;; \gt :== Null;; WHILE \tl\Null DO \hd :== \tl;; \tl :== \tl\\next;; IF \hd\\cont \ \p\\cont THEN \hd\\next :== \le;; \le :== \hd ELSE \hd\\next :== \gt;; \gt :== \hd FI OD;; \le :== CALL quickSort(\le);; \gt :== CALL quickSort(\gt);; \p\\next :== \gt;; \le :== CALL append(\le,\p);; \p :== \le FI" quickSort_spec: "\\ Ps. \\ \\. List \p \next Ps\ \p :== PROC quickSort(\p) \(\sortedPs. List \p \next sortedPs \ sorted (\) (map \<^bsup>\\<^esup>cont sortedPs) \ mset Ps = mset sortedPs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" quickSort_modifies: "\\. \\ {\} \p :== PROC quickSort(\p) {t. t may_only_modify_globals \ in [next]}" lemma (in quickSort_impl) quickSort_modifies: shows "\\. \\ {\} \p :== PROC quickSort(\p) {t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcRec1) apply (vcg spec=modifies) done lemma (in quickSort_impl) quickSort_spec: shows "\\ Ps. \\ \\. List \p \next Ps\ \p :== PROC quickSort(\p) \(\sortedPs. List \p \next sortedPs \ sorted (\) (map \<^bsup>\\<^esup>cont sortedPs) \ mset Ps = mset sortedPs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoarePartial.ProcRec1) apply (hoare_rule anno = "IF \p=Null THEN SKIP ELSE \tl :== \p\\next;; \le :== Null;; \gt :== Null;; WHILE \tl\Null INV \ (\les grs tls. List \le \next les \ List \gt \next grs \ List \tl \next tls \ mset Ps = mset (\p#tls@les@grs) \ distinct(\p#tls@les@grs) \ (\x\set les. x\\cont \ \p\\cont) \ (\x\set grs. \p\\cont < x\\cont)) \ \p=\<^bsup>\\<^esup>p \ \cont=\<^bsup>\\<^esup>cont \ List \<^bsup>\\<^esup>p \<^bsup>\\<^esup>next Ps \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\ DO \hd :== \tl;; \tl :== \tl\\next;; IF \hd\\cont \ \p\\cont THEN \hd\\next :== \le;; \le :== \hd ELSE \hd\\next :== \gt;; \gt :== \hd FI OD;; \le :== CALL quickSort(\le);; \gt :== CALL quickSort(\gt);; \p\\next :== \gt;; \le :== CALL append(\le,\p);; \p :== \le FI" in HoarePartial.annotateI) apply vcg apply fastforce apply clarsimp apply (rule conjI) apply clarify apply (rule conjI) apply (rule_tac x="tl#les" in exI) apply simp apply (rule_tac x="grs" in exI) apply simp apply (rule_tac x="ps" in exI) apply simp apply (metis insertCI set_mset_add_mset_insert set_mset_mset) apply clarify apply (rule conjI) apply (rule_tac x="les" in exI) apply simp apply (rule_tac x="tl#grs" in exI) apply simp apply (rule_tac x="ps" in exI) apply simp apply (metis insertCI set_mset_add_mset_insert set_mset_mset) apply clarsimp apply (rule_tac ?x=grs in exI) apply (rule conjI) apply (erule heap_eq_ListI1) apply clarify apply (erule_tac x=x in allE) back apply blast apply clarsimp apply (rule_tac x="sortedPs" in exI) apply (rule conjI) apply (erule heap_eq_ListI1) apply (clarsimp) apply (erule_tac x=x in allE) back back apply (metis IntI empty_iff set_mset_mset) apply (rule_tac x="p#sortedPsa" in exI) apply (rule conjI) apply (metis List_cons List_updateI Null_notin_List fun_upd_same insert_iff set_mset_add_mset_insert set_mset_mset) apply (rule conjI) apply (metis disjoint_iff mset_eq_setD set_ConsD) apply clarsimp apply (rule conjI) apply (metis less_or_eq_imp_le mset_eq_setD) apply (rule conjI) apply (metis leD less_le_trans mset_eq_setD nat_le_linear) apply clarsimp apply (erule_tac x=x in allE)+ apply (metis Un_iff insert_iff list.set(2) mset.simps(2) mset_append set_append set_mset_mset) done end diff --git a/thys/Simpl/ex/VcgEx.thy b/thys/Simpl/ex/VcgEx.thy --- a/thys/Simpl/ex/VcgEx.thy +++ b/thys/Simpl/ex/VcgEx.thy @@ -1,1511 +1,1490 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: VcgEx.thy - Author: Norbert Schirmer, TU Muenchen - -Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (C) 2006-2008 Norbert Schirmer *) section \Examples using the Verification Environment\ theory VcgEx imports "../HeapList" "../Vcg" begin text \Some examples, especially the single-step Isar proofs are taken from \texttt{HOL/Isar\_examples/HoareEx.thy}. \ subsection \State Spaces\ text \ First of all we provide a store of program variables that occur in the programs considered later. Slightly unexpected things may happen when attempting to work with undeclared variables. \ record 'g vars = "'g state" + A_' :: nat I_' :: nat M_' :: nat N_' :: nat R_' :: nat S_' :: nat B_' :: bool Arr_' :: "nat list" Abr_':: string text \We decorate the state components in the record with the suffix \_'\, to avoid cluttering the namespace with the simple names that could no longer be used for logical variables otherwise. \ text \We will first consider programs without procedures, later on we will regard procedures without global variables and finally we will get the full pictures: mutually recursive procedures with global variables (including heap). \ subsection \Basic Examples\ text \ We look at few trivialities involving assignment and sequential composition, in order to get an idea of how to work with our formulation of Hoare Logic. \ text \ Using the basic rule directly is a bit cumbersome. \ lemma "\\ {|\N = 5|} \N :== 2 * \N {|\N = 10|}" apply (rule HoarePartial.Basic) apply simp done text \ If we refer to components (variables) of the state-space of the program we always mark these with \\\. It is the acute-symbol and is present on most keyboards. So all program variables are marked with the acute and all logical variables are not. The assertions of the Hoare tuple are ordinary Isabelle sets. As we usually want to refer to the state space in the assertions, we provide special brackets for them. They can be written as {\verb+{| |}+} in ASCII or \\ \\ with symbols. Internally marking variables has two effects. First of all we refer to the implicit state and secondary we get rid of the suffix \_'\. So the assertion @{term "{|\N = 5|}"} internally gets expanded to \{s. N_' s = 5}\ written in ordinary set comprehension notation of Isabelle. It describes the set of states where the \N_'\ component is equal to \5\. \ text \ Certainly we want the state modification already done, e.g.\ by simplification. The \vcg\ method performs the basic state update for us; we may apply the Simplifier afterwards to achieve ``obvious'' consequences as well. \ lemma "\\ \True\ \N :== 10 \\N = 10\" by vcg lemma "\\ \2 * \N = 10\ \N :== 2 * \N \\N = 10\" by vcg lemma "\\ \\N = 5\ \N :== 2 * \N \\N = 10\" apply vcg apply simp done lemma "\\ \\N + 1 = a + 1\ \N :== \N + 1 \\N = a + 1\" by vcg lemma "\\ \\N = a\ \N :== \N + 1 \\N = a + 1\" by vcg lemma "\\ \a = a \ b = b\ \M :== a;; \N :== b \\M = a \ \N = b\" by vcg lemma "\\ \True\ \M :== a;; \N :== b \\M = a \ \N = b\" by vcg lemma "\\ \\M = a \ \N = b\ \I :== \M;; \M :== \N;; \N :== \I \\M = b \ \N = a\" by vcg text \ We can also perform verification conditions generation step by step by using the \vcg_step\ method. \ lemma "\\ \\M = a \ \N = b\ \I :== \M;; \M :== \N;; \N :== \I \\M = b \ \N = a\" apply vcg_step apply vcg_step apply vcg_step apply vcg_step done text \ It is important to note that statements like the following one can only be proven for each individual program variable. Due to the extra-logical nature of record fields, we cannot formulate a theorem relating record selectors and updates schematically. \ lemma "\\ \\N = a\ \N :== \N \\N = a\" by vcg (* lemma "\\ \\x = a\ \x :== \x \\x = a\" apply (rule HoarePartial.Basic) -- {* We can't proof this since we don't know what @{text "x_'_update"} is. *} oops *) lemma "\\{s. x_' s = a} (Basic (\s. x_'_update (x_' s) s)) {s. x_' s = a}" oops text \ In the following assignments we make use of the consequence rule in order to achieve the intended precondition. Certainly, the \vcg\ method is able to handle this case, too. \ lemma "\\ \\M = \N\ \M :== \M + 1 \\M \ \N\" proof - have "\\M = \N\ \ \\M + 1 \ \N\" by auto also have "\\ \ \M :== \M + 1 \\M \ \N\" by vcg finally show ?thesis . qed lemma "\\ \\M = \N\ \M :== \M + 1 \\M \ \N\" proof - have "\m n::nat. m = n \ m + 1 \ n" \ \inclusion of assertions expressed in ``pure'' logic,\ \ \without mentioning the state space\ by simp also have "\\ \\M + 1 \ \N\ \M :== \M + 1 \\M \ \N\" by vcg finally show ?thesis . qed lemma "\\ \\M = \N\ \M :== \M + 1 \\M \ \N\" apply vcg apply simp done subsection \Multiplication by Addition\ text \ We now do some basic examples of actual \texttt{WHILE} programs. This one is a loop for calculating the product of two natural numbers, by iterated addition. We first give detailed structured proof based on single-step Hoare rules. \ lemma "\\ \\M = 0 \ \S = 0\ WHILE \M \ a DO \S :== \S + b;; \M :== \M + 1 OD \\S = a * b\" proof - let "\\ _ ?while _" = ?thesis let "\\?inv\" = "\\S = \M * b\" have "\\M = 0 & \S = 0\ \ \\?inv\" by auto also have "\\ \ ?while \\?inv \ \ (\M \ a)\" proof let ?c = "\S :== \S + b;; \M :== \M + 1" have "\\?inv \ \M \ a\ \ \\S + b = (\M + 1) * b\" by auto also have "\\ \ ?c \\?inv\" by vcg finally show "\\ \\?inv \ \M \ a\ ?c \\?inv\" . qed also have "\\?inv \ \ (\M \ a)\ \ \\S = a * b\" by auto finally show ?thesis by blast qed text \ The subsequent version of the proof applies the \vcg\ method to reduce the Hoare statement to a purely logical problem that can be solved fully automatically. Note that we have to specify the \texttt{WHILE} loop invariant in the original statement. \ lemma "\\ \\M = 0 \ \S = 0\ WHILE \M \ a INV \\S = \M * b\ DO \S :== \S + b;; \M :== \M + 1 OD \\S = a * b\" apply vcg apply auto done text \Here some examples of ``breaking'' out of a loop\ lemma "\\ \\M = 0 \ \S = 0\ TRY WHILE True INV \\S = \M * b\ DO IF \M = a THEN THROW ELSE \S :== \S + b;; \M :== \M + 1 FI OD CATCH SKIP END \\S = a * b\" apply vcg apply auto done lemma "\\ \\M = 0 \ \S = 0\ TRY WHILE True INV \\S = \M * b\ DO IF \M = a THEN \Abr :== ''Break'';;THROW ELSE \S :== \S + b;; \M :== \M + 1 FI OD CATCH IF \Abr = ''Break'' THEN SKIP ELSE Throw FI END \\S = a * b\" apply vcg apply auto done text \Some more syntactic sugar, the label statement \\ \ \\ as shorthand for the \TRY-CATCH\ above, and the \RAISE\ for an state-update followed by a \THROW\. \ lemma "\\ \\M = 0 \ \S = 0\ \\Abr = ''Break''\\ WHILE True INV \\S = \M * b\ DO IF \M = a THEN RAISE \Abr :== ''Break'' ELSE \S :== \S + b;; \M :== \M + 1 FI OD \\S = a * b\" apply vcg apply auto done lemma "\\ \\M = 0 \ \S = 0\ TRY WHILE True INV \\S = \M * b\ DO IF \M = a THEN RAISE \Abr :== ''Break'' ELSE \S :== \S + b;; \M :== \M + 1 FI OD CATCH IF \Abr = ''Break'' THEN SKIP ELSE Throw FI END \\S = a * b\" apply vcg apply auto done lemma "\\ \\M = 0 \ \S = 0\ \\Abr = ''Break''\ \ WHILE True INV \\S = \M * b\ DO IF \M = a THEN RAISE \Abr :== ''Break'' ELSE \S :== \S + b;; \M :== \M + 1 FI OD \\S = a * b\" apply vcg apply auto done text \Blocks\ lemma "\\\\I = i\ LOC \I;; \I :== 2 COL \\I \ i\" apply vcg by simp lemma "\\ \\N = n\ LOC \N :== 10;; \N :== \N + 2 COL \\N = n\" by vcg lemma "\\ \\N = n\ LOC \N :== 10, \M;; \N :== \N + 2 COL \\N = n\" by vcg subsection \Summing Natural Numbers\ text \ We verify an imperative program to sum natural numbers up to a given limit. First some functional definition for proper specification of the problem. \ primrec sum :: "(nat => nat) => nat => nat" where "sum f 0 = 0" | "sum f (Suc n) = f n + sum f n" syntax "_sum" :: "idt => nat => nat => nat" ("SUMM _<_. _" [0, 0, 10] 10) translations "SUMM jj. b) k" text \ The following proof is quite explicit in the individual steps taken, with the \vcg\ method only applied locally to take care of assignment and sequential composition. Note that we express intermediate proof obligation in pure logic, without referring to the state space. \ theorem "\\ \True\ \S :== 0;; \I :== 1;; WHILE \I \ n DO \S :== \S + \I;; \I :== \I + 1 OD \\S = (SUMM j" (is "\\ _ (_;; ?while) _") proof - let ?sum = "\k. SUMM j\ \True\ \S :== 0;; \I :== 1 \?inv \S \I\" proof - have "True \ 0 = ?sum 1" by simp also have "\\ \\\ \S :== 0;; \I :== 1 \?inv \S \I\" by vcg finally show ?thesis . qed also have "\\ \?inv \S \I\ ?while \?inv \S \I \ \ \I \ n\" proof let ?body = "\S :== \S + \I;; \I :== \I + 1" have "\s i. ?inv s i \ i \ n \ ?inv (s + i) (i + 1)" by simp also have "\\ \\S + \I = ?sum (\I + 1)\ ?body \?inv \S \I\" by vcg finally show "\\ \?inv \S \I \ \I \ n\ ?body \?inv \S \I\" . qed also have "\s i. s = ?sum i \ \ i \ n \ s = ?sum n" by simp finally show ?thesis . qed text \ The next version uses the \vcg\ method, while still explaining the resulting proof obligations in an abstract, structured manner. \ theorem "\\ \True\ \S :== 0;; \I :== 1;; WHILE \I \ n INV \\S = (SUMM j<\I. j)\ DO \S :== \S + \I;; \I :== \I + 1 OD \\S = (SUMM j" proof - let ?sum = "\k. SUMM j n" thus "?inv (s + i) (i + 1)" by simp next fix i s assume x: "?inv s i" "\ i \ n" thus "s = ?sum n" by simp qed qed text \ Certainly, this proof may be done fully automatically as well, provided that the invariant is given beforehand. \ theorem "\\ \True\ \S :== 0;; \I :== 1;; WHILE \I \ n INV \\S = (SUMM j<\I. j)\ DO \S :== \S + \I;; \I :== \I + 1 OD \\S = (SUMM j" apply vcg apply auto done subsection \SWITCH\ lemma "\\ \\N = 5\ SWITCH \B {True} \ \N :== 6 | {False} \ \N :== 7 END \\N > 5\" apply vcg apply simp done lemma "\\ \\N = 5\ SWITCH \N {v. v < 5} \ \N :== 6 | {v. v \ 5} \ \N :== 7 END \\N > 5\" apply vcg apply simp done subsection \(Mutually) Recursive Procedures\ subsubsection \Factorial\ text \We want to define a procedure for the factorial. We first define a HOL functions that calculates it to specify the procedure later on. \ primrec fac:: "nat \ nat" where "fac 0 = 1" | "fac (Suc n) = (Suc n) * fac n" lemma fac_simp [simp]: "0 < i \ fac i = i * fac (i - 1)" by (cases i) simp_all text \Now we define the procedure\ procedures Fac (N|R) = "IF \N = 0 THEN \R :== 1 ELSE \R :== CALL Fac(\N - 1);; \R :== \N * \R FI" text \A procedure is given by the signature of the procedure followed by the procedure body. The signature consists of the name of the procedure and a list of parameters. The parameters in front of the pipe \|\ are value parameters and behind the pipe are the result parameters. Value parameters model call by value semantics. The value of a result parameter at the end of the procedure is passed back to the caller. \ text \ Behind the scenes the \procedures\ command provides us convenient syntax for procedure calls, defines a constant for the procedure body (named @{term "Fac_body"}) and creates some locales. The purpose of locales is to set up logical contexts to support modular reasoning. A locale is named \Fac_impl\ and extends the \hoare\ locale with a theorem @{term "\ ''Fac'' = Fac_body"} that simply states how the procedure is defined in the procedure context. Check out the locales. The purpose of the locales is to give us easy means to setup the context in which we will prove programs correct. In these locales the procedure context @{term "\"} is fixed. So always use this letter in procedure specifications. This is crucial, if we later on prove some tuples under the assumption of some procedure specifications. \ thm Fac_body.Fac_body_def print_locale Fac_impl text \ To see how a call is syntactically translated you can switch off the printing translation via the configuration option \hoare_use_call_tr'\ \ context Fac_impl begin text \ @{term "CALL Fac(\N,\M)"} is internally: \ declare [[hoare_use_call_tr' = false]] text \ @{term "CALL Fac(\N,\M)"} \ term "CALL Fac(\N,\M)" declare [[hoare_use_call_tr' = true]] end text \ Now let us prove that @{term "Fac"} meets its specification. \ text \ Procedure specifications are ordinary Hoare tuples. We use the parameterless call for the specification; \\R :== PROC Fac(\N)\ is syntactic sugar for \Call ''Fac''\. This emphasises that the specification describes the internal behaviour of the procedure, whereas parameter passing corresponds to the procedure call. \ lemma (in Fac_impl) shows "\n. \,\\\\N=n\ PROC Fac(\N,\R) \\R = fac n\" apply (hoare_rule HoarePartial.ProcRec1) apply vcg apply simp done text \ Since the factorial was implemented recursively, the main ingredient of this proof is, to assume that the specification holds for the recursive call of @{term Fac} and prove the body correct. The assumption for recursive calls is added to the context by the rule @{thm [source] HoarePartial.ProcRec1} (also derived from general rule for mutually recursive procedures): @{thm [display] HoarePartial.ProcRec1 [no_vars]} The verification condition generator will infer the specification out of the context when it encounters a recursive call of the factorial. \ text \We can also step through verification condition generation. When the verification condition generator encounters a procedure call it tries to use the rule \ProcSpec\. To be successful there must be a specification of the procedure in the context. \ lemma (in Fac_impl) shows "\n. \\\\N=n\ \R :== PROC Fac(\N) \\R = fac n\" apply (hoare_rule HoarePartial.ProcRec1) apply vcg_step apply vcg_step apply vcg_step apply vcg_step apply vcg_step apply simp done text \Here some Isar style version of the proof\ lemma (in Fac_impl) shows "\n. \\\\N=n\ \R :== PROC Fac(\N) \\R = fac n\" proof (hoare_rule HoarePartial.ProcRec1) have Fac_spec: "\n. \,(\n. {(\\N=n\, Fac_'proc, \\R = fac n\,{})}) \ \\N=n\ \R :== PROC Fac(\N) \\R = fac n\" apply (rule allI) apply (rule hoarep.Asm) by auto show "\n. \,(\n. {(\\N=n\, Fac_'proc, \\R = fac n\,{})}) \ \\N=n\ IF \N = 0 THEN \R :== 1 ELSE \R :== CALL Fac(\N - 1);; \R :== \N * \R FI \\R = fac n\" apply vcg apply simp done qed text \To avoid retyping of potentially large pre and postconditions in the previous proof we can use the casual term abbreviations of the Isar language. \ lemma (in Fac_impl) shows "\n. \\\\N=n\ \R :== PROC Fac(\N) \\R = fac n\" (is "\n. \\(?Pre n) ?Fac (?Post n)") proof (hoare_rule HoarePartial.ProcRec1) have Fac_spec: "\n. \,(\n. {(?Pre n, Fac_'proc, ?Post n,{})}) \(?Pre n) ?Fac (?Post n)" apply (rule allI) apply (rule hoarep.Asm) by auto show "\n. \,(\n. {(?Pre n, Fac_'proc, ?Post n,{})}) \ (?Pre n) IF \N = 0 THEN \R :== 1 ELSE \R :== CALL Fac(\N - 1);; \R :== \N * \R FI (?Post n)" apply vcg apply simp done qed text \The previous proof pattern has still some kind of inconvenience. The augmented context is always printed in the proof state. That can mess up the state, especially if we have large specifications. This may be annoying if we want to develop single step or structured proofs. In this case it can be a good idea to introduce a new variable for the augmented context. \ lemma (in Fac_impl) Fac_spec: shows "\n. \\\\N=n\ \R :== PROC Fac(\N) \\R = fac n\" (is "\n. \\(?Pre n) ?Fac (?Post n)") proof (hoare_rule HoarePartial.ProcRec1) define \' where "\' = (\n. {(?Pre n, Fac_'proc, ?Post n,{}::('a, 'b) vars_scheme set)})" have Fac_spec: "\n. \,\'\(?Pre n) ?Fac (?Post n)" by (unfold \'_def, rule allI, rule hoarep.Asm) auto txt \We have to name the fact \Fac_spec\, so that the vcg can use the specification for the recursive call, since it cannot infer it from the opaque @{term "\'"}.\ show "\\. \,\'\ (?Pre \) IF \N = 0 THEN \R :== 1 ELSE \R :== CALL Fac(\N - 1);; \R :== \N * \R FI (?Post \)" apply vcg apply simp done qed text \There are different rules available to prove procedure calls, depending on the kind of postcondition and whether or not the procedure is recursive or even mutually recursive. See for example @{thm [source] HoarePartial.ProcRec1}, @{thm [source] HoarePartial.ProcNoRec1}. They are all derived from the most general rule @{thm [source] HoarePartial.ProcRec}. All of them have some side-condition concerning definedness of the procedure. They can be solved in a uniform fashion. Thats why we have created the method \hoare_rule\, which behaves like the method \rule\ but automatically tries to solve the side-conditions. \ subsubsection \Odd and Even\ text \Odd and even are defined mutually recursive here. In the \procedures\ command we conjoin both definitions with \and\. \ procedures odd(N | A) = "IF \N=0 THEN \A:==0 ELSE IF \N=1 THEN CALL even (\N - 1,\A) ELSE CALL odd (\N - 2,\A) FI FI" and even(N | A) = "IF \N=0 THEN \A:==1 ELSE IF \N=1 THEN CALL odd (\N - 1,\A) ELSE CALL even (\N - 2,\A) FI FI" print_theorems thm odd_body.odd_body_def thm even_body.even_body_def print_locale odd_even_clique text \To prove the procedure calls to @{term "odd"} respectively @{term "even"} correct we first derive a rule to justify that we can assume both specifications to verify the bodies. This rule can be derived from the general @{thm [source] HoarePartial.ProcRec} rule. An ML function does this work: \ ML \ML_Thms.bind_thm ("ProcRec2", Hoare.gen_proc_rec @{context} Hoare.Partial 2)\ lemma (in odd_even_clique) shows odd_spec: "\n. \\\\N=n\ \A :== PROC odd(\N) \(\b. n = 2 * b + \A) \ \A < 2 \" (is ?P1) and even_spec: "\n. \\\\N=n\ \A :== PROC even(\N) \(\b. n + 1 = 2 * b + \A) \ \A < 2 \" (is ?P2) proof - have "?P1 \ ?P2" apply (hoare_rule ProcRec2) apply vcg apply clarsimp apply (rule_tac x="b + 1" in exI) apply arith apply vcg apply clarsimp apply arith done thus "?P1" "?P2" by iprover+ qed subsection \Expressions With Side Effects\ text \\texttt{R := N++ + M++}\ lemma "\\ \True\ \N \ n. \N :== \N + 1 \ \M \ m. \M :== \M + 1 \ \R :== n + m \\R = \N + \M - 2\" apply vcg apply simp done text \\texttt{R := Fac (N) + Fac (M)}\ lemma (in Fac_impl) shows "\\ \True\ CALL Fac(\N) \ n. CALL Fac(\M) \ m. \R :== n + m \\R = fac \N + fac \M\" apply vcg done text \\texttt{ R := (Fac(Fac (N)))}\ lemma (in Fac_impl) shows "\\ \True\ CALL Fac(\N) \ n. CALL Fac(n) \ m. \R :== m \\R = fac (fac \N)\" apply vcg done subsection \Global Variables and Heap\ text \ Now we define and verify some procedures on heap-lists. We consider list structures consisting of two fields, a content element @{term "cont"} and a reference to the next list element @{term "next"}. We model this by the following state space where every field has its own heap. \ record globals_list = next_' :: "ref \ ref" cont_' :: "ref \ nat" record 'g list_vars = "'g state" + p_' :: "ref" q_' :: "ref" r_' :: "ref" root_' :: "ref" tmp_' :: "ref" text \Updates to global components inside a procedure will always be propagated to the caller. This is implicitly done by the parameter passing syntax translations. The record containing the global variables must begin with the prefix "globals". \ text \We first define an append function on lists. It takes two references as parameters. It appends the list referred to by the first parameter with the list referred to by the second parameter, and returns the result right into the first parameter. \ procedures append(p,q|p) = "IF \p=Null THEN \p :== \q ELSE \p \\next:== CALL append(\p\\next,\q) FI" (* append_spec: "\\ Ps Qs. \\ \\. List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {}\ \p :== PROC append(\p,\q) \List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" append_modifies: "\\. \\ {\} \p :== PROC append(\p,\q){t. t may_only_modify_globals \ in [next]}" *) context append_impl begin declare [[hoare_use_call_tr' = false]] term "CALL append(\p,\q,\p\\next)" declare [[hoare_use_call_tr' = true]] end text \Below we give two specifications this time. One captures the functional behaviour and focuses on the entities that are potentially modified by the procedure, the other one is a pure frame condition. The list in the modifies clause has to list all global state components that may be changed by the procedure. Note that we know from the modifies clause that the @{term cont} parts of the lists will not be changed. Also a small side note on the syntax. We use ordinary brackets in the postcondition of the modifies clause, and also the state components do not carry the acute, because we explicitly note the state @{term t} here. The functional specification now introduces two logical variables besides the state space variable @{term "\"}, namely @{term "Ps"} and @{term "Qs"}. They are universally quantified and range over both the pre and the postcondition, so that we are able to properly instantiate the specification during the proofs. The syntax \\\. \\\ is a shorthand to fix the current state: \{s. \ = s \}\. \ lemma (in append_impl) append_spec: shows "\\ Ps Qs. \\ \\. List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {}\ \p :== PROC append(\p,\q) \List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoarePartial.ProcRec1) apply vcg apply fastforce done text \The modifies clause is equal to a proper record update specification of the following form. \ lemma "{t. t may_only_modify_globals Z in [next]} = {t. \next. globals t=next_'_update (\_. next) (globals Z)}" apply (unfold mex_def meq_def) apply (simp) done text \If the verification condition generator works on a procedure call it checks whether it can find a modified clause in the context. If one is present the procedure call is simplified before the Hoare rule @{thm [source] HoarePartial.ProcSpec} is applied. Simplification of the procedure call means, that the ``copy back'' of the global components is simplified. Only those components that occur in the modifies clause will actually be copied back. This simplification is justified by the rule @{thm [source] HoarePartial.ProcModifyReturn}. So after this simplification all global components that do not appear in the modifies clause will be treated as local variables. \ text \You can study the effect of the modifies clause on the following two examples, where we want to prove that @{term "append"} does not change the @{term "cont"} part of the heap. \ lemma (in append_impl) shows "\\ \\p=Null \ \cont=c\ \p :== CALL append(\p,Null) \\cont=c\" apply vcg oops text \To prove the frame condition, we have to tell the verification condition generator to use only the modifies clauses and not to search for functional specifications by the parameter \spec=modifies\ It will also try to solve the verification conditions automatically. \ lemma (in append_impl) append_modifies: shows "\\. \\ {\} \p :== PROC append(\p,\q){t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcRec1) apply (vcg spec=modifies) done lemma (in append_impl) shows "\\ \\p=Null \ \cont=c\ \p\\next :== CALL append(\p,Null) \\cont=c\" apply vcg apply simp done text \ Of course we could add the modifies clause to the functional specification as well. But separating both has the advantage that we split up the verification work. We can make use of the modifies clause before we apply the functional specification in a fully automatic fashion. \ text \To verify the body of @{term "append"} we do not need the modifies clause, since the specification does not talk about @{term "cont"} at all, and we don't access @{term "cont"} inside the body. This may be different for more complex procedures. \ text \ To prove that a procedure respects the modifies clause, we only need the modifies clauses of the procedures called in the body. We do not need the functional specifications. So we can always prove the modifies clause without functional specifications, but me may need the modifies clause to prove the functional specifications. \ subsubsection \Insertion Sort\ primrec sorted:: "('a \ 'a \ bool) \ 'a list \ bool" where "sorted le [] = True" | "sorted le (x#xs) = ((\y\set xs. le x y) \ sorted le xs)" procedures insert(r,p | p) = "IF \r=Null THEN SKIP ELSE IF \p=Null THEN \p :== \r;; \p\\next :== Null ELSE IF \r\\cont \ \p\\cont THEN \r\\next :== \p;; \p:==\r ELSE \p\\next :== CALL insert(\r,\p\\next) FI FI FI" text \ In the postcondition of the functional specification there is a small but important subtlety. Whenever we talk about the @{term "cont"} part we refer to the one of the pre-state, even in the conclusion of the implication. The reason is, that we have separated out, that @{term "cont"} is not modified by the procedure, to the modifies clause. So whenever we talk about unmodified parts in the postcondition we have to use the pre-state part, or explicitly state an equality in the postcondition. The reason is simple. If the postcondition would talk about \\cont\ instead of \\<^bsup>\\<^esup>cont\, we get a new instance of \cont\ during verification and the postcondition would only state something about this new instance. But as the verification condition generator uses the modifies clause the caller of \insert\ instead still has the old \cont\ after the call. Thats the very reason for the modifies clause. So the caller and the specification will simply talk about two different things, without being able to relate them (unless an explicit equality is added to the specification). \ lemma (in insert_impl) insert_modifies: "\\. \\ {\} \p :== PROC insert(\r,\p){t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcRec1) apply (vcg spec=modifies) done lemma (in insert_impl) insert_spec: "\\ Ps . \\ \\. List \p \next Ps \ sorted (\) (map \cont Ps) \ \r \ Null \ \r \ set Ps\ \p :== PROC insert(\r,\p) \\Qs. List \p \next Qs \ sorted (\) (map \<^bsup>\\<^esup>cont Qs) \ set Qs = insert \<^bsup>\\<^esup>r (set Ps) \ (\x. x \ set Qs \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoarePartial.ProcRec1) apply vcg apply (intro conjI impI) apply fastforce apply fastforce apply fastforce apply (clarsimp) apply force done procedures insertSort(p | p) = "\r:==Null;; WHILE (\p \ Null) DO \q :== \p;; \p :== \p\\next;; \r :== CALL insert(\q,\r) OD;; \p:==\r" lemma (in insertSort_impl) insertSort_modifies: shows "\\. \\ {\} \p :== PROC insertSort(\p) {t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcRec1) apply (vcg spec=modifies) done text \Insertion sort is not implemented recursively here but with a while loop. Note that the while loop is not annotated with an invariant in the procedure definition. The invariant only comes into play during verification. Therefore we will annotate the body during the proof with the rule @{thm [source] HoarePartial.annotateI}. \ lemma (in insertSort_impl) insertSort_body_spec: shows "\\ Ps. \,\\ \\. List \p \next Ps \ \p :== PROC insertSort(\p) \\Qs. List \p \next Qs \ sorted (\) (map \<^bsup>\\<^esup>cont Qs) \ set Qs = set Ps\" apply (hoare_rule HoarePartial.ProcRec1) apply (hoare_rule anno= "\r :== Null;; WHILE \p \ Null INV \\Qs Rs. List \p \next Qs \ List \r \next Rs \ set Qs \ set Rs = {} \ sorted (\) (map \cont Rs) \ set Qs \ set Rs = set Ps \ \cont = \<^bsup>\\<^esup>cont \ DO \q :== \p;; \p :== \p\\next;; \r :== CALL insert(\q,\r) OD;; \p :== \r" in HoarePartial.annotateI) apply vcg apply fastforce prefer 2 apply fastforce apply (clarsimp) apply (rule_tac x=ps in exI) apply (intro conjI) apply (rule heap_eq_ListI1) apply assumption apply clarsimp apply (subgoal_tac "x\p \ x \ set Rs") apply auto done subsubsection "Memory Allocation and Deallocation" text \The basic idea of memory management is to keep a list of allocated references in the state space. Allocation of a new reference adds a new reference to the list deallocation removes a reference. Moreover we keep a counter "free" for the free memory. \ record globals_list_alloc = globals_list + alloc_'::"ref list" free_'::nat record 'g list_vars' = "'g list_vars" + i_'::nat first_'::ref definition "sz = (2::nat)" text \Restrict locale \hoare\ to the required type.\ locale hoare_ex = hoare \ for \ :: "'c \ (('a globals_list_alloc_scheme, 'b) list_vars'_scheme, 'c, 'd) com" lemma (in hoare_ex) "\\ \\i = 0 \ \first = Null \ n*sz \ \free\ WHILE \i < n INV \\Ps. List \first \next Ps \ length Ps = \i \ \i \ n \ set Ps \ set \alloc \ (n - \i)*sz \ \free\ DO \p :== NEW sz [\cont:==0,\next:== Null];; \p\\next :== \first;; \first :== \p;; \i :== \i+ 1 OD \\Ps. List \first \next Ps \ length Ps = n \ set Ps \ set \alloc\" apply (vcg) apply simp apply clarsimp apply (rule conjI) apply clarsimp apply (rule_tac x="new (set alloc)#Ps" in exI) apply clarsimp apply (rule conjI) apply fastforce apply (simp add: sz_def) apply (simp add: sz_def) apply fastforce done lemma (in hoare_ex) "\\ \\i = 0 \ \first = Null \ n*sz \ \free\ WHILE \i < n INV \\Ps. List \first \next Ps \ length Ps = \i \ \i \ n \ set Ps \ set \alloc \ (n - \i)*sz \ \free\ DO \p :== NNEW sz [\cont:==0,\next:== Null];; \p\\next :== \first;; \first :== \p;; \i :== \i+ 1 OD \\Ps. List \first \next Ps \ length Ps = n \ set Ps \ set \alloc\" apply (vcg) apply simp apply clarsimp apply (rule conjI) apply clarsimp apply (rule_tac x="new (set alloc)#Ps" in exI) apply clarsimp apply (rule conjI) apply fastforce apply (simp add: sz_def) apply (simp add: sz_def) apply fastforce done subsection \Fault Avoiding Semantics\ text \ If we want to ensure that no runtime errors occur we can insert guards into the code. We will not be able to prove any nontrivial Hoare triple about code with guards, if we cannot show that the guards will never fail. A trivial hoare triple is one with an empty precondition. \ lemma "\\ \True\ \\p\Null\\ \p\\next :== \p \True\" apply vcg oops lemma "\\ {} \\p\Null\\ \p\\next :== \p \True\" apply vcg done text \Let us consider this small program that reverts a list. At first without guards. \ lemma (in hoare_ex) rev_strip: "\\ \List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {} \ set Ps \ set \alloc \ set Qs \ set \alloc\ WHILE \p \ Null INV \\ps qs. List \p \next ps \ List \q \next qs \ set ps \ set qs = {} \ rev ps @ qs = rev Ps @ Qs \ set ps \ set \alloc \ set qs \ set \alloc\ DO \r :== \p;; \p :== \p\ \next;; \r\\next :== \q;; \q :== \r OD \List \q \next (rev Ps @ Qs) \ set Ps\ set \alloc \ set Qs \ set \alloc\" apply (vcg) apply fastforce+ done text \If we want to ensure that we do not dereference @{term "Null"} or access unallocated memory, we have to add some guards. \ locale hoare_ex_guard = hoare \ for \ :: "'c \ (('a globals_list_alloc_scheme, 'b) list_vars'_scheme, 'c, bool) com" lemma (in hoare_ex_guard) "\\ \List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {} \ set Ps \ set \alloc \ set Qs \ set \alloc\ WHILE \p \ Null INV \\ps qs. List \p \next ps \ List \q \next qs \ set ps \ set qs = {} \ rev ps @ qs = rev Ps @ Qs \ set ps \ set \alloc \ set qs \ set \alloc\ DO \r :== \p;; \\p\Null \ \p\set \alloc\\ \p :== \p\ \next;; \\r\Null \ \r\set \alloc\\ \r\\next :== \q;; \q :== \r OD \List \q \next (rev Ps @ Qs) \ set Ps \ set \alloc \ set Qs \ set \alloc\" apply (vcg) apply fastforce+ done text \We can also just prove that no faults will occur, by giving the trivial postcondition. \ lemma (in hoare_ex_guard) rev_noFault: "\\ \List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {} \ set Ps \ set \alloc \ set Qs \ set \alloc\ WHILE \p \ Null INV \\ps qs. List \p \next ps \ List \q \next qs \ set ps \ set qs = {} \ rev ps @ qs = rev Ps @ Qs \ set ps \ set \alloc \ set qs \ set \alloc\ DO \r :== \p;; \\p\Null \ \p\set \alloc\\ \p :== \p\ \next;; \\r\Null \ \r\set \alloc\\ \r\\next :== \q;; \q :== \r OD UNIV,UNIV" apply (vcg) apply fastforce+ done lemma (in hoare_ex_guard) rev_moduloGuards: "\\\<^bsub>/{True}\<^esub> \List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {} \ set Ps \ set \alloc \ set Qs \ set \alloc\ WHILE \p \ Null INV \\ps qs. List \p \next ps \ List \q \next qs \ set ps \ set qs = {} \ rev ps @ qs = rev Ps @ Qs \ set ps \ set \alloc \ set qs \ set \alloc\ DO \r :== \p;; \\p\Null \ \p\set \alloc\\ \ \p :== \p\ \next;; \\r\Null \ \r\set \alloc\\ \ \r\\next :== \q;; \q :== \r OD \List \q \next (rev Ps @ Qs) \ set Ps \ set \alloc \ set Qs \ set \alloc\" apply vcg apply fastforce+ done lemma CombineStrip': assumes deriv: "\,\\\<^bsub>/F\<^esub> P c' Q,A" assumes deriv_strip: "\,\\\<^bsub>/{}\<^esub> P c'' UNIV,UNIV" assumes c'': "c''= mark_guards False (strip_guards (-F) c')" assumes c: "c = mark_guards False c'" shows "\,\\\<^bsub>/{}\<^esub> P c Q,A" proof - from deriv_strip [simplified c''] have "\,\\ P (strip_guards (- F) c') UNIV,UNIV" by (rule HoarePartialProps.MarkGuardsD) with deriv have "\,\\ P c' Q,A" by (rule HoarePartialProps.CombineStrip) hence "\,\\ P mark_guards False c' Q,A" by (rule HoarePartialProps.MarkGuardsI) thus ?thesis by (simp add: c) qed text \We can then combine the prove that no fault will occur with the functional proof of the programme without guards to get the full prove by the rule @{thm HoarePartialProps.CombineStrip} \ lemma (in hoare_ex_guard) "\\ \List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {} \ set Ps \ set \alloc \ set Qs \ set \alloc\ WHILE \p \ Null INV \\ps qs. List \p \next ps \ List \q \next qs \ set ps \ set qs = {} \ rev ps @ qs = rev Ps @ Qs \ set ps \ set \alloc \ set qs \ set \alloc\ DO \r :== \p;; \\p\Null \ \p\set \alloc\\ \p :== \p\ \next;; \\r\Null \ \r\set \alloc\\ \r\\next :== \q;; \q :== \r OD \List \q \next (rev Ps @ Qs) \ set Ps \ set \alloc \ set Qs \ set \alloc\" apply (rule CombineStrip' [OF rev_moduloGuards rev_noFault]) apply simp apply simp done text \In the previous example the effort to split up the prove did not really pay off. But when we think of programs with a lot of guards and complicated specifications it may be better to first focus on a prove without the messy guards. Maybe it is possible to automate the no fault proofs so that it suffices to focus on the stripped program. \ text \ The purpose of guards is to watch for faults that can occur during evaluation of expressions. In the example before we watched for null pointer dereferencing or memory faults. We can also look for array index bounds or division by zero. As the condition of a while loop is evaluated in each iteration we cannot just add a guard before the while loop. Instead we need a special guard for the condition. Example: @{term "WHILE \\p\Null\\ \p\\next\Null DO SKIP OD"} \ subsection \Circular Lists\ definition distPath :: "ref \ (ref \ ref) \ ref \ ref list \ bool" where "distPath x next y as = (Path x next y as \ distinct as)" lemma neq_dP: "\p \ q; Path p h q Ps; distinct Ps\ \ \Qs. p\Null \ Ps = p#Qs \ p \ set Qs" by (cases Ps, auto) lemma circular_list_rev_I: "\\ \\root = r \ distPath \root \next \root (r#Ps)\ \p :== \root;; \q :== \root\\next;; WHILE \q \ \root INV \\ ps qs. distPath \p \next \root ps \ distPath \q \next \root qs \ \root = r \ r\Null \ r \ set Ps \ set ps \ set qs = {} \ Ps = (rev ps) @ qs \ DO \tmp :== \q;; \q :== \q\\next;; \tmp\\next :== \p;; \p:==\tmp OD;; \root\\next :== \p \\root = r \ distPath \root \next \root (r#rev Ps)\" apply (simp only:distPath_def) apply vcg apply (rule_tac x="[]" in exI) apply fastforce apply clarsimp apply (drule (2) neq_dP) apply (rule_tac x="q # ps" in exI) apply clarsimp apply fastforce done lemma path_is_list:"\a next b. \Path b next a Ps ; a \ set Ps; a\Null\ \ List b (next(a := Null)) (Ps @ [a])" apply (induct Ps) apply (auto simp add:fun_upd_apply) done text \ The simple algorithm for acyclic list reversal, with modified annotations, works for cyclic lists as well.: \ lemma circular_list_rev_II: "\\ \\p = r \ distPath \p \next \p (r#Ps)\ \q:==Null;; WHILE \p \ Null INV \ ((\q = Null) \ (\ps. distPath \p \next r ps \ ps = r#Ps)) \ ((\q \ Null) \ (\ps qs. distPath \q \next r qs \ List \p \next ps \ set ps \ set qs = {} \ rev qs @ ps = Ps@[r])) \ \ (\p = Null \ \q = Null \ r = Null ) \ DO \tmp :== \p;; \p :== \p\\next;; \tmp\\next :== \q;; \q:==\tmp OD \\q = r \ distPath \q \next \q (r # rev Ps)\" apply (simp only:distPath_def) apply vcg apply clarsimp apply clarsimp apply (case_tac "(q = Null)") apply (fastforce intro: path_is_list) apply clarify apply (rule_tac x="psa" in exI) apply (rule_tac x=" p # qs" in exI) apply force apply fastforce done text\Although the above algorithm is more succinct, its invariant looks more involved. The reason for the case distinction on @{term q} is due to the fact that during execution, the pointer variables can point to either cyclic or acyclic structures. \ text \ When working on lists, its sometimes better to remove @{thm[source] fun_upd_apply} from the simpset, and instead include @{thm[source] fun_upd_same} and @{thm[source] fun_upd_other} to the simpset \ (* declare fun_upd_apply[simp del]fun_upd_same[simp] fun_upd_other[simp] *) lemma "\\ {\} \I :== \M;; ANNO \. \\. \I = \<^bsup>\\<^esup>M\ \M :== \N;; \N :== \I \\M = \<^bsup>\\<^esup>N \ \N = \<^bsup>\\<^esup>I\ \\M = \<^bsup>\\<^esup>N \ \N = \<^bsup>\\<^esup>M\" apply vcg apply auto done lemma "\\ ({\} \ \\M = 0 \ \S = 0\) (ANNO \. ({\} \ \\A=\<^bsup>\\<^esup>A \ \I=\<^bsup>\\<^esup>I \ \M=0 \ \S=0\) WHILE \M \ \A INV \\S = \M * \I \ \A=\<^bsup>\\<^esup>A \ \I=\<^bsup>\\<^esup>I\ DO \S :== \S + \I;; \M :== \M + 1 OD \\S = \<^bsup>\\<^esup>A * \<^bsup>\\<^esup>I\) \\S = \<^bsup>\\<^esup>A * \<^bsup>\\<^esup>I\" apply vcg_step apply vcg_step apply simp apply vcg_step apply vcg_step apply simp apply vcg apply simp apply simp apply vcg_step apply auto done text \Instead of annotations one can also directly use previously proven lemmas.\ lemma foo_lemma: "\n m. \\ \\N = n \ \M = m\ \N :== \N + 1;; \M :== \M + 1 \\N = n + 1 \ \M = m + 1\" by vcg lemma "\\ \\N = n \ \M = m\ LEMMA foo_lemma \N :== \N + 1;; \M :== \M + 1 END;; \N :== \N + 1 \\N = n + 2 \ \M = m + 1\" apply vcg apply simp done lemma "\\ \\N = n \ \M = m\ LEMMA foo_lemma \N :== \N + 1;; \M :== \M + 1 END;; LEMMA foo_lemma \N :== \N + 1;; \M :== \M + 1 END \\N = n + 2 \ \M = m + 2\" apply vcg apply simp done lemma "\\ \\N = n \ \M = m\ \N :== \N + 1;; \M :== \M + 1;; \N :== \N + 1;; \M :== \M + 1 \\N = n + 2 \ \M = m + 2\" apply (hoare_rule anno= "LEMMA foo_lemma \N :== \N + 1;; \M :== \M + 1 END;; LEMMA foo_lemma \N :== \N + 1;; \M :== \M + 1 END" in HoarePartial.annotate_normI) apply vcg apply simp done text \Just some test on marked, guards\ lemma "\\\True\ WHILE \P \N \\, \Q \M\#, \R \N\\ \N < \M INV \\N < 2\ DO \N :== \M OD \hard\" apply vcg oops lemma "\\\<^bsub>/{True}\<^esub> \True\ WHILE \P \N \\, \Q \M\#, \R \N\\ \N < \M INV \\N < 2\ DO \N :== \M OD \hard\" apply vcg oops term "\\\<^bsub>/{True}\<^esub> \True\ WHILE\<^sub>g \N < \Arr!i FIX Z. INV \\N < 2\ DO \N :== \M OD \hard\" lemma "\\\<^bsub>/{True}\<^esub> \True\ WHILE\<^sub>g \N < \Arr!i FIX Z. INV \\N < 2\ VAR arbitrary DO \N :== \M OD \hard\" apply vcg oops lemma "\\\<^bsub>/{True}\<^esub> \True\ WHILE \P \N \\, \Q \M\#, \R \N\\ \N < \M FIX Z. INV \\N < 2\ VAR arbitrary DO \N :== \M OD \hard\" apply vcg oops end diff --git a/thys/Simpl/ex/VcgExSP.thy b/thys/Simpl/ex/VcgExSP.thy --- a/thys/Simpl/ex/VcgExSP.thy +++ b/thys/Simpl/ex/VcgExSP.thy @@ -1,1388 +1,1367 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: VcgEx.thy - Author: Norbert Schirmer, TU Muenchen - -Copyright (C) 2004-2008 Norbert Schirmer -Some rights reserved, TU Muenchen -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (C) 2006-2008 Norbert Schirmer *) section \Examples using Statespaces\ theory VcgExSP imports "../HeapList" "../Vcg" begin subsection \State Spaces\ text \ First of all we provide a store of program variables that occur in the programs considered later. Slightly unexpected things may happen when attempting to work with undeclared variables. \ hoarestate state_space = A :: nat I :: nat M :: nat N :: nat R :: nat S :: nat B :: bool Abr:: string lemma (in state_space)"\\ \\N = n\ LOC \N :== 10;; \N :== \N + 2 COL \\N = n\" by vcg text \Internally we decorate the state components in the statespace with the suffix \_'\, to avoid cluttering the namespace with the simple names that could no longer be used for logical variables otherwise. \ text \We will first consider programs without procedures, later on we will regard procedures without global variables and finally we will get the full pictures: mutually recursive procedures with global variables (including heap). \ subsection \Basic Examples\ text \ We look at few trivialities involving assignment and sequential composition, in order to get an idea of how to work with our formulation of Hoare Logic. \ text \ Using the basic rule directly is a bit cumbersome. \ lemma (in state_space) "\\ {|\N = 5|} \N :== 2 * \N {|\N = 10|}" apply (rule HoarePartial.Basic) apply simp done lemma (in state_space) "\\ \True\ \N :== 10 \\N = 10\" by vcg lemma (in state_space) "\\ \2 * \N = 10\ \N :== 2 * \N \\N = 10\" by vcg lemma (in state_space) "\\ \\N = 5\ \N :== 2 * \N \\N = 10\" apply vcg apply simp done lemma (in state_space) "\\ \\N + 1 = a + 1\ \N :== \N + 1 \\N = a + 1\" by vcg lemma (in state_space) "\\ \\N = a\ \N :== \N + 1 \\N = a + 1\" apply vcg apply simp done lemma (in state_space) shows "\\ \a = a \ b = b\ \M :== a;; \N :== b \\M = a \ \N = b\" by vcg lemma (in state_space) shows "\\ \True\ \M :== a;; \N :== b \\M = a \ \N = b\" by vcg lemma (in state_space) shows "\\ \\M = a \ \N = b\ \I :== \M;; \M :== \N;; \N :== \I \\M = b \ \N = a\" apply vcg apply simp done text \ We can also perform verification conditions generation step by step by using the \vcg_step\ method. \ lemma (in state_space) shows "\\ \\M = a \ \N = b\ \I :== \M;; \M :== \N;; \N :== \I \\M = b \ \N = a\" apply vcg_step apply vcg_step apply vcg_step apply vcg_step apply simp done text \ In the following assignments we make use of the consequence rule in order to achieve the intended precondition. Certainly, the \vcg\ method is able to handle this case, too. \ lemma (in state_space) shows "\\ \\M = \N\ \M :== \M + 1 \\M \ \N\" proof - have "\\M = \N\ \ \\M + 1 \ \N\" by auto also have "\\ \ \M :== \M + 1 \\M \ \N\" by vcg finally show ?thesis . qed lemma (in state_space) shows "\\ \\M = \N\ \M :== \M + 1 \\M \ \N\" proof - have "\m n::nat. m = n \ m + 1 \ n" \ \inclusion of assertions expressed in ``pure'' logic,\ \ \without mentioning the state space\ by simp also have "\\ \\M + 1 \ \N\ \M :== \M + 1 \\M \ \N\" by vcg finally show ?thesis . qed lemma (in state_space) shows "\\ \\M = \N\ \M :== \M + 1 \\M \ \N\" apply vcg apply simp done subsection \Multiplication by Addition\ text \ We now do some basic examples of actual \texttt{WHILE} programs. This one is a loop for calculating the product of two natural numbers, by iterated addition. We first give detailed structured proof based on single-step Hoare rules. \ lemma (in state_space) shows "\\ \\M = 0 \ \S = 0\ WHILE \M \ a DO \S :== \S + b;; \M :== \M + 1 OD \\S = a * b\" proof - let "\\ _ ?while _" = ?thesis let "\\?inv\" = "\\S = \M * b\" have "\\M = 0 & \S = 0\ \ \\?inv\" by auto also have "\\ \ ?while \\?inv \ \ (\M \ a)\" proof let ?c = "\S :== \S + b;; \M :== \M + 1" have "\\?inv \ \M \ a\ \ \\S + b = (\M + 1) * b\" by auto also have "\\ \ ?c \\?inv\" by vcg finally show "\\ \\?inv \ \M \ a\ ?c \\?inv\" . qed also have "\\?inv \ \ (\M \ a)\ \ \\S = a * b\" by auto finally show ?thesis by blast qed text \ The subsequent version of the proof applies the \vcg\ method to reduce the Hoare statement to a purely logical problem that can be solved fully automatically. Note that we have to specify the \texttt{WHILE} loop invariant in the original statement. \ lemma (in state_space) shows "\\ \\M = 0 \ \S = 0\ WHILE \M \ a INV \\S = \M * b\ DO \S :== \S + b;; \M :== \M + 1 OD \\S = a * b\" apply vcg apply auto done text \Here some examples of ``breaking'' out of a loop\ lemma (in state_space) shows "\\ \\M = 0 \ \S = 0\ TRY WHILE True INV \\S = \M * b\ DO IF \M = a THEN THROW ELSE \S :== \S + b;; \M :== \M + 1 FI OD CATCH SKIP END \\S = a * b\" apply vcg apply auto done lemma (in state_space) shows "\\ \\M = 0 \ \S = 0\ TRY WHILE True INV \\S = \M * b\ DO IF \M = a THEN \Abr :== ''Break'';;THROW ELSE \S :== \S + b;; \M :== \M + 1 FI OD CATCH IF \Abr = ''Break'' THEN SKIP ELSE Throw FI END \\S = a * b\" apply vcg apply auto done text \Some more syntactic sugar, the label statement \\ \ \\ as shorthand for the \TRY-CATCH\ above, and the \RAISE\ for an state-update followed by a \THROW\. \ lemma (in state_space) shows "\\ \\M = 0 \ \S = 0\ \\Abr = ''Break''\\ WHILE True INV \\S = \M * b\ DO IF \M = a THEN RAISE \Abr :== ''Break'' ELSE \S :== \S + b;; \M :== \M + 1 FI OD \\S = a * b\" apply vcg apply auto done lemma (in state_space) shows "\\ \\M = 0 \ \S = 0\ TRY WHILE True INV \\S = \M * b\ DO IF \M = a THEN RAISE \Abr :== ''Break'' ELSE \S :== \S + b;; \M :== \M + 1 FI OD CATCH IF \Abr = ''Break'' THEN SKIP ELSE Throw FI END \\S = a * b\" apply vcg apply auto done lemma (in state_space) shows "\\ \\M = 0 \ \S = 0\ \\Abr = ''Break''\ \ WHILE True INV \\S = \M * b\ DO IF \M = a THEN RAISE \Abr :== ''Break'' ELSE \S :== \S + b;; \M :== \M + 1 FI OD \\S = a * b\" apply vcg apply auto done text \Blocks\ lemma (in state_space) shows "\\\\I = i\ LOC \I;; \I :== 2 COL \\I \ i\" apply vcg by simp subsection \Summing Natural Numbers\ text \ We verify an imperative program to sum natural numbers up to a given limit. First some functional definition for proper specification of the problem. \ primrec sum :: "(nat => nat) => nat => nat" where "sum f 0 = 0" | "sum f (Suc n) = f n + sum f n" syntax "_sum" :: "idt => nat => nat => nat" ("SUMM _<_. _" [0, 0, 10] 10) translations "SUMM jj. b) k" text \ The following proof is quite explicit in the individual steps taken, with the \vcg\ method only applied locally to take care of assignment and sequential composition. Note that we express intermediate proof obligation in pure logic, without referring to the state space. \ theorem (in state_space) shows "\\ \True\ \S :== 0;; \I :== 1;; WHILE \I \ n DO \S :== \S + \I;; \I :== \I + 1 OD \\S = (SUMM j" (is "\\ _ (_;; ?while) _") proof - let ?sum = "\k. SUMM j\ \True\ \S :== 0;; \I :== 1 \?inv \S \I\" proof - have "True \ 0 = ?sum 1" by simp also have "\\ \\\ \S :== 0;; \I :== 1 \?inv \S \I\" by vcg finally show ?thesis . qed also have "\\ \?inv \S \I\ ?while \?inv \S \I \ \ \I \ n\" proof let ?body = "\S :== \S + \I;; \I :== \I + 1" have "\s i. ?inv s i \ i \ n \ ?inv (s + i) (i + 1)" by simp also have "\\ \\S + \I = ?sum (\I + 1)\ ?body \?inv \S \I\" by vcg finally show "\\ \?inv \S \I \ \I \ n\ ?body \?inv \S \I\" . qed also have "\s i. s = ?sum i \ \ i \ n \ s = ?sum n" by simp finally show ?thesis . qed text \ The next version uses the \vcg\ method, while still explaining the resulting proof obligations in an abstract, structured manner. \ theorem (in state_space) shows "\\ \True\ \S :== 0;; \I :== 1;; WHILE \I \ n INV \\S = (SUMM j<\I. j)\ DO \S :== \S + \I;; \I :== \I + 1 OD \\S = (SUMM j" proof - let ?sum = "\k. SUMM j n" thus "?inv (s + i) (i + 1)" by simp next fix i s assume x: "?inv s i" "\ i \ n" thus "s = ?sum n" by simp qed qed text \ Certainly, this proof may be done fully automatically as well, provided that the invariant is given beforehand. \ theorem (in state_space) shows "\\ \True\ \S :== 0;; \I :== 1;; WHILE \I \ n INV \\S = (SUMM j<\I. j)\ DO \S :== \S + \I;; \I :== \I + 1 OD \\S = (SUMM j" apply vcg apply auto done subsection \SWITCH\ lemma (in state_space) shows "\\ \\N = 5\ SWITCH \B {True} \ \N :== 6 | {False} \ \N :== 7 END \\N > 5\" apply vcg apply simp done lemma (in state_space) shows "\\ \\N = 5\ SWITCH \N {v. v < 5} \ \N :== 6 | {v. v \ 5} \ \N :== 7 END \\N > 5\" apply vcg apply simp done subsection \(Mutually) Recursive Procedures\ subsubsection \Factorial\ text \We want to define a procedure for the factorial. We first define a HOL functions that calculates it to specify the procedure later on. \ primrec fac:: "nat \ nat" where "fac 0 = 1" | "fac (Suc n) = (Suc n) * fac n" lemma fac_simp [simp]: "0 < i \ fac i = i * fac (i - 1)" by (cases i) simp_all text \Now we define the procedure\ procedures Fac (N::nat|R::nat) "IF \N = 0 THEN \R :== 1 ELSE \R :== CALL Fac(\N - 1);; \R :== \N * \R FI" print_locale Fac_impl text \ To see how a call is syntactically translated you can switch off the printing translation via the configuration option \hoare_use_call_tr'\ \ context Fac_impl begin text \ @{term "CALL Fac(\N,\R)"} is internally: \ declare [[hoare_use_call_tr' = false]] text \ @{term "CALL Fac(\N,\R)"} \ term "CALL Fac(\N,\R)" declare [[hoare_use_call_tr' = true]] text \ Now let us prove that @{term "Fac"} meets its specification. \ end lemma (in Fac_impl) Fac_spec': shows "\\. \,\\{\} PROC Fac(\N,\R) \\R = fac \<^bsup>\\<^esup>N\" apply (hoare_rule HoarePartial.ProcRec1) apply vcg apply simp done text \ Since the factorial was implemented recursively, the main ingredient of this proof is, to assume that the specification holds for the recursive call of @{term Fac} and prove the body correct. The assumption for recursive calls is added to the context by the rule @{thm [source] HoarePartial.ProcRec1} (also derived from general rule for mutually recursive procedures): @{thm [display] HoarePartial.ProcRec1 [no_vars]} The verification condition generator will infer the specification out of the context when it encounters a recursive call of the factorial. \ text \We can also step through verification condition generation. When the verification condition generator encounters a procedure call it tries to use the rule \ProcSpec\. To be successful there must be a specification of the procedure in the context. \ lemma (in Fac_impl) Fac_spec1: shows "\\. \,\\{\} \R :== PROC Fac(\N) \\R = fac \<^bsup>\\<^esup>N\" apply (hoare_rule HoarePartial.ProcRec1) apply vcg_step apply vcg_step apply vcg_step apply vcg_step apply vcg_step apply simp done text \Here some Isar style version of the proof\ lemma (in Fac_impl) Fac_spec2: shows "\\. \,\\{\} \R :== PROC Fac(\N) \\R = fac \<^bsup>\\<^esup>N\" proof (hoare_rule HoarePartial.ProcRec1) have Fac_spec: "\\. \,(\\(\\. {({\}, Fac_'proc, \\R = fac \<^bsup>\\<^esup>N\,{})})) \ {\} \R :== PROC Fac(\N) \\R = fac \<^bsup>\\<^esup>N\" apply (rule allI) apply (rule hoarep.Asm) by simp show "\\. \,(\\(\\. {({\}, Fac_'proc, \\R = fac \<^bsup>\\<^esup>N\,{})})) \ {\} IF \N = 0 THEN \R :== 1 ELSE \R :== CALL Fac(\N - 1);; \R :== \N * \R FI \\R = fac \<^bsup>\\<^esup>N\" apply vcg apply simp done qed text \To avoid retyping of potentially large pre and postconditions in the previous proof we can use the casual term abbreviations of the Isar language. \ lemma (in Fac_impl) Fac_spec3: shows "\\. \,\\{\} \R :== PROC Fac(\N) \\R = fac \<^bsup>\\<^esup>N\" (is "\\. \,\\(?Pre \) ?Fac (?Post \)") proof (hoare_rule HoarePartial.ProcRec1) have Fac_spec: "\\. \,(\\(\\. {(?Pre \, Fac_'proc, ?Post \,{})})) \(?Pre \) ?Fac (?Post \)" apply (rule allI) apply (rule hoarep.Asm) by simp show "\\. \,(\\(\\. {(?Pre \, Fac_'proc, ?Post \,{})})) \ (?Pre \) IF \N = 0 THEN \R :== 1 ELSE \R :== CALL Fac(\N - 1);; \R :== \N * \R FI (?Post \)" apply vcg apply simp done qed text \The previous proof pattern has still some kind of inconvenience. The augmented context is always printed in the proof state. That can mess up the state, especially if we have large specifications. This may be annoying if we want to develop single step or structured proofs. In this case it can be a good idea to introduce a new variable for the augmented context. \ lemma (in Fac_impl) Fac_spec4: shows "\\. \,\\{\} \R :== PROC Fac(\N) \\R = fac \<^bsup>\\<^esup>N\" (is "\\. \,\\(?Pre \) ?Fac (?Post \)") proof (hoare_rule HoarePartial.ProcRec1) define \' where "\' = \ \ (\\. {(?Pre \, Fac_'proc, ?Post \,{})})" have Fac_spec: "\\. \,\'\(?Pre \) ?Fac (?Post \)" by (unfold \'_def, rule allI, rule hoarep.Asm) simp txt \We have to name the fact \Fac_spec\, so that the vcg can use the specification for the recursive call, since it cannot infer it from the opaque @{term "\'"}.\ show "\\. \,\'\ (?Pre \) IF \N = 0 THEN \R :== 1 ELSE \R :== CALL Fac(\N - 1);; \R :== \N * \R FI (?Post \)" apply vcg apply simp done qed text \There are different rules available to prove procedure calls, depending on the kind of postcondition and whether or not the procedure is recursive or even mutually recursive. See for example @{thm [source] HoareTotal.ProcRec1}, @{thm [source] HoareTotal.ProcNoRec1}. They are all derived from the most general rule @{thm [source] HoareTotal.ProcRec}. All of them have some side-conditions concerning the parameter passing protocol and its relation to the pre and postcondition. They can be solved in a uniform fashion. Thats why we have created the method \hoare_rule\, which behaves like the method \rule\ but automatically tries to solve the side-conditions. \ subsubsection \Odd and Even\ text \Odd and even are defined mutually recursive here. In the \procedures\ command we conjoin both definitions with \and\. \ procedures odd(N::nat | A::nat) "IF \N=0 THEN \A:==0 ELSE IF \N=1 THEN CALL even (\N - 1,\A) ELSE CALL odd (\N - 2,\A) FI FI" and even(N::nat | A::nat) "IF \N=0 THEN \A:==1 ELSE IF \N=1 THEN CALL odd (\N - 1,\A) ELSE CALL even (\N - 2,\A) FI FI" print_theorems print_locale! odd_even_clique text \To prove the procedure calls to @{term "odd"} respectively @{term "even"} correct we first derive a rule to justify that we can assume both specifications to verify the bodies. This rule can be derived from the general @{thm [source] HoareTotal.ProcRec} rule. An ML function will do this work: \ ML \ML_Thms.bind_thm ("ProcRec2", Hoare.gen_proc_rec @{context} Hoare.Partial 2)\ lemma (in odd_even_clique) shows odd_spec: "\\. \\{\} \A :== PROC odd(\N) \(\b. \<^bsup>\\<^esup>N = 2 * b + \A) \ \A < 2 \" (is ?P1) and even_spec: "\\. \\{\} \A :== PROC even(\N) \(\b. \<^bsup>\\<^esup>N + 1 = 2 * b + \A) \ \A < 2 \" (is ?P2) proof - have "?P1 \ ?P2" apply (hoare_rule ProcRec2) apply vcg apply clarsimp apply (rule_tac x="b + 1" in exI) apply arith apply vcg apply clarsimp apply arith done thus "?P1" "?P2" by iprover+ qed subsection \Expressions With Side Effects\ (* R := N++ + M++*) lemma (in state_space) shows "\\ \True\ \N \ n. \N :== \N + 1 \ \M \ m. \M :== \M + 1 \ \R :== n + m \\R = \N + \M - 2\" apply vcg apply simp done (* R := Fac (N) + Fac (N) *) lemma (in Fac_impl) shows "\\ \True\ CALL Fac(\N) \ n. CALL Fac(\N) \ m. \R :== n + m \\R = fac \N + fac \N\" proof - note Fac_spec = Fac_spec4 show ?thesis apply vcg done qed (* R := Fac (N) + Fac (M) *) lemma (in Fac_impl) shows "\\ \True\ CALL Fac(\N) \ n. CALL Fac(n) \ m. \R :== m \\R = fac (fac \N)\" proof - note Fac_spec = Fac_spec4 show ?thesis apply vcg done qed subsection \Global Variables and Heap\ text \ Now we will define and verify some procedures on heap-lists. We consider list structures consisting of two fields, a content element @{term "cont"} and a reference to the next list element @{term "next"}. We model this by the following state space where every field has its own heap. \ hoarestate globals_list = "next" :: "ref \ ref" cont :: "ref \ nat" text \Updates to global components inside a procedure will always be propagated to the caller. This is implicitly done by the parameter passing syntax translations. The record containing the global variables must begin with the prefix "globals". \ text \We will first define an append function on lists. It takes two references as parameters. It appends the list referred to by the first parameter with the list referred to by the second parameter, and returns the result right into the first parameter. \ procedures (imports globals_list) append(p::ref,q::ref|p::ref) "IF \p=Null THEN \p :== \q ELSE \p \\next:== CALL append(\p\\next,\q) FI" declare [[hoare_use_call_tr' = false]] context append_impl begin term "CALL append(\p,\q,\p\\next)" end declare [[hoare_use_call_tr' = true]] text \Below we give two specifications this time.. The first one captures the functional behaviour and focuses on the entities that are potentially modified by the procedure, the second one is a pure frame condition. The list in the modifies clause has to list all global state components that may be changed by the procedure. Note that we know from the modifies clause that the @{term cont} parts of the lists will not be changed. Also a small side note on the syntax. We use ordinary brackets in the postcondition of the modifies clause, and also the state components do not carry the acute, because we explicitly note the state @{term t} here. The functional specification now introduces two logical variables besides the state space variable @{term "\"}, namely @{term "Ps"} and @{term "Qs"}. They are universally quantified and range over both the pre and the postcondition, so that we are able to properly instantiate the specification during the proofs. The syntax \\\. \\\ is a shorthand to fix the current state: \{s. \ = s \}\. \ lemma (in append_impl) append_spec: shows "\\ Ps Qs. \\ \\. List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {}\ \p :== PROC append(\p,\q) \List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoarePartial.ProcRec1) apply vcg apply fastforce done text \The modifies clause is equal to a proper record update specification of the following form. \ lemma (in append_impl) shows "{t. t may_only_modify_globals Z in [next]} = {t. \next. globals t=update id id next_' (K_statefun next) (globals Z)}" apply (unfold mex_def meq_def) apply simp done text \If the verification condition generator works on a procedure call it checks whether it can find a modifies clause in the context. If one is present the procedure call is simplified before the Hoare rule @{thm [source] HoareTotal.ProcSpec} is applied. Simplification of the procedure call means, that the ``copy back'' of the global components is simplified. Only those components that occur in the modifies clause will actually be copied back. This simplification is justified by the rule @{thm [source] HoareTotal.ProcModifyReturn}. So after this simplification all global components that do not appear in the modifies clause will be treated as local variables. \ text \You can study the effect of the modifies clause on the following two examples, where we want to prove that @{term "append"} does not change the @{term "cont"} part of the heap. \ lemma (in append_impl) shows "\\ \\p=Null \ \cont=c\ \p :== CALL append(\p,Null) \\cont=c\" apply vcg oops text \To prove the frame condition, we have to tell the verification condition generator to use only the modifies clauses and not to search for functional specifications by the parameter \spec=modifies\ It will also try to solve the verification conditions automatically. \ lemma (in append_impl) append_modifies: shows "\\. \\ {\} \p :== PROC append(\p,\q){t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcRec1) apply (vcg spec=modifies) done lemma (in append_impl) shows "\\ \\p=Null \ \cont=c\ \p\\next :== CALL append(\p,Null) \\cont=c\" apply vcg apply simp done text \ Of course we could add the modifies clause to the functional specification as well. But separating both has the advantage that we split up the verification work. We can make use of the modifies clause before we apply the functional specification in a fully automatic fashion. \ text \To verify the body of @{term "append"} we do not need the modifies clause, since the specification does not talk about @{term "cont"} at all, and we don't access @{term "cont"} inside the body. This may be different for more complex procedures. \ text \ To prove that a procedure respects the modifies clause, we only need the modifies clauses of the procedures called in the body. We do not need the functional specifications. So we can always prove the modifies clause without functional specifications, but me may need the modifies clause to prove the functional specifications. \ subsubsection \Insertion Sort\ primrec sorted:: "('a \ 'a \ bool) \ 'a list \ bool" where "sorted le [] = True" | "sorted le (x#xs) = ((\y\set xs. le x y) \ sorted le xs)" procedures (imports globals_list) insert(r::ref,p::ref | p::ref) "IF \r=Null THEN SKIP ELSE IF \p=Null THEN \p :== \r;; \p\\next :== Null ELSE IF \r\\cont \ \p\\cont THEN \r\\next :== \p;; \p:==\r ELSE \p\\next :== CALL insert(\r,\p\\next) FI FI FI" text \ In the postcondition of the functional specification there is a small but important subtlety. Whenever we talk about the @{term "cont"} part we refer to the one of the pre-state, even in the conclusion of the implication. The reason is, that we have separated out, that @{term "cont"} is not modified by the procedure, to the modifies clause. So whenever we talk about unmodified parts in the postcondition we have to use the pre-state part, or explicitely state an equality in the postcondition. The reason is simple. If the postcondition would talk about \\cont\ instead of \\<^bsup>\\<^esup>cont\, we will get a new instance of \cont\ during verification and the postcondition would only state something about this new instance. But as the verification condition generator will use the modifies clause the caller of \insert\ instead will still have the old \cont\ after the call. Thats the sense of the modifies clause. So the caller and the specification will simply talk about two different things, without being able to relate them (unless an explicit equality is added to the specification). \ lemma (in insert_impl) insert_modifies: "\\. \\ {\} \p :== PROC insert(\r,\p){t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcRec1) apply (vcg spec=modifies) done lemma (in insert_impl) insert_spec: "\\ Ps . \\ \\. List \p \next Ps \ sorted (\) (map \cont Ps) \ \r \ Null \ \r \ set Ps\ \p :== PROC insert(\r,\p) \\Qs. List \p \next Qs \ sorted (\) (map \<^bsup>\\<^esup>cont Qs) \ set Qs = insert \<^bsup>\\<^esup>r (set Ps) \ (\x. x \ set Qs \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoarePartial.ProcRec1) apply vcg apply (intro conjI impI) apply fastforce apply fastforce apply fastforce apply (clarsimp) apply force done procedures (imports globals_list) insertSort(p::ref | p::ref) where r::ref q::ref in "\r:==Null;; WHILE (\p \ Null) DO \q :== \p;; \p :== \p\\next;; \r :== CALL insert(\q,\r) OD;; \p:==\r" print_locale insertSort_impl lemma (in insertSort_impl) insertSort_modifies: shows "\\. \\ {\} \p :== PROC insertSort(\p) {t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcRec1) apply (vcg spec=modifies) done text \Insertion sort is not implemented recursively here but with a while loop. Note that the while loop is not annotated with an invariant in the procedure definition. The invariant only comes into play during verification. Therefore we will annotate the body during the proof with the rule @{thm [source] HoareTotal.annotateI}. \ lemma (in insertSort_impl) insertSort_body_spec: shows "\\ Ps. \,\\ \\. List \p \next Ps \ \p :== PROC insertSort(\p) \\Qs. List \p \next Qs \ sorted (\) (map \<^bsup>\\<^esup>cont Qs) \ set Qs = set Ps\" apply (hoare_rule HoarePartial.ProcRec1) apply (hoare_rule anno= "\r :== Null;; WHILE \p \ Null INV \\Qs Rs. List \p \next Qs \ List \r \next Rs \ set Qs \ set Rs = {} \ sorted (\) (map \cont Rs) \ set Qs \ set Rs = set Ps \ \cont = \<^bsup>\\<^esup>cont \ DO \q :== \p;; \p :== \p\\next;; \r :== CALL insert(\q,\r) OD;; \p :== \r" in HoarePartial.annotateI) apply vcg apply fastforce prefer 2 apply fastforce apply (clarsimp) apply (rule_tac x=ps in exI) apply (intro conjI) apply (rule heap_eq_ListI1) apply assumption apply clarsimp apply (subgoal_tac "x\p \ x \ set Rs") apply auto done subsubsection "Memory Allocation and Deallocation" text \The basic idea of memory management is to keep a list of allocated references in the state space. Allocation of a new reference adds a new reference to the list deallocation removes a reference. Moreover we keep a counter "free" for the free memory. \ (* record globals_list_alloc = globals_list + alloc_'::"ref list" free_'::nat record 'g list_vars' = "'g list_vars" + i_'::nat first_'::ref *) hoarestate globals_list_alloc = alloc::"ref list" free::nat "next"::"ref \ ref" cont::"ref \ nat" hoarestate locals_list_alloc = i::nat first::ref p::ref q::ref r::ref root::ref tmp::ref locale list_alloc = globals_list_alloc + locals_list_alloc definition "sz = (2::nat)" lemma (in list_alloc) shows "\,\\ \\i = 0 \ \first = Null \ n*sz \ \free\ WHILE \i < n INV \\Ps. List \first \next Ps \ length Ps = \i \ \i \ n \ set Ps \ set \alloc \ (n - \i)*sz \ \free\ DO \p :== NEW sz [\cont:==0,\next:== Null];; \p\\next :== \first;; \first :== \p;; \i :== \i+ 1 OD \\Ps. List \first \next Ps \ length Ps = n \ set Ps \ set \alloc\" apply (vcg) apply simp apply clarsimp apply (rule conjI) apply clarsimp apply (rule_tac x="new (set alloc)#Ps" in exI) apply clarsimp apply (rule conjI) apply fastforce apply (simp add: sz_def) apply (simp add: sz_def) apply fastforce done lemma (in list_alloc) shows "\\ \\i = 0 \ \first = Null \ n*sz \ \free\ WHILE \i < n INV \\Ps. List \first \next Ps \ length Ps = \i \ \i \ n \ set Ps \ set \alloc \ (n - \i)*sz \ \free\ DO \p :== NNEW sz [\cont:==0,\next:== Null];; \p\\next :== \first;; \first :== \p;; \i :== \i+ 1 OD \\Ps. List \first \next Ps \ length Ps = n \ set Ps \ set \alloc\" apply (vcg) apply simp apply clarsimp apply (rule conjI) apply clarsimp apply (rule_tac x="new (set alloc)#Ps" in exI) apply clarsimp apply (rule conjI) apply fastforce apply (simp add: sz_def) apply (simp add: sz_def) apply fastforce done subsection \Fault Avoiding Semantics\ text \ If we want to ensure that no runtime errors occur we can insert guards into the code. We will not be able to prove any nontrivial Hoare triple about code with guards, if we cannot show that the guards will never fail. A trivial Hoare triple is one with an empty precondtion. \ lemma (in list_alloc) "\,\\ \True\ \\p\Null\\ \p\\next :== \p \True\" apply vcg oops lemma (in list_alloc) "\,\\ {} \\p\Null\\ \p\\next :== \p \True\" apply vcg done text \Let us consider this small program that reverts a list. At first without guards. \ lemma (in list_alloc) shows "\,\\ \List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {} \ set Ps \ set \alloc \ set Qs \ set \alloc\ WHILE \p \ Null INV \\ps qs. List \p \next ps \ List \q \next qs \ set ps \ set qs = {} \ rev ps @ qs = rev Ps @ Qs \ set ps \ set \alloc \ set qs \ set \alloc\ DO \r :== \p;; \p :== \p\ \next;; \r\\next :== \q;; \q :== \r OD \List \q \next (rev Ps @ Qs) \ set Ps\ set \alloc \ set Qs \ set \alloc\" apply (vcg) apply fastforce+ done text \If we want to ensure that we do not dereference @{term "Null"} or access unallocated memory, we have to add some guards. \ lemma (in list_alloc) shows "\,\\ \List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {} \ set Ps \ set \alloc \ set Qs \ set \alloc\ WHILE \p \ Null INV \\ps qs. List \p \next ps \ List \q \next qs \ set ps \ set qs = {} \ rev ps @ qs = rev Ps @ Qs \ set ps \ set \alloc \ set qs \ set \alloc\ DO \r :== \p;; \\p\Null \ \p\set \alloc\\ \p :== \p\ \next;; \\r\Null \ \r\set \alloc\\ \r\\next :== \q;; \q :== \r OD \List \q \next (rev Ps @ Qs) \ set Ps \ set \alloc \ set Qs \ set \alloc\" apply (vcg) apply fastforce+ done text \We can also just prove that no faults will occur, by giving the trivial postcondition. \ lemma (in list_alloc) rev_noFault: shows "\,\\ \List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {} \ set Ps \ set \alloc \ set Qs \ set \alloc\ WHILE \p \ Null INV \\ps qs. List \p \next ps \ List \q \next qs \ set ps \ set qs = {} \ rev ps @ qs = rev Ps @ Qs \ set ps \ set \alloc \ set qs \ set \alloc\ DO \r :== \p;; \\p\Null \ \p\set \alloc\\ \p :== \p\ \next;; \\r\Null \ \r\set \alloc\\ \r\\next :== \q;; \q :== \r OD UNIV,UNIV" apply (vcg) apply fastforce+ done lemma (in list_alloc) rev_moduloGuards: shows "\,\\\<^bsub>/{True}\<^esub> \List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {} \ set Ps \ set \alloc \ set Qs \ set \alloc\ WHILE \p \ Null INV \\ps qs. List \p \next ps \ List \q \next qs \ set ps \ set qs = {} \ rev ps @ qs = rev Ps @ Qs \ set ps \ set \alloc \ set qs \ set \alloc\ DO \r :== \p;; \\p\Null \ \p\set \alloc\\ \ \p :== \p\ \next;; \\r\Null \ \r\set \alloc\\ \ \r\\next :== \q;; \q :== \r OD \List \q \next (rev Ps @ Qs) \ set Ps \ set \alloc \ set Qs \ set \alloc\" apply vcg apply fastforce+ done lemma CombineStrip': assumes deriv: "\,\\\<^bsub>/F\<^esub> P c' Q,A" assumes deriv_strip: "\,\\\<^bsub>/{}\<^esub> P c'' UNIV,UNIV" assumes c'': "c''= mark_guards False (strip_guards (-F) c')" assumes c: "c = mark_guards False c'" shows "\,\\\<^bsub>/{}\<^esub> P c Q,A" proof - from deriv_strip [simplified c''] have "\,\\ P (strip_guards (- F) c') UNIV,UNIV" by (rule HoarePartialProps.MarkGuardsD) with deriv have "\,\\ P c' Q,A" by (rule HoarePartialProps.CombineStrip) hence "\,\\ P mark_guards False c' Q,A" by (rule HoarePartialProps.MarkGuardsI) thus ?thesis by (simp add: c) qed text \We can then combine the prove that no fault will occur with the functional prove of the programm without guards to get the full proove by the rule @{thm HoarePartialProps.CombineStrip} \ lemma (in list_alloc) shows "\,\\ \List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {} \ set Ps \ set \alloc \ set Qs \ set \alloc\ WHILE \p \ Null INV \\ps qs. List \p \next ps \ List \q \next qs \ set ps \ set qs = {} \ rev ps @ qs = rev Ps @ Qs \ set ps \ set \alloc \ set qs \ set \alloc\ DO \r :== \p;; \\p\Null \ \p\set \alloc\\ \p :== \p\ \next;; \\r\Null \ \r\set \alloc\\ \r\\next :== \q;; \q :== \r OD \List \q \next (rev Ps @ Qs) \ set Ps \ set \alloc \ set Qs \ set \alloc\" apply (rule CombineStrip' [OF rev_moduloGuards rev_noFault]) apply simp apply simp done text \In the previous example the effort to split up the prove did not really pay off. But when we think of programs with a lot of guards and complicated specifications it may be better to first focus on a prove without the messy guards. Maybe it is possible to automate the no fault proofs so that it suffices to focus on the stripped program. \ context list_alloc begin text \ The purpose of guards is to watch for faults that can occur during evaluation of expressions. In the example before we watched for null pointer dereferencing or memory faults. We can also look for array index bounds or division by zero. As the condition of a while loop is evaluated in each iteration we cannot just add a guard before the while loop. Instead we need a special guard for the condition. Example: @{term "WHILE \\p\Null\\ \p\\next\Null DO SKIP OD"} \ end subsection \Cicular Lists\ definition distPath :: "ref \ (ref \ ref) \ ref \ ref list \ bool" where "distPath x next y as = (Path x next y as \ distinct as)" lemma neq_dP: "\p \ q; Path p h q Ps; distinct Ps\ \ \Qs. p\Null \ Ps = p#Qs \ p \ set Qs" by (cases Ps, auto) lemma (in list_alloc) circular_list_rev_I: "\,\\ \\root = r \ distPath \root \next \root (r#Ps)\ \p :== \root;; \q :== \root\\next;; WHILE \q \ \root INV \\ ps qs. distPath \p \next \root ps \ distPath \q \next \root qs \ \root = r \ r\Null \ r \ set Ps \ set ps \ set qs = {} \ Ps = (rev ps) @ qs \ DO \tmp :== \q;; \q :== \q\\next;; \tmp\\next :== \p;; \p:==\tmp OD;; \root\\next :== \p \\root = r \ distPath \root \next \root (r#rev Ps)\" apply (simp only:distPath_def) apply vcg apply (rule_tac x="[]" in exI) apply fastforce apply clarsimp apply (drule (2) neq_dP) apply (rule_tac x="q # ps" in exI) apply clarsimp apply fastforce done lemma path_is_list:"\a next b. \Path b next a Ps ; a \ set Ps; a\Null\ \ List b (next(a := Null)) (Ps @ [a])" apply (induct Ps) apply (auto simp add:fun_upd_apply) done text \ The simple algorithm for acyclic list reversal, with modified annotations, works for cyclic lists as well.: \ lemma (in list_alloc) circular_list_rev_II: "\,\\ \\p = r \ distPath \p \next \p (r#Ps)\ \q:==Null;; WHILE \p \ Null INV \ ((\q = Null) \ (\ps. distPath \p \next r ps \ ps = r#Ps)) \ ((\q \ Null) \ (\ps qs. distPath \q \next r qs \ List \p \next ps \ set ps \ set qs = {} \ rev qs @ ps = Ps@[r])) \ \ (\p = Null \ \q = Null \ r = Null ) \ DO \tmp :== \p;; \p :== \p\\next;; \tmp\\next :== \q;; \q:==\tmp OD \\q = r \ distPath \q \next \q (r # rev Ps)\" apply (simp only:distPath_def) apply vcg apply clarsimp apply clarsimp apply (case_tac "(q = Null)") apply (fastforce intro: path_is_list) apply clarify apply (rule_tac x="psa" in exI) apply (rule_tac x=" p # qs" in exI) apply force apply fastforce done text\Although the above algorithm is more succinct, its invariant looks more involved. The reason for the case distinction on @{term q} is due to the fact that during execution, the pointer variables can point to either cyclic or acyclic structures. \ text \ When working on lists, its sometimes better to remove @{thm[source] fun_upd_apply} from the simpset, and instead include @{thm[source] fun_upd_same} and @{thm[source] fun_upd_other} to the simpset \ (* declare fun_upd_apply[simp del]fun_upd_same[simp] fun_upd_other[simp] *) lemma (in state_space) "\\ {\} \I :== \M;; ANNO \. \\. \I = \<^bsup>\\<^esup>M\ \M :== \N;; \N :== \I \\M = \<^bsup>\\<^esup>N \ \N = \<^bsup>\\<^esup>I\ \\M = \<^bsup>\\<^esup>N \ \N = \<^bsup>\\<^esup>M\" apply vcg apply auto done context state_space begin term "ANNO (\,m,k). (\\. \M = m\) \M :== \N;; \N :== \I \\M = \<^bsup> \\<^esup>N & \N = \<^bsup>\\<^esup>I\,{}" end lemma (in state_space) "\\ ({\} \ \\M = 0 \ \S = 0\) (ANNO \. ({\} \ \\A=\<^bsup>\\<^esup>A \ \I=\<^bsup>\\<^esup>I \ \M=0 \ \S=0\) WHILE \M \ \A INV \\S = \M * \I \ \A=\<^bsup>\\<^esup>A \ \I=\<^bsup>\\<^esup>I\ DO \S :== \S + \I;; \M :== \M + 1 OD \\S = \<^bsup>\\<^esup>A * \<^bsup>\\<^esup>I\) \\S = \<^bsup>\\<^esup>A * \<^bsup>\\<^esup>I\" apply vcg_step apply vcg_step apply simp apply vcg_step apply vcg_step apply simp apply vcg apply simp apply simp apply vcg_step apply auto done text \Just some test on marked, guards\ lemma (in state_space) "\\\True\ WHILE \P \N \\, \Q \M\#, \R \N\\ \N < \M INV \\N < 2\ DO \N :== \M OD \hard\" apply vcg oops lemma (in state_space) "\\\<^bsub>/{True}\<^esub> \True\ WHILE \P \N \\, \Q \M\#, \R \N\\ \N < \M INV \\N < 2\ DO \N :== \M OD \hard\" apply vcg oops end diff --git a/thys/Simpl/ex/VcgExTotal.thy b/thys/Simpl/ex/VcgExTotal.thy --- a/thys/Simpl/ex/VcgExTotal.thy +++ b/thys/Simpl/ex/VcgExTotal.thy @@ -1,370 +1,366 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: VcgExTotal.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2006-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA *) section \Examples for Total Correctness\ theory VcgExTotal imports "../HeapList" "../Vcg" begin record 'g vars = "'g state" + A_' :: nat I_' :: nat M_' :: nat N_' :: nat R_' :: nat S_' :: nat Abr_':: string lemma "\\\<^sub>t \\M = 0 \ \S = 0\ WHILE \M \ a INV \\S = \M * b \ \M \ a\ VAR MEASURE a - \M DO \S :== \S + b;; \M :== \M + 1 OD \\S = a * b\" apply vcg apply (auto) done lemma "\\\<^sub>t \\I \ 3\ WHILE \I < 10 INV \\I\ 10\ VAR MEASURE 10 - \I DO \I :== \I + 1 OD \\I = 10\" apply vcg apply auto done text \Total correctness of a nested loop. In the inner loop we have to express that the loop variable of the outer loop is not changed. We use \FIX\ to introduce a new logical variable \ lemma "\\\<^sub>t \\M=0 \ \N=0\ WHILE (\M < i) INV \\M \ i \ (\M \ 0 \ \N = j) \ \N \ j\ VAR MEASURE (i - \M) DO \N :== 0;; WHILE (\N < j) FIX m. INV \\M=m \ \N \ j\ VAR MEASURE (j - \N) DO \N :== \N + 1 OD;; \M :== \M + 1 OD \\M=i \ (\M\0 \ \N=j)\" apply vcg apply simp_all apply arith done primrec fac:: "nat \ nat" where "fac 0 = 1" | "fac (Suc n) = (Suc n) * fac n" lemma fac_simp [simp]: "0 < i \ fac i = i * fac (i - 1)" by (cases i) simp_all procedures Fac (N | R) = "IF \N = 0 THEN \R :== 1 ELSE CALL Fac(\N - 1,\R);; \R :== \N * \R FI" lemma (in Fac_impl) Fac_spec: shows "\n. \\\<^sub>t \\N=n\ \R :== PROC Fac(\N) \\R = fac n\" apply (hoare_rule HoareTotal.ProcRec1 [where r="measure (\(s,p). \<^bsup>s\<^esup>N)"]) apply vcg apply simp done procedures p91(R,N | R) = "IF 100 < \N THEN \R :== \N - 10 ELSE \R :== CALL p91(\R,\N+11);; \R :== CALL p91(\R,\R) FI" p91_spec: "\n. \\\<^sub>t \\N=n\ \R :== PROC p91(\R,\N) \if 100 < n then \R = n - 10 else \R = 91\,{}" lemma (in p91_impl) p91_spec: shows "\\. \\\<^sub>t {\} \R :== PROC p91(\R,\N) \if 100 < \<^bsup>\\<^esup>N then \R = \<^bsup>\\<^esup>N - 10 else \R = 91\,{}" apply (hoare_rule HoareTotal.ProcRec1 [where r="measure (\(s,p). 101 - \<^bsup>s\<^esup>N)"]) apply vcg apply clarsimp apply arith done record globals_list = next_' :: "ref \ ref" cont_' :: "ref \ nat" record 'g list_vars = "'g state" + p_' :: "ref" q_' :: "ref" r_' :: "ref" root_' :: "ref" tmp_' :: "ref" procedures append(p,q|p) = "IF \p=Null THEN \p :== \q ELSE \p\\next :== CALL append(\p\\next,\q) FI" lemma (in append_impl) shows "\\ Ps Qs. \\\<^sub>t \\. List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {} \ \p :== PROC append(\p,\q) \List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoareTotal.ProcRec1 [where r="measure (\(s,p). length (list \<^bsup>s\<^esup>p \<^bsup>s\<^esup>next))"]) apply vcg apply (fastforce simp add: List_list) done lemma (in append_impl) shows "\\ Ps Qs. \\\<^sub>t \\. List \p \next Ps \ List \q \next Qs \ set Ps \ set Qs = {} \ \p :== PROC append(\p,\q) \List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoareTotal.ProcRec1 [where r="measure (\(s,p). length (list \<^bsup>s\<^esup>p \<^bsup>s\<^esup>next))"]) apply vcg apply (fastforce simp add: List_list) done lemma (in append_impl) shows append_spec: "\\. \\\<^sub>t ({\} \ \islist \p \next\) \p :== PROC append(\p,\q) \\Ps Qs. List \<^bsup>\\<^esup>p \<^bsup>\\<^esup>next Ps \ List \<^bsup>\\<^esup>q \<^bsup>\\<^esup>next Qs \ set Ps \ set Qs = {} \ List \p \next (Ps@Qs) \ (\x. x\set Ps \ \next x = \<^bsup>\\<^esup>next x)\" apply (hoare_rule HoareTotal.ProcRec1 [where r="measure (\(s,p). length (list \<^bsup>s\<^esup>p \<^bsup>s\<^esup>next))"]) apply vcg apply fastforce done lemma "\\\List \p \next Ps\ \q :== Null;; WHILE \p \ Null INV \\Ps' Qs'. List \p \next Ps' \ List \q \next Qs' \ set Ps' \ set Qs' = {} \ rev Ps' @ Qs' = rev Ps\ DO \r :== \p;; \p :== \p\\next;; \r\\next :== \q;; \q :== \r OD;; \p :==\q \List \p \next (rev Ps)\ " apply vcg apply clarsimp apply clarsimp apply force apply simp done lemma conjI2: "\Q; Q \ P\ \ P \ Q" by blast procedures Rev(p|p) = "\q :== Null;; WHILE \p \ Null DO \r :== \p;; \\p \ Null\\ \p :== \p\\next;; \\r \ Null\\ \r\\next :== \q;; \q :== \r OD;; \p :==\q" Rev_spec: "\Ps. \\\<^sub>t \List \p \next Ps\ \p :== PROC Rev(\p) \List \p \next (rev Ps)\" Rev_modifies: "\\. \\\<^bsub>/UNIV\<^esub> {\} \p :== PROC Rev(\p) {t. t may_only_modify_globals \ in [next]}" text \We only need partial correctness of modifies clause!\ lemma upd_hd_next: assumes p_ps: "List p next (p#ps)" shows "List (next p) (next(p := q)) ps" proof - from p_ps have "p \ set ps" by auto with p_ps show ?thesis by simp qed lemma (in Rev_impl) shows Rev_spec: "\Ps. \\\<^sub>t \List \p \next Ps\ \p :== PROC Rev(\p) \List \p \next (rev Ps)\" apply (hoare_rule HoareTotal.ProcNoRec1) apply (hoare_rule anno = "\q :== Null;; WHILE \p \ Null INV \\Ps' Qs'. List \p \next Ps' \ List \q \next Qs' \ set Ps' \ set Qs' = {} \ rev Ps' @ Qs' = rev Ps\ VAR MEASURE (length (list \p \next) ) DO \r :== \p;; \\p \ Null\\\p :== \p\\next;; \\r \ Null\\\r\\next :== \q;; \q :== \r OD;; \p :==\q" in HoareTotal.annotateI) apply vcg apply clarsimp apply clarsimp apply (rule conjI2) apply force apply clarsimp apply (subgoal_tac "List p next (p#ps)") prefer 2 apply simp apply (frule_tac q=q in upd_hd_next) apply (simp only: List_list) apply simp apply simp done lemma (in Rev_impl) shows Rev_modifies: "\\. \\\<^bsub>/UNIV \<^esub>{\} \p :== PROC Rev(\p) {t. t may_only_modify_globals \ in [next]}" apply (hoare_rule HoarePartial.ProcNoRec1) apply (vcg spec=modifies) done lemma "\\\<^sub>t \List \p \next Ps\ \q :== Null;; WHILE \p \ Null INV \\Ps' Qs'. List \p \next Ps' \ List \q \next Qs' \ set Ps' \ set Qs' = {} \ rev Ps' @ Qs' = rev Ps\ VAR MEASURE (length (list \p \next) ) DO \r :== \p;; \p :== \p\\next;; \r\\next :== \q;; \q :== \r OD;; \p :==\q \List \p \next (rev Ps)\ " apply vcg apply clarsimp apply clarsimp apply (rule conjI2) apply force apply clarsimp apply (subgoal_tac "List p next (p#ps)") prefer 2 apply simp apply (frule_tac q=q in upd_hd_next) apply (simp only: List_list) apply simp apply simp done procedures pedal(N,M) = "IF 0 < \N THEN IF 0 < \M THEN CALL coast(\N- 1,\M- 1) FI;; CALL pedal(\N- 1,\M) FI" and coast(N,M) = "CALL pedal(\N,\M);; IF 0 < \M THEN CALL coast(\N,\M- 1) FI" ML \ML_Thms.bind_thm ("HoareTotal_ProcRec2", Hoare.gen_proc_rec @{context} Hoare.Total 2)\ lemma (in pedal_coast_clique) shows "(\\\<^sub>t \True\ PROC pedal(\N,\M) \True\) \ (\\\<^sub>t \True\ PROC coast(\N,\M) \True\)" apply (hoare_rule HoareTotal_ProcRec2 [where ?r= "inv_image (measure (\m. m) <*lex*> measure (\p. if p = coast_'proc then 1 else 0)) (\(s,p). (\<^bsup>s\<^esup>N + \<^bsup>s\<^esup>M,p))"]) apply simp_all - apply (vcg,force)+ + apply vcg + apply simp + apply vcg + apply simp done lemma (in pedal_coast_clique) shows "(\\\<^sub>t \True\ PROC pedal(\N,\M) \True\) \ (\\\<^sub>t \True\ PROC coast(\N,\M) \True\)" apply (hoare_rule HoareTotal_ProcRec2 [where ?r= "inv_image (measure (\m. m) <*lex*> measure (\p. if p = coast_'proc then 1 else 0)) (\(s,p). (\<^bsup>s\<^esup>N + \<^bsup>s\<^esup>M,p))"]) apply simp_all - apply (vcg,force)+ + apply vcg + apply simp + apply vcg + apply simp done lemma (in pedal_coast_clique) shows "(\\\<^sub>t \True\ PROC pedal(\N,\M) \True\) \ (\\\<^sub>t \True\ PROC coast(\N,\M) \True\)" apply(hoare_rule HoareTotal_ProcRec2 [where ?r= "measure (\(s,p). \<^bsup>s\<^esup>N + \<^bsup>s\<^esup>M + (if p = coast_'proc then 1 else 0))"]) apply simp_all - apply (vcg,force)+ + apply vcg + apply simp + apply arith + apply vcg + apply simp done lemma (in pedal_coast_clique) shows "(\\\<^sub>t \True\ PROC pedal(\N,\M) \True\) \ (\\\<^sub>t \True\ PROC coast(\N,\M) \True\)" apply(hoare_rule HoareTotal_ProcRec2 [where ?r= "(\(s,p). \<^bsup>s\<^esup>N) <*mlex*> (\(s,p). \<^bsup>s\<^esup>M) <*mlex*> measure (\(s,p). if p = coast_'proc then 1 else 0)"]) apply simp_all - apply (vcg,force)+ - done + apply vcg + apply simp + apply vcg + apply simp + done lemma (in pedal_coast_clique) shows "(\\\<^sub>t \True\ PROC pedal(\N,\M) \True\) \ (\\\<^sub>t \True\ PROC coast(\N,\M) \True\)" apply(hoare_rule HoareTotal_ProcRec2 [where ?r= "measure (\s. \<^bsup>s\<^esup>N + \<^bsup>s\<^esup>M) <*lex*> measure (\p. if p = coast_'proc then 1 else 0)"]) apply simp_all - apply (vcg,force)+ - done + apply vcg + apply simp + apply vcg + apply simp + done end diff --git a/thys/Simpl/ex/XVcgEx.thy b/thys/Simpl/ex/XVcgEx.thy --- a/thys/Simpl/ex/XVcgEx.thy +++ b/thys/Simpl/ex/XVcgEx.thy @@ -1,109 +1,88 @@ (* Author: Norbert Schirmer Maintainer: Norbert Schirmer, norbert.schirmer at web de - License: LGPL -*) - -(* Title: XVcgEx.thy - Author: Norbert Schirmer, TU Muenchen Copyright (C) 2006-2008 Norbert Schirmer -Some rights reserved, TU Muenchen - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA *) section "Examples for Parallel Assignments" theory XVcgEx imports "../XVcg" begin record "globals" = "G_'"::"nat" "H_'"::"nat" record 'g vars = "'g state" + A_' :: nat B_' :: nat C_' :: nat I_' :: nat M_' :: nat N_' :: nat R_' :: nat S_' :: nat Arr_' :: "nat list" Abr_':: string term "BASIC \A :== x, \B :== y END" term "BASIC \G :== \H, \H :== \G END" term "BASIC LET (x,y) = (\A,b); z = \B IN \A :== x, \G :== \A + y + z END" lemma "\\ \\A = 0\ \\A < 0\ \ BASIC LET (a,b,c) = foo \A IN \A :== a, \B :== b, \C :== c END \\A = x \ \B = y \ \C = c\" apply vcg oops lemma "\\ \\A = 0\ \\A < 0\ \ BASIC LET (a,b,c) = foo \A IN \A :== a, \G :== b + \B, \H :== c END \\A = x \ \G = y \ \H = c\" apply vcg oops definition foo:: "nat \ (nat \ nat \ nat)" where "foo n = (n,n+1,n+2)" lemma "\\ \\A = 0\ \\A < 0\ \ BASIC LET (a,b,c) = foo \A IN \A :== a, \G :== b + \B, \H :== c END \\A = x \ \G = y \ \H = c\" apply (vcg add: foo_def snd_conv fst_conv) oops end diff --git a/thys/Simpl/generalise_state.ML b/thys/Simpl/generalise_state.ML --- a/thys/Simpl/generalise_state.ML +++ b/thys/Simpl/generalise_state.ML @@ -1,301 +1,286 @@ (* Title: generalise_state.ML - Author: Norbert Schirmer, TU Muenchen -Copyright (C) 2005-2007 Norbert Schirmer - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (C) 2006-2008 Norbert Schirmer +Copyright (c) 2022 Apple Inc. All rights reserved. *) signature SPLIT_STATE = -sig - val isState: term -> bool - val abs_state: term -> term option +sig val isState: Proof.context -> term -> bool + val abs_state: Proof.context -> term -> term option val abs_var: Proof.context -> term -> (string * typ) val split_state: Proof.context -> string -> typ -> term -> (term * term list) - val ex_tac: Proof.context -> term list -> tactic + val ex_tac: Proof.context -> term list -> int -> tactic (* the term-list is the list of selectors as returned by split_state. They may be used to - construct the instatiation of the existentially + construct the instantiation of the existentially quantified state. *) end; -functor GeneraliseFun (structure SplitState: SPLIT_STATE) = +signature GENERALISE = +sig + val GENERALISE: Proof.context -> int -> tactic +end + +functor Generalise (structure SplitState: SPLIT_STATE) : GENERALISE = struct val genConj = @{thm generaliseConj}; val genImp = @{thm generaliseImp}; val genImpl = @{thm generaliseImpl}; val genAll = @{thm generaliseAll}; val gen_all = @{thm generalise_all}; val genEx = @{thm generaliseEx}; val genRefl = @{thm generaliseRefl}; val genRefl' = @{thm generaliseRefl'}; val genTrans = @{thm generaliseTrans}; val genAllShift = @{thm generaliseAllShift}; val gen_allShift = @{thm generalise_allShift}; val meta_spec = @{thm meta_spec}; val protectRefl = @{thm protectRefl}; val protectImp = @{thm protectImp}; fun gen_thm decomp (t,ct) = let val (ts,cts,recomb) = decomp (t,ct) in recomb (map (gen_thm decomp) (ts~~cts)) end; fun dest_prop (Const (@{const_name Pure.prop}, _) $ P) = P | dest_prop t = raise TERM ("dest_prop", [t]); fun prem_of thm = #1 (Logic.dest_implies (dest_prop (Thm.prop_of thm))); fun conc_of thm = #2 (Logic.dest_implies (dest_prop (Thm.prop_of thm))); fun dest_All (Const (@{const_name "All"},_)$t) = t | dest_All t = raise TERM ("dest_All",[t]); fun SIMPLE_OF ctxt rule prems = let val mx = fold (fn thm => fn i => Int.max (Thm.maxidx_of thm,i)) prems 0; in DistinctTreeProver.discharge ctxt prems (Thm.incr_indexes (mx + 1) rule) end; infix 0 OF_RAW fun tha OF_RAW thb = thb COMP (Drule.incr_indexes thb tha); fun SIMPLE_OF_RAW ctxt tha thb = SIMPLE_OF ctxt tha [thb]; datatype qantifier = Meta_all | Hol_all | Hol_ex fun list_exists (vs, x) = fold_rev (fn (x, T) => fn P => HOLogic.exists_const T $ Abs (x, T, P)) vs x; fun spec' cv thm = let (* thm = Pure.prop ((all x. P x) ==> Q), where "all" is meta or HOL *) val (ct1,ct2) = thm |> Thm.cprop_of |> Thm.dest_comb |> #2 |> Thm.dest_comb |> #2 |> Thm.dest_comb; in (case Thm.term_of ct1 of Const (@{const_name "Trueprop"},_) => let val (Var (sP,_)$Var (sV,sVT)) = HOLogic.dest_Trueprop (Thm.concl_of spec); val cvT = Thm.ctyp_of_cterm cv; val vT = Thm.typ_of cvT; in Thm.instantiate (TVars.make [(dest_TVar sVT, cvT)], Vars.make [((sP, vT --> HOLogic.boolT), #2 (Thm.dest_comb ct2)), ((sV, vT), cv)]) spec end | Const (@{const_name Pure.all},_) => let val (Var (sP,_)$Var (sV,sVT)) = Thm.concl_of meta_spec; val cvT = Thm.ctyp_of_cterm cv; val vT = Thm.typ_of cvT; in Thm.instantiate (TVars.make [(dest_TVar sVT, cvT)], Vars.make [((sP, vT --> propT), ct2), ((sV, vT),cv)]) meta_spec end | _ => raise THM ("spec'",0,[thm])) end; fun split_thm qnt ctxt s T t = let val (t',vars) = SplitState.split_state ctxt s T t; val vs = map (SplitState.abs_var ctxt) vars; val prop = (case qnt of Meta_all => Logic.list_all (vs,t') | Hol_all => HOLogic.mk_Trueprop (HOLogic.list_all (vs, t')) | Hol_ex => Logic.mk_implies (HOLogic.mk_Trueprop (list_exists (vs, t')), HOLogic.mk_Trueprop (HOLogic.mk_exists (s,T,t)))) in (case qnt of - Hol_ex => Goal.prove ctxt [] [] prop (fn _ => SplitState.ex_tac ctxt vars) + Hol_ex => Goal.prove ctxt [] [] prop (fn _ => SplitState.ex_tac ctxt vars 1) | _ => let val rP = conc_of genRefl'; val thm0 = Thm.instantiate (TVars.empty, Vars.make [(dest_Var rP, Thm.cterm_of ctxt prop)]) genRefl'; fun elim_all v thm = let val cv = Thm.cterm_of ctxt v; val spc = Goal.protect 0 (spec' cv thm); in SIMPLE_OF ctxt genTrans [thm,spc] end; val thm = fold elim_all vars thm0; in thm end) end; fun eta_expand ctxt ct = let val mi = Thm.maxidx_of_cterm ct; val T = domain_type (Thm.typ_of_cterm ct); val x = Thm.cterm_of ctxt (Var (("s",mi+1),T)); in Thm.lambda x (Thm.apply ct x) end; fun split_abs ctxt ct = (case Thm.term_of ct of Abs x => (x, Thm.dest_abs_global ct) | _ => split_abs ctxt (eta_expand ctxt ct)) fun decomp ctxt (Const (@{const_name HOL.conj}, _) $ t $ t', ct) = ([t,t'],snd (Drule.strip_comb ct), fn [thm,thm'] => SIMPLE_OF ctxt genConj [thm,thm']) | decomp ctxt ((allc as Const (@{const_name "All"},aT)) $ f, ct) = let val cf = snd (Thm.dest_comb ct); val (abst as (x,T,_),(cx',cb)) = split_abs ctxt cf; val Free (x',_) = Thm.term_of cx'; val (Const (@{const_name Pure.all},_)$Abs (s,_,_)) = genAll |> Thm.prems_of |> hd |> dest_prop; val genAll' = Drule.rename_bvars [(s,x)] genAll; val (Const (@{const_name Pure.all},_)$Abs (s',_,_)) = genAllShift |> Thm.prems_of |> hd |> dest_prop; val genAllShift' = Drule.rename_bvars [(s',x)] genAllShift; - in if SplitState.isState (allc$Abs abst) + in if SplitState.isState ctxt (allc$Abs abst) then ([Thm.term_of cb],[cb], fn [thm] => let val P = HOLogic.dest_Trueprop (dest_prop (prem_of thm)); val thm' = split_thm Hol_all ctxt x' T P; val thm1 = genAllShift' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm')); val thm2 = genAll' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm)); in SIMPLE_OF ctxt genTrans [thm1,thm2] end) else ([Thm.term_of cb],[cb], fn [thm] => genAll' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm))) end | decomp ctxt ((exc as Const (@{const_name "Ex"},_)) $ f, ct) = let val cf = snd (Thm.dest_comb ct); val (abst as (x,T,_),(cx',cb)) = split_abs ctxt cf; val Free (x',_) = Thm.term_of cx'; val (Const (@{const_name Pure.all},_)$Abs (s,_,_)) = genEx |> Thm.prems_of |> hd |> dest_prop; val genEx' = Drule.rename_bvars [(s,x)] genEx; - in if SplitState.isState (exc$Abs abst) + in if SplitState.isState ctxt (exc$Abs abst) then ([Thm.term_of cb],[cb], fn [thm] => let val P = HOLogic.dest_Trueprop (dest_prop (prem_of thm)); val thm' = split_thm Hol_ex ctxt x' T P; in SIMPLE_OF_RAW ctxt protectImp (Goal.protect 0 thm') end ) else ([Thm.term_of cb],[cb], fn [thm] => genEx' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm))) end | decomp ctxt (Const (@{const_name HOL.implies},_)$P$Q, ct) = let val [cP,cQ] = (snd (Drule.strip_comb ct)); in ([Q],[cQ],fn [thm] => let val X = genImp |> Thm.concl_of |> dest_prop |> Logic.dest_implies |> #1 |> dest_prop |> HOLogic.dest_Trueprop |> HOLogic.dest_imp |> #1 |> dest_Var; val genImp' = Thm.instantiate (TVars.empty, Vars.make [(X,cP)]) genImp; in SIMPLE_OF ctxt genImp' [thm] end) end | decomp ctxt (Const (@{const_name Pure.imp},_)$P$Q, ct) = let val [cP,cQ] = (snd (Drule.strip_comb ct)); in ([Q],[cQ],fn [thm] => let val X = genImpl |> Thm.concl_of |> dest_prop |> Logic.dest_implies |> #1 |> dest_prop |> Logic.dest_implies |> #1 |> dest_Var; val genImpl' = Thm.instantiate (TVars.empty, Vars.make [(X,cP)]) genImpl; in SIMPLE_OF ctxt genImpl' [thm] end) end | decomp ctxt ((allc as Const (@{const_name Pure.all},_)) $ f, ct) = let val cf = snd (Thm.dest_comb ct); val (abst as (x,T,_),(cx',cb)) = split_abs ctxt cf; val Free (x',_) = Thm.term_of cx'; val (Const (@{const_name Pure.all},_)$Abs (s,_,_)) = gen_all |> Thm.prems_of |> hd |> dest_prop; val gen_all' = Drule.rename_bvars [(s,x)] gen_all; val (Const (@{const_name Pure.all},_)$Abs (s',_,_)) = gen_allShift |> Thm.prems_of |> hd |> dest_prop; val gen_allShift' = Drule.rename_bvars [(s',x)] gen_allShift; - in if SplitState.isState (allc$Abs abst) + in if SplitState.isState ctxt (allc$Abs abst) then ([Thm.term_of cb],[cb], fn [thm] => let val P = dest_prop (prem_of thm); val thm' = split_thm Meta_all ctxt x' T P; val thm1 = gen_allShift' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm')); val thm2 = gen_all' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm)); in SIMPLE_OF ctxt genTrans [thm1,thm2] end) else ([Thm.term_of cb],[cb], fn [thm] => gen_all' OF_RAW Goal.protect 0 (Thm.forall_intr cx' (Goal.conclude thm))) end | decomp ctxt (Const (@{const_name "Trueprop"},_)$P, ct) = ([P],snd (Drule.strip_comb ct),fn [thm] => thm) | decomp ctxt (t, ct) = ([],[], fn [] => let val rP = HOLogic.dest_Trueprop (dest_prop (conc_of genRefl)); in Thm.instantiate (TVars.empty, Vars.make [(dest_Var rP, ct)]) genRefl end) fun generalise ctxt ct = gen_thm (decomp ctxt) (Thm.term_of ct,ct); (* -------- (init) #C ==> #C *) fun init ct = Thm.instantiate' [] [SOME ct] protectRefl; -fun generalise_over_tac ctxt P i st = - let - val t = List.nth (Thm.prems_of st, i - 1); (* FIXME !? *) - in (case P t of - SOME t' => - let - val ct = Thm.cterm_of ctxt t'; - val meta_spec_protect' = infer_instantiate ctxt [(("x", 0), ct)] @{thm meta_spec_protect}; - in - (init (Thm.adjust_maxidx_cterm 0 (List.nth (Drule.cprems_of st, i - 1))) - |> resolve_tac ctxt [meta_spec_protect'] 1 - |> Seq.maps (fn st' => - Thm.bicompose NONE {flatten = true, match = false, incremented = false} - (false, Goal.conclude st', Thm.nprems_of st') i st)) - end - | NONE => no_tac st) - end; +fun generalise_over_tac ctxt P = SUBGOAL (fn (t, i) => fn st => + (case P t of + SOME t' => + let + val ct = Thm.cterm_of ctxt t'; + val meta_spec_protect' = infer_instantiate ctxt [(("x", 0), ct)] @{thm meta_spec_protect}; + in + (init (Thm.adjust_maxidx_cterm 0 (List.nth (Drule.cprems_of st, i - 1))) + |> resolve_tac ctxt [meta_spec_protect'] 1 + |> Seq.maps (fn st' => + Thm.bicompose NONE {flatten = true, match = false, incremented = false} + (false, Goal.conclude st', Thm.nprems_of st') i st)) + end + | NONE => no_tac st)) fun generalise_over_all_states_tac ctxt i = - REPEAT (generalise_over_tac ctxt SplitState.abs_state i); + REPEAT (generalise_over_tac ctxt (SplitState.abs_state ctxt) i); -fun generalise_tac ctxt i st = +fun generalise_tac ctxt = CSUBGOAL (fn (ct, i) => fn st => let - val ct = List.nth (Drule.cprems_of st, i - 1); val ct' = Thm.dest_equals_rhs (Thm.cprop_of (Thm.eta_conversion ct)); val r = Goal.conclude (generalise ctxt ct'); in (init (Thm.adjust_maxidx_cterm 0 (List.nth (Drule.cprems_of st, i - 1))) |> (resolve_tac ctxt [r] 1 THEN resolve_tac ctxt [Drule.protectI] 1) |> Seq.maps (fn st' => Thm.bicompose NONE {flatten = true, match = false, incremented = false} (false, Goal.conclude st', Thm.nprems_of st') i st)) - end + end) fun GENERALISE ctxt i = generalise_over_all_states_tac ctxt i THEN generalise_tac ctxt i end; diff --git a/thys/Simpl/hoare.ML b/thys/Simpl/hoare.ML --- a/thys/Simpl/hoare.ML +++ b/thys/Simpl/hoare.ML @@ -1,3403 +1,3609 @@ (* Title: hoare.ML Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2007 Norbert Schirmer - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (c) 2022 Apple Inc. All rights reserved. *) signature HOARE = sig datatype hoareMode = Partial | Total val gen_proc_rec: Proof.context -> hoareMode -> int -> thm - datatype state_kind = Record | Function + datatype state_kind = Record | Function | Other of string (* future extensions of state space *) datatype par_kind = In | Out val deco: string val proc_deco: string val par_deco: string -> string val chopsfx: string -> string -> string val is_state_var: string -> bool val extern: Proof.context -> string -> string val remdeco: Proof.context -> string -> string - val remdeco': string -> string + val remdeco': Proof.context -> string -> string val undeco: Proof.context -> term -> term val varname: string -> string val resuffix: string -> string -> string -> string + type state_space = { + name: string, + is_state_type: Proof.context -> typ -> bool, + generalise: Proof.context -> int -> tactic, + state_simprocs: simproc list, + state_upd_simprocs: simproc list, + state_ex_sel_eq_simprocs: simproc list, + + (* syntax *) + read_function_name: Proof.context -> xstring -> term, + is_defined: Proof.context -> xstring -> bool, + lookup_tr: Proof.context -> xstring -> term, + update_tr: Proof.context -> xstring -> term, + is_lookup: Proof.context -> term -> bool, + lookup_tr': Proof.context -> term -> term, + dest_update_tr': Proof.context -> term -> term * term * term option, + update_tr': Proof.context -> term -> term + } + type lense = {lookup: term, update : term} type proc_info = - {params: ((par_kind * string) list), + {params: ((par_kind * string * lense option) list), recursive: bool, state_kind: state_kind} type hoare_tac = (bool -> int -> tactic) -> Proof.context -> hoareMode -> int -> tactic type hoare_data = {proc_info: proc_info Symtab.table, active_procs: string list list, default_state_kind: state_kind, - generate_guard: (stamp * (Proof.context -> term -> term option)), - wp_tacs: (string * hoare_tac) list, + generate_guard: (stamp option * (Proof.context -> term -> term option)), + name_tr: (stamp option * (Proof.context -> bool -> string -> string)), (* bool indicates input tr: true vs. output tr: false *) hoare_tacs: (string * hoare_tac) list, - vcg_simps: thm list} + vcg_simps: thm list, + state_spaces: (string * state_space) list (* registry for state space extensions *) + } val get_data: Proof.context -> hoare_data - val get_params: string -> Proof.context -> (par_kind * string) list option + val get_params: string -> Proof.context -> (par_kind * string * lense option) list option val get_default_state_kind: Proof.context -> state_kind val get_state_kind: string -> Proof.context -> state_kind option + val get_default_state_space: Proof.context -> state_space option val clique_name: string list -> string val install_generate_guard: (Proof.context -> term -> term option) -> Context.generic -> Context.generic val generate_guard: Proof.context -> term -> term option + val install_name_tr: (Proof.context -> bool -> string -> string) -> + Context.generic -> Context.generic + val name_tr: Proof.context -> bool -> string -> string + val install_state_space: state_space -> Context.generic -> Context.generic val BasicSimpTac: Proof.context -> state_kind -> bool -> thm list -> (int -> tactic) -> int -> tactic val hoare: (Proof.context -> Proof.method) context_parser val hoare_raw: (Proof.context -> Proof.method) context_parser val vcg: (Proof.context -> Proof.method) context_parser val vcg_step: (Proof.context -> Proof.method) context_parser val hoare_rule: (Proof.context -> Proof.method) context_parser val add_foldcongsimps: thm list -> theory -> theory val get_foldcong_ss : theory -> simpset val add_foldcongs : thm list -> theory -> theory val modeqN : string val modexN : string val implementationN : string val specL : string val vcg_tac : string -> string -> string list -> Proof.context -> int -> tactic val hoare_rule_tac : Proof.context -> thm list -> int -> tactic + val solve_modifies_tac: Proof.context -> state_kind -> (term -> int) -> int -> tactic + + val add_hoare_tacs: (string * hoare_tac) list -> Context.generic -> Context.generic datatype 'a bodykind = BodyTyp of 'a | BodyTerm of 'a val proc_specs : (bstring * string) list parser - val add_params : morphism -> string -> (par_kind * string) list -> + val add_params : morphism -> string -> (par_kind * string * lense option) list -> Context.generic -> Context.generic val set_default_state_kind : state_kind -> Context.generic -> Context.generic val add_state_kind : morphism -> string -> state_kind -> Context.generic -> Context.generic val add_recursive : morphism -> string -> Context.generic -> Context.generic + + structure FunSplitState : SPLIT_STATE + val first_subterm: (term -> bool) -> term -> ((string * typ) list * term) option + val dest_string: term -> string + val dest_hoare_raw: term -> term * term * term * term * hoareMode * term * term * term + val idx: ('a -> string -> bool) -> 'a list -> string -> int + val sort_variables: bool Config.T + val destr_to_constr: term -> term end; structure Hoare: HOARE = struct (* Misc *) val record_vanish = Attrib.setup_config_bool @{binding hoare_record_vanish} (K true); val use_generalise = Attrib.setup_config_bool @{binding hoare_use_generalise} (K false); val sort_variables = Attrib.setup_config_bool @{binding hoare_sort_variables} (K true); val use_cond_inv_modifies = Attrib.setup_config_bool @{binding hoare_use_cond_inv_modifies} (K true); -val hoare_trace = Attrib.setup_config_bool @{binding hoare_trace} (K false); +val hoare_trace = Attrib.setup_config_int @{binding hoare_trace} (K 0); val body_def_sfx = "_body"; val programN = "\"; val hoare_ctxtL = "hoare"; val specL = "_spec"; val procL = "_proc"; val bodyP = "_impl"; val modifysfx = "_modifies"; val modexN = "Hoare.mex"; val modeqN = "Hoare.meq"; val KNF = @{const_name StateFun.K_statefun}; (* Some abstract syntax operations *) val Trueprop = HOLogic.mk_Trueprop; infix 0 ===; val (op ===) = Trueprop o HOLogic.mk_eq; fun is_empty_set (Const (@{const_name Orderings.bot}, _)) = true | is_empty_set _ = false; fun mk_Int Ts A B = let val T = fastype_of1 (Ts, A) in Const (@{const_name Lattices.inf}, T --> T --> T) $ A $ B end; fun mk_Un T (A, B) = Const (@{const_name Lattices.sup}, T --> T --> T) $ A $ B; fun dest_Un (Const (@{const_name Lattices.sup}, _) $ t1 $ t2) = dest_Un t1 @ dest_Un t2 | dest_Un t = [t] fun mk_UN' dT rT t = let val dTs = HOLogic.mk_setT dT; val rTs = HOLogic.mk_setT rT; in Const (@{const_name Complete_Lattices.Sup}, rTs --> rT) $ (Const (@{const_name image}, (dT --> rT) --> dTs --> rTs) $ t $ Const (@{const_name Orderings.top}, dTs)) end; fun mk_UN ((x, T), P) = mk_UN' T (fastype_of P) (absfree (x, T) P); fun dest_UN (Const (@{const_name Complete_Lattices.Sup}, _) $ (Const (@{const_name Set.image}, _) $ Abs (x, T, t) $ Const (@{const_name Orderings.top}, _))) = let val (vars, body) = dest_UN t in ((x, T) :: vars, body) end | dest_UN t = ([], t); fun tap_UN (Const (@{const_name Complete_Lattices.Sup}, _) $ (Const (@{const_name Set.image}, _) $ t $ Const (@{const_name Orderings.top}, _))) = SOME t | tap_UN _ = NONE; (* Fetching the rules *) datatype hoareMode = Partial | Total fun get_rule p t Partial = p | get_rule p t Total = t +fun get_rule' p t m Partial true = m + | get_rule' p t m Partial false = p + | get_rule' p t m Total _ = t + +fun get_call_rule p t p_exn t_exn Partial NONE = p + | get_call_rule p t p_exn t_exn Partial (SOME _) = p_exn + | get_call_rule p t p_exn t_exn Total NONE = t + | get_call_rule p t p_exn t_exn Total (SOME _) = t_exn + val Guard = get_rule @{thm HoarePartial.Guard} @{thm HoareTotal.Guard}; val GuardStrip = get_rule @{thm HoarePartial.GuardStrip} @{thm HoareTotal.GuardStrip}; val GuaranteeAsGuard = get_rule @{thm HoarePartial.GuaranteeAsGuard} @{thm HoareTotal.GuaranteeAsGuard}; val Guarantee = get_rule @{thm HoarePartial.Guarantee} @{thm HoareTotal.Guarantee}; val GuaranteeStrip = get_rule @{thm HoarePartial.GuaranteeStrip} @{thm HoareTotal.GuaranteeStrip}; val GuardsNil = get_rule @{thm HoarePartial.GuardsNil} @{thm HoareTotal.GuardsNil}; val GuardsCons = get_rule @{thm HoarePartial.GuardsCons} @{thm HoareTotal.GuardsCons}; val GuardsConsGuaranteeStrip = get_rule @{thm HoarePartial.GuardsConsGuaranteeStrip} @{thm HoareTotal.GuardsConsGuaranteeStrip}; val Skip = get_rule @{thm HoarePartial.Skip} @{thm HoareTotal.Skip}; val Basic = get_rule @{thm HoarePartial.Basic} @{thm HoareTotal.Basic}; val BasicCond = get_rule @{thm HoarePartial.BasicCond} @{thm HoareTotal.BasicCond}; val Spec = get_rule @{thm HoarePartial.Spec} @{thm HoareTotal.Spec}; val SpecIf = get_rule @{thm HoarePartial.SpecIf} @{thm HoareTotal.SpecIf}; val Throw = get_rule @{thm HoarePartial.Throw} @{thm HoareTotal.Throw}; val Raise = get_rule @{thm HoarePartial.raise} @{thm HoareTotal.raise}; -val Catch = get_rule @{thm HoarePartial.Catch} @{thm HoareTotal.Catch}; - -val CondCatch = get_rule @{thm HoarePartial.condCatch} @{thm HoareTotal.condCatch}; +val Catch = get_rule' @{thm HoarePartial.Catch} @{thm HoareTotal.Catch} @{thm HoarePartial.CatchSame}; + +val CondCatch = get_rule' @{thm HoarePartial.condCatch} @{thm HoareTotal.condCatch} @{thm HoarePartial.condCatchSame}; val CatchSwap = get_rule @{thm HoarePartial.CatchSwap} @{thm HoareTotal.CatchSwap}; val CondCatchSwap = get_rule @{thm HoarePartial.condCatchSwap} @{thm HoareTotal.condCatchSwap}; -val Seq = get_rule @{thm HoarePartial.Seq} @{thm HoareTotal.Seq}; +val Seq = get_rule' @{thm HoarePartial.Seq} @{thm HoareTotal.Seq} @{thm HoarePartial.SeqSame}; val SeqSwap = get_rule @{thm HoarePartial.SeqSwap} @{thm HoareTotal.SeqSwap}; -val BSeq = get_rule @{thm HoarePartial.BSeq} @{thm HoareTotal.BSeq}; +val BSeq = get_rule' @{thm HoarePartial.BSeq} @{thm HoareTotal.BSeq} @{thm HoarePartial.BSeqSame}; val Cond = get_rule @{thm HoarePartial.Cond} @{thm HoareTotal.Cond}; val CondInv'Partial = @{thm HoarePartial.CondInv'}; val CondInv'Total = @{thm HoareTotal.CondInv'}; val CondInv' = get_rule CondInv'Partial CondInv'Total; val SwitchNil = get_rule @{thm HoarePartial.switchNil} @{thm HoareTotal.switchNil}; val SwitchCons = get_rule @{thm HoarePartial.switchCons} @{thm HoareTotal.switchCons}; val CondSwap = get_rule @{thm HoarePartial.CondSwap} @{thm HoareTotal.CondSwap}; val While = get_rule @{thm HoarePartial.While} @{thm HoareTotal.While}; val WhileAnnoG = get_rule @{thm HoarePartial.WhileAnnoG} @{thm HoareTotal.WhileAnnoG}; val WhileAnnoFix = get_rule @{thm HoarePartial.WhileAnnoFix'} @{thm HoareTotal.WhileAnnoFix'}; val WhileAnnoGFix = get_rule @{thm HoarePartial.WhileAnnoGFix} @{thm HoareTotal.WhileAnnoGFix}; val BindR = get_rule @{thm HoarePartial.Bind} @{thm HoareTotal.Bind}; val Block = get_rule @{thm HoarePartial.Block} @{thm HoareTotal.Block}; val BlockSwap = get_rule @{thm HoarePartial.BlockSwap} @{thm HoareTotal.BlockSwap}; -val Proc = get_rule @{thm HoarePartial.ProcSpec} @{thm HoareTotal.ProcSpec}; - -val ProcNoAbr = get_rule @{thm HoarePartial.ProcSpecNoAbrupt} @{thm HoareTotal.ProcSpecNoAbrupt}; +val Proc = get_call_rule + @{thm HoarePartial.ProcSpec} @{thm HoareTotal.ProcSpec} + @{thm HoarePartial.Proc_exnSpec} @{thm HoareTotal.Proc_exnSpec}; + +val ProcNoAbr = get_call_rule + @{thm HoarePartial.ProcSpecNoAbrupt} @{thm HoareTotal.ProcSpecNoAbrupt} + @{thm HoarePartial.Proc_exnSpecNoAbrupt} @{thm HoareTotal.Proc_exnSpecNoAbrupt}; val ProcBody = get_rule @{thm HoarePartial.ProcBody} @{thm HoareTotal.ProcBody}; -val CallBody = get_rule @{thm HoarePartial.CallBody} @{thm HoareTotal.CallBody}; +val CallBody = get_call_rule + @{thm HoarePartial.CallBody} @{thm HoareTotal.CallBody} + @{thm HoarePartial.Call_exnBody} @{thm HoareTotal.Call_exnBody}; val FCall = get_rule @{thm HoarePartial.FCall} @{thm HoareTotal.FCall}; val ProcRecSpecs = get_rule @{thm HoarePartial.ProcRecSpecs} @{thm HoareTotal.ProcRecSpecs}; -val ProcModifyReturnSameFaults = - get_rule @{thm HoarePartial.ProcModifyReturnSameFaults} @{thm HoareTotal.ProcModifyReturnSameFaults}; - -val ProcModifyReturn = get_rule @{thm HoarePartial.ProcModifyReturn} @{thm HoareTotal.ProcModifyReturn}; - -val ProcModifyReturnNoAbr = get_rule @{thm HoarePartial.ProcModifyReturnNoAbr} @{thm HoareTotal.ProcModifyReturnNoAbr}; +val ProcModifyReturnSameFaults = get_call_rule + @{thm HoarePartial.ProcModifyReturnSameFaults} @{thm HoareTotal.ProcModifyReturnSameFaults} + @{thm HoarePartial.Proc_exnModifyReturnSameFaults} @{thm HoareTotal.Proc_exnModifyReturnSameFaults}; + +val ProcModifyReturn = get_call_rule + @{thm HoarePartial.ProcModifyReturn} @{thm HoareTotal.ProcModifyReturn} + @{thm HoarePartial.Proc_exnModifyReturn} @{thm HoareTotal.Proc_exnModifyReturn}; + +val ProcModifyReturnNoAbr = get_call_rule + @{thm HoarePartial.ProcModifyReturnNoAbr} @{thm HoareTotal.ProcModifyReturnNoAbr} + @{thm HoarePartial.Proc_exnModifyReturnNoAbr} @{thm HoareTotal.Proc_exnModifyReturnNoAbr}; val ProcModifyReturnNoAbrSameFaultsPartial = @{thm HoarePartial.ProcModifyReturnNoAbrSameFaults}; val ProcModifyReturnNoAbrSameFaultsTotal = @{thm HoareTotal.ProcModifyReturnNoAbrSameFaults}; -val ProcModifyReturnNoAbrSameFaults = - get_rule ProcModifyReturnNoAbrSameFaultsPartial ProcModifyReturnNoAbrSameFaultsTotal; +val ProcModifyReturnNoAbrSameFaults = get_call_rule + ProcModifyReturnNoAbrSameFaultsPartial ProcModifyReturnNoAbrSameFaultsTotal + @{thm HoarePartial.Proc_exnModifyReturnNoAbrSameFaults} @{thm HoareTotal.Proc_exnModifyReturnNoAbrSameFaults}; val TrivPost = get_rule @{thm HoarePartial.TrivPost} @{thm HoareTotal.TrivPost}; val TrivPostNoAbr = get_rule @{thm HoarePartial.TrivPostNoAbr} @{thm HoareTotal.TrivPostNoAbr}; -val DynProcProcPar = get_rule @{thm HoarePartial.DynProcProcPar} @{thm HoareTotal.DynProcProcPar}; - -val DynProcProcParNoAbr = get_rule @{thm HoarePartial.DynProcProcParNoAbrupt} @{thm HoareTotal.DynProcProcParNoAbrupt}; - -val ProcProcParModifyReturn = get_rule @{thm HoarePartial.ProcProcParModifyReturn} @{thm HoareTotal.ProcProcParModifyReturn}; +val DynProcProcPar = get_call_rule + @{thm HoarePartial.DynProcProcPar} @{thm HoareTotal.DynProcProcPar} + @{thm HoarePartial.DynProc_exnProcPar} @{thm HoareTotal.DynProc_exnProcPar}; + +val DynProcProcParNoAbr = get_call_rule + @{thm HoarePartial.DynProcProcParNoAbrupt} @{thm HoareTotal.DynProcProcParNoAbrupt} + @{thm HoarePartial.DynProc_exnProcParNoAbrupt} @{thm HoareTotal.DynProc_exnProcParNoAbrupt}; + +val ProcProcParModifyReturn = get_call_rule + @{thm HoarePartial.ProcProcParModifyReturn} @{thm HoareTotal.ProcProcParModifyReturn} + @{thm HoarePartial.Proc_exnProcParModifyReturn} @{thm HoareTotal.Proc_exnProcParModifyReturn}; val ProcProcParModifyReturnSameFaultsPartial = @{thm HoarePartial.ProcProcParModifyReturnSameFaults}; val ProcProcParModifyReturnSameFaultsTotal = @{thm HoareTotal.ProcProcParModifyReturnSameFaults}; -val ProcProcParModifyReturnSameFaults = - get_rule ProcProcParModifyReturnSameFaultsPartial - ProcProcParModifyReturnSameFaultsTotal; - -val ProcProcParModifyReturnNoAbr = - get_rule @{thm HoarePartial.ProcProcParModifyReturnNoAbr} @{thm HoareTotal.ProcProcParModifyReturnNoAbr}; +val ProcProcParModifyReturnSameFaults = get_call_rule + ProcProcParModifyReturnSameFaultsPartial ProcProcParModifyReturnSameFaultsTotal + @{thm HoarePartial.ProcProcParModifyReturnSameFaults} @{thm HoareTotal.ProcProcParModifyReturnSameFaults}; + +val ProcProcParModifyReturnNoAbr = get_call_rule + @{thm HoarePartial.ProcProcParModifyReturnNoAbr} @{thm HoareTotal.ProcProcParModifyReturnNoAbr} + @{thm HoarePartial.Proc_exnProcParModifyReturnNoAbr} @{thm HoareTotal.Proc_exnProcParModifyReturnNoAbr}; val ProcProcParModifyReturnNoAbrSameFaultsPartial = @{thm HoarePartial.ProcProcParModifyReturnNoAbrSameFaults}; val ProcProcParModifyReturnNoAbrSameFaultsTotal = @{thm HoareTotal.ProcProcParModifyReturnNoAbrSameFaults}; -val ProcProcParModifyReturnNoAbrSameFaults = - get_rule ProcProcParModifyReturnNoAbrSameFaultsPartial - ProcProcParModifyReturnNoAbrSameFaultsTotal; +val ProcProcParModifyReturnNoAbrSameFaults = get_call_rule + ProcProcParModifyReturnNoAbrSameFaultsPartial ProcProcParModifyReturnNoAbrSameFaultsTotal + @{thm HoarePartial.Proc_exnProcParModifyReturnNoAbrSameFaults} @{thm HoareTotal.Proc_exnProcParModifyReturnNoAbrSameFaults}; val DynCom = get_rule @{thm HoarePartial.DynComConseq} @{thm HoareTotal.DynComConseq}; val AugmentContext = get_rule @{thm HoarePartial.augment_context'} @{thm HoareTotal.augment_context'}; val AugmentEmptyFaults = get_rule @{thm HoarePartial.augment_emptyFaults} @{thm HoareTotal.augment_emptyFaults}; val AsmUN = get_rule @{thm HoarePartial.AsmUN} @{thm HoareTotal.AsmUN}; val SpecAnno = get_rule @{thm HoarePartial.SpecAnno'} @{thm HoareTotal.SpecAnno'}; val SpecAnnoNoAbrupt = get_rule @{thm HoarePartial.SpecAnnoNoAbrupt} @{thm HoareTotal.SpecAnnoNoAbrupt}; val LemAnno = get_rule @{thm HoarePartial.LemAnno} @{thm HoareTotal.LemAnno}; val LemAnnoNoAbrupt = get_rule @{thm HoarePartial.LemAnnoNoAbrupt} @{thm HoareTotal.LemAnnoNoAbrupt}; val singleton_conv_sym = @{thm Set.singleton_conv2} RS sym; val anno_defs = [@{thm Language.whileAnno_def},@{thm Language.whileAnnoG_def},@{thm Language.specAnno_def}, @{thm Language.whileAnnoGFix_def},@{thm Language.whileAnnoFix_def},@{thm Language.lem_def}]; val strip_simps = @{thm Language.strip_simp} :: @{thm Option.option.map(2)} :: @{thms Language.strip_guards_simps}; val normalize_simps = [@{thm Language.while_def}, @{thm Language.bseq_def}, @{thm List.append_Nil}, @{thm List.append_Cons}] @ @{thms List.list.cases} @ @{thms Language.flatten_simps} @ @{thms Language.sequence.simps} @ @{thms Language.normalize_simps} @ @{thms Language.guards.simps} @ [@{thm fst_conv}, @{thm snd_conv}]; val K_rec_convs = []; val K_fun_convs = [@{thm StateFun.K_statefun_apply}, @{thm StateFun.K_statefun_comp}]; val K_convs = K_rec_convs @ K_fun_convs; val K_rec_congs = []; val K_fun_congs = [@{thm StateFun.K_statefun_cong}]; val K_congs = K_rec_congs @ K_fun_congs; (* misc. aux. functions *) (* first_subterm * yields result x of P for first subterm for which P is (SOME x), and all bound variables on the path * to that term *) fun first_subterm_dest P = let fun first abs_vars t = (case P t of SOME x => SOME (abs_vars,x) |_=> (case t of u $ v => (case first abs_vars u of NONE => first abs_vars v | SOME x => SOME x) | Abs (c,T,u) => first (abs_vars @ [(c,T)]) u | _ => NONE)) in first [] end; (* first_subterm * yields first subterm for which P holds, and all bound variables on the path * to that term *) fun first_subterm P = let fun P' t = if P t then SOME t else NONE; in first_subterm_dest P' end; (* max_subterm_dest * yields results of P for all maximal subterms for which P is (SOME x), * and all bound variables on the path to that subterm *) fun max_subterms_dest P = let fun collect abs_vars t = (case P t of SOME x => [(abs_vars,x)] |_=> (case t of u $ v => collect abs_vars u @ collect abs_vars v | Abs (c,T,u) => collect (abs_vars @ [(c,T)]) u | _ => [])) in collect [] end; fun last [] = raise Empty | last [x] = x | last (_::xs) = last xs; fun dest_splits (Const (@{const_name case_prod},_)$Abs (n,T,t)) = (n,T)::dest_splits t | dest_splits (Const (@{const_name case_prod},_)$Abs (n,T,t)$_) = (n,T)::dest_splits t | dest_splits (Abs (n,T,_)) = [(n,T)] | dest_splits _ = []; fun idx eq [] x = ~1 | idx eq (x::rs) y = if eq x y then 0 else let val i = idx eq rs y in if i < 0 then i else i+1 end; fun resuffix sfx1 sfx2 s = suffix sfx2 (unsuffix sfx1 s) handle Fail _ => s; -(* state space representation dependent functions *) - -datatype state_kind = Record | Function -fun state_simprocs Record = [Record.simproc] - | state_simprocs Function = [Record.simproc, StateFun.lookup_simproc]; - -fun state_upd_simproc Record = Record.upd_simproc - | state_upd_simproc Function = StateFun.update_simproc; - -fun state_ex_sel_eq_simproc Record = Record.ex_sel_eq_simproc - | state_ex_sel_eq_simproc Function = StateFun.ex_lookup_eq_simproc; - -val state_split_simp_tac = Record.split_simp_tac -val state_hierarchy = Record.dest_recTs - - -fun stateT_id T = case (state_hierarchy T) of [] => NONE | Ts => SOME (last Ts); - -fun globalsT (Type (_, T :: _)) = SOME T - | globalsT _ = NONE; - -fun stateT_ids T = - (case stateT_id T of - NONE => NONE - | SOME sT => (case globalsT T of - NONE => SOME [sT] - | SOME gT => (case stateT_id gT of - NONE => SOME [sT] - | SOME gT' => SOME [sT,gT']))); - datatype par_kind = In | Out +datatype state_kind = Record | Function | Other of string; + (*** utilities ***) (* utils for variable name decorations *) val deco = "_'"; val proc_deco = "_'proc"; fun par_deco name = deco ^ name ^ deco; fun chopsfx sfx str = (case try (unsuffix sfx) str of SOME s => s | NONE => str) val is_state_var = can (unsuffix deco); -(* removes the suffix of the string beginning with deco. - * "xys_'a" --> "xys"; - * The a is also chopped, since sometimes the bound variables - * are renamed, I think SELECT_GOAL in rename_goal is to blame - *) -fun remdeco' str = - let - fun chop (p::ps) (x::xs) = chop ps xs - | chop [] xs = [] - | chop (p::ps) [] = error "remdeco: code should never be reached"; - - fun remove prf (s as (x::xs)) = if is_prefix (op =) prf s then chop prf s - else (x::remove prf xs) - | remove prf [] = []; - - in String.implode (remove (String.explode deco) (String.explode str)) end; fun extern ctxt s = (case try (Proof_Context.extern_const ctxt o Lexicon.unmark_const) s of NONE => s | SOME s' => s'); -fun remdeco ctxt s = remdeco' (extern ctxt s); - -fun undeco ctxt (Const (c, T)) = Const (remdeco ctxt c, T) - | undeco ctxt ((f as Const (@{syntax_const "_free"},_)) $ Free (x, T)) = - (*f$*)Const (remdeco' x, T) - | undeco ctxt (Const _ $ _ $ ((Const (@{syntax_const "_free"},_)) $ Free (x, T))) = - (*f$*)Const (remdeco' x, T) - | undeco ctxt (Free (c, T)) = Const (remdeco' c, T) - | undeco ctxt x = x - fun varname x = x ^ deco -val dest_string = map (chr o HOLogic.dest_char) o HOLogic.dest_list; +val dest_string = implode o map (chr o HOLogic.dest_char) o HOLogic.dest_list; fun dest_string' t = (case try dest_string t of - SOME s => implode s + SOME s => s | NONE => (case t of Free (s,_) => s | Const (s,_) => Long_Name.base_name s | _ => raise TERM ("dest_string'",[t]))) - -fun is_state_space_var Tids t = - let - fun is_stateT T = (case stateT_id T of NONE => 0 - | SOME id => if member (op =) Tids id then ~1 else 0); - in - (case t of - Const _ $ Abs (_,T,_) => is_stateT T - | Free (_,T) => is_stateT T - | _ => 0) - end; - - -datatype callMode = Static | Parameter - -fun proc_name Static (Const (p,_)$_) = resuffix deco proc_deco (Long_Name.base_name p) - | proc_name Static (Const (@{const_name StateFun.lookup},_)$_$Free (p,_)$_) = - suffix proc_deco (remdeco' (Long_Name.base_name p)) - | proc_name Static p = dest_string' p - | proc_name Parameter (Const (p,_)) = resuffix deco proc_deco (Long_Name.base_name p) - | proc_name Parameter (Abs (_,_,Const (p,_)$Bound 0)) = - resuffix deco proc_deco (Long_Name.base_name p) - | proc_name Parameter (Abs (_,_,Const (@{const_name StateFun.lookup},_)$_$Free (p,_)$_)) = - suffix proc_deco (remdeco' (Long_Name.base_name p)) - | proc_name _ t = raise TERM ("proc_name",[t]); - - - -fun dest_call (Const (@{const_name Language.call},_)$init$pname$return$c) = - (init,pname,return,c,Static,true) - | dest_call (Const (@{const_name Language.fcall},_)$init$pname$return$_$c) = - (init,pname,return,c,Static,true) - | dest_call (Const (@{const_name Language.com.Call},_)$pname) = - (Bound 0,pname,Bound 0,Bound 0,Static,false) - | dest_call (Const (@{const_name Language.dynCall},_)$init$pname$return$c) = - (init,pname,return,c,Parameter,true) - | dest_call t = raise TERM ("Hoare.dest_call: unexpected term",[t]); - -fun dest_whileAnno (Const (@{const_name Language.whileAnnoG},_) $gs$b$I$V$c) = - (SOME gs,b,I,V,c,false) - | dest_whileAnno (Const (@{const_name Language.whileAnno},_) $b$I$V$c) = (NONE,b,I,V,c,false) - | dest_whileAnno (Const (@{const_name Language.whileAnnoGFix},_)$gs$b$I$V$c) = - (SOME gs,b,I,V,c,true) - | dest_whileAnno (Const (@{const_name Language.whileAnnoFix},_) $b$I$V$c) = (NONE,b,I,V,c,true) - | dest_whileAnno t = raise TERM ("Hoare.dest_while: unexpected term",[t]); - -fun dest_Guard (Const (@{const_name Language.com.Guard},_)$f$g$c) = (f,g,c,false) - | dest_Guard (Const (@{const_name Language.guaranteeStrip},_)$f$g$c) = (f,g,c,true) - | dest_Guard t = raise TERM ("Hoare.dest_guard: unexpected term",[t]); - +val state_hierarchy = Record.dest_recTs +fun stateT_id T = case (state_hierarchy T) of [] => NONE | Ts => SOME (last Ts); + +fun globalsT (Type (_, T :: _)) = SOME T + | globalsT _ = NONE; + +fun stateT_ids T = + (case stateT_id T of + NONE => NONE + | SOME sT => (case globalsT T of + NONE => SOME [sT] + | SOME gT => (case stateT_id gT of + NONE => SOME [sT] + | SOME gT' => SOME [sT,gT']))); (*** extend theory by procedure definition ***) fun add_declaration name decl thy = thy |> Named_Target.init [] name |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} decl |> Local_Theory.exit |> Proof_Context.theory_of; (* data kind 'HOL/hoare' *) +type lense = {lookup: term, update : term} + type proc_info = - {params: ((par_kind * string) list), - recursive: bool, - state_kind: state_kind} + { + params: ((par_kind * string * lense option) list), + recursive: bool, + state_kind: state_kind + }; + +type state_space = + { + name: string, + is_state_type: Proof.context -> typ -> bool, + generalise: Proof.context -> int -> tactic, + state_simprocs: simproc list, + state_upd_simprocs: simproc list, + state_ex_sel_eq_simprocs: simproc list, + is_defined: Proof.context -> xstring -> bool, + read_function_name: Proof.context -> xstring -> term, + lookup_tr: Proof.context -> xstring -> term, + update_tr: Proof.context -> xstring -> term, + is_lookup: Proof.context -> term -> bool, + lookup_tr': Proof.context -> term -> term, + dest_update_tr': Proof.context -> term -> term * term * term option, + update_tr': Proof.context -> term -> term + }; + type hoare_tac = (bool -> int -> tactic) -> Proof.context -> hoareMode -> int -> tactic; type hoare_data = - {proc_info: proc_info Symtab.table, - active_procs: string list list, - default_state_kind: state_kind, - generate_guard: (stamp * (Proof.context -> term -> term option)), - wp_tacs: (string * hoare_tac) list, - hoare_tacs: (string * hoare_tac) list, - vcg_simps: thm list}; + { + proc_info: proc_info Symtab.table, + active_procs: string list list, + default_state_kind: state_kind, + generate_guard: (stamp option * (Proof.context -> term -> term option)), + name_tr: (stamp option * (Proof.context -> bool -> string -> string)), + hoare_tacs: (string * hoare_tac) list, + vcg_simps: thm list, + state_spaces: (string * state_space) list + }; fun make_hoare_data proc_info active_procs default_state_kind generate_guard - wp_tacs hoare_tacs vcg_simps = + name_tr hoare_tacs vcg_simps state_spaces = {proc_info = proc_info, active_procs = active_procs, default_state_kind = default_state_kind, generate_guard = generate_guard, - wp_tacs = wp_tacs, hoare_tacs = hoare_tacs, vcg_simps = vcg_simps}; + name_tr = name_tr, hoare_tacs = hoare_tacs, vcg_simps = vcg_simps, state_spaces = state_spaces}; + +fun merge_stamped err_msg ((NONE, _), p) = p + | merge_stamped err_msg (p, (NONE,_)) = p + | merge_stamped err_msg ((SOME (stamp1:stamp), x), (SOME stamp2, _)) = + if stamp1 = stamp2 then (SOME stamp1, x) + else error err_msg; + +fun fast_merge merge (x, y) = if pointer_eq (x, y) then x else merge (x, y) structure Hoare_Data = Generic_Data ( type T = hoare_data; val empty = make_hoare_data (Symtab.empty: proc_info Symtab.table) ([]:string list list) (Function) - (stamp (),(K (K NONE)): Proof.context -> term -> term option) - ([]:(string * hoare_tac) list) + (NONE,(K (K NONE)): Proof.context -> term -> term option) + (NONE,(K (K I)): Proof.context -> bool -> string -> string) ([]:(string * hoare_tac) list) - ([]:thm list); - - (* FIXME exponential blowup due to append !? *) - fun merge ({proc_info = proc_info1, active_procs = active_procs1, + ([]:thm list) + ([]:(string * state_space) list); + + val merge = fast_merge (fn ({proc_info = proc_info1, active_procs = active_procs1, default_state_kind = _, - generate_guard = (stmp1,generate_gaurd1), - wp_tacs = wp_tacs1, hoare_tacs = hoare_tacs1, vcg_simps = vcg_simps1}, + generate_guard = generate_guard1, + name_tr = name_tr1, hoare_tacs = hoare_tacs1, vcg_simps = vcg_simps1, state_spaces=state_spaces1}, {proc_info = proc_info2, active_procs = active_procs2, default_state_kind = default_state_kind2, - generate_guard = (stmp2, _), - wp_tacs = wp_tacs2, hoare_tacs = hoare_tacs2, vcg_simps=vcg_simps2}) : T = - if stmp1=stmp2 then + generate_guard = generate_guard2, + name_tr = name_tr2, hoare_tacs = hoare_tacs2, vcg_simps=vcg_simps2, state_spaces=state_spaces2}) => make_hoare_data (Symtab.merge (K true) (proc_info1,proc_info2)) (active_procs1 @ active_procs2) (default_state_kind2) - (stmp1,generate_gaurd1) - (wp_tacs1 @ wp_tacs2) - (hoare_tacs1 @ hoare_tacs2) + (merge_stamped + "Theories have different aux. functions to generate guards, please resolve before merge" + (generate_guard1, generate_guard2)) + (merge_stamped + "Theories have different aux. functions to translate names, please resolve before merge" + (name_tr1, name_tr2)) + (AList.merge (op =) (K true) (hoare_tacs1, hoare_tacs2)) (Thm.merge_thms (vcg_simps1,vcg_simps2)) - else error ("Theories have different aux. functions to generate guards") + (AList.merge (op =) (K true) (state_spaces1, state_spaces2))) ); val get_data = Hoare_Data.get o Context.Proof; +(* state space representation dependent functions *) + +fun get_state_space_comps sel ctxt n = + AList.lookup (op =) (#state_spaces (Hoare_Data.get (Context.Proof ctxt))) n + |> Option.map sel |> these; + + +fun state_simprocs ctxt Record = [Record.simproc] + | state_simprocs ctxt Function = [Record.simproc, StateFun.lookup_simproc] + | state_simprocs ctxt (Other n) = get_state_space_comps (#state_simprocs) ctxt n; + + +fun state_upd_simprocs ctxt Record = [Record.upd_simproc] + | state_upd_simprocs ctxt Function = [StateFun.update_simproc] + | state_upd_simprocs ctxt (Other n) = get_state_space_comps (#state_upd_simprocs) ctxt n; + +fun state_ex_sel_eq_simprocs ctxt Record = [Record.ex_sel_eq_simproc] + | state_ex_sel_eq_simprocs ctxt Function = [StateFun.ex_lookup_eq_simproc] + | state_ex_sel_eq_simprocs ctxt (Other n) = get_state_space_comps (#state_ex_sel_eq_simprocs) ctxt n; + +val state_split_simp_tac = Record.split_simp_tac +val state_hierarchy = Record.dest_recTs + + +fun stateT_id T = case (state_hierarchy T) of [] => NONE | Ts => SOME (last Ts); + +fun globalsT (Type (_, T :: _)) = SOME T + | globalsT _ = NONE; + +fun stateT_ids T = + (case stateT_id T of + NONE => NONE + | SOME sT => (case globalsT T of + NONE => SOME [sT] + | SOME gT => (case stateT_id gT of + NONE => SOME [sT] + | SOME gT' => SOME [sT,gT']))); + (* access 'params' *) fun mk_free context name = let val ctxt = Context.proof_of context; val n' = Variable.intern_fixed ctxt name |> perhaps Long_Name.dest_hidden; val T' = Proof_Context.infer_type ctxt (n', dummyT) handle ERROR _ => dummyT in (Free (n',T')) end; fun morph_name context phi name = (case Morphism.term phi (mk_free context name) of Free (x,_) => x | _ => name); datatype 'a bodykind = BodyTyp of 'a | BodyTerm of 'a fun set_default_state_kind sk context = let - val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, - vcg_simps,...} + val {proc_info,active_procs,default_state_kind,generate_guard,name_tr,hoare_tacs, + vcg_simps,state_spaces, ...} = Hoare_Data.get context; val data = make_hoare_data proc_info active_procs sk - generate_guard wp_tacs hoare_tacs vcg_simps; + generate_guard name_tr hoare_tacs vcg_simps state_spaces; in Hoare_Data.put data context end; val get_default_state_kind = #default_state_kind o get_data; +fun get_default_state_space ctxt = + case get_default_state_kind ctxt of + Other sp => AList.lookup (op =) (#state_spaces (Hoare_Data.get (Context.Proof ctxt))) sp + | _ => NONE + fun add_active_procs phi ps context = let - val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, - vcg_simps,...} + val {proc_info,active_procs,default_state_kind,generate_guard,name_tr,hoare_tacs, + vcg_simps,state_spaces, ...} = Hoare_Data.get context; val data = make_hoare_data proc_info ((map (morph_name context phi) ps)::active_procs) default_state_kind - generate_guard wp_tacs hoare_tacs vcg_simps; + generate_guard name_tr hoare_tacs vcg_simps state_spaces; in Hoare_Data.put data context end; fun add_hoare_tacs tacs context = let - val {proc_info,active_procs, default_state_kind, generate_guard,wp_tacs,hoare_tacs, - vcg_simps,...} + val {proc_info,active_procs, default_state_kind, generate_guard,name_tr,hoare_tacs, + vcg_simps,state_spaces,...} = Hoare_Data.get context; val data = make_hoare_data proc_info active_procs default_state_kind generate_guard - wp_tacs (hoare_tacs@tacs) vcg_simps; + name_tr (AList.merge (op =) (K true) (hoare_tacs, tacs)) vcg_simps state_spaces; in Hoare_Data.put data context end; fun map_vcg_simps f context = let - val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, - vcg_simps,...} + val {proc_info,active_procs,default_state_kind,generate_guard,name_tr,hoare_tacs, + vcg_simps,state_spaces,...} = Hoare_Data.get context; val data = make_hoare_data proc_info active_procs default_state_kind generate_guard - wp_tacs hoare_tacs (f vcg_simps); + name_tr hoare_tacs (f vcg_simps) state_spaces; in Hoare_Data.put data context end; -fun thy_attrib f = Thm.declaration_attribute (fn thm => map_vcg_simps (f thm)); - -val vcg_simpadd = Thm.add_thm -val vcg_simpdel = Thm.del_thm - -val vcg_simp_add = thy_attrib vcg_simpadd; -val vcg_simp_del = thy_attrib vcg_simpdel; +val vcg_simp_add = Thm.declaration_attribute (map_vcg_simps o Thm.add_thm o Thm.trim_context); +val vcg_simp_del = Thm.declaration_attribute (map_vcg_simps o Thm.del_thm); (* add 'procedure' *) fun mk_proc_info params recursive state_kind = {params=params,recursive=recursive,state_kind=state_kind}; val empty_proc_info = {params=[],recursive=false,state_kind=Record}; fun map_proc_info_params f {params,recursive,state_kind} = mk_proc_info (f params) recursive state_kind; fun map_proc_info_recursive f {params,recursive,state_kind} = mk_proc_info params (f recursive) state_kind; fun map_proc_info_state_kind f {params,recursive,state_kind} = mk_proc_info params recursive (f state_kind); +fun morph_lense phi ({lookup, update}:lense) = + {lookup = Morphism.term phi lookup, update = Morphism.term phi update}:lense; fun add_params phi name frmls context = let - val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, - vcg_simps,...} + val {proc_info,active_procs,default_state_kind,generate_guard,name_tr,hoare_tacs, + vcg_simps,state_spaces, ...} = Hoare_Data.get context; - val params = map (apsnd (morph_name context phi)) frmls; + val params = map (fn (kind, name, lense_opt) => + (kind, morph_name context phi name, Option.map (morph_lense phi) lense_opt)) frmls; val f = map_proc_info_params (K params); val default = f empty_proc_info; - val proc_info' = Symtab.map_default (morph_name context phi name,default) f proc_info; + val proc_info' = Symtab.map_default (morph_name context phi name, default) f proc_info; val data = make_hoare_data proc_info' active_procs default_state_kind - generate_guard wp_tacs hoare_tacs vcg_simps; + generate_guard name_tr hoare_tacs vcg_simps state_spaces; in Hoare_Data.put data context end; fun get_params name ctxt = Option.map #params (Symtab.lookup (#proc_info (get_data ctxt)) name); fun add_recursive phi name context = let - val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, - vcg_simps,...} + val {proc_info,active_procs,default_state_kind,generate_guard,name_tr,hoare_tacs, + vcg_simps,state_spaces, ...} = Hoare_Data.get context; val f = map_proc_info_recursive (K true); val default = f empty_proc_info; val proc_info'= Symtab.map_default (morph_name context phi name,default) f proc_info; val data = make_hoare_data proc_info' active_procs default_state_kind - generate_guard wp_tacs hoare_tacs vcg_simps; + generate_guard name_tr hoare_tacs vcg_simps state_spaces; in Hoare_Data.put data context end; fun get_recursive name ctxt = Option.map #recursive (Symtab.lookup (#proc_info (get_data ctxt)) name); fun add_state_kind phi name sk context = let - val {proc_info,active_procs,default_state_kind,generate_guard,wp_tacs,hoare_tacs, - vcg_simps,...} + val {proc_info,active_procs,default_state_kind,generate_guard,name_tr,hoare_tacs, + vcg_simps,state_spaces,...} = Hoare_Data.get context; val f = map_proc_info_state_kind (K sk); val default = f empty_proc_info; val proc_info'= Symtab.map_default (morph_name context phi name,default) f proc_info; val data = make_hoare_data proc_info' active_procs default_state_kind - generate_guard wp_tacs hoare_tacs vcg_simps; + generate_guard name_tr hoare_tacs vcg_simps state_spaces; in Hoare_Data.put data context end; fun get_state_kind name ctxt = Option.map #state_kind (Symtab.lookup (#proc_info (get_data ctxt)) name); fun install_generate_guard f context = let - val {proc_info,active_procs, default_state_kind, generate_guard,wp_tacs,hoare_tacs, - vcg_simps,...} = + val {proc_info,active_procs, default_state_kind, generate_guard,name_tr,hoare_tacs, + vcg_simps,state_spaces,...} = Hoare_Data.get context; - val data = make_hoare_data proc_info active_procs default_state_kind (stamp (), f) - wp_tacs hoare_tacs vcg_simps + val data = make_hoare_data proc_info active_procs default_state_kind (SOME (stamp ()), f) + name_tr hoare_tacs vcg_simps state_spaces in Hoare_Data.put data context end; fun generate_guard ctxt = snd (#generate_guard (get_data ctxt)) ctxt; +fun install_state_space sp ctxt = + let + val {proc_info,active_procs, default_state_kind, generate_guard,name_tr,hoare_tacs, + vcg_simps,state_spaces,...} = + Hoare_Data.get ctxt; + val data = make_hoare_data proc_info active_procs default_state_kind generate_guard + name_tr hoare_tacs vcg_simps (AList.update (op =) (#name sp, sp) state_spaces) + in Hoare_Data.put data ctxt end; + +fun generalise_other ctxt name = + Option.map #generalise (AList.lookup (op =) (#state_spaces (get_data ctxt)) name); + +fun install_name_tr f ctxt = + let + val {proc_info,active_procs, default_state_kind, generate_guard,name_tr,hoare_tacs, + vcg_simps,state_spaces,...} = + Hoare_Data.get ctxt; + val data = make_hoare_data proc_info active_procs default_state_kind generate_guard + (SOME (stamp ()), f) hoare_tacs vcg_simps state_spaces + in Hoare_Data.put data ctxt end; + +fun name_tr ctxt = snd (#name_tr (get_data ctxt)) ctxt; + + +(* utils for variable name decorations *) + + +(* removes the suffix of the string beginning with deco. + * "xys_'a" --> "xys"; + * The a is also chopped, since sometimes the bound variables + * are renamed, I think SELECT_GOAL in rename_goal is to blame + *) +fun remdeco' ctxt str = + let + fun chop (p::ps) (x::xs) = chop ps xs + | chop [] xs = [] + | chop (p::ps) [] = error "remdeco: code should never be reached"; + + fun remove prf (s as (x::xs)) = if is_prefix (op =) prf s then chop prf s + else (x::remove prf xs) + | remove prf [] = []; + + in name_tr ctxt false (String.implode (remove (String.explode deco) (String.explode str))) end; + + +fun remdeco ctxt s = remdeco' ctxt (extern ctxt s); + +fun undeco ctxt (Const (c, T)) = Const (remdeco ctxt c, T) + | undeco ctxt ((f as Const (@{syntax_const "_free"},_)) $ Free (x, T)) = + (*f$*)Const (remdeco' ctxt x, T) + | undeco ctxt (Const _ $ _ $ ((Const (@{syntax_const "_free"},_)) $ Free (x, T))) = + (*f$*)Const (remdeco' ctxt x, T) + | undeco ctxt (Free (c, T)) = Const (remdeco' ctxt c, T) + | undeco ctxt x = x + +fun is_state_space_var Tids t = + let + fun is_stateT T = (case stateT_id T of NONE => 0 + | SOME id => if member (op =) Tids id then ~1 else 0); + in + (case t of + Const _ $ Abs (_,T,_) => is_stateT T + | Free (_,T) => is_stateT T + | _ => 0) + end; + + +datatype callMode = Static | Parameter + +fun proc_name ctxt Static (Const (p,_)$_) = resuffix deco proc_deco (Long_Name.base_name p) + | proc_name ctxt Static (Const (@{const_name StateFun.lookup},_)$_$Free (p,_)$_) = + suffix proc_deco (remdeco' ctxt (Long_Name.base_name p)) + | proc_name ctxt Static p = dest_string' p + | proc_name ctxt Parameter (Const (p,_)) = resuffix deco proc_deco (Long_Name.base_name p) + | proc_name ctxt Parameter (Abs (_,_,Const (p,_)$Bound 0)) = + resuffix deco proc_deco (Long_Name.base_name p) + | proc_name ctxt Parameter (Abs (_,_,Const (@{const_name StateFun.lookup},_)$_$Free (p,_)$_)) = + suffix proc_deco (remdeco' ctxt (Long_Name.base_name p)) + | proc_name _ _ t = raise TERM ("proc_name",[t]); + + +fun dest_call (Const (@{const_name Language.call},_)$init$pname$return$c) = + (init,pname,return,c,Static,true,NONE) + | dest_call (Const (@{const_name Language.fcall},_)$init$pname$return$_$c) = + (init,pname,return,c,Static,true,NONE) + | dest_call (Const (@{const_name Language.com.Call},_)$pname) = + (Bound 0,pname,Bound 0,Bound 0,Static,false,NONE) + | dest_call (Const (@{const_name Language.dynCall},_)$init$pname$return$c) = + (init,pname,return,c,Parameter,true,NONE) + | dest_call (Const (@{const_name Language.call_exn},_)$init$pname$return$result_exn$c) = + (init,pname,return,c,Static,true,SOME result_exn) + | dest_call (Const (@{const_name Language.dynCall_exn},_)$init$pname$return$result_exn$c) = + (init,pname,return,c,Parameter,true,SOME result_exn) + | dest_call t = raise TERM ("Hoare.dest_call: unexpected term",[t]); + + +fun dest_whileAnno (Const (@{const_name Language.whileAnnoG},_) $gs$b$I$V$c) = + (SOME gs,b,I,V,c,false) + | dest_whileAnno (Const (@{const_name Language.whileAnno},_) $b$I$V$c) = (NONE,b,I,V,c,false) + | dest_whileAnno (Const (@{const_name Language.whileAnnoGFix},_)$gs$b$I$V$c) = + (SOME gs,b,I,V,c,true) + | dest_whileAnno (Const (@{const_name Language.whileAnnoFix},_) $b$I$V$c) = (NONE,b,I,V,c,true) + | dest_whileAnno t = raise TERM ("Hoare.dest_while: unexpected term",[t]); + +fun dest_Guard (Const (@{const_name Language.com.Guard},_)$f$g$c) = (f,g,c,false) + | dest_Guard (Const (@{const_name Language.guaranteeStrip},_)$f$g$c) = (f,g,c,true) + | dest_Guard t = raise TERM ("Hoare.dest_guard: unexpected term",[t]); + + fun check_procedures_definition procs thy = let val ctxt = Proof_Context.init_global thy; fun already_defined name = if is_some (get_params name ctxt) then ["procedure " ^ quote name ^ " already defined"] else [] val err_already_defined = maps (already_defined o #1) procs; fun duplicate_procs names = (case duplicates (op =) names of [] => [] | dups => ["Duplicate procedures " ^ commas_quote dups]); val err_duplicate_procs = duplicate_procs (map #1 procs); fun duplicate_pars name pars = (case duplicates (op =) (map fst pars) of [] => [] | dups => ["Duplicate parameters in procedure " ^ quote name ^ ": " ^ commas_quote dups]); val err_duplicate_pars = maps (fn (name,inpars,outpars,locals,_,_,_) => duplicate_pars name (inpars @ locals) @ duplicate_pars name (outpars @ locals)) procs; - (* FIXME: Check that no global variables are used as result parameters *) + (* fixme: Check that no global variables are used as result parameters *) val errs = err_already_defined @ err_duplicate_procs @ err_duplicate_pars; in if null errs then () else error (cat_lines errs) end; fun add_parameter_info phi cname (name,(inpars,outpars,state_kind)) context = let fun par_deco' T = if T = "" then deco else par_deco (cname name); - val pars = map (fn (par,T) => (In,suffix (par_deco' T) par)) inpars@ - map (fn (par,T) => (Out,suffix (par_deco' T) par)) outpars; - in - context - |> add_params phi name pars - |> add_state_kind phi name state_kind + val pars = map (fn (par,T) => (In,suffix (par_deco' T) par, NONE)) inpars@ + map (fn (par,T) => (Out,suffix (par_deco' T) par, NONE)) outpars; + + val ctxt_decl = context + |> add_params phi name pars + |> add_state_kind phi name state_kind + in ctxt_decl end; fun mk_loc_exp xs = let fun mk_expr s = (s,(("",false),(Expression.Named [],[]))) in (map mk_expr xs,[]) end; val parametersN = "_parameters"; val variablesN = "_variables"; val signatureN = "_signature"; val bodyN = "_body"; val implementationN = "_impl"; val cliqueN = "_clique"; val clique_namesN = "_clique_names"; val NoBodyN = @{const_name Vcg.NoBody}; val statetypeN = "StateType"; val proc_nameT = HOLogic.stringT; fun expression_no_pos (expr, fixes) : Expression.expression = (map (fn (name, inst) => ((name, Position.none), inst)) expr, fixes); fun add_locale name expr elems thy = thy |> Expression.add_locale (Binding.name name) (Binding.name name) [] expr elems |> snd |> Local_Theory.exit; fun add_locale' name expr elems thy = thy |> Expression.add_locale (Binding.name name) (Binding.name name) [] expr elems ||> Local_Theory.exit; fun add_locale_cmd name expr elems thy = thy |> Expression.add_locale_cmd (Binding.name name) (Binding.name name) [] (expression_no_pos expr) elems |> snd |> Local_Theory.exit; fun read_typ thy raw_T env = let val ctxt' = Proof_Context.init_global thy |> fold (Variable.declare_typ o TFree) env; val T = Syntax.read_typ ctxt' raw_T; val env' = Term.add_tfreesT T env; in (T, env') end; fun add_variable_statespaces (cname, (inpars, outpars, locvars)) thy = let val inpars' = if forall (fn (_,T) => T = "") inpars then [] else inpars; val outpars' = if forall (fn (_,T) => T = "") outpars then [] else outpars; fun prep_comp (n, T) env = let val (T', env') = read_typ thy T env handle ERROR msg => cat_error msg ("The error(s) above occurred in component " ^ quote n) in ((n, T'), env') end; val (in_outs,in_out_env) = fold_map prep_comp (distinct (op =) (inpars'@outpars')) []; val (locs,var_env) = fold_map prep_comp locvars in_out_env; val parSP = cname ^ parametersN; val in_outs' = map (apfst (suffix (par_deco cname))) in_outs; val in_out_args = map fst in_out_env; val varSP = cname ^ variablesN; val locs' = map (apfst (suffix (par_deco cname))) locs; val var_args = map fst var_env; in if null inpars' andalso null outpars' andalso null locvars then thy |> add_locale_cmd parSP ([],[]) [] |> Proof_Context.theory_of |> add_locale_cmd varSP ([],[]) [] |> Proof_Context.theory_of else thy |> StateSpace.define_statespace_i (SOME statetypeN) in_out_args parSP [] in_outs' |> StateSpace.define_statespace_i (SOME statetypeN) var_args varSP [((cname, false), ((map TFree in_out_env),parSP,[]))] locs' end; fun intern_locale thy = Locale.intern thy #> perhaps Long_Name.dest_hidden; fun apply_in_context thy lexp f t = let fun name_variant lname = if intern_locale thy lname = lname then lname else name_variant (lname ^ "'"); in thy (* Create a dummy locale in dummy theory just to read the term *) |> add_locale_cmd (name_variant "foo") lexp [] |> (fn ctxt => f ctxt t) end; fun add_abbrev loc mode name spec thy = thy |> Named_Target.init [] loc |> (fn lthy => let val t = Syntax.read_term (Local_Theory.target_of lthy) spec; in Local_Theory.abbrev mode ((Binding.name name, NoSyn), t) lthy end) |> #2 |> Local_Theory.exit |> Proof_Context.theory_of; exception TOPSORT of string fun topsort less [] = [] | topsort less xs = let fun list_all P xs = fold (fn x => fn b => b andalso P x) xs true; fun split_min n (x::xs) = if n=0 then raise TOPSORT "no minimum in list" else if list_all (less x) xs then (x,xs) else split_min (n-1) (xs@[x]); fun tsort [] = [] | tsort xs = let val (x,xs') = split_min (length xs) xs; in x::tsort xs' end; in tsort xs end; fun clique_name clique = (foldr1 (fn (a,b) => a ^ "_" ^ b) (map (unsuffix proc_deco) clique)); fun error_to_warning msg f thy = f thy handle ERROR msg' => (warning (msg' ^ "\n" ^ msg); thy); fun procedures_definition locname procs thy = let val procs' = map (fn (name,a,b,c,d,e,f) => (suffix proc_deco name,a,b,c,d,e,f)) procs; val _ = check_procedures_definition procs' thy; val name_pars = map (fn (name,inpars,outpars,_,_,_,sk) => (name,(inpars,outpars,sk))) procs'; val name_vars = map (fn (name,inpars,outpars,locals,_,_,_) => (name,(inpars,outpars,locals))) procs'; val name_body = map (fn (name,_,_,_,body,_,_) => (name,body)) procs'; val name_pars_specs = map (fn (name,inpars,outpars,_,_,specs,sk) => (name,(inpars,outpars,sk),specs)) procs'; val names = map #1 procs'; val sk = #7 (hd procs'); val thy = thy |> Context.theory_map (set_default_state_kind sk); val (all_callss,cliques,is_recursive,has_body) = let val context = Context.Theory thy |> fold (add_parameter_info Morphism.identity (unsuffix proc_deco)) name_pars |> Config.put_generic StateSpace.silent true fun read_body (_, body) = Syntax.read_term (Context.proof_of context) body; val bodies = map read_body name_body; fun dcall t = (case try dest_call t of - SOME (_,p,_,_,m,_) => SOME (proc_name m p) + SOME (_,p,_,_,m,_,_) => SOME (proc_name (Context.proof_of context) m p) | _ => NONE); fun in_names x = if member (op =) names x then SOME x else NONE; fun add_edges n = fold (fn x => Graph.add_edge (n, x)); val all_callss = map (map snd o max_subterms_dest dcall) bodies; val callss = map (map_filter in_names) all_callss; val graph = fold (fn n => Graph.new_node (n, ())) names Graph.empty; val graph' = fold2 add_edges names callss graph; fun idx x = find_index (fn y => x=y) names; fun name_ord (a,b) = int_ord (idx a, idx b); val cliques = Graph.strong_conn graph'; val cliques' = map (sort name_ord) cliques; val my_calls = these o AList.lookup (op =) (names ~~ map (distinct (op =)) callss); val my_body = AList.lookup (op =) (names ~~ bodies); fun is_recursive n = exists (fn [_] => false | ns => member (op =) ns n) (Graph.strong_conn graph'); fun has_body n = (case my_body n of SOME (Const (c,_)) => c <> NoBodyN | _ => true) fun clique_less c1 c2 = null (inter (op =) (distinct (op =) (maps my_calls c1)) c2); val cliques'' = topsort clique_less cliques'; in (all_callss,cliques'',is_recursive,has_body) end; (* cliques may only depend on ones to the left, so it is safe to * add the locales from the left to the right. *) fun my_clique n = Library.find_first (fn c => member (op =) c n) cliques; fun lname sfx clique = suffix sfx (clique_name clique); fun cname n = clique_name (the (my_clique n)); fun parameter_info_decl phi = fold (add_parameter_info phi cname) name_pars; fun get_loc sfx clique n = if member (op =) clique n then NONE else SOME (resuffix proc_deco sfx n); fun parent_locales thy sfx clique = let val calls = distinct (op =) (flat (map_filter (AList.lookup (op =) (names ~~ all_callss)) clique)); in map (intern_locale thy) (distinct (op =) (map_filter (get_loc sfx clique) calls)) end; val names_all_callss = names ~~ map (distinct (op =)) all_callss; val get_calls = the o AList.lookup (op =) names_all_callss; fun clique_vars clique = let fun add name (ins,outs,locs) = let val (nins,nouts,nlocs) = the (AList.lookup (op =) name_vars name) in (ins@nins,outs@nouts,locs@nlocs) end; val (is,os,ls) = fold add clique ([],[],[]); in (lname "" clique, (distinct (op =) is, distinct (op =) os, distinct (op =) ls)) end; fun add_signature_locale (cname, name) thy = let val name' = unsuffix proc_deco name; val fixes = [Element.Fixes [(Binding.name name, SOME proc_nameT, NoSyn)]]; - (* FIXME: may use HOLogic.typeT as soon as locale type-inference works properly *) + (* fixme: may use HOLogic.typeT as soon as locale type-inference works properly *) val pE = mk_loc_exp [intern_locale thy (suffix parametersN cname)]; val sN = suffix signatureN name'; in thy |> add_locale sN pE fixes |> Proof_Context.theory_of |> (fn thy => add_declaration (intern_locale thy sN) parameter_info_decl thy) end; fun mk_bdy_def read_term name = let val name' = unsuffix proc_deco name; val bdy = read_term (the (AList.lookup (op =) name_body name)); val bdy_defN = suffix body_def_sfx name'; val b = Binding.name bdy_defN; in ((b, NoSyn), ((Thm.def_binding b, []), bdy)) end; fun add_body_locale (name, _) thy = let val name' = unsuffix proc_deco name; val callees = filter_out (fn n => n = name) (get_calls name) val fixes = [Element.Fixes [(Binding.name name, SOME proc_nameT, NoSyn)]]; - (* FIXME: may use HOLogic.typeT as soon as locale type-inference works properly *) + (* fixme: may use HOLogic.typeT as soon as locale type-inference works properly *) val pE = mk_loc_exp (map (intern_locale thy) ([lname variablesN (the (my_clique name))]@ the_list locname@ map (resuffix proc_deco signatureN) callees)); fun def lthy = let val read = Syntax.read_term (Context.proof_map (add_active_procs Morphism.identity (the (my_clique name))) (Local_Theory.target_of lthy)) in mk_bdy_def read name end; fun add_decl_and_def lname ctxt = ctxt |> Proof_Context.theory_of |> Named_Target.init [] lname |> Local_Theory.declaration {syntax = false, pervasive = false, pos = \<^here>} parameter_info_decl |> (fn lthy => if has_body name then snd (Local_Theory.define (def lthy) lthy) else lthy) |> Local_Theory.exit |> Proof_Context.theory_of; in thy |> add_locale' (suffix bodyN name') pE fixes |-> add_decl_and_def end; fun mk_def_eq thy read_term name = if has_body name then let - (* FIXME: All the read_term stuff is just because type-inference/abbrevs for + (* fixme: All the read_term stuff is just because type-inference/abbrevs for * new locale elements does not work right now; * We read the term to expand the abbreviations, then we print it again * (without folding the abbreviation) and reread as string *) val name' = unsuffix proc_deco name; val bdy_defN = suffix body_def_sfx name'; val rhs = read_term ("Some " ^ bdy_defN); val nt = read_term name; val Free (gamma,_) = read_term programN; val eq = HOLogic.Trueprop$ HOLogic.mk_eq (Free (gamma,fastype_of nt --> fastype_of rhs)$nt,rhs) val consts = Sign.consts_of thy; val eqs = - YXML.string_of_body (Term_XML.Encode.term consts (Consts.dummy_types consts eq)); + YXML.string_of_body (Term_XML.Encode.term consts (Consts.dummy_types consts eq)); val assms = Element.Assumes [((Binding.name (suffix bodyP name'), []),[(eqs,[])])] in [assms] end else []; fun add_impl_locales clique thy = let val cliqN = lname cliqueN clique; val cnamesN = lname clique_namesN clique; val multiple_procs = length clique > 1; val add_distinct_procs_namespace = if multiple_procs then StateSpace.namespace_definition cnamesN proc_nameT ([],[]) [] clique else I; val bodyLs = map (suffix bodyN o unsuffix proc_deco) clique; fun pE thy = mk_loc_exp (map (intern_locale thy) (hoare_ctxtL::bodyLs) @ (parent_locales thy implementationN clique) @ (if multiple_procs then [intern_locale thy cnamesN] else [])); fun read_term thy = apply_in_context thy (pE thy) Syntax.read_term; fun elems thy = maps (mk_def_eq thy (read_term thy)) clique; fun add_recursive_info phi name = if is_recursive name then (add_recursive phi name) else I; fun proc_declaration phi = add_active_procs phi clique; fun recursive_declaration phi context = context |> fold (add_recursive_info phi) clique; fun add_impl_locale name thy = let val implN = suffix implementationN (unsuffix proc_deco name); val parentN = intern_locale thy cliqN val parent = mk_loc_exp [parentN]; in thy |> add_locale_cmd implN parent [] |> Proof_Context.theory_of |> (fn thy => Interpretation.global_sublocale parentN (mk_loc_exp [intern_locale thy implN]) [] thy) |> Proof.global_terminal_proof ((Method.Basic (fn ctxt => Method.SIMPLE_METHOD (Locale.intro_locales_tac {strict = true, eager = false} ctxt [])), Position.no_range), NONE) |> Proof_Context.theory_of end; in thy |> add_distinct_procs_namespace |> (fn thy => add_locale_cmd cliqN (pE thy) (elems thy) thy) |> Proof_Context.theory_of |> fold add_impl_locale clique |> (fn thy => add_declaration (intern_locale thy cliqN) proc_declaration thy) |> (fn thy => add_declaration (intern_locale thy cliqN) recursive_declaration thy) end; fun add_spec_locales (name, _, specs) thy = let val name' = unsuffix proc_deco name; val ps = (suffix signatureN name' :: the_list locname); val ps' = hoare_ctxtL :: ps ; val pE = mk_loc_exp (map (intern_locale thy) ps) val pE' = mk_loc_exp (map (intern_locale thy) ps') fun read thy = apply_in_context thy (mk_loc_exp [intern_locale thy (suffix cliqueN (cname name))]) (Syntax.read_prop); fun proc_declaration phi = (*parameter_info_decl phi o already in signature *) add_active_procs phi (the (my_clique name)); fun add_locale'' (thm_name,spec) thy = let val spec' = read thy spec; val elem = Element.Assumes [((Binding.name thm_name, []), [(spec', [])])]; in thy |> add_locale thm_name pE' [elem] |> Proof_Context.theory_of |> (fn thy => add_declaration (intern_locale thy thm_name) proc_declaration thy) |> error_to_warning ("abbreviation: '" ^ thm_name ^ "' not added") (add_abbrev (intern_locale thy (suffix cliqueN (cname name))) Syntax.mode_input thm_name spec) end; in thy |> fold add_locale'' specs end; in thy |> fold (add_variable_statespaces o clique_vars) cliques |> fold (fn c => fold (fn n => add_signature_locale (lname "" c, n)) c) cliques |> fold add_body_locale name_pars |> fold add_impl_locales cliques |> fold add_spec_locales name_pars_specs end; (********************* theory extender interface ********************************) (** package setup **) (* outer syntax *) val var_declP = Parse.name -- (@{keyword "::"} |-- Parse.!!! Parse.embedded); val var_declP' = Parse.name >> (fn n => (n,"")); val localsP = Scan.repeat var_declP; val argP = var_declP; val argP' = var_declP'; val not_eqP = Scan.ahead (Scan.unless @{keyword "="} (Scan.one (K true))) val proc_decl_statespace = (Parse.short_ident --| @{keyword "("}) -- ((Parse.list argP) -- (Scan.optional (@{keyword "|"} |-- Parse.list argP) []) --| @{keyword ")"}) --| not_eqP val proc_decl_record = (Parse.short_ident --| @{keyword "("}) -- ((Parse.list argP') -- (Scan.optional (@{keyword "|"} |-- Parse.list argP') []) --| @{keyword ")"}) --| Scan.option @{keyword "="} val proc_decl = proc_decl_statespace >> pair Function || proc_decl_record >> pair Record; val loc_decl = Scan.optional (@{keyword "where"} |-- localsP --| @{keyword "in"}) [] val proc_body = Parse.embedded (*>> BodyTerm*) fun proc_specs x = (Parse.!!! (Scan.repeat (Parse_Spec.thm_name ":" -- Parse.embedded)) >> map (fn ((thm_name, _), prop) => (Binding.name_of thm_name, prop))) x val par_loc = Scan.option (@{keyword "("} |-- @{keyword "imports"} |-- Parse.name --| @{keyword ")"}); val _ = Outer_Syntax.command @{command_keyword procedures} "define procedures" (par_loc -- (Parse.and_list1 (proc_decl -- loc_decl -- proc_body -- proc_specs)) >> (fn (loc,decls) => let val decls' = map (fn ((((state_kind,(name,(ins,outs))),ls),bdy),specs) => (name,ins,outs,ls,bdy,specs,state_kind)) decls in Toplevel.theory (procedures_definition loc decls') end)); val _ = Outer_Syntax.command @{command_keyword hoarestate} "define state space for hoare logic" (StateSpace.statespace_decl >> (fn ((args,name),(parents,comps)) => Toplevel.theory (StateSpace.define_statespace args name parents (map (apfst (suffix deco)) comps)))); (*************************** Auxiliary Functions for integration of ********************) (*************************** automatic program analysers ********************) fun dest_conjs t = (case HOLogic.dest_conj t of [t1,t2] => dest_conjs t1 @ dest_conjs t2 | ts => ts); fun split_guard (Const (@{const_name Collect},CT)$(Abs (s,T,t))) = let fun mkCollect t = Const (@{const_name Collect},CT)$(Abs (s,T,t)); in map mkCollect (dest_conjs t) end | split_guard t = [t]; fun split_guards gs = let fun norm c f g = map (fn g => c$f$g) (split_guard g); fun norm_guard ((c as Const (@{const_name Pair},_))$f$g) = norm c f g | norm_guard ((c as Const (@{const_name Language.guaranteeStripPair},_))$f$g) = norm c f g | norm_guard t = [t]; in maps norm_guard (HOLogic.dest_list gs) end fun fold_com f t = let (* traverse does not descend into abstractions, like in DynCom, call, etc. *) fun traverse cnt (c as Const (@{const_name Language.com.Skip},_)) = (cnt,f cnt c [] []) | traverse cnt ((c as Const (@{const_name Language.com.Basic},_))$g) = (cnt, f cnt c [g] []) | traverse cnt ((c as Const (@{const_name Language.com.Spec},_))$r) = (cnt, f cnt c [r] []) | traverse cnt ((c as Const (@{const_name Language.com.Seq},_))$c1$c2) = let val (cnt1,v1) = traverse cnt c1; val (cnt2,v2) = traverse cnt1 c2; in (cnt2, f cnt c [] [v1,v2]) end | traverse cnt ((c as Const (@{const_name Language.com.Cond},_))$b$c1$c2) = let val (cnt1,v1) = traverse cnt c1; val (cnt2,v2) = traverse cnt1 c2; in (cnt2, f cnt c [b] [v1,v2]) end | traverse cnt ((c as Const (@{const_name Language.com.While},_))$b$c1) = let val (cnt1,v1) = traverse cnt c1 in (cnt1, f cnt c [b] [v1]) end | traverse cnt ((c as Const (@{const_name Language.com.Call},_))$p) = (cnt, f cnt c [p] []) | traverse cnt ((c as Const (@{const_name Language.com.DynCom},_))$c1) = (cnt, f cnt c [c1] []) | traverse cnt ((c as Const (@{const_name Language.com.Guard},_))$flt$g$c1) = let val (cnt1,v1) = traverse (cnt + length (split_guard g)) c1 in (cnt1, f cnt c [flt,g] [v1]) end | traverse cnt (c as Const (@{const_name Language.com.Throw},_)) = (cnt,f cnt c [] []) | traverse cnt ((c as Const (@{const_name Language.com.Catch},_))$c1$c2) = let val (cnt1,v1) = traverse cnt c1; val (cnt2,v2) = traverse cnt1 c2; in (cnt2, f cnt c [] [v1,v2]) end | traverse cnt ((c as Const (@{const_name Language.guards},_))$gs$c1) = let val (cnt1,v1) = traverse (cnt + length (split_guards gs)) c1; in (cnt1, f cnt c [gs] [v1]) end | traverse cnt ((c as Const (@{const_name Language.block},_))$init$c1$return$c2) = let val (cnt1,v1) = traverse cnt c1 in (cnt1, f cnt c [init,return,c2] [v1]) end | traverse cnt ((c as Const (@{const_name Language.call},_))$init$p$return$c1) = (cnt, f cnt c [init,p,return,c1] []) | traverse cnt ((c as Const (@{const_name Language.whileAnno},_))$b$I$V$c1) = let val (cnt1,v1) = traverse cnt c1 in (cnt1, f cnt c [b,I,V] [v1]) end | traverse cnt ((c as Const (@{const_name Language.whileAnnoG},_))$gs$b$I$V$c1) = let val (cnt1,v1) = traverse (cnt + length (split_guards gs)) c1 in (cnt1, f cnt c [gs,b,I,V] [v1]) end | traverse _ t = raise TERM ("fold_com: unknown command",[t]); in snd (traverse 0 t) end; (*************************** Tactics ****************************************************) (*** Aux. tactics ***) fun cond_rename_bvars cond name thm = let fun rename (tm as (Abs (x, T, t))) = if cond tm then Abs (name x, T, rename t) else Abs (x, T, rename t) | rename (t $ u) = rename t $ rename u | rename a = a; in Thm.renamed_prop (rename (Thm.prop_of thm)) thm end; val rename_bvars = cond_rename_bvars (K true); -fun trace_tac ctxt str st = (if Config.get ctxt hoare_trace then tracing str else (); all_tac st); +fun trace_msg ctxt str = if Config.get ctxt hoare_trace > 0 then tracing str else () +fun trace_tac ctxt str st = (trace_msg ctxt str; all_tac st); + +fun trace_subgoal_tac ctxt s i = + SUBGOAL (fn (prem, _) => trace_tac ctxt (s ^ (Syntax.string_of_term ctxt prem))) i + fun error_tac str st = (error str;no_tac st); fun rename_goal ctxt name = EVERY' [K (trace_tac ctxt "rename_goal -- START"), SELECT_GOAL (PRIMITIVE (rename_bvars name)), K (trace_tac ctxt "rename_goal -- STOP")]; (* splits applications of tupled arguments to a schematic Variables, e.g. * ALL a b. ?P (a,b) --> ?Q (a,b) gets * ALL a b. ?P a b --> ?Q a b * only tuples nested to the right are splitted. *) fun split_pair_apps ctxt thm = let val t = Thm.prop_of thm; fun mk_subst subst (Abs (x,T,t)) = mk_subst subst t | mk_subst subst (t as (t1$t2)) = (case strip_comb t of (var as Var (v,vT),args) => (if not (AList.defined (op =) subst var) then let val len = length args; val (argTs,bdyT) = strip_type vT; val (z, _) = Name.variant "z" (fold Term.declare_term_frees args Name.context); val frees = map (apfst (fn i => z^string_of_int i)) (0 upto (len - 1) ~~ argTs); fun splitT (Type (@{type_name Product_Type.prod}, [T1, T2])) = T1::splitT T2 | splitT T = [T]; fun pair_depth (Const (@{const_name Pair},aT)$t1$t2) = 1 + (pair_depth t2) | pair_depth _ = 0; fun mk_sel max free i = let val snds = funpow i HOLogic.mk_snd (Free free) in if i=max then snds else HOLogic.mk_fst snds end; fun split (free,arg) = let val depth = (pair_depth arg); in if depth = 0 then [Free free] else map (mk_sel depth free) (0 upto depth) end; val args' = maps split (frees ~~ args); val argTs' = maps splitT argTs; val inst = fold_rev absfree frees (list_comb (Var (v,argTs' ---> bdyT), args')) in subst@[(var,inst)] end else subst) | _ => mk_subst (mk_subst subst t1) t2) | mk_subst subst t = subst; val subst = map (fn (v,t) => (dest_Var v, Thm.cterm_of ctxt t)) (mk_subst [] t); in full_simplify (put_simpset HOL_basic_ss ctxt addsimps [@{thm fst_conv}, @{thm snd_conv}]) (Drule.instantiate_normalize (TVars.empty, Vars.make subst) thm) end; (* Generates split theorems, for !!,!,? quantifiers and for UN, e.g. * ALL x. P x = ALL a b. P a b *) fun mk_split_thms ctxt (vars as _::_) = let val thy = Proof_Context.theory_of ctxt; val names = map fst vars; val types = map snd vars; val free_vars = map Free vars; val pT = foldr1 HOLogic.mk_prodT types; val x = (singleton (Name.variant_list names) "x", pT); val xp = foldr1 HOLogic.mk_prod free_vars; val tfree_names = fold Term.add_tfree_names free_vars []; val zeta = TFree (singleton (Name.variant_list tfree_names) "z", Sign.defaultS thy); val split_meta_prop = let val P = Free (singleton (Name.variant_list names) "P", pT --> Term.propT) in Logic.mk_equals (Logic.all (Free x) (P $ Free x), fold_rev Logic.all free_vars (P $ xp)) end; val P = Free (singleton (Name.variant_list names) "P", pT --> HOLogic.boolT); val split_object_prop = let fun ALL vs t = Library.foldr (fn ((v,T),t) => HOLogic.mk_all (v,T,t)) (vs,t) in (ALL [x] (P $ Free x)) === (ALL vars (P $ xp)) end; val split_ex_prop = let fun EX vs t = Library.foldr (fn ((v,T),t) => HOLogic.mk_exists (v,T,t)) (vs,t) in (EX [x] (P $ Free x)) === (EX vars (P $ xp)) end; val split_UN_prop = let val P = Free (singleton (Name.variant_list names) "P", pT --> HOLogic.mk_setT zeta); fun UN vs t = Library.foldr mk_UN (vs, t) in (UN [x] (P $ Free x)) === (UN vars (P $ xp)) end; fun prove_simp simps prop = - let val ([prop'], _) = Variable.importT_terms [prop] ctxt (* FIXME continue context!? *) + let val ([prop'], _) = Variable.importT_terms [prop] ctxt (* fixme continue context!? *) in Goal.prove_global thy [] [] prop' (fn {context = goal_ctxt, ...} => ALLGOALS (simp_tac (put_simpset HOL_basic_ss goal_ctxt addsimps simps))) end; val split_meta = prove_simp [@{thm split_paired_all}] split_meta_prop; val split_object = prove_simp [@{thm split_paired_All}] split_object_prop; val split_ex = prove_simp [@{thm split_paired_Ex}] split_ex_prop; val split_UN = prove_simp [@{thm Hoare.split_paired_UN}] split_UN_prop; in [split_meta,split_object,split_ex,split_UN] end | mk_split_thms _ _ = raise Match; fun rename_aux_var name rule = let fun is_aux_var (Abs ("Z",TVar(_,_),_)) = true | is_aux_var _ = false; in cond_rename_bvars is_aux_var (K name) rule end; (* adapts single auxiliary variable in a rule to potentialy multiple auxiliary * variables in actual specification, e.g. if vars are a b, * split_app=false: ALL Z. ?P Z gets to ALL a b. ?P (a,b) * split_app=true: ALL Z. ?P Z gets to ALL a b. ?P a b * If only one auxiliary variable is given, the variables are just renamed, * If no auxiliary is given, unit is inserted for Z: * ALL Z. ?P Z gets P () *) fun adapt_aux_var ctxt split_app (vars as (_::_::_)) tvar_rules = let val thy = Proof_Context.theory_of ctxt; val max_idx = fold Integer.max (map (Thm.maxidx_of o snd) tvar_rules) 0; val types = map (fn i => TVar (("z",i),Sign.defaultS thy)) (max_idx + 1 upto (max_idx + length vars)); fun tvar n = (n, Sign.defaultS thy); val pT = Thm.ctyp_of ctxt (foldr1 HOLogic.mk_prodT types); val rules' = map (fn (z,r) => (Drule.instantiate_normalize (TVars.make [(tvar z,pT)], Vars.empty) r)) tvar_rules; val splits = mk_split_thms ctxt (vars ~~ types); val rules'' = map (full_simplify (put_simpset HOL_basic_ss ctxt addsimps splits)) rules'; in if split_app then (map (split_pair_apps ctxt) rules'') else rules'' end | adapt_aux_var _ _ ([name]) tvar_rules = map (rename_aux_var name o snd) tvar_rules | adapt_aux_var ctxt _ ([]) tvar_rules = let val thy = Proof_Context.theory_of ctxt; fun tvar n = (n, Sign.defaultS thy); val uT = Thm.ctyp_of ctxt HOLogic.unitT; val rules' = map (fn (z,r) => (Drule.instantiate_normalize (TVars.make [(tvar z,uT)], Vars.empty) r)) tvar_rules; val splits = [@{thm Hoare.unit_meta},@{thm Hoare.unit_object},@{thm Hoare.unit_ex},@{thm Hoare.unit_UN}]; val rules'' = map (full_simplify (put_simpset HOL_basic_ss ctxt addsimps splits)) rules'; in rules'' end (* Generates a rule for recursion for n procedures out of general recursion rule *) fun gen_call_rec_rule ctxt specs_name n rule = let val thy = Proof_Context.theory_of ctxt; val maxidx = Thm.maxidx_of rule; val vars = Term.add_vars (Thm.prop_of rule) []; fun get_type n = the (AList.lookup (op =) vars (n, 0)); val (Type (_, [Type (_, [assT, Type (_, [pT,_])])])) = get_type specs_name; val zT = TVar (("z",maxidx+1),Sign.defaultS thy) fun mk_var i n T = Var ((n ^ string_of_int i,0),T); val quadT = HOLogic.mk_prodT (assT, HOLogic.mk_prodT (pT, HOLogic.mk_prodT (assT,assT))); val quadT_set = HOLogic.mk_setT quadT; fun mk_spec i = let val quadruple = HOLogic.mk_ptuple (HOLogic.flat_tupleT_paths quadT) quadT [mk_var i "P" (zT --> assT)$Bound 0, mk_var i "p" pT, mk_var i "Q" (zT --> assT)$Bound 0, mk_var i "A" (zT --> assT)$Bound 0]; val single = HOLogic.mk_set quadT [quadruple]; in mk_UN' zT quadT_set (Abs ("Z", zT, single)) end; val Specs = foldr1 (mk_Un quadT_set) (map mk_spec (1 upto n)); val rule' = Thm.instantiate' [] [SOME (Thm.cterm_of ctxt Specs)] rule |> full_simplify (put_simpset (simpset_of @{theory_context Main}) ctxt addsimps [@{thm Hoare.conjE_simp},@{thm Hoare.in_Specs_simp},@{thm Hoare.in_set_Un_simp},@{thm split_all_conj}, @{thm image_Un},@{thm image_Un_single_simp}] ) |> rename_bvars (fn s => if member (op =) ["s","\"] s then s else "Z") in rule' end; fun gen_proc_rec ctxt mode n = gen_call_rec_rule ctxt "Specs" n (ProcRecSpecs mode); (*** verification condition generator ***) (* simplifications on "Collect" sets, like {s. P s} Int {s. Q s} = {s. P s & Q s} *) fun merge_assertion_simp_tac ctxt thms = simp_tac (put_simpset HOL_basic_ss ctxt addsimps ([@{thm Hoare.CollectInt_iff},@{thm HOL.conj_assoc},@{thm Hoare.Compl_Collect},singleton_conv_sym, @{thm Set.Int_empty_right},@{thm Set.Int_empty_left},@{thm Un_empty_right},@{thm Un_empty_left}]@thms)) ; (* The following probably shouldn't live here, but refactoring so that Hoare could depend on recursive_records does not look feasible. The upshot is that there's a duplicate foldcong_ss set here. *) structure FoldCongData = Theory_Data ( type T = simpset; val empty = HOL_basic_ss; val merge = merge_ss; ) val get_foldcong_ss = FoldCongData.get fun add_foldcongs congs thy = FoldCongData.map (fn ss => Proof_Context.init_global thy |> put_simpset ss |> fold Simplifier.add_cong congs |> simpset_of) thy fun add_foldcongsimps simps thy = FoldCongData.map (fn ss => Proof_Context.init_global thy |> put_simpset ss |> (fn ctxt => ctxt addsimps simps) |> simpset_of) thy (* propagates state into "Collect" sets and simplifies selections updates like: * s:{s. P s} = P s *) fun in_assertion_simp_tac ctxt state_kind thms i = let - val vcg_simps = #vcg_simps (get_data ctxt); - val fold_simps = get_foldcong_ss (Proof_Context.theory_of ctxt) + val thy = Proof_Context.theory_of ctxt + val vcg_simps = map (Thm.transfer thy) (#vcg_simps (get_data ctxt)); + val fold_simps = get_foldcong_ss thy + val state_simps = Named_Theorems.get ctxt @{named_theorems "state_simp"} in EVERY [simp_tac (put_simpset HOL_basic_ss ctxt addsimps ([mem_Collect_eq,@{thm Set.Un_iff},@{thm Set.Int_iff}, @{thm Set.empty_subsetI}, @{thm Set.empty_iff}, UNIV_I, - @{thm Hoare.Collect_False}]@thms@K_convs@vcg_simps) - addsimprocs (state_simprocs state_kind) + @{thm Hoare.Collect_False}]@state_simps@thms@K_convs@vcg_simps) + addsimprocs (state_simprocs ctxt state_kind) |> fold Simplifier.add_cong K_congs) i THEN_MAYBE - (simp_tac (put_simpset fold_simps ctxt addsimprocs [state_upd_simproc state_kind]) i) + (simp_tac (put_simpset fold_simps ctxt addsimprocs (state_upd_simprocs ctxt state_kind)) i) ] end; fun assertion_simp_tac ctxt state_kind thms i = merge_assertion_simp_tac ctxt [] i THEN_MAYBE in_assertion_simp_tac ctxt state_kind thms i (* simplify equality test on strings (and datatype-constructors) and propagate result*) fun string_eq_simp_tac ctxt = simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms list.inject list.distinct char.inject cong_exp_iff_simps simp_thms}); fun assertion_string_eq_simp_tac ctxt state_kind thms i = assertion_simp_tac ctxt state_kind thms i THEN_MAYBE string_eq_simp_tac ctxt i; fun before_set2pred_simp_tac ctxt = (simp_tac (put_simpset HOL_basic_ss ctxt addsimps [singleton_conv_sym, @{thm Hoare.CollectInt_iff}, @{thm Hoare.Compl_Collect}])); (*****************************************************************************) (** set2pred transforms sets inclusion into predicates implication, **) (** maintaining the original variable names. **) (** Ex. "{x. x=0} <= {x. x <= 1}" -set2pred-> "x=0 --> x <= 1" **) (** Subgoals containing intersections (A Int B) or complement sets (-A) **) (** are first simplified by "before_set2pred_simp_tac", that returns only **) (** subgoals of the form "{x. P x} <= {x. Q x}", which are easily **) (** transformed. **) (** This transformation may solve very easy subgoals due to a ligth **) (** simplification done by full_simp_tac **) (*****************************************************************************) +val Collect_subset_to_pred = +@{lemma \(\x. A x \ P x) + \ {x. A x} \ {x. P x}\ + by (rule subsetI, rule CollectI, drule CollectD, assumption)} + + fun set2pred_tac ctxt i thm = ((before_set2pred_simp_tac ctxt i) THEN_MAYBE (EVERY [trace_tac ctxt "set2pred", - resolve_tac ctxt [subsetI] i, - resolve_tac ctxt [CollectI] i, - dresolve_tac ctxt [CollectD] i, + resolve_tac ctxt [Collect_subset_to_pred] i, full_simp_tac (put_simpset HOL_basic_ss ctxt) i ])) thm (*****************************************************************************) (** BasicSimpTac is called to simplify all verification conditions. It does **) (** a light simplification by applying "mem_Collect_eq" **) (** then it tries to solve subgoals of the form "A <= A" and then if **) (** set2pred is true it **) (** transforms any other into predicates, applying then **) (** the tactic chosen by the user, which may solve the subgoal completely **) (** (MaxSimpTac). **) (*****************************************************************************) fun MaxSimpTac ctxt tac i = TRY (FIRST[resolve_tac ctxt [subset_refl] i, - set2pred_tac ctxt i THEN_MAYBE tac i, + (set2pred_tac ctxt i THEN_MAYBE tac i) ORELSE tac i, trace_tac ctxt "final_tac failed" ]); fun BasicSimpTac ctxt state_kind set2pred thms tac i = EVERY [(trace_tac ctxt "BasicSimpTac -- START --"), assertion_simp_tac ctxt state_kind thms i THEN_MAYBE (if set2pred then MaxSimpTac ctxt tac i else TRY (resolve_tac ctxt [subset_refl] i)), (trace_tac ctxt "BasicSimpTac -- STOP --")]; (* EVERY [(trace_tac ctxt "BasicSimpTac -- START --"), simp_tac (HOL_basic_ss addsimps [mem_Collect_eq,@{thm Hoare.CollectInt_iff}, @{thm Set.empty_subsetI}, @{thm Set.empty_iff}, UNIV_I] addsimprocs [state_simproc sk]) i THEN_MAYBE simp_tac (HOL_basic_ss addsimprocs [state_upd_simproc sk]) i THEN_MAYBE (if set2pred then MaxSimpTac ctxt tac i else TRY (rtac subset_refl i)), (trace_tac ctxt "BasicSimpTac -- STOP --")]; *) (* fun simp_state_eq_tac Record state_space = full_simp_tac (HOL_basic_ss addsimprocs (state_simprocs Record)) THEN_MAYBE' full_simp_tac (HOL_basic_ss addsimprocs [state_upd_simproc Record]) THEN_MAYBE' (state_split_simp_tac [] state_space) | simp_state_eq_tac StateFun state_space = *) fun post_conforms_tac ctxt state_kind i = EVERY [REPEAT1 (resolve_tac ctxt [allI,impI] i), ((fn i => TRY (resolve_tac ctxt [conjI] i)) THEN_ALL_NEW (fn i => (REPEAT (resolve_tac ctxt [allI,impI] i)) THEN (full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [mem_Collect_eq,@{thm Set.singleton_iff},@{thm Set.empty_iff},UNIV_I] - addsimprocs (state_simprocs state_kind)) i))) i]; + addsimprocs (state_simprocs ctxt state_kind)) i))) i]; fun dest_hoare_raw (Const(@{const_name HoarePartialDef.hoarep},_)$G$T$F$P$C$Q$A) = (P,C,Q,A,Partial,G,T,F) | dest_hoare_raw (Const(@{const_name HoareTotalDef.hoaret},_)$G$T$F$P$C$Q$A) = (P,C,Q,A,Total,G,T,F) | dest_hoare_raw t = raise TERM ("Hoare.dest_hoare_raw: unexpected term",[t]) fun mk_hoare_abs Ts (P,C,Q,A,mode,G,T,F) = let val hoareT = map (curry fastype_of1 Ts) [G,T,F,P,C,Q,A] ---> HOLogic.boolT; val hoareC = (case mode of Partial => Const (@{const_name HoarePartialDef.hoarep},hoareT) | Total => Const (@{const_name HoareTotalDef.hoaret},hoareT)); in hoareC$G$T$F$P$C$Q$A end; val is_hoare = can dest_hoare_raw fun dest_hoare t = let val triple = (strip_qnt_body @{const_name "All"} o - HOLogic.dest_Trueprop o strip_qnt_body @{const_name Pure.all}) t; + HOLogic.dest_Trueprop o Logic.strip_assums_concl) t; in dest_hoare_raw triple end; fun get_aux_tvar rule = let fun aux_hoare (Abs ("Z",TVar (z,_),t)) = if is_hoare (strip_qnt_body @{const_name All} t) then SOME z else NONE | aux_hoare _ = NONE; in (case first_subterm_dest (aux_hoare) (Thm.prop_of rule) of SOME (_,z) => (z,rule) | NONE => raise TERM ("get_aux_tvar: No auxiliary variable of hoare-rule found", [Thm.prop_of rule])) end; fun strip_vars t = let val bdy = (HOLogic.dest_Trueprop o Logic.strip_assums_concl) t; in strip_qnt_vars @{const_name Pure.all} t @ strip_qnt_vars @{const_name All} bdy end; local (* ex_simps are necessary in case of multiple logical variables. The state will usually be the first variable. EX s a b. s=s' ... . We have to transport EX s to s=s' to perform the substitution *) val conseq1_ss_base = simpset_of (put_simpset HOL_basic_ss @{context} addsimps ([mem_Collect_eq,@{thm Set.singleton_iff},@{thm Set.Int_iff}, @{thm Set.empty_iff},UNIV_I, @{thm HOL.conj_assoc}, @{thm disj_assoc}] @ @{thms Hoare.all_imp_eq_triv} @K_convs @ @{thms simp_thms} @ @{thms ex_simps} @ @{thms all_simps}) delsimps [@{thm Hoare.all_imp_to_ex}] |> fold Simplifier.add_cong K_congs) -val conseq1_ss_record = - simpset_of (put_simpset conseq1_ss_base @{context} addsimprocs (state_simprocs Record)); -val conseq1_ss_fun = - simpset_of (put_simpset conseq1_ss_base @{context} addsimprocs (state_simprocs Function)); -fun conseq1_ss Record = conseq1_ss_record - | conseq1_ss Function = conseq1_ss_fun; val conseq2_ss_base = simpset_of (put_simpset HOL_basic_ss @{context} addsimps (@{thms Hoare.all_imp_eq_triv} @ @{thms simp_thms} @ @{thms ex_simps} @ @{thms all_simps}) delsimps [@{thm Hoare.all_imp_to_ex}] |> Simplifier.add_cong @{thm imp_cong}); - -val conseq2_ss_record = - simpset_of (put_simpset conseq2_ss_base @{context} - addsimprocs [state_upd_simproc Record, state_ex_sel_eq_simproc Record]); -val conseq2_ss_fun = - simpset_of (put_simpset conseq2_ss_base @{context} - addsimprocs [state_upd_simproc Function, state_ex_sel_eq_simproc Function]); -fun conseq2_ss Record = conseq2_ss_record - | conseq2_ss Function = conseq2_ss_fun; - in fun raw_conseq_simp_tac ctxt state_kind thms i = let val ctxt' = Config.put simp_depth_limit 0 ctxt; in - simp_tac (put_simpset (conseq1_ss state_kind) ctxt' addsimps thms) i + simp_tac (put_simpset conseq1_ss_base ctxt' + addsimprocs (state_simprocs ctxt' state_kind) + addsimps thms) i THEN_MAYBE - simp_tac (put_simpset (conseq2_ss state_kind) ctxt') i + simp_tac (put_simpset conseq2_ss_base ctxt' + addsimprocs (state_upd_simprocs ctxt' state_kind @ state_ex_sel_eq_simprocs ctxt' state_kind)) i end end val conseq_simp_tac = raw_conseq_simp_tac; (* Generates the hoare-quadruples that can be derived out of the hoare-context T *) fun gen_context_thms ctxt mode params G T F = let val Type (_,[comT]) = range_type (fastype_of G); fun destQuadruple (Const (@{const_name Set.insert},_) $ PpQA $ Const (@{const_name Orderings.bot}, _)) = PpQA | destQuadruple t = raise Match; fun mkCallQuadruple (Const (@{const_name Pair}, _) $ P $ (Const (@{const_name Pair}, _) $ p $ (Const(@{const_name Pair}, _) $ Q $ A))) = let val Call_p = Const (@{const_name Language.com.Call}, fastype_of p --> comT) $ p; in (P, Call_p, Q, A) end; fun mkHoare mode G T F (vars,PpQA) = let val hoare = (case mode of Partial => @{const_name HoarePartialDef.hoarep} | Total => @{const_name HoareTotalDef.hoaret}); - (* FIXME: Use future Proof_Context.rename_vars or make closed term and remove by hand *) + (* fixme: Use future Proof_Context.rename_vars or make closed term and remove by hand *) (* fun free_params ps t = foldr (fn ((x,xT),t) => snd (variant_abs (x,xT,t))) (ps,t); val PpQA' = mkCallQuadruple (strip_qnt_body @{const_name Pure.all} (free_params params (Term.list_all (vars,PpQA)))); *) val params' = (Variable.variant_frees ctxt [PpQA] params); val bnds = map Bound (0 upto (length vars - 1)); fun free_params_vars t = subst_bounds (bnds @ rev (map Free params' ), t) fun free_params t = subst_bounds (rev (map Free params' ), t) val (P',p',Q',A') = mkCallQuadruple (free_params_vars PpQA); val T' = free_params T; val G' = free_params G; val F' = free_params F; val bdy = mk_hoare_abs [] (P',p',Q',A',mode,G',T',F'); in (HOLogic.mk_Trueprop (HOLogic.list_all (vars,bdy)), map fst params') end; fun hoare_context_specs mode G T F = let fun mk t = try (mkHoare mode G T F o apsnd destQuadruple o dest_UN) t; in map_filter mk (dest_Un T) end; fun mk_prove mode (prop,params) = let val vars = map fst (strip_qnt_vars @{const_name All} (HOLogic.dest_Trueprop (Logic.strip_assums_concl prop))); + val [asmUN'] = adapt_aux_var ctxt true vars [get_aux_tvar (AsmUN mode)] in Goal.prove ctxt params [] prop (fn {context = ctxt', ...} => EVERY[trace_tac ctxt' "extracting specifications from hoare context", - resolve_tac ctxt' (adapt_aux_var ctxt' true vars [get_aux_tvar (AsmUN mode)]) 1, + resolve_tac ctxt' [asmUN'] 1, DEPTH_SOLVE_1 (resolve_tac ctxt' [subset_refl,refl] 1 ORELSE ((resolve_tac ctxt' [@{thm Hoare.subset_unI1}] 1 APPEND resolve_tac ctxt' [@{thm Hoare.subset_unI2}] 1) ORELSE (resolve_tac ctxt' [@{thm Hoare.subset_singleton_insert1}] 1 APPEND resolve_tac ctxt' [@{thm Hoare.subset_singleton_insert2}] 1))) ORELSE error_tac ("vcg: cannot extract specifications from context") ]) end; val specs = hoare_context_specs mode G T F; in map (mk_prove mode) specs end; +fun is_modifies_assertion t = + exists_subterm (fn (Const (@{const_name Hoare.meq},_)) => true| _ => false) t fun is_modifies_clause t = - exists_subterm (fn (Const (@{const_name Hoare.meq},_)) => true| _ => false) - (#3 (dest_hoare (Logic.strip_assums_concl t))) + is_modifies_assertion (#3 (dest_hoare (Logic.strip_assums_concl t))) handle (TERM _) => false; val is_spec_clause = not o is_modifies_clause; (* e.g: Intg => the_Intg lift Intg => lift the_Intg map Ingt => map the_Intg Hp o lift Intg => lift the_Intg o the_Hp *) fun swap_constr_destr f (t as (Const (@{const_name Fun.id},_))) = t | swap_constr_destr f (t as (Const (c,Type ("fun",[T,valT])))) = (Const (f c, Type ("fun",[valT,T])) handle Empty => raise TERM ("Hoare.swap_constr_destr",[t])) - | swap_constr_destr f (Const ("StateFun.map_fun",Type ("fun", (* FIXME unknown "StateFun.map_fun" !? *) + | swap_constr_destr f (Const ("StateFun.map_fun",Type ("fun", (* fixme: unknown "StateFun.map_fun" !? *) [Type ("fun",[T,valT]), Type ("fun",[Type ("fun",[xT,T']), Type ("fun",[xT',valT'])])]))$g) = Const ("StateFun.map_fun",Type("fun",[Type ("fun",[valT,T]), Type ("fun",[Type ("fun",[xT,valT']), Type ("fun",[xT',T'])])]))$ swap_constr_destr f g | swap_constr_destr f (Const (@{const_name Fun.comp},Type ("fun", [Type ("fun",[bT',cT]), Type ("fun",[Type ("fun",[aT ,bT]), Type ("fun",[aT',cT'])])]))$h$g) = let val h'=swap_constr_destr f h; val g'=swap_constr_destr f g; in Const (@{const_name Fun.comp},Type ("fun", [Type ("fun",[bT,aT]), Type ("fun",[Type ("fun",[cT,bT']), Type ("fun",[cT',aT'])])]))$g'$h' end | swap_constr_destr f (Const (@{const_name List.map},Type ("fun", [Type ("fun",[aT,bT]), Type ("fun",[asT,bsT])]))$g) = (Const (@{const_name List.map},Type ("fun", [Type ("fun",[bT,aT]), Type ("fun",[bsT,asT])]))$swap_constr_destr f g) | swap_constr_destr f t = raise TERM ("Hoare.swap_constr_destr",[t]); -(* FIXME: unused? *) +(* fixme: unused? *) val destr_to_constr = let fun convert c = let val (path,base) = split_last (Long_Name.explode c); in Long_Name.implode (path @ ["val",unprefix "the_" base]) end; in swap_constr_destr convert end; fun gen_call_tac cont_tac mode cmode state_kind state_space ctxt asms spec_sfx - pname return has_args _ = + pname return has_args result_exn _ = let val thy = Proof_Context.theory_of ctxt; - val pname' = unsuffix proc_deco pname; + val pname' = chopsfx proc_deco pname; val spec = (case AList.lookup (op =) asms pname of SOME s => SOME s | NONE => try (Proof_Context.get_thm ctxt) (suffix spec_sfx pname')); - fun auxvars_for p t = (case first_subterm_dest (try dest_call) t of - SOME (vars,(_,p',_,_,m,_)) => (if m=Static andalso + SOME (vars,(_,p',_,_,m,_,_)) => (if m=Static andalso p=(dest_string' p') then SOME vars else NONE) | NONE => NONE); fun get_auxvars_for p t = (case (map_filter ((auxvars_for p) o snd) (max_subterms_dest tap_UN t)) of (vars::_) => vars | _ => []); fun spec_tac ctxt' augment_rule augment_emptyFaults _ spec i = let val spec' = augment_emptyFaults OF [spec] handle THM _ => spec; in EVERY [resolve_tac ctxt' [augment_rule] i, resolve_tac ctxt' [spec'] (i+1), TRY (resolve_tac ctxt' [subset_refl, @{thm Set.empty_subsetI}, @{thm Set.Un_upper1}, @{thm Set.Un_upper2}] i)] end; fun check_spec name P thm = (case try dest_hoare (Thm.concl_of thm) of SOME spc => (case try dest_call (#2 spc) of - SOME (_,p,_,_,m,_) => if proc_name m p = name andalso + SOME (_,p,_,_,m,_,_) => if proc_name ctxt m p = name andalso P (Thm.concl_of thm) then SOME (#5 spc,thm) else NONE | _ => NONE) | _ => NONE) fun find_dyn_specs name P thms = map_filter (check_spec name P) thms; fun get_spec name P thms = case find_dyn_specs name P thms of (spec_mode,spec)::_ => SOME (spec_mode,spec) | _ => NONE; fun solve_spec ctxt' augment_rule _ augment_emptyFaults mode _ (SOME spec_mode) (SOME spec) i= if mode=Partial andalso spec_mode=Total then resolve_tac ctxt' [@{thm HoareTotal.hoaret_to_hoarep'}] i THEN spec_tac ctxt' augment_rule augment_emptyFaults mode spec i else if mode=spec_mode then spec_tac ctxt' augment_rule augment_emptyFaults mode spec i else error("vcg: cannot use a partial correctness specification of " ^ pname' ^ " for a total correctness proof") | solve_spec ctxt' _ asmUN_rule _ mode Static _ _ i =(* try to infer spec out of context *) EVERY[trace_tac ctxt' "inferring specification from hoare context1", resolve_tac ctxt' [asmUN_rule] i, DEPTH_SOLVE_1 (resolve_tac ctxt' [subset_refl,refl] i ORELSE ((resolve_tac ctxt' [@{thm Hoare.subset_unI1}] i APPEND resolve_tac ctxt' [@{thm Hoare.subset_unI2}] i) ORELSE (resolve_tac ctxt' [@{thm Hoare.subset_singleton_insert1}] i APPEND resolve_tac ctxt' [@{thm Hoare.subset_singleton_insert2}] i))) ORELSE error_tac ("vcg: cannot infer specification of " ^ pname' ^ " from context") (* if tactic for DEPTH_SOLVE_1 would create new subgoals, use SELECT_GOAL and DEPTH_SOLVE *) ] | solve_spec ctxt' augment_rule asmUN_rule augment_emptyFaults mode Parameter _ _ i = (* try to infer spec out of assumptions *) let - fun tac ({context = ctxt'', prems, ...}: Subgoal.focus) = - (case (find_dyn_specs pname is_spec_clause prems) of + fun tac thms = + (case (find_dyn_specs pname is_spec_clause thms) of (spec_mode,spec)::_ - => solve_spec ctxt'' augment_rule asmUN_rule augment_emptyFaults mode Parameter + => solve_spec ctxt' augment_rule asmUN_rule augment_emptyFaults mode Parameter (SOME spec_mode) (SOME spec) 1 | _ => all_tac) - in Subgoal.FOCUS tac ctxt' i end + in Subgoal.FOCUS (tac o #prems) ctxt i end val strip_spec_vars = strip_qnt_vars @{const_name All} o HOLogic.dest_Trueprop; fun apply_call_tac ctxt' pname mode cmode spec_mode spec_goal is_abr spec (subgoal,i) = let val spec_vars = map fst (case spec of SOME sp => (strip_spec_vars (Thm.concl_of sp)) | NONE => (case try (dest_hoare) subgoal of SOME (_,_,_,_,_,_,Theta,_) => get_auxvars_for pname Theta | _ => [])); - fun get_call_rule Static mode is_abr = - if is_abr then Proc mode else ProcNoAbr mode - | get_call_rule Parameter mode is_abr = - if is_abr then DynProcProcPar mode else DynProcProcParNoAbr mode; + fun get_call_rule' Static mode is_abr result_exn = + if is_abr then Proc mode result_exn else ProcNoAbr mode result_exn + | get_call_rule' Parameter mode is_abr result_exn = + if is_abr then DynProcProcPar mode result_exn else DynProcProcParNoAbr mode result_exn; val [call_rule,augment_ctxt_rule,asmUN_rule, augment_emptyFaults] = adapt_aux_var ctxt' true spec_vars (map get_aux_tvar - [get_call_rule cmode mode is_abr, + [get_call_rule' cmode mode is_abr result_exn, AugmentContext mode, AsmUN mode, AugmentEmptyFaults mode]); - in EVERY [resolve_tac ctxt' [call_rule] i, trace_tac ctxt' "call_tac -- basic_tac -- solving spec", - solve_spec ctxt' augment_ctxt_rule asmUN_rule augment_emptyFaults + solve_spec ctxt' augment_ctxt_rule asmUN_rule augment_emptyFaults mode cmode spec_mode spec spec_goal] end; fun basic_tac ctxt' spec i = let val msg ="Theorem " ^pname'^spec_sfx ^ " is no proper specification for procedure " ^pname'^ "; trying to infer specification from hoare context"; fun spec' s mode abr = if is_modifies_clause (Thm.concl_of s) then if abr then (TrivPost mode) OF [s] else (TrivPostNoAbr mode) OF [s] else s; val (is_abr,spec_mode,spec,spec_has_args) = (* is there a proper specification fact? *) case spec of NONE => (true,NONE,NONE,false) | SOME s => case try dest_hoare (Thm.concl_of s) of NONE => (warning msg;(true,NONE,NONE,false)) | SOME (_,c,Q,spec_abr,spec_mode,_,_,_) => case try dest_call c of NONE => (warning msg;(true,NONE,NONE,false)) - | SOME (_,p,_,_,m,spec_has_args) - => if proc_name m p = pname + | SOME (_,p,_,_,m,spec_has_args,_) + => if proc_name ctxt m p = pname then if (mode=Total andalso spec_mode=Partial) then (warning msg;(true,NONE,NONE,false)) else if is_empty_set spec_abr then (false,SOME spec_mode, SOME (spec' s spec_mode false),spec_has_args) else (true,SOME spec_mode, SOME (spec' s spec_mode true),spec_has_args) else (warning msg;(true,NONE,NONE,false)); val () = if spec_has_args then error "procedure call in specification must be parameterless!" else (); val spec_goal = i+2; in EVERY[trace_tac ctxt' "call_tac -- basic_tac -- START --", SUBGOAL (apply_call_tac ctxt' pname mode cmode spec_mode spec_goal is_abr spec) i, resolve_tac ctxt' [allI] (i+1), resolve_tac ctxt' [allI] (i+1), cont_tac ctxt' (i+1), trace_tac ctxt' "call_tac -- basic_tac -- simplify", - conseq_simp_tac ctxt' state_kind [@{thm StateSpace.upd_globals_def}] i, + conseq_simp_tac ctxt' state_kind (Named_Theorems.get ctxt @{named_theorems "state_simp"}) i, trace_tac ctxt' "call_tac -- basic_tac -- STOP --"] end; fun get_modifies (Const (@{const_name Collect},_) $ Abs (_,_,m)) = m | get_modifies t = raise TERM ("gen_call_tac.get_modifies",[t]); fun get_updates (Abs (_,_,t)) = get_updates t | get_updates (Const (@{const_name Hoare.mex},_) $ t) = get_updates t | get_updates (Const (@{const_name Hoare.meq},T) $ _ $ upd) = (T,upd) | get_updates t = raise TERM ("gen_call_tac.get_updates",[t]); (* return has the form: %s t. s(|globals:=globals t,...|) * should be translated to %s t. s(|globals := globals s(|m := m (globals t),...|),...|) * for all m in the modifies list. *) fun mk_subst gT meqT = fst (Sign.typ_unify thy (gT,domain_type meqT) (Vartab.empty,0)); fun mk_selR subst gT (upd,uT) = let val vT = range_type (hd (binder_types uT)); in Const (unsuffix Record.updateN upd,gT --> (Envir.norm_type subst vT)) end; (* lookup:: "('v => 'a) => 'n => ('n => 'v) => 'a" * update:: "('v => 'a) => ('a => 'v) => 'n => ('a => 'a) => ('n => 'v) => ('n => 'v)" *) fun mk_selF subst uT d n = let val vT_a::a_vT::nT::aT_aT::n_vT::_ = binder_types uT; val lT = (Envir.norm_type subst (vT_a --> nT --> n_vT --> (domain_type aT_aT))); val d' = map_types (Envir.norm_type subst) d; in Const (@{const_name StateFun.lookup},lT)$d'$n end; fun mk_rupdR subst gT (upd,uT) = let val [vT,_] = binder_types uT in Const (upd,(Envir.norm_type subst vT) --> gT --> gT) end; fun K_fun kn uT = let val T=range_type (hd (binder_types uT)) in Const (kn,T --> T --> T) end; fun K_rec uT t = let val T=range_type (hd (binder_types uT)) in Abs ("_", T, incr_boundvars 1 t) end; fun mk_supdF subst uT d c n = let val uT' = Envir.norm_type subst uT; val c' = map_types (Envir.norm_type subst) c; val d' = map_types (Envir.norm_type subst) d; in Const (@{const_name StateFun.update},uT')$d'$c'$n end; fun modify_updatesR subst gT glob ((Const (upd,uT))$_$(Const _$Z)) = mk_rupdR subst gT (upd,uT)$ (K_rec uT (mk_selR subst gT (upd,uT)$(glob$Bound 0)))$(glob$Bound 1) | modify_updatesR subst gT glob ((Const (upd,uT))$_$s) = mk_rupdR subst gT (upd,uT)$ (K_rec uT (mk_selR subst gT (upd,uT)$(glob$Bound 0)))$ modify_updatesR subst gT glob s | modify_updatesR subst gT glob ((_$Z)) = (glob$Bound 1) (* may_not_modify *) | modify_updatesR _ _ _ t = raise TERM ("gen_call_tac.modify_updatesR",[t]); fun modify_updatesF subst _ glob (Const (@{const_name StateFun.update},uT)$d$c$n$_$(Const globs$Z)) = mk_supdF subst uT d c n$ (K_fun KNF uT$(mk_selF subst uT d n$(glob$Bound 0)))$(glob$Bound 1) | modify_updatesF subst gT glob (Const (@{const_name StateFun.update},uT)$d$c$n$_$s) = mk_supdF subst uT d c n$ (K_fun KNF uT$(mk_selF subst uT d n$(glob$Bound 0)))$modify_updatesF subst gT glob s | modify_updatesF subst _ glob ((globs$Z)) = (glob$Bound 1) (* may_not_modify *) | modify_updatesF _ _ _ t = raise TERM ("gen_call_tac.modify_updatesF",[t]); - fun modify_updates Record = modify_updatesR - | modify_updates _ = modify_updatesF + fun modify_updates Function = modify_updatesF + | modify_updates _ (* Record and Other *) = modify_updatesR fun globalsT (Const (gupd,T)) = domain_type (hd (binder_types T)) | globalsT t = raise TERM ("gen_call_tac.globalsT",[t]); fun mk_upd meqT mods (gupd$(Abs (dmy,dmyT,(glob$Bound 1)))$Bound 1) = let val gT = (globalsT gupd); val subst = mk_subst gT meqT; in (gupd$(Abs (dmy,dmyT,incr_boundvars 1 (modify_updates state_kind subst gT glob mods)))$Bound 1) end | mk_upd meqT mods (upd$v$s) = upd$v$mk_upd meqT mods s | mk_upd _ _ t = raise TERM ("gen_call_tac.mk_upd",[t]); fun modify_return (meqT,mods) (Abs (s,T,Abs (t,U,upd))) = (Abs (s,T,Abs (t,U,mk_upd meqT mods upd))) | modify_return _ t = raise TERM ("get_call_tac.modify_return",[t]); fun modify_tac ctxt' spec modifies_thm i = try (fn () => let val (_,call,modif_spec_nrm,modif_spec_abr,modif_spec_mode,G,Theta,_) = dest_hoare (Thm.concl_of modifies_thm); val is_abr = not (is_empty_set modif_spec_abr); val emptyTheta = is_empty_set Theta; (*val emptyFaults = is_empty_set F;*) - val spec_has_args = #6 (dest_call call); + val (_,_,_,_,_,spec_has_args,_) = dest_call call; val () = if spec_has_args then error "procedure call in modifies-specification must be parameterless!" else (); val (mxprem,ModRet) = (case cmode of Static => (8,if is_abr - then if emptyTheta then (ProcModifyReturn mode) - else (ProcModifyReturnSameFaults mode) - else if emptyTheta then (ProcModifyReturnNoAbr mode) - else (ProcModifyReturnNoAbrSameFaults mode)) + then if emptyTheta then (ProcModifyReturn mode result_exn) + else (ProcModifyReturnSameFaults mode result_exn) + else if emptyTheta then (ProcModifyReturnNoAbr mode result_exn) + else (ProcModifyReturnNoAbrSameFaults mode result_exn)) | Parameter => (9,if is_abr - then if emptyTheta then (ProcProcParModifyReturn mode) - else (ProcProcParModifyReturnSameFaults mode) - else if emptyTheta then (ProcProcParModifyReturnNoAbr mode) - else (ProcProcParModifyReturnNoAbrSameFaults mode))); + then if emptyTheta then (ProcProcParModifyReturn mode result_exn) + else (ProcProcParModifyReturnSameFaults mode result_exn) + else if emptyTheta then (ProcProcParModifyReturnNoAbr mode result_exn) + else (ProcProcParModifyReturnNoAbrSameFaults mode result_exn))); val to_prove_prem = (case cmode of Static => 0 | Parameter => 1); val spec_goal = if is_abr then i + mxprem - 5 else i + mxprem - 6 val mods_nrm = modif_spec_nrm |> get_modifies |> get_updates; val return' = modify_return mods_nrm return; (* val return' = if is_abr then let val mods_abr = modif_spec_abr |> get_modifies |> get_updates; in modify_return mods_abr return end else return;*) val cret = Thm.cterm_of ctxt' return'; - val (_,_,return'_var,_,_,_) = nth (Thm.prems_of ModRet) to_prove_prem + val (_,_,return'_var,_,_,_,_) = nth (Thm.prems_of ModRet) to_prove_prem |> dest_hoare |> #2 |> dest_call; val ModRet' = infer_instantiate ctxt' [(#1 (dest_Var return'_var), cret)] ModRet; val modifies_thm_partial = if modif_spec_mode = Total then @{thm HoareTotal.hoaret_to_hoarep'} OF [modifies_thm] else modifies_thm; fun solve_modifies_tac i = (clarsimp_tac ((ctxt' |> put_claset (claset_of @{theory_context Set}) |> put_simpset (simpset_of @{theory_context Set})) - addsimps ([@{thm Hoare.mex_def},@{thm Hoare.meq_def},@{thm StateSpace.upd_globals_def}]@K_convs) - addsimprocs (state_upd_simproc Record::(state_simprocs state_kind)) + addsimps ([@{thm Hoare.mex_def},@{thm Hoare.meq_def}]@K_convs@ + (Named_Theorems.get ctxt @{named_theorems "state_simp"})) + addsimprocs (state_upd_simprocs ctxt Record @ state_simprocs ctxt state_kind) |> fold Simplifier.add_cong K_congs) i) THEN_MAYBE EVERY [trace_tac ctxt' "modify_tac: splitting record", state_split_simp_tac ctxt' [] state_space i]; val cnt = i + mxprem; in EVERY[trace_tac ctxt' "call_tac -- modifies_tac --", resolve_tac ctxt' [ModRet'] i, solve_spec ctxt' (AugmentContext Partial) (AsmUN Partial) (AugmentEmptyFaults Partial) Partial Static (SOME Partial) (SOME modifies_thm_partial) spec_goal, if is_abr then EVERY [trace_tac ctxt' "call_tac -- Solving abrupt modifies --", solve_modifies_tac (cnt - 6)] else all_tac, trace_tac ctxt' "call_tac -- Solving Modifies --", solve_modifies_tac (cnt - 7), basic_tac ctxt' spec (cnt - 8), if cmode = Parameter then EVERY [resolve_tac ctxt' [subset_refl] (cnt - 8), simp_tac (put_simpset HOL_basic_ss ctxt' addsimps (@{thm Hoare.CollectInt_iff} :: @{thms simp_thms})) 1] else all_tac] end) () |> (fn SOME res => res | NONE => raise TERM ("get_call_tac.modify_tac: no proper modifies spec", [])); - fun specs_of_assms_tac ({context = ctxt', prems, ...}: Subgoal.focus) = + fun specs_of_assms_tac ({context = ctxt', prems, ...}: Subgoal.focus) = (case get_spec pname is_spec_clause prems of SOME (_,spec) => (case get_spec pname is_modifies_clause prems of SOME (_,modifies_thm) => modify_tac ctxt' (SOME spec) modifies_thm 1 | NONE => basic_tac ctxt' (SOME spec) 1) | NONE => (warning ("no proper specification for procedure " ^pname^ " in assumptions"); all_tac)); val test_modify_in_ctxt_tac = let val mname = (suffix modifysfx pname'); val mspec = (case try (Proof_Context.get_thm ctxt) mname of SOME s => SOME s | NONE => (case AList.lookup (op =) asms pname of SOME s => if is_modifies_clause (Thm.concl_of s) then SOME s else NONE | NONE => NONE)); in (case mspec of NONE => basic_tac ctxt spec | SOME modifies_thm => (case check_spec pname is_modifies_clause modifies_thm of SOME _ => modify_tac ctxt spec modifies_thm | NONE => (warning ("ignoring theorem " ^ (suffix modifysfx pname') ^ "; no proper modifies specification for procedure "^pname'); basic_tac ctxt spec))) end; - fun inline_bdy_tac has_args i = + fun inline_bdy_tac has_args result_exn i = (case try (Proof_Context.get_thm ctxt) (suffix bodyP pname') of NONE => no_tac | SOME impl => (case try (Proof_Context.get_thm ctxt) (suffix (body_def_sfx^"_def") pname') of NONE => no_tac | SOME bdy => (tracing ("No specification found for procedure \"" ^ pname' ^ "\". Inlining procedure!"); if has_args then EVERY [trace_tac ctxt "inline_bdy_tac args", - resolve_tac ctxt [CallBody mode] i, + resolve_tac ctxt [CallBody mode result_exn] i, resolve_tac ctxt [impl] (i+3), resolve_tac ctxt [allI] (i+2), resolve_tac ctxt [allI] (i+2), in_assertion_simp_tac ctxt state_kind [] (i+2), cont_tac ctxt (i+2), resolve_tac ctxt [allI] (i+1),in_assertion_simp_tac ctxt state_kind [bdy] (i+1), cont_tac ctxt (i+1), in_assertion_simp_tac ctxt state_kind [@{thm StateSpace.upd_globals_def}] i] else EVERY [trace_tac ctxt "inline_bdy_tac no args", resolve_tac ctxt [ProcBody mode] i, resolve_tac ctxt [impl] (i+2), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [bdy]) (i+1), cont_tac ctxt (i+1)]))); in (case cmode of Static => if get_recursive pname ctxt = SOME false andalso is_none spec - then inline_bdy_tac has_args + then inline_bdy_tac has_args result_exn else test_modify_in_ctxt_tac | Parameter => (case spec of - NONE => (tracing "no spec found!"; Subgoal.FOCUS specs_of_assms_tac ctxt) + NONE => (trace_msg ctxt "no spec found!"; Subgoal.FOCUS specs_of_assms_tac ctxt) | SOME spec => - (tracing "found spec!"; case check_spec pname is_spec_clause spec of + (trace_msg ctxt "found spec!"; case check_spec pname is_spec_clause spec of SOME _ => test_modify_in_ctxt_tac | NONE => (warning ("ignoring theorem " ^ (suffix spec_sfx pname') ^ "; no proper specification for procedure " ^pname'); Subgoal.FOCUS specs_of_assms_tac ctxt)))) end; fun call_tac cont_tac mode state_kind state_space ctxt asms spec_sfx t = let - val (_,c,_,_,_,_,_,F) = dest_hoare (Logic.strip_assums_concl t); - fun gen_tac (_,pname,return,c,cmode,has_args) = + val (_,c,_,_,_,_,_,F) = dest_hoare t; + fun gen_tac (_,pname,return,c,cmode,has_args,result_exn) = gen_call_tac cont_tac mode cmode state_kind state_space ctxt asms spec_sfx - (proc_name cmode pname) return has_args F; + (proc_name ctxt cmode pname) return has_args result_exn F; in gen_tac (dest_call c) end handle TERM _ => K no_tac; fun solve_in_Faults_tac ctxt i = resolve_tac ctxt [UNIV_I, @{thm in_insert_hd}] i ORELSE SELECT_GOAL (SOLVE (simp_tac (put_simpset (simpset_of @{theory_context Set}) ctxt) 1)) i; local fun triv_simp ctxt = merge_assertion_simp_tac ctxt [mem_Collect_eq] THEN' simp_tac (put_simpset HOL_basic_ss ctxt addsimps @{thms simp_thms} |> fold Simplifier.add_cong [@{thm conj_cong}, @{thm imp_cong}]); (* a guarded while produces stupid things, since the guards are put at the end of the body and in the invariant (rule WhileAnnoG): - guard: g /\ g - guarantee: g --> g *) in fun guard_tac ctxt strip cont_tac mode (t,i) = let val (_,c,_,_,_,_,_,F) = dest_hoare (Logic.strip_assums_concl t); val (_,_,_,doStrip) = dest_Guard c; val guarantees = if strip orelse doStrip then [GuardStrip mode, GuaranteeStrip mode] else [Guarantee mode] fun basic_tac i = EVERY [resolve_tac ctxt [Guard mode, GuaranteeAsGuard mode] i, trace_tac ctxt "Guard", cont_tac ctxt (i+1), triv_simp ctxt i] fun guarantee_tac i = EVERY [resolve_tac ctxt guarantees i, solve_in_Faults_tac ctxt (i+2), cont_tac ctxt (i+1), triv_simp ctxt i] + in if is_empty_set F then EVERY [trace_tac ctxt "Guard: basic_tac", basic_tac i] else EVERY [trace_tac ctxt "Guard: trying guarantee_tac", guarantee_tac i ORELSE basic_tac i] end handle TERM _ => no_tac end; fun wf_tac ctxt = simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm Wellfounded.wf_measure},@{thm Wellfounded.wf_lex_prod},@{thm Wfrec.wf_same_fst}, @{thm Hoare.wf_measure_lex_prod},@{thm Wellfounded.wf_inv_image}]); fun in_rel_simp ctxt = simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm Hoare.in_measure_iff},@{thm Hoare.in_lex_iff},@{thm Hoare.in_mlex_iff},@{thm Hoare.in_inv_image_iff}, @{thm split_conv}]); -fun while_annotate_tac ctxt inv i st = +fun while_annotate_tac ctxt inv state_space i st = let val annotateWhile = Thm.lift_rule (Thm.cprem_of st i) @{thm HoarePartial.reannotateWhileNoGuard}; - val lifted_inv = fold_rev Term.abs (Logic.strip_params (Logic.get_goal (Thm.prop_of st) i)) inv; + val params = Logic.strip_params (Logic.get_goal (Thm.prop_of st) i) + val first_state_idx = find_index (fn x => state_space (Free x) <> 0) (rev params) + val inv = if first_state_idx > 0 then incr_boundvars first_state_idx inv else inv + val lifted_inv = fold_rev Term.abs params inv; val invVar = (#1 o strip_comb o #3 o dest_whileAnno o #2 o dest_hoare) (List.last (Thm.prems_of annotateWhile)); val annotate = infer_instantiate ctxt [(#1 (dest_Var invVar), Thm.cterm_of ctxt lifted_inv)] annotateWhile; - in ((trace_tac ctxt ("annotating While with: " ^ Syntax.string_of_term ctxt lifted_inv )) + in ((trace_tac ctxt ("try annotating While with: " ^ Syntax.string_of_term ctxt lifted_inv )) THEN compose_tac ctxt (false,annotate,1) i) st end; -fun cond_annotate_tac ctxt inv mode (_,i) st = +fun cond_annotate_tac ctxt inv mode state_space (_,i) st = let val annotateCond = Thm.lift_rule (Thm.cprem_of st i) (CondInv' mode); - val lifted_inv = fold_rev Term.abs (Logic.strip_params (Logic.get_goal (Thm.prop_of st) i)) inv; + val params = Logic.strip_params (Logic.get_goal (Thm.prop_of st) i) + val first_state_idx = find_index (fn x => state_space (Free x) <> 0) (rev params) + val inv = if first_state_idx > 0 then incr_boundvars first_state_idx inv else inv + val lifted_inv = fold_rev Term.abs params inv; val invVar = List.last (Thm.prems_of annotateCond) |> dest_hoare |> #3 |> strip_comb |> #1; val annotate = infer_instantiate ctxt [(#1 (dest_Var invVar), Thm.cterm_of ctxt lifted_inv)] annotateCond; - in ((trace_tac ctxt ("annotating Cond with: "^ Syntax.string_of_term ctxt lifted_inv)) + in ((trace_tac ctxt ("try annotating Cond with: "^ Syntax.string_of_term ctxt lifted_inv)) THEN compose_tac ctxt (false,annotate,5) i) st end; fun basic_while_tac ctxt state_kind cont_tac tac mode i = let fun common_tac i = EVERY[if mode=Total then wf_tac ctxt (i+3) else all_tac, BasicSimpTac ctxt state_kind true [] tac (i+2), if mode=Total then in_rel_simp ctxt (i+1) THEN (resolve_tac ctxt [allI] (i+1)) else all_tac, cont_tac ctxt (i+1) ] fun basic_tac i = EVERY [resolve_tac ctxt [While mode] i, common_tac i] in EVERY [trace_tac ctxt "basic_while_tac: basic_tac", basic_tac i] end; -fun while_tac ctxt state_kind inv cont_tac tac mode t i= +fun while_tac ctxt state_kind state_space inv cont_tac tac mode t i= let val basic_tac = basic_while_tac ctxt state_kind cont_tac tac mode; in (case inv of NONE => basic_tac i - | SOME I => EVERY [while_annotate_tac ctxt I i, basic_tac i]) + | SOME I => EVERY [while_annotate_tac ctxt I state_space i, basic_tac i]) end handle TERM _ => no_tac fun dest_split (Abs (x,T,t)) = let val (vs,recomb,bdy) = dest_split t; in ((x,T)::vs,fn t' => Abs (x,T,recomb t'),bdy) end | dest_split (c as Const (@{const_name case_prod},_)$Abs(x,T,t)) = let val (vs,recomb,bdy) = dest_split t; in ((x,T)::vs,fn t' => c$Abs (x,T,recomb t'),bdy) end | dest_split t = ([],I,t); fun whileAnnoG_tac ctxt strip_guards mode t i st = let val (_,c,_,_,_,_,_,F) = dest_hoare (Logic.strip_assums_concl t); val (SOME grds,_,I,_,_,fix) = dest_whileAnno c; val Rule = if fix then WhileAnnoGFix mode else WhileAnnoG mode; fun extract_faults (Const (@{const_name Set.insert}, _) $ t $ _) = [t] | extract_faults _ = []; fun leave_grd fs (Const (@{const_name Pair}, _) $ f $ g) = if member (op =) fs f andalso strip_guards then NONE else SOME g | leave_grd fs (Const (@{const_name Language.guaranteeStripPair}, _) $ f $ g) = if member (op =) fs f then NONE else SOME g | leave_grd fs _ = NONE; val (I_vs,I_recomb,I') = dest_split I; val grds' = map_filter (leave_grd (extract_faults F)) (HOLogic.dest_list grds); val pars = (Logic.strip_params (Logic.get_goal (Thm.prop_of st) i)); val J = fold_rev Term.abs pars (I_recomb (fold_rev (mk_Int (map snd (pars@I_vs))) grds' I')); val WhileG = Thm.lift_rule (Thm.cprem_of st i) Rule; val invVar = (fst o strip_comb o #3 o dest_whileAnno o (fn xs => nth xs 1) o snd o strip_comb o #2 o dest_hoare) (List.last (Thm.prems_of WhileG)); val WhileGInst = infer_instantiate ctxt [(#1 (dest_Var invVar), Thm.cterm_of ctxt J)] WhileG; in ((trace_tac ctxt ("WhileAnnoG, adding guards to invariant: " ^ Syntax.string_of_term ctxt J)) THEN compose_tac ctxt (false,WhileGInst,1) i) st end handle TERM _ => no_tac st | Bind => no_tac st (* renames bound state variable according to name given in goal, * before rule specAnno is applied, and solves sidecondition *) fun gen_Anno_tac dest rules tac cont_tac ctxt state_kind (t,i) st = let val vars = (dest o #2 o dest_hoare) (Logic.strip_assums_concl t); val rules' = (case (List.filter (not o null) (map dest_splits vars)) of [] => rules |(xs::_) => adapt_aux_var ctxt false (map fst xs) (map get_aux_tvar rules)); in EVERY [resolve_tac ctxt rules' i, tac, simp_tac (put_simpset HOL_basic_ss ctxt addsimps ([@{thm split_conv}, refl, @{thm Hoare.triv_All_eq}]@anno_defs) addsimprocs [@{simproc case_prod_beta}]) (i+2), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]) (i+1), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]) i, REPEAT (resolve_tac ctxt [allI] (i+1)), cont_tac ctxt (i+1), conseq_simp_tac ctxt state_kind [] i] st end handle TERM _ => no_tac st; fun specAnno_tac ctxt state_kind cont_tac mode = let fun dest_specAnno (Const (@{const_name Language.specAnno},_)$P$c$Q$A) = [P,c,Q,A] | dest_specAnno t = raise TERM ("dest_specAnno",[t]); val rules = [SpecAnnoNoAbrupt mode,SpecAnno mode]; in gen_Anno_tac dest_specAnno rules all_tac cont_tac ctxt state_kind end; fun whileAnnoFix_tac ctxt state_kind cont_tac mode (t,i) = let fun dest (Const (@{const_name Language.whileAnnoFix},_)$b$I$V$c) = [I,V,c] | dest t = raise TERM ("dest_whileAnnoFix",[t]); val rules = [WhileAnnoFix mode]; fun wf_tac' i = EVERY [REPEAT (resolve_tac ctxt [allI] i), simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm split_conv}]) i, wf_tac ctxt i]; val tac = if mode=Total then EVERY [wf_tac' (i+3),in_rel_simp ctxt (i+1)] else all_tac in gen_Anno_tac dest rules tac cont_tac ctxt state_kind (t,i) end; fun lemAnno_tac ctxt state_kind mode (t,i) st = let fun dest_name (Const (x,_)) = x | dest_name (Free (x,_)) = x | dest_name t = raise TERM ("dest_name",[t]); fun dest_lemAnno (Const (@{const_name Language.lem},_)$n$c) = let val x = Long_Name.base_name (dest_name n); in (case try (Proof_Context.get_thm ctxt) x of NONE => error ("No lemma: '" ^ x ^ "' found.") | SOME spec => (strip_qnt_vars @{const_name All} (HOLogic.dest_Trueprop (Thm.concl_of spec)),spec)) end | dest_lemAnno t = raise TERM ("dest_lemAnno",[t]); val (vars,spec) = dest_lemAnno (#2 (dest_hoare t)); val rules = [LemAnnoNoAbrupt mode,LemAnno mode]; val rules' = (case vars of [] => rules | xs => adapt_aux_var ctxt true (map fst xs) (map get_aux_tvar rules)); in EVERY [resolve_tac ctxt rules' i, resolve_tac ctxt [spec] (i+1), conseq_simp_tac ctxt state_kind [] i] st end handle TERM _ => no_tac st; fun prems_tac ctxt i = TRY (resolve_tac ctxt (Assumption.all_prems_of ctxt) i); -fun mk_proc_assoc thms = +fun mk_proc_assoc ctxt thms = let - fun name (_,p,_,_,cmode,_) = proc_name cmode p; + fun name (_,p,_,_,cmode,_,_) = proc_name ctxt cmode p; fun proc_name thm = thm |> Thm.concl_of |> dest_hoare |> #2 |> dest_call |> name; in map (fn thm => (proc_name thm,thm)) thms end; fun mk_hoare_tac cont ctxt mode i (name,tac) = EVERY [trace_tac ctxt ("trying: " ^ name),tac cont ctxt mode i]; (* the main hoare tactic *) fun HoareTac annotate_inv exspecs strip_guards mode state_kind state_space spec_sfx ctxt tac st = let val (P,c,Q,A,_,G,T,F) = dest_hoare (Logic.strip_assums_concl (Logic.get_goal (Thm.prop_of st) 1)); - val wp_tacs = #wp_tacs (get_data ctxt); + val solve_modifies = spec_sfx = modifysfx andalso annotate_inv andalso mode = Partial andalso + is_modifies_assertion Q andalso is_modifies_assertion A + + val hoare_tacs = #hoare_tacs (get_data ctxt); val params = (strip_vars (Logic.get_goal (Thm.prop_of st) 1)); val Inv = (if annotate_inv then (* Take the postcondition of the triple as invariant for all *) (* while loops (makes sense for the modifies clause) *) SOME Q else NONE); val exspecthms = map (Proof_Context.get_thm ctxt) exspecs; val asms = try (fn () => - mk_proc_assoc (gen_context_thms ctxt mode params G T F @ exspecthms)) () + mk_proc_assoc ctxt (gen_context_thms ctxt mode params G T F @ exspecthms)) () |> the_default []; fun while_annoG_tac (t,i) = whileAnnoG_tac ctxt (annotate_inv orelse strip_guards) mode t i; fun WlpTac tac i = (* WlpTac does not end with subset_refl *) FIRST - ([EVERY [resolve_tac ctxt [Seq mode] i,trace_tac ctxt "Seq",HoareRuleTac tac false ctxt (i+1)], - EVERY [resolve_tac ctxt [Catch mode] i,trace_tac ctxt "Catch",HoareRuleTac tac false ctxt (i+1)], - EVERY [resolve_tac ctxt [CondCatch mode] i,trace_tac ctxt "CondCatch",HoareRuleTac tac false ctxt (i+1)], - EVERY [resolve_tac ctxt [BSeq mode] i,trace_tac ctxt "BSeq",HoareRuleTac tac false ctxt (i+1)], + + ([EVERY [resolve_tac ctxt [Seq mode solve_modifies] i,trace_tac ctxt "Seq", HoareRuleTac tac false ctxt (i+1)], + EVERY [resolve_tac ctxt [Catch mode solve_modifies] i,trace_tac ctxt "Catch",HoareRuleTac tac false ctxt (i+1)], + EVERY [resolve_tac ctxt [CondCatch mode solve_modifies] i,trace_tac ctxt "CondCatch",HoareRuleTac tac false ctxt (i+1)], + EVERY [resolve_tac ctxt [BSeq mode solve_modifies] i,trace_tac ctxt "BSeq",HoareRuleTac tac false ctxt (i+1)], EVERY [resolve_tac ctxt [FCall mode] i,trace_tac ctxt "FCall"], EVERY [resolve_tac ctxt [GuardsNil mode] i,trace_tac ctxt "GuardsNil"], EVERY [resolve_tac ctxt [GuardsConsGuaranteeStrip mode] i, trace_tac ctxt "GuardsConsGuaranteeStrip"], EVERY [resolve_tac ctxt [GuardsCons mode] i,trace_tac ctxt "GuardsCons"], EVERY [SUBGOAL while_annoG_tac i] - ] - @ - map (mk_hoare_tac (fn p => HoareRuleTac tac p ctxt) ctxt mode i) wp_tacs) + ]) and HoareRuleTac tac pre_cond ctxt i st = - let fun call (t,i) = call_tac (HoareRuleTac tac false) + let + val _ = if Config.get ctxt hoare_trace > 1 then + print_tac ctxt ("HoareRuleTac (" ^ @{make_string} (pre_cond, i) ^ "):") st + else all_tac st + fun call (t,i) = call_tac (HoareRuleTac tac false) mode state_kind state_space ctxt asms spec_sfx t i fun cond_tac i = if annotate_inv andalso Config.get ctxt use_cond_inv_modifies then - EVERY[SUBGOAL (cond_annotate_tac ctxt (the Inv) mode) i, + EVERY[SUBGOAL (cond_annotate_tac ctxt (the Inv) mode state_space) i, HoareRuleTac tac false ctxt (i+4), HoareRuleTac tac false ctxt (i+3), BasicSimpTac ctxt state_kind true [] tac (i+2), BasicSimpTac ctxt state_kind true [] tac (i+1) ] else EVERY[resolve_tac ctxt [Cond mode] i,trace_tac ctxt "Cond", HoareRuleTac tac false ctxt (i+2), HoareRuleTac tac false ctxt (i+1)]; fun switch_tac i = EVERY[resolve_tac ctxt [SwitchNil mode] i, trace_tac ctxt "SwitchNil"] ORELSE EVERY[resolve_tac ctxt [SwitchCons mode] i,trace_tac ctxt "SwitchCons", HoareRuleTac tac false ctxt (i+2), HoareRuleTac tac false ctxt (i+1)]; - fun while_tac' (t,i) = while_tac ctxt state_kind Inv + fun while_tac' (t,i) = while_tac ctxt state_kind state_space Inv (HoareRuleTac tac true) tac mode t i; in st |> ( (WlpTac tac i THEN HoareRuleTac tac pre_cond ctxt i) ORELSE - (FIRST([EVERY[resolve_tac ctxt [Skip mode] i,trace_tac ctxt "Skip"], + (TRY (FIRST([EVERY[resolve_tac ctxt [Skip mode] i, trace_tac ctxt "Skip"], EVERY[resolve_tac ctxt [BasicCond mode] i, trace_tac ctxt "BasicCond", assertion_simp_tac ctxt state_kind [] i], (resolve_tac ctxt [Basic mode] i THEN trace_tac ctxt "Basic") THEN_MAYBE (assertion_simp_tac ctxt state_kind [] i), (* we don't really need simplificaton here. The question is if it is better to simplify the assertion after each Basic step, so that intermediate assertions stay "small", or if we just accumulate the raw assertions and leave the simplification to the final BasicSimpTac *) EVERY[resolve_tac ctxt [Throw mode] i, trace_tac ctxt "Throw"], (resolve_tac ctxt [Raise mode] i THEN trace_tac ctxt "Raise") THEN_MAYBE (assertion_string_eq_simp_tac ctxt state_kind [] i), - cond_tac i, - switch_tac i, + EVERY[cond_tac i], + EVERY[switch_tac i], EVERY[resolve_tac ctxt [Block mode] i, trace_tac ctxt "Block", resolve_tac ctxt [allI] (i+2), resolve_tac ctxt [allI] (i+2), HoareRuleTac tac false ctxt (i+2), resolve_tac ctxt [allI] (i+1), in_assertion_simp_tac ctxt state_kind [] (i+1), HoareRuleTac tac false ctxt (i+1)], SUBGOAL while_tac' i, SUBGOAL (guard_tac ctxt (annotate_inv orelse strip_guards) - (HoareRuleTac tac false) mode) i, + (HoareRuleTac tac false) mode THEN' (K (trace_tac ctxt "guard_tac succeeded"))) i, EVERY[SUBGOAL (specAnno_tac ctxt state_kind (HoareRuleTac tac true) mode) i], EVERY[SUBGOAL (whileAnnoFix_tac ctxt state_kind (HoareRuleTac tac true) mode) i], EVERY[resolve_tac ctxt [SpecIf mode] i, trace_tac ctxt "SpecIf", assertion_simp_tac ctxt state_kind [] i], (resolve_tac ctxt [Spec mode] i THEN trace_tac ctxt "Spec") THEN_MAYBE (assertion_simp_tac ctxt state_kind [@{thm split_conv}] i), EVERY[resolve_tac ctxt [BindR mode] i, trace_tac ctxt "Bind", resolve_tac ctxt [allI] (i+1), HoareRuleTac tac false ctxt (i+1)], EVERY[resolve_tac ctxt [DynCom mode] i, trace_tac ctxt "DynCom"], EVERY[trace_tac ctxt "calling call_tac",SUBGOAL call i], EVERY[trace_tac ctxt "LemmaAnno",SUBGOAL (lemAnno_tac ctxt state_kind mode) i]] @ - map (mk_hoare_tac (fn p => HoareRuleTac tac p ctxt) ctxt mode i) hoare_tacs) - THEN (if pre_cond - then EVERY [trace_tac ctxt "pre_cond", - TRY (BasicSimpTac ctxt state_kind true [] tac i), - (* FIXME: Do we need TRY *) - trace_tac ctxt "after BasicSimpTac"] + map (mk_hoare_tac (fn p => HoareRuleTac tac p ctxt) ctxt mode i) hoare_tacs)) + THEN (if pre_cond orelse solve_modifies + then EVERY [trace_tac ctxt ("pre_cond / solve_modfies: " ^ @{make_string} (pre_cond, solve_modifies)), + TRY (BasicSimpTac ctxt state_kind true (Named_Theorems.get ctxt @{named_theorems "state_simp"}) tac i), + trace_tac ctxt ("after BasicSimpTac " ^ string_of_int i)] else (resolve_tac ctxt [subset_refl] i)))) end; in ((K (EVERY [REPEAT (resolve_tac ctxt [allI] 1), HoareRuleTac tac true ctxt 1])) THEN_ALL_NEW (prems_tac ctxt)) 1 st (*Procedure specifications may have an locale assumption as premise. These are accumulated by the vcg and are be solved afterward by prems_tac *) end; fun prefer_tac i = (Tactic.defer_tac i THEN PRIMITIVE (Thm.permute_prems 0 ~1)); fun HoareStepTac strip_guards mode state_kind state_space spec_sfx ctxt tac st = let val asms = try (fn () => let val (_,_,_,_,_,G,T,F) = dest_hoare (Logic.strip_assums_concl (Logic.get_goal (Thm.prop_of st) 1)); val params = (strip_vars (Logic.get_goal (Thm.prop_of st) 1)); - in mk_proc_assoc (gen_context_thms ctxt mode params G T F) + in mk_proc_assoc ctxt (gen_context_thms ctxt mode params G T F) end) () |> the_default []; fun result_tac ctxt' i = TRY (EVERY [resolve_tac ctxt' [Basic mode] i, resolve_tac ctxt' [subset_refl] i]); fun call (t,i) = call_tac result_tac mode state_kind state_space ctxt asms spec_sfx t i fun final_simp_tac i = EVERY [full_simp_tac (put_simpset HOL_basic_ss ctxt addsimps [mem_Collect_eq]) i, REPEAT (eresolve_tac ctxt [conjE] i), TRY (hyp_subst_tac_thin true ctxt i), BasicSimpTac ctxt state_kind true [] tac i] fun while_annoG_tac (t,i) = whileAnnoG_tac ctxt strip_guards mode t i; - - in st |> + val hoare_tacs = #hoare_tacs (get_data ctxt); + in st |> CHANGED ( (REPEAT (resolve_tac ctxt [allI] 1) THEN - FIRST [resolve_tac ctxt [subset_refl] 1, + FIRST ([resolve_tac ctxt [subset_refl] 1, EVERY[resolve_tac ctxt [Skip mode] 1,TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [BasicCond mode] 1,trace_tac ctxt "BasicCond", TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [Basic mode] 1,TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [Throw mode] 1,TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [Raise mode] 1,TRY (assertion_string_eq_simp_tac ctxt state_kind [] 1)], resolve_tac ctxt [SeqSwap mode] 1 - THEN_MAYBE HoareStepTac strip_guards mode state_kind state_space spec_sfx - ctxt tac, - EVERY[resolve_tac ctxt [BSeq mode] 1, + THEN_MAYBE TRY (HoareStepTac strip_guards mode state_kind state_space spec_sfx + ctxt tac), + EVERY[resolve_tac ctxt [BSeq mode false] 1, prefer_tac 2 THEN_MAYBE HoareStepTac strip_guards mode state_kind state_space spec_sfx ctxt tac], resolve_tac ctxt [CondSwap mode] 1, resolve_tac ctxt [SwitchNil mode] 1, resolve_tac ctxt [SwitchCons mode] 1, EVERY [SUBGOAL while_annoG_tac 1], EVERY[resolve_tac ctxt [While mode] 1, if mode=Total then wf_tac ctxt 4 else all_tac, BasicSimpTac ctxt state_kind false [] tac 3, if mode=Total then in_rel_simp ctxt 2 THEN (resolve_tac ctxt [allI] 2) else all_tac, BasicSimpTac ctxt state_kind false [] tac 1], resolve_tac ctxt [CatchSwap mode] 1, resolve_tac ctxt [CondCatchSwap mode] 1, EVERY[resolve_tac ctxt [BlockSwap mode] 1, resolve_tac ctxt [allI] 1, resolve_tac ctxt [allI] 1, resolve_tac ctxt [allI] 2, BasicSimpTac ctxt state_kind false [] tac 2], resolve_tac ctxt [GuardsNil mode] 1, resolve_tac ctxt [GuardsConsGuaranteeStrip mode] 1, resolve_tac ctxt [GuardsCons mode] 1, SUBGOAL (guard_tac ctxt strip_guards (K (K all_tac)) mode) 1, EVERY[SUBGOAL (specAnno_tac ctxt state_kind (K (K all_tac)) mode) 1], EVERY[SUBGOAL (whileAnnoFix_tac ctxt state_kind (K (K all_tac)) mode) 1], EVERY[resolve_tac ctxt [SpecIf mode] 1,trace_tac ctxt "SpecIf", TRY (BasicSimpTac ctxt state_kind false [] tac 1)], EVERY[resolve_tac ctxt [Spec mode] 1, TRY (BasicSimpTac ctxt state_kind false [@{thm split_conv}] tac 1)], EVERY[resolve_tac ctxt [BindR mode] 1, resolve_tac ctxt [allI] 2, prefer_tac 2], EVERY[resolve_tac ctxt [FCall mode] 1], EVERY[resolve_tac ctxt [DynCom mode] 1], EVERY[SUBGOAL call 1, BasicSimpTac ctxt state_kind false [] tac 1], EVERY[SUBGOAL (lemAnno_tac ctxt state_kind mode) 1, - BasicSimpTac ctxt state_kind false [] tac 1], - final_simp_tac 1 - ]) + BasicSimpTac ctxt state_kind false [] tac 1] + + ] @ + map (mk_hoare_tac (K (K all_tac)) ctxt mode 1) hoare_tacs @ + [final_simp_tac 1]))) end; (*****************************************************************************) (** Generalise verification condition **) (*****************************************************************************) structure RecordSplitState : SPLIT_STATE = struct val globals = @{const_name StateSpace.state.globals}; -fun isState (Const _$Abs (s,T,t)) = +fun isState _ (Const _$Abs (s,T,t)) = (case (state_hierarchy T) of ((n,_)::_) => n = "StateSpace.state.state" andalso is_none (try dest_hoare_raw (strip_qnt_body @{const_name All} t)) | _ => false) - | isState _ = false; + | isState _ _ = false; fun isFreeState (Free (_,T)) = (case (state_hierarchy T) of ((n,_)::_) => n = "StateSpace.state.state" | _ => false) | isFreeState _ = false; -val abs_state = Option.map snd o first_subterm isFreeState; +fun abs_state _ = Option.map snd o first_subterm isFreeState; fun sel_eq (Const (x,_)$_) y = (x=y) | sel_eq t y = raise TERM ("RecordSplitState.sel_eq",[t]); val sel_idx = idx sel_eq; fun bound xs (t as (Const (x,_)$_)) = let val i = sel_idx xs x in if i < 0 then (length xs, xs@[t]) else (i,xs) end | bound xs t = raise TERM ("RecordSplitState.bound",[t]); -fun abs_var _ (Const (x,T)$_) = - (remdeco' (Long_Name.base_name x),range_type T) +fun abs_var ctxt (Const (x,T)$_) = + (remdeco' ctxt (Long_Name.base_name x),range_type T) | abs_var _ t = raise TERM ("RecordSplitState.abs_var",[t]); fun fld_eq (x, _) y = (x = y) fun fld_idx xs x = idx fld_eq xs x; fun sort_vars ctxt T vars = let val thy = Proof_Context.theory_of ctxt; val (flds,_) = Record.get_recT_fields thy T; val gT = the (AList.lookup (fn (x:string,y) => x=y) flds globals); val (gflds,_) = (Record.get_recT_fields thy gT handle TYPE _ => ([],("",dummyT))); fun compare (Const _$Free _, Const _$(Const _$Free _)) = GREATER | compare (Const (s1,_)$Free _, Const (s2,_)$Free _) = int_ord (fld_idx flds s1, fld_idx flds s2) | compare (Const (s1,_)$(Const _$Free _), Const (s2,_)$(Const _$Free _)) = int_ord (fld_idx gflds s1, fld_idx gflds s2) | compare _ = LESS; in sort (rev_order o compare) vars end; fun fold_state_prop loc glob app abs other inc s res (t as (Const (sel,_)$Free (s',_))) = if s'=s then if is_state_var sel then loc inc res t else raise TERM ("RecordSplitState.fold_state_prop",[t]) else other res t | fold_state_prop loc glob app abs other inc s res (t as ((t1 as (Const (sel,_)))$(t2 as (Const (glb,_)$Free (s',_))))) = if s'=s andalso is_state_var sel andalso (glb=globals) then glob inc res t else let val res1 = fold_state_prop loc glob app abs other inc s res t1 val res2 = fold_state_prop loc glob app abs other inc s res1 t2 in app res1 res2 end | fold_state_prop loc glob app abs other inc s res (t as (Free (s',_))) = if s'=s then raise TERM ("RecordSplitState.fold_state_prop",[t]) else other res t | fold_state_prop loc glob app abs other inc s res (t1$t2) = let val res1 = fold_state_prop loc glob app abs other inc s res t1 val res2 = fold_state_prop loc glob app abs other inc s res1 t2 in app res1 res2 end | fold_state_prop loc glob app abs other inc s res (Abs (x,T,t)) = let val res1 = fold_state_prop loc glob app abs other (inc+1) s res t in abs x T res1 end | fold_state_prop loc glob app abs other inc s res t = other res t fun collect_vars s t = let fun loc _ vars t = snd (bound vars t); fun glob _ vars t = snd (bound vars t); fun app _ vars2 = vars2; fun abs _ _ vars = vars; fun other vars _ = vars; in fold_state_prop loc glob app abs other 0 s [] t end; fun abstract_vars vars s t = let fun loc inc _ t = let val i = fst (bound vars t) in Bound (i+inc) end; fun glob inc _ t = let val i = fst (bound vars t) in Bound (i+inc) end; fun app t1 t2 = t1$t2; fun abs x T t = Abs (x,T,t); fun other _ t = t; val dummy = Bound 0; in fold_state_prop loc glob app abs other 0 s dummy t end; fun split_state ctxt s T t = let val vars = collect_vars s t; val vars' = if Config.get ctxt sort_variables then sort_vars ctxt T vars else vars; in (abstract_vars vars' s t,rev vars') end; -fun ex_tac ctxt _ st = Record.split_simp_tac ctxt @{thms simp_thms} (K ~1) 1 st; +fun ex_tac ctxt _ i = Record.split_simp_tac ctxt @{thms simp_thms} (K ~1) i; end; structure FunSplitState : SPLIT_STATE = struct val full_globalsN = @{const_name StateSpace.state.globals}; -fun isState (Const _$Abs (s,T,t)) = +fun isState _ (Const _$Abs (s,T,t)) = (case (state_hierarchy T) of ((n,_)::_) => n = "StateSpace.state.state" andalso is_none (try dest_hoare_raw (strip_qnt_body @{const_name All} t)) | _ => false) - | isState _ = false; + | isState _ _ = false; fun isFreeState (Free (_,T)) = (case (state_hierarchy T) of ((n,_)::_) => n = "StateSpace.state.state" | _ => false) | isFreeState _ = false; -val abs_state = Option.map snd o first_subterm isFreeState; +fun abs_state _ = Option.map snd o first_subterm isFreeState; fun comp_name t = - case try (implode o dest_string) t of + case try dest_string t of SOME str => str | NONE => (case t of Free (s,_) => s | Const (s,_) => s | t => raise TERM ("FunSplitState.comp_name",[t])) fun sel_name (Const _$_$name$_) = comp_name name | sel_name t = raise TERM ("FunSplitState.sel_name",[t]); fun sel_raw_name (Const _$_$name$_) = name | sel_raw_name t = raise TERM ("FunSplitState.sel_raw_name",[t]); fun component_type (Const _$_$_$(sel$_)) = range_type (fastype_of sel) | component_type t = raise TERM ("FunSplitState.component_type",[t]); fun component_name (Const _$_$_$((Const (sel,_)$_))) = sel | component_name t = raise TERM ("FunSplitState.component_name",[t]); fun sel_type (Const _$destr$_$_) = range_type (fastype_of destr) | sel_type t = raise TERM ("FunSplitState.sel_type",[t]); fun sel_destr (Const _$destr$_$_) = destr | sel_destr t = raise TERM ("FunSplitState.sel_destr",[t]); fun sel_eq t y = (sel_name t = y) | sel_eq t y = raise TERM ("FunSplitState.sel_eq",[t]); val sel_idx = idx sel_eq; fun bound xs t = let val i = sel_idx xs (sel_name t) in if i < 0 then (length xs, xs@[t]) else (i,xs) end | bound xs t = raise TERM ("FunSplitState.bound",[t]); fun fold_state_prop var app abs other inc s res (t as (Const (@{const_name StateFun.lookup},_)$destr$name$(Const _$Free (s',_)))) = if s'=s then var inc res t else other res t (*raise TERM ("FunSplitState.fold_state_prop",[t])*) | fold_state_prop var app abs other inc s res (t as (Free (s',_))) = if s'=s then raise TERM ("FunSplitState.fold_state_prop",[t]) else other res t | fold_state_prop var app abs other inc s res (t1$t2) = let val res1 = fold_state_prop var app abs other inc s res t1 val res2 = fold_state_prop var app abs other inc s res1 t2 in app res1 res2 end | fold_state_prop var app abs other inc s res (Abs (x,T,t)) = let val res1 = fold_state_prop var app abs other (inc+1) s res t in abs x T res1 end | fold_state_prop var app abs other inc s res t = other res t fun collect_vars s t = let fun var _ vars t = snd (bound vars t); fun app _ vars2 = vars2; fun abs _ _ vars = vars; fun other vars _ = vars; in fold_state_prop var app abs other 0 s [] t end; fun abstract_vars vars s t = let fun var inc _ t = let val i = fst (bound vars t) in Bound (i+inc) end; fun app t1 t2 = t1$t2; fun abs x T t = Abs (x,T,t); fun other _ t = t; val dummy = Bound 0; in fold_state_prop var app abs other 0 s dummy t end; -fun sort_vars _ vars = +fun sort_vars ctxt vars = let val fld_idx = idx (fn s1:string => fn s2 => s1 = s2); fun compare (_$_$n$(Const (s1,_)$_),_$_$m$(Const (s2,_)$_)) = let - val n' = remdeco' (comp_name n); - val m' = remdeco' (comp_name m); + val n' = remdeco' ctxt (comp_name n); + val m' = remdeco' ctxt (comp_name m); in if s1 = full_globalsN then if s2 = full_globalsN then string_ord (n',m') else LESS else if s2 = full_globalsN then GREATER else string_ord (n',m') end | compare (t1,t2) = raise TERM ("FunSplitState.sort_vars.compare",[t1,t2]); in sort (rev_order o compare) vars end; fun split_state ctxt s _ t = let val vars = collect_vars s t; val vars' = if Config.get ctxt sort_variables then sort_vars ctxt vars else vars; in (abstract_vars vars' s t,rev vars') end; -fun abs_var _ t = (remdeco' (sel_name t), sel_type t); +fun abs_var ctxt t = (remdeco' ctxt (sel_name t), sel_type t); (* Proof for: EX x_1 ... x_n. P x_1 ... x_n * ==> EX s. P (lookup destr_1 "x_1" s) ... (lookup destr_n "x_n" s) * Implementation: * 1. Eliminate existential quantifiers in premise * 2. Instantiate s with: (%x. undefined)("x_1" := constr_1 x_1, ..., "x_n" := constr_n x_n) * 3. Simplify *) local val ss = simpset_of (put_simpset (simpset_of @{theory_context Fun}) @{context} addsimps (@{thm StateFun.lookup_def} :: @{thm StateFun.id_id_cancel} :: @{thms list.inject list.distinct char.inject cong_exp_iff_simps simp_thms}) addsimprocs [Record.simproc, StateFun.lazy_conj_simproc] |> fold Simplifier.add_cong @{thms block_conj_cong}); in -fun ex_tac ctxt vs st = +fun ex_tac ctxt vs = SUBGOAL (fn (goal, i) => let val vs' = rev vs; - val (Const (_,exT)$_) = HOLogic.dest_Trueprop - (Logic.strip_imp_concl (Logic.get_goal (Thm.prop_of st) 1)); + val (Const (_,exT)$_) = HOLogic.dest_Trueprop (Logic.strip_imp_concl goal); val sT = domain_type (domain_type exT); val s0 = Const (@{const_name HOL.undefined},sT); fun streq (s1:string,s2) = s1=s2 ; fun mk_init [] = [] | mk_init (t::ts) = let val xs = mk_init ts; val n = component_name t; val T = component_type t; in if AList.defined streq xs n then xs else (n,(T,Const (n,sT --> component_type t)$s0))::xs end; fun mk_upd (i,t) xs = let val selN = component_name t; val selT = component_type t; val (_,s) = the (AList.lookup streq xs selN); val strT = domain_type selT; val valT = range_type selT; val constr = destr_to_constr (sel_destr t); val name = (sel_raw_name t); val upd = Const (@{const_name Fun.fun_upd}, (strT --> valT)-->strT-->valT--> (strT --> valT)) $ s $ name $ (constr $ Bound i) in AList.update streq (selN,(selT,upd)) xs end; val upds = fold_index mk_upd vs' (mk_init vs'); val upd = fold (fn (n,(T,upd)) => fn s => Const (n ^ Record.updateN, T --> sT --> sT)$upd$s) upds s0; val inst = fold_rev (Term.abs o (fn t => (sel_name t, sel_type t))) vs upd; fun lift_inst_ex_tac i st = let val rule = Thm.lift_rule (Thm.cprem_of st i) (Drule.incr_indexes st exI); val (_$x) = HOLogic.dest_Trueprop (Logic.strip_assums_concl (hd (Thm.prems_of rule))); val inst_rule = infer_instantiate ctxt [(#1 (dest_Var (head_of x)), Thm.cterm_of ctxt inst)] rule; in (compose_tac ctxt (false,inst_rule, Thm.nprems_of exI) i st) end; - in EVERY [REPEAT_DETERM_N (length vs) (eresolve_tac ctxt [exE] 1), - lift_inst_ex_tac 1, - simp_tac (put_simpset ss ctxt) 1 - ] st end + in EVERY + [REPEAT_DETERM_N (length vs) (eresolve_tac ctxt [exE] i), + lift_inst_ex_tac i, + simp_tac (put_simpset ss ctxt) i + ] + end +) + end (* Test: What happens when there are no lookups., EX s. True *) end; -structure GeneraliseRecord = GeneraliseFun (structure SplitState=RecordSplitState); -structure GeneraliseStateFun = GeneraliseFun (structure SplitState=FunSplitState); - -fun generalise Record = GeneraliseRecord.GENERALISE - | generalise Function = GeneraliseStateFun.GENERALISE; +structure GeneraliseRecord = Generalise (structure SplitState=RecordSplitState); +structure GeneraliseStateFun = Generalise (structure SplitState=FunSplitState); + +fun generalise _ Record = GeneraliseRecord.GENERALISE + | generalise _ Function = GeneraliseStateFun.GENERALISE + | generalise ctxt (Other i) = the (generalise_other ctxt i); (*****************************************************************************) (** record_vanish_tac splits up the records of a verification condition, **) (** trying to generate a predicate without records. **) (** A typical verification condition with a procedure call will have the **) (** form "!!s Z. s=Z ==> ..., where s and Z are records **) (*****************************************************************************) -(* FIXME: Check out if removing the useless vars is a performance issue. +(* fixme: Check out if removing the useless vars is a performance issue. If so, maybe we can remove all useless vars at once (no iterated simplification) or try to avoid introducing them. Bevore splitting the gaol we can simplifiy the goal with state_simproc this may leed to better performance... *) fun record_vanish_tac ctxt state_kind state_space i = if Config.get ctxt record_vanish then let val rem_useless_vars_simps = [Drule.triv_forall_equality,@{thm Hoare.triv_All_eq},@{thm Hoare.triv_Ex_eq}]; val rem_useless_vars_simpset = empty_simpset ctxt addsimps rem_useless_vars_simps; fun no_spec (t as (Const (@{const_name All},_)$_)) = is_none (try dest_hoare_raw (strip_qnt_body @{const_name All} t)) | no_spec _ = true; fun state_space_no_spec t = if state_space t <> 0 andalso no_spec t then ~1 else 0; + val state_split_tac = state_split_simp_tac ctxt rem_useless_vars_simps state_space_no_spec i + fun generalise_tac split_record st = + DETERM (generalise ctxt state_kind ctxt i) st + handle (exn as (TERM _)) => + let + val _ = warning ("record_vanish_tac: generalise subgoal " ^ string_of_int i ^ + " failed" ^ (if split_record then ", fallback to record split:\n " else "") ^ + Runtime.exn_message (Runtime.exn_context (SOME ctxt) exn)); + in + if split_record then + EVERY [ + state_split_tac, + full_simp_tac (ctxt + addsimprocs (state_simprocs ctxt state_kind @ + state_upd_simprocs ctxt state_kind) + addsimps (Named_Theorems.get ctxt @{named_theorems "state_simp"})) i, + trace_subgoal_tac ctxt "after record split and simp" i, + generalise_tac false, + trace_subgoal_tac ctxt "after 'generalise_tac false'" i + ] st + else all_tac st + end; + in EVERY [trace_tac ctxt "record_vanish_tac -- START --", REPEAT (eresolve_tac ctxt [conjE] i), trace_tac ctxt "record_vanish_tac -- hyp_subst_tac ctxt --", TRY (hyp_subst_tac_thin true ctxt i), full_simp_tac rem_useless_vars_simpset i, (* hyp_subst_tac may have made some state variables unnecessary. We do not want to split them to avoid naming conflicts and increase performance *) trace_tac ctxt "record_vanish_tac -- Splitting records --", if Config.get ctxt use_generalise orelse state_kind = Function - then generalise state_kind ctxt i - else state_split_simp_tac ctxt rem_useless_vars_simps state_space_no_spec i - (*THEN_MAYBE - EVERY [trace_tac ctxt "record_vanish_tac -- removing useless vars --", - full_simp_tac rem_useless_vars_simpset i, - trace_tac ctxt "record_vanish_tac -- STOP --"]*) + then EVERY [generalise_tac true] + else state_split_tac ] end else all_tac; (* solve_modifies_tac tries to solve modifies-clauses automatically; * The following strategy is followed: * After clar-simplifying the modifies clause we remain with a goal of the form * * EX a b. s(|A := x|) = s(|A:=a,B:=b|) * * (or because of conditional statements conjunctions of these kind of goals) * We split up the state-records and simplify (record_vanish_tac) and get to a goal of the form: * * EX a b. (|A=x,B=B|) = (|A=a,B=b|). * * If the modifies clause was correct we just have to introduce the existential quantifies * and apply reflexivity. * If not we just simplify the goal. *) local val state_fun_update_ss = simpset_of (put_simpset HOL_basic_ss @{context} addsimps ([@{thm StateFun.update_def}, @{thm id_def}, @{thm fun_upd_apply}, @{thm if_True}, @{thm if_False}] @ @{thms list.inject list.distinct char.inject cong_exp_iff_simps simp_thms} @ K_fun_convs) addsimprocs [DistinctTreeProver.distinct_simproc ["distinct_fields", "distinct_fields_globals"]] - |> Simplifier.add_cong @{thm imp_cong} (* K_fun_congs FIXME: Stefan fragen*) + |> Simplifier.add_cong @{thm imp_cong} |> Splitter.add_split @{thm if_split}); in fun solve_modifies_tac ctxt state_kind state_space i st = let val thy = Proof_Context.theory_of ctxt; fun is_split_state (trm as (Const (@{const_name Pure.all},_)$Abs(x,T,t))) = if state_space trm <> 0 then try (fn () => let fun seed (_ $ v $ st) = seed st | seed (_ $ t) = (1,t) (* split only state pair *) | seed t = (~1,t) (* split globals completely *) val all_vars = strip_qnt_vars @{const_name Pure.all} t; val concl = HOLogic.dest_Trueprop (Logic.strip_assums_concl t); val ex_vars = strip_qnt_vars @{const_name Ex} concl; val state = Bound (length all_vars + length ex_vars); val (Const (@{const_name HOL.eq},_)$x_upd$x_upd') = strip_qnt_body @{const_name Ex} concl; val (split,sd) = seed x_upd; in if sd = state then split else 0 end) () |> the_default 0 else 0 | is_split_state t = 0; val simp_ctxt = put_simpset HOL_ss ctxt addsimps (@{thm Ex_True} :: @{thm Ex_False} :: Record.get_extinjects thy); - fun try_solve Record i = (*(SOLVE*) - (((fn k => (TRY (REPEAT_ALL_NEW (resolve_tac ctxt [conjI, impI, allI]) k))) - THEN_ALL_NEW - (fn k => EVERY [state_split_simp_tac ctxt [] is_split_state k, - simp_tac simp_ctxt k THEN_MAYBE rename_goal ctxt remdeco' k - ])) i) (*)*) - | try_solve _ i = + fun try_solve Function i = ((fn k => (TRY (REPEAT_ALL_NEW (resolve_tac ctxt [conjI, impI, allI]) k))) THEN_ALL_NEW (fn k => REPEAT (resolve_tac ctxt [exI] k) THEN resolve_tac ctxt [ext] k THEN simp_tac (put_simpset state_fun_update_ss ctxt) k THEN_MAYBE (REPEAT_ALL_NEW (resolve_tac ctxt [conjI,impI,refl]) k))) i + | try_solve _ i = (*(SOLVE*) (* Record and Others *) + (((fn k => (TRY (REPEAT_ALL_NEW (resolve_tac ctxt [conjI, impI, allI]) k))) + THEN_ALL_NEW + (fn k => EVERY [state_split_simp_tac ctxt [] is_split_state k, + simp_tac simp_ctxt k THEN_MAYBE rename_goal ctxt (remdeco' ctxt) k + ])) i) (*)*) in ((trace_tac ctxt "solve_modifies_tac" THEN clarsimp_tac ((ctxt |> put_claset (claset_of @{theory_context HOL}) |> put_simpset (simpset_of @{theory_context Set})) - addsimps ([@{thm Hoare.mex_def},@{thm Hoare.meq_def}]@K_convs) - addsimprocs (state_upd_simproc Record::(state_simprocs Record)) + addsimps (@{thms Hoare.mex_def Hoare.meq_def} @K_convs@(Named_Theorems.get ctxt @{named_theorems "state_simp"})) + addsimprocs (state_upd_simprocs ctxt Record @ state_simprocs ctxt Record) |> fold Simplifier.add_cong K_congs) i) THEN_MAYBE - try_solve state_kind i) st + (try_solve state_kind i)) st end; end fun proc_lookup_simp_tac ctxt i st = try (fn () => let val name = (Logic.concl_of_goal (Thm.prop_of st) i) |> dest_hoare |> #2 |> strip_comb |> #2 |> last |> strip_comb |> #2 |> last; (* the$(Gamma$name) or the$(strip$Gamma$name) *) - val pname = (unsuffix proc_deco (dest_string' name)); + val pname = chopsfx proc_deco (dest_string' name); val thms = map_filter I (map (try (Proof_Context.get_thm ctxt)) [suffix bodyP pname, suffix (body_def_sfx^"_def") pname, suffix procL pname^"."^ (suffix (body_def_sfx^"_def") pname)]); in simp_tac (put_simpset HOL_basic_ss ctxt addsimps thms @ strip_simps @ @{thms option.sel}) i st end) () |> the_default (Seq.single st); fun proc_lookup_in_dom_simp_tac ctxt i st = try (fn () => let val _$name$_ = (HOLogic.dest_Trueprop (Logic.concl_of_goal (Thm.prop_of st) i)); (* name : Gamma *) - val pname = (unsuffix proc_deco (dest_string' name)); + val pname = chopsfx proc_deco (dest_string' name); val thms = map_filter I (map (try (Proof_Context.get_thm ctxt)) - [suffix bodyP pname]); + [suffix bodyP pname, suffix "_def" pname]); in simp_tac (put_simpset HOL_basic_ss ctxt addsimps (@{thm Hoare.lookup_Some_in_dom} :: @{thm dom_strip} :: thms)) i st end) () |> the_default (Seq.single st); fun HoareRuleTac ctxt insts fixes st = let val annotate_simp_tac = simp_tac (put_simpset HOL_basic_ss ctxt addsimps (anno_defs@normalize_simps) addsimprocs [@{simproc case_prod_beta}]); fun is_com_eq (Const (@{const_name Trueprop},_)$(Const (@{const_name HOL.eq},T)$_$_)) = (case (binder_types T) of (Type (@{type_name Language.com},_)::_) => true | _ => false) | is_com_eq _ = false; fun annotate_tac i st = if is_com_eq (Logic.concl_of_goal (Thm.prop_of st) i) then annotate_simp_tac i st else all_tac st; in ((fn i => REPEAT (resolve_tac ctxt [allI] i)) THEN' Rule_Insts.res_inst_tac ctxt insts fixes st) THEN_ALL_NEW annotate_tac end; fun HoareCallRuleTac state_kind state_space ctxt thms i st = let fun dest_All (Const (@{const_name All},_)$t) = SOME t | dest_All _ = NONE; fun auxvars t = (case (map_filter ((first_subterm is_hoare) o snd) (max_subterms_dest dest_All t)) of ((vars,_)::_) => vars | _ => []); fun auxtype rule = (case (auxvars (Thm.prop_of rule)) of [] => NONE | vs => (case (last vs) of (_,TVar (z,_)) => SOME (z,rule) | _ => NONE)); val thms' = let val auxvs = map fst (auxvars (Logic.concl_of_goal (Thm.prop_of st) i)); val tvar_thms = map_filter auxtype thms in if length thms = length tvar_thms then adapt_aux_var ctxt true auxvs tvar_thms else thms end; val is_sidecondition = not o can dest_hoare; fun solve_sidecondition_tac (t,i) = if is_sidecondition t then FIRST [CHANGED_PROP (wf_tac ctxt i), (*init_conforms_tac state_kind state_space i,*) post_conforms_tac ctxt state_kind i THEN_MAYBE (if is_modifies_clause t then solve_modifies_tac ctxt state_kind state_space i else all_tac), proc_lookup_in_dom_simp_tac ctxt i ] else in_rel_simp ctxt i THEN simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm Un_empty_left},@{thm Un_empty_right}]) i THEN proc_lookup_simp_tac ctxt i fun basic_tac i = (((resolve_tac ctxt thms') THEN_ALL_NEW - (fn k => (SUBGOAL solve_sidecondition_tac k))) i) + (fn k => + (SUBGOAL solve_sidecondition_tac k))) i) in (basic_tac ORELSE' (fn k => (REPEAT (resolve_tac ctxt [allI] k)) THEN EVERY [resolve_tac ctxt thms' k])) i st end; (* vcg_polish_tac tries to solve modifies-clauses automatically; for other specifications the * records are only splitted and simplified. *) fun vcg_polish_tac solve_modifies ctxt state_kind state_space i = if solve_modifies then solve_modifies_tac ctxt state_kind state_space i else record_vanish_tac ctxt state_kind state_space i - THEN_MAYBE EVERY [rename_goal ctxt remdeco' i(*, + THEN_MAYBE EVERY [rename_goal ctxt (remdeco' ctxt) i(*, simp_tac (HOL_basic_ss addsimps @{thms simp_thms})) i*)]; fun is_funtype (Type ("fun", _)) = true | is_funtype _ = false; +fun get_state_kind_extension ctxt T = + let + val sps = #state_spaces (Hoare_Data.get (Context.Proof ctxt)) + in + case find_first (fn (n, sp) => (#is_state_type sp) ctxt T) sps of + SOME (n, sp) => SOME n + | NONE => NONE + end + fun state_kind_of ctxt T = let val thy = Proof_Context.theory_of ctxt; val (s,sT) = nth (fst (Record.get_recT_fields thy T)) 1; - in if Long_Name.base_name s = "locals" andalso is_funtype sT then Function else Record end + in + case get_state_kind_extension ctxt T of + SOME n => Other n + | _ => if Long_Name.base_name s = "locals" andalso is_funtype sT then + Function + else + Record + end handle Subscript => Record; fun find_state_space_in_triple ctxt t = try (fn () => (case first_subterm is_hoare t of NONE => NONE | SOME (abs_vars,triple) => let val (_,com,_,_,mode,_,_,_) = dest_hoare_raw triple; val T = fastype_of1 (map snd abs_vars,com) val Type(_,state_spaceT::_) = T; val SOME Tids = stateT_ids state_spaceT; in SOME (Tids,mode, state_kind_of ctxt state_spaceT) end)) () |> Option.join; fun get_state_space_in_subset_eq ctxt t = (* get state type from the following kind of terms: P <= Q, s: P *) try (fn () => let val (subset_eq,_) = (strip_comb o HOLogic.dest_Trueprop o strip_qnt_body @{const_name Pure.all}) t; val Ts = map snd (strip_vars t); val T = fastype_of1 (Ts,subset_eq); val Type (_, [_,Type (_, [Type (_, [state_spaceT]), _])]) = T; (* also works for "in": x : P *) val SOME Tids = stateT_ids state_spaceT; in (Tids,Partial, state_kind_of ctxt state_spaceT) end) (); fun get_state_space ctxt i st = (case try (Logic.concl_of_goal (Thm.prop_of st)) i of SOME t => (case find_state_space_in_triple ctxt t of SOME sp => SOME sp | NONE => get_state_space_in_subset_eq ctxt t) | NONE => NONE); fun mk_hoare_tac hoare_tac finish_tac annotate_inv exnames strip_guards spec_sfx ctxt i st = case get_state_space ctxt i st of SOME (Tids,mode,kind) => SELECT_GOAL (hoare_tac annotate_inv exnames strip_guards mode kind (is_state_space_var Tids) spec_sfx ctxt (finish_tac kind (is_state_space_var Tids))) i st | NONE => no_tac st fun vcg_tac spec_sfx strip_guards exnames ctxt i st = mk_hoare_tac HoareTac (vcg_polish_tac (spec_sfx="_modifies") ctxt) (spec_sfx="_modifies") exnames (strip_guards="true") spec_sfx ctxt i st; fun hoare_tac spec_sfx strip_guards _ ctxt i st = let fun tac state_kind state_space i = if spec_sfx="_modifies" then solve_modifies_tac ctxt state_kind state_space i else all_tac; in mk_hoare_tac HoareTac tac (spec_sfx="_modifies") [] (strip_guards="true") spec_sfx ctxt i st end; fun hoare_raw_tac spec_sfx strip_guards exnames ctxt i st = mk_hoare_tac HoareTac (K (K (K all_tac))) (spec_sfx="_modifies") [] (strip_guards="true") spec_sfx ctxt i st; fun hoare_step_tac spec_sfx strip_guards exnames ctxt i st = mk_hoare_tac (K (K HoareStepTac)) (vcg_polish_tac (spec_sfx="_modifies") ctxt) false [] (strip_guards="true") spec_sfx ctxt i st; fun hoare_rule_tac ctxt thms i st = SUBGOAL (fn _ => (case get_state_space ctxt i st of SOME (Tids,_,kind) => HoareCallRuleTac kind (is_state_space_var Tids) ctxt thms i | NONE => error "could not find proper state space type (structure or record) in goal")) i st; (*** Methods ***) val hoare_rule = Rule_Insts.method HoareRuleTac hoare_rule_tac; val argP = Args.name --| @{keyword "="} -- Args.name val argsP = Scan.repeat argP val default_args = [("spec","spec"),("strip_guards","false")] val vcg_simp_modifiers = [Args.add -- Args.colon >> K (Method.modifier vcg_simp_add \<^here>), Args.del -- Args.colon >> K (Method.modifier vcg_simp_del \<^here>)]; fun assocs2 key = map snd o filter (curry (op =) key o fst); fun gen_simp_method tac = Scan.lift (argsP >> (fn args => args @ default_args)) --| Method.sections vcg_simp_modifiers >> (fn args => fn ctxt => Method.SIMPLE_METHOD' (tac ("_" ^ the (AList.lookup (op =) args "spec")) (the (AList.lookup (op =) args "strip_guards")) (assocs2 "exspec" args) ctxt)); val hoare = gen_simp_method hoare_tac; val hoare_raw = gen_simp_method hoare_raw_tac; val vcg = gen_simp_method vcg_tac; val vcg_step = gen_simp_method hoare_step_tac; val trace_hoare_users = Unsynchronized.ref false -fun print_subgoal_tac ctxt s i = - SUBGOAL (fn (prem, _) => trace_tac ctxt (s ^ (Syntax.string_of_term ctxt prem))) i - fun mk_hoare_thm thm _ ctxt _ i = - EVERY [resolve_tac ctxt [thm] i, - if !trace_hoare_users then print_subgoal_tac ctxt "Tracing: " i + EVERY [resolve_tac ctxt [Thm.transfer' ctxt thm] i, + if !trace_hoare_users then trace_subgoal_tac ctxt "Tracing: " i else all_tac] val vcg_hoare_add = - Thm.declaration_attribute (fn thm => add_hoare_tacs [(Thm.derivation_name thm, mk_hoare_thm thm)]) + let + fun get_name thm = + case Properties.get (Thm.get_tags thm) Markup.nameN of + SOME n => n + | NONE => error ("theorem with attribute 'vcg_hoare' must have a name") + in + Thm.declaration_attribute (fn thm => + add_hoare_tacs [(get_name thm, mk_hoare_thm (Thm.trim_context thm))]) + end exception UNDEF val vcg_hoare_del = Thm.declaration_attribute (fn _ => fn _ => raise UNDEF) (* setup theory *) val _ = Theory.setup (Attrib.setup @{binding vcg_simp} (Attrib.add_del vcg_simp_add vcg_simp_del) "declaration of Simplifier rule for vcg" #> Attrib.setup @{binding vcg_hoare} (Attrib.add_del vcg_hoare_add vcg_hoare_del) "declaration of wp rule for vcg") - (*#> add_wp_tacs initial_wp_tacs*) end; diff --git a/thys/Simpl/hoare_syntax.ML b/thys/Simpl/hoare_syntax.ML --- a/thys/Simpl/hoare_syntax.ML +++ b/thys/Simpl/hoare_syntax.ML @@ -1,1625 +1,1743 @@ (* Title: hoare_syntax.ML Author: Norbert Schirmer, TU Muenchen Copyright (C) 2004-2007 Norbert Schirmer - -This library is free software; you can redistribute it and/or modify -it under the terms of the GNU Lesser General Public License as -published by the Free Software Foundation; either version 2.1 of the -License, or (at your option) any later version. - -This library is distributed in the hope that it will be useful, but -WITHOUT ANY WARRANTY; without even the implied warranty of -MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU -Lesser General Public License for more details. - -You should have received a copy of the GNU Lesser General Public -License along with this library; if not, write to the Free Software -Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307 -USA +Copyright (c) 2022 Apple Inc. All rights reserved. *) -(* FIXME: Adapt guard generation to new syntax of op + etc. *) +(* fixme: Adapt guard generation to new syntax of op + etc. *) signature HOARE_SYNTAX = sig val antiquoteCur: string val antiquoteOld: string val antiquoteOld_tr: Proof.context -> term list -> term val antiquote_applied_only_to: (term -> bool) -> term -> bool val antiquote_varname_tr: string -> term list -> term val app_quote_tr': Proof.context -> term -> term list -> term val assert_tr': Proof.context -> term list -> term val assign_tr': Proof.context -> term list -> term val assign_tr: Proof.context -> term list -> term val basic_assigns_tr: Proof.context -> term list -> term val basic_tr': Proof.context -> term list -> term val basic_tr: Proof.context -> term list -> term val bexp_tr': string -> Proof.context -> term list -> term val bind_tr': Proof.context -> term list -> term - val call_ass_tr: bool -> bool -> Proof.context -> term list -> term + val call_ass_tr: bool -> bool -> term list -> Proof.context -> term list -> term val call_tr': Proof.context -> term list -> term - val call_tr: bool -> bool -> Proof.context -> term list -> term + val call_exn_tr': Proof.context -> term list -> term + val call_tr: bool -> bool -> term list -> Proof.context -> term list -> term val dyn_call_tr': Proof.context -> term list -> term + val dyn_call_exn_tr': Proof.context -> term list -> term val fcall_tr': Proof.context -> term list -> term val fcall_tr: Proof.context -> term list -> term val guarded_Assign_tr: Proof.context -> term list -> term val guarded_Cond_tr: Proof.context -> term list -> term val guarded_NNew_tr: Proof.context -> term list -> term val guarded_New_tr: Proof.context -> term list -> term val guarded_WhileFix_tr: Proof.context -> term list -> term val guarded_While_tr: Proof.context -> term list -> term val guards_tr': Proof.context -> term list -> term val hide_guards: bool Config.T val init_tr': Proof.context -> term list -> term val init_tr: Proof.context -> term list -> term val loc_tr': Proof.context -> term list -> term val loc_tr: Proof.context -> term list -> term val new_tr : Proof.context -> term list -> term val new_tr': Proof.context -> term list -> term val nnew_tr : Proof.context -> term list -> term val nnew_tr': Proof.context -> term list -> term val proc_ass_tr: Proof.context -> term list -> term val proc_tr': Proof.context -> term list -> term val proc_tr: Proof.context -> term list -> term val quote_mult_tr': Proof.context -> (term -> bool) -> string -> string -> term -> term val quote_tr': Proof.context -> string -> term -> term val quote_tr: Proof.context -> string -> term -> term val raise_tr': Proof.context -> term list -> term val raise_tr: Proof.context -> term list -> term val switch_tr': Proof.context -> term list -> term - val update_comp: Proof.context -> string list -> bool -> bool -> xstring -> term -> term -> term + val update_comp: Proof.context -> Hoare.lense option -> string list -> bool -> bool -> xstring -> term -> term -> term val use_call_tr': bool Config.T val whileAnnoGFix_tr': Proof.context -> term list -> term val whileAnnoG_tr': Proof.context -> term list -> term end; structure Hoare_Syntax: HOARE_SYNTAX = struct val use_call_tr' = Attrib.setup_config_bool @{binding hoare_use_call_tr'} (K true); val hide_guards = Attrib.setup_config_bool @{binding hoare_hide_guards} (K false); val globalsN = "globals"; val localsN = "locals"; val globals_updateN = suffix Record.updateN globalsN; val locals_updateN = suffix Record.updateN localsN; -val upd_globalsN = "upd_globals"; (* FIXME authentic syntax !? *) +val upd_globalsN = "upd_globals"; (* fixme authentic syntax !? *) val allocN = "alloc_'"; val freeN = "free_'"; -val Null = Syntax.free "Simpl_Heap.Null"; (* FIXME ?? *) +val Null = Syntax.free "Simpl_Heap.Null"; (* fixme ?? *) (** utils **) (* transpose [[a,b],[c,d],[e,f]] = [[a,c,d],[b,d,f]] *) fun transpose [[]] = [[]] | transpose ([]::xs) = [] | transpose ((y::ys)::xs) = (y::map hd xs)::transpose (ys::map tl xs) fun maxprefix eq ([], ys) = [] | maxprefix eq (xs, []) = [] | maxprefix eq ((x::xs),(y::ys)) = if eq (x,y) then x::maxprefix eq (xs,ys) else [] fun maxprefixs eq [] = [] | maxprefixs eq [[]] = [] | maxprefixs eq xss = foldr1 (maxprefix eq) xss; fun mk_list [] = Syntax.const @{const_syntax Nil} | mk_list (x::xs) = Syntax.const @{const_syntax Cons} $ x $ mk_list xs; (* convert Fail to Match, useful for print translations *) fun unsuffix' sfx a = unsuffix sfx a handle Fail _ => raise Match; fun unsuffixI sfx a = unsuffix sfx a handle Fail _ => a; fun is_prefix_or_suffix s t = can (unprefix s) t orelse can (unsuffix s) t; +fun is_other ctxt = case Hoare.get_default_state_kind ctxt of Hoare.Other _ => true | _ => false + (** hoare data **) fun is_global_comp ctxt name = - (case StateSpace.get_comp (Context.Proof ctxt) name of - SOME (_, ln) => is_prefix_or_suffix "globals" (Long_Name.base_name ln) - | NONE => false); - + let + val res = case Hoare.get_default_state_space ctxt of + SOME {is_defined, ...} => not (is_defined ctxt name) + | _ => + (case StateSpace.get_comp' (Context.Proof ctxt) name of + SOME (_, lns) => forall (fn ln => is_prefix_or_suffix "globals" (Long_Name.base_name ln)) lns + | NONE => false) + in + res + end (** parsing and printing **) (* quote/antiquote *) val antiquoteCur = @{syntax_const "_antiquoteCur"}; val antiquoteOld = @{syntax_const "_antiquoteOld"}; fun intern_const_syntax consts = Consts.intern_syntax consts #> perhaps Long_Name.dest_hidden; fun is_global ctxt name = let val thy = Proof_Context.theory_of ctxt; val consts = Proof_Context.consts_of ctxt; in (case Sign.const_type thy (intern_const_syntax consts name) of NONE => is_global_comp ctxt name | SOME T => String.isPrefix globalsN (Long_Name.base_name (fst (dest_Type (domain_type T))))) handle Match => false end; exception UNDEFINED of string -(* FIXME: if is_state_var etc. is reimplemented, rethink about when adding the deco to +(* fixme: if is_state_var etc. is reimplemented, rethink about when adding the deco to the records *) fun first_successful_tr _ [] = raise TERM ("first_successful_tr: no success",[]) | first_successful_tr f [x] = f x | first_successful_tr f (x::xs) = f x handle TERM _ => first_successful_tr f xs; fun statespace_lookup_tr ctxt ps s n = - let - val cn = map Hoare.clique_name (#active_procs (Hoare.get_data ctxt)); - val procs = ps @ cn; - val names = - n :: map (fn p => (suffix (Hoare.par_deco p) (unsuffixI Hoare.deco n))) procs; - in first_successful_tr (StateSpace.gen_lookup_tr ctxt s) names - end; + case Hoare.get_default_state_space ctxt of + SOME {lookup_tr, ...} => lookup_tr ctxt n $ s + | _ => + let + val cn = map Hoare.clique_name (#active_procs (Hoare.get_data ctxt)); + val procs = ps @ cn; + val names = + (Hoare.name_tr ctxt true n) :: map (fn p => (suffix (Hoare.par_deco p) (unsuffixI Hoare.deco n))) procs; + in + first_successful_tr (StateSpace.gen_lookup_tr ctxt s) names + end -fun statespace_update_tr ctxt ps id n v s = - let - val cn = map Hoare.clique_name (#active_procs (Hoare.get_data ctxt)); - val procs = ps @ cn; - val names = - n :: map (fn p => (suffix (Hoare.par_deco p) (unsuffixI Hoare.deco n))) procs; - in first_successful_tr (fn n => StateSpace.gen_update_tr id ctxt n v s) names - end; +fun K_rec_syntax v = Abs ("_", dummyT, incr_boundvars 1 v); + +fun statespace_update_tr ctxt NONE ps id n v = + (case Hoare.get_default_state_space ctxt of + SOME {update_tr, ...} => update_tr ctxt n $ K_rec_syntax v + | _ => + let + fun gen_update_tr id ctxt n v = + StateSpace.gen'_update_tr true id ctxt n v (Bound 0) |> dest_comb |> fst + + val cn = map Hoare.clique_name (#active_procs (Hoare.get_data ctxt)); + val procs = ps @ cn; + val names = + (Hoare.name_tr ctxt true n) :: map (fn p => (suffix (Hoare.par_deco p) (unsuffixI Hoare.deco n))) procs; + in first_successful_tr (fn n => gen_update_tr id ctxt n v) names + end) + | statespace_update_tr ctxt (SOME {lookup, update}) ps id n v = + update $ K_rec_syntax v + local fun is_record_sel ctxt nm = let - val consts = Proof_Context.consts_of ctxt; - val exists_const = can (Consts.the_const consts) o intern_const_syntax consts; - val exists_abbrev = can (Consts.the_abbreviation consts) o intern_const_syntax consts; - in (exists_const nm) orelse (exists_abbrev nm) end; + val SOME (Const (c, T)) = try (Syntax.read_term ctxt) nm + val recT = domain_type T + val (flds, _) = Record.get_recT_fields (Proof_Context.theory_of ctxt) recT + in member (op =) (map fst flds) c end + handle TYPE _ => false + | Bind => false + | Match => false in fun lookup_comp ctxt ps name = if is_record_sel ctxt name then if is_global ctxt name then (fn s => Syntax.free name $ (Syntax.free "globals" $ s)) else (fn s => Syntax.free name $ s) else let val sel = Syntax.const (if is_global_comp ctxt name then "globals" else "locals"); in (fn s => statespace_lookup_tr ctxt ps (sel $ s) name) end; (* -FIXME: +fixme: update of global and local components: One should generally provide functions: glob_upd:: ('g => 'g) => 's => 's loc_upd:: ('l => 'l) => 's => 's so that global and local updates can nicely be composed. loc_upd for the record implementation is vacuous. Right now for example an assignment of NEW to a global variable returns a funny repeated update of global components... This would make the composition more straightforward... Basically one wants the map on a component rather then the update. Maps can be composed nicer... *) -fun K_rec_syntax v = Abs ("_", dummyT, incr_boundvars 1 v); -fun update_comp ctxt ps atomic id name value = + +fun update_comp ctxt lense_opt ps atomic id name value = if is_record_sel ctxt name then let val upd = Syntax.free (suffix Record.updateN name) $ K_rec_syntax value; in if atomic andalso is_global ctxt name then (fn s => - Syntax.free globals_updateN $ (K_rec_syntax (upd $ (Syntax.free globalsN $ s))) $ s) + Syntax.free globals_updateN $ (*(K_rec_syntax*) upd $ s) else (fn s => upd $ s) end else let val reg = if is_global_comp ctxt name then "globals" else "locals"; val upd = Syntax.free (reg ^ Record.updateN); val sel = Syntax.free reg; in fn s => if atomic then - upd $ (K_rec_syntax (statespace_update_tr ctxt ps id name value (sel $ s))) $ s - else statespace_update_tr ctxt ps id name value s + upd $ statespace_update_tr ctxt lense_opt ps id name value $ s + else statespace_update_tr ctxt lense_opt ps id name value $ s end; end; fun antiquote_global_tr ctxt off i t = let fun mk n t = lookup_comp ctxt [] n (Bound (i + off n)); (*if is_global ctxt n then t$(Free ("globals",dummyT)$Bound (i + off n)) else t$Bound (i + off n)*) in (case t of Free (n, _) => mk n t | Const (n, _) => mk n t | _ => t $ Bound i) end; fun antiquote_off_tr offset ctxt name = let fun tr i ((t as Const (c, _)) $ u) = if c = name then antiquote_global_tr ctxt offset i (tr i u) else tr i t $ tr i u | tr i (t $ u) = tr i t $ tr i u | tr i (Abs (x, T, t)) = Abs (x, T, tr (i + 1) t) | tr _ a = a; in tr 0 end; val antiquote_tr = antiquote_off_tr (K 0) fun quote_tr ctxt name t = Abs ("s", dummyT, antiquote_tr ctxt name (Term.incr_boundvars 1 t)); fun antiquoteCur_tr ctxt t = antiquote_tr ctxt antiquoteCur (Term.incr_boundvars 1 t); fun antiquote_varname_tr anti [n] = (case n of Free (v, T) => Syntax.const anti $ Free (Hoare.varname v, T) | Const (c, T) => Syntax.const anti $ Const (Hoare.varname c, T) | _ => Syntax.const anti $ n); fun antiquoteOld_tr ctxt [s, n] = (case n of Free (v, T) => lookup_comp ctxt [] (Hoare.varname v) s | Const (c, T) => lookup_comp ctxt [] (Hoare.varname c) s | _ => n $ s); +fun first_match [] t = raise Match + | first_match (f::fs) t = f t handle Match => first_match fs t + +fun lookup_tr' ctxt t = t |> first_match [ + fn t => + case Hoare.get_default_state_space ctxt of + SOME {lookup_tr', ...} => lookup_tr' ctxt t + | NONE => raise Match, + fn t => Hoare.undeco ctxt t] + + fun antiquote_tr' ctxt name = let fun is_state i t = (case t of Bound j => i = j | Const (g,_) $ Bound j => i = j andalso member (op =) [globalsN, localsN] (Long_Name.base_name g) | _ => false); fun tr' i (t $ u) = - if is_state i u then Syntax.const name $ tr' i (Hoare.undeco ctxt t) + if is_state i u then Syntax.const name $ tr' i (lookup_tr' ctxt t) else tr' i t $ tr' i u | tr' i (Abs (x, T, t)) = Abs (x, T, tr' (i + 1) t) | tr' i a = if a = Bound i then raise Match else a; in tr' 0 end; fun quote_tr' ctxt name (Abs (_, _, t)) = Term.incr_boundvars ~1 (antiquote_tr' ctxt name t) | quote_tr' ctxt name (t as (Const _)) (* eta contracted *) = Syntax.const name $ Hoare.undeco ctxt t | quote_tr' _ _ _ = raise Match; local fun state_test (t as Const (g,_) $ u) f = if member (op =) [localsN, globalsN] (Long_Name.base_name g) then f u else f t | state_test u f = f u; in fun antiquote_applied_only_to P = let fun test i (t $ u) = state_test u (fn Bound j => if j=i then P t else test i t andalso test i u | u => test i t andalso test i u) | test i (Abs (x, T, t)) = test (i + 1) t | test i _ = true; in test 0 end; + fun antiquote_mult_tr' ctxt is_selector current old = let fun tr' i (t $ u) = state_test u (fn Bound j => - if j = i then Syntax.const current $ tr' i (Hoare.undeco ctxt t) + if j = i then Syntax.const current $ tr' i (lookup_tr' ctxt t) else if is_selector t (* other quantified states *) - then Syntax.const old $ Bound j $ tr' i (Hoare.undeco ctxt t) + then Syntax.const old $ Bound j $ tr' i (lookup_tr' ctxt t) else tr' i t $ tr' i u | pre as ((Const (m,_) $ Free _)) (* pre state *) => if (m = @{syntax_const "_bound"} orelse m = @{syntax_const "_free"}) andalso is_selector t - then Syntax.const old $ pre $ tr' i (Hoare.undeco ctxt t) + then Syntax.const old $ pre $ tr' i (lookup_tr' ctxt t) else tr' i t $ pre | pre as ((Const (m,_) $ Var _)) (* pre state *) => if m = @{syntax_const "_var"} andalso is_selector t - then Syntax.const old $ pre $ tr' i (Hoare.undeco ctxt t) + then Syntax.const old $ pre $ tr' i (lookup_tr' ctxt t) else tr' i t $ pre | u => tr' i t $ tr' i u) | tr' i (Abs (x, T, t)) = Abs (x, T, tr' (i + 1) t) | tr' i a = if a = Bound i then raise Match else a; in tr' 0 end; end; fun quote_mult_tr' ctxt is_selector current old (Abs (_, _, t)) = Term.incr_boundvars ~1 (antiquote_mult_tr' ctxt is_selector current old t) | quote_mult_tr' _ _ _ _ _ = raise Match; fun app_quote_tr' ctxt f (t :: ts) = Term.list_comb (f $ quote_tr' ctxt antiquoteCur t, ts) | app_quote_tr' _ _ _ = raise Match; fun app_quote_mult_tr' ctxt is_selector f (t :: ts) = Term.list_comb (f $ quote_mult_tr' ctxt is_selector antiquoteCur antiquoteOld t, ts) | app_quote_mult_tr' _ _ _ _ = raise Match; + fun atomic_var_tr ctxt ps name value = - update_comp ctxt ps true false name value; + update_comp ctxt NONE ps true false name value; fun heap_var_tr ctxt hp p value = let fun upd s = - update_comp ctxt [] true false hp + update_comp ctxt NONE [] true false hp (Syntax.const @{const_syntax fun_upd} $ lookup_comp ctxt [] hp s $ p $ value) s; in upd end; fun get_arr_var (Const (@{const_syntax List.nth},_) $ arr $ i) = (case get_arr_var arr of SOME (name,p,is) => SOME (name,p,i::is) | NONE => NONE) | get_arr_var (Const (@{syntax_const "_antiquoteCur"},_) $ Free (var,_)) = if Hoare.is_state_var var then SOME (var,NONE,[]) else NONE | get_arr_var (Const (@{syntax_const "_antiquoteCur"},_) $ Const (var,_)) = if Hoare.is_state_var var then SOME (var,NONE,[]) else NONE | get_arr_var ((Const (@{syntax_const "_antiquoteCur"},_) $ Free (hp,_)) $ p) = if Hoare.is_state_var hp then SOME (hp,SOME p,[]) else NONE | get_arr_var ((Const (@{syntax_const "_antiquoteCur"},_) $ Const (hp,_)) $ p) = if Hoare.is_state_var hp then SOME (hp,SOME p,[]) else NONE | get_arr_var _ = NONE fun arr_var_tr ctxt ps name arr pos value idxs = let fun sel_tr [] = arr | sel_tr (i::is) = Syntax.const @{const_syntax nth} $ sel_tr is $ i; fun lupd_tr value [] _ = value | lupd_tr value (i::is) idxs = Syntax.const @{const_syntax list_update} $ sel_tr idxs $ i $ lupd_tr value is (i::idxs); val value' = lupd_tr value idxs []; in case pos of NONE => atomic_var_tr ctxt ps name value' | SOME p => heap_var_tr ctxt name p value' end; fun get_arr_mult_var (Const (@{syntax_const "_antiquoteCur"},_) $ Free (var,_)) = if Hoare.is_state_var var then SOME (var,NONE) else NONE | get_arr_mult_var (Const (@{syntax_const "_antiquoteCur"},_) $ Const (var,_)) = if Hoare.is_state_var var then SOME (var,NONE) else NONE | get_arr_mult_var ((Const (@{syntax_const "_antiquoteCur"},_) $ Free (hp,_)) $ p) = if Hoare.is_state_var hp then SOME (hp,SOME p) else NONE | get_arr_mult_var ((Const (@{syntax_const "_antiquoteCur"},_) $ Const (hp,_)) $ p) = if Hoare.is_state_var hp then SOME (hp,SOME p) else NONE | get_arr_mult_var _ = NONE fun arr_mult_var_tr ctxt ps name arr pos vals idxs = let val value' = Syntax.const @{const_syntax list_multupd} $ arr $ idxs $ vals; in case pos of NONE => atomic_var_tr ctxt ps name value' | SOME p => heap_var_tr ctxt name p value' end; fun update_tr ctxt ps off_var off_val e (v as Const (@{syntax_const "_antiquoteCur"},_) $ Free (var,_)) = if Hoare.is_state_var var then atomic_var_tr ctxt ps var e else raise TERM ("no proper lvalue", [v]) | update_tr ctxt ps off_var off_val e ((v as Const (@{syntax_const "_antiquoteCur"},_) $ Free (hp, _)) $ p) = if Hoare.is_state_var hp then heap_var_tr ctxt hp (antiquote_off_tr off_val ctxt antiquoteCur p) e else raise TERM ("no proper lvalue",[v]) | update_tr ctxt ps off_var off_val e (v as Const (@{const_syntax list_multsel}, _) $ arr $ idxs) = (case get_arr_mult_var arr of SOME (var, pos) => let val pos' = Option.map (antiquote_off_tr off_val ctxt antiquoteCur) pos; val var' = lookup_comp ctxt ps var (Bound (off_var var)); val arr' = case pos' of NONE => var' | SOME p => var' $ p; val idxs' = antiquote_off_tr off_val ctxt antiquoteCur idxs; in arr_mult_var_tr ctxt ps var arr' pos' e idxs' end | NONE => raise TERM ("no proper lvalue", [v])) | update_tr ctxt ps off_var off_val e v = (case get_arr_var v of SOME (var,pos,idxs) => let val pos' = Option.map (antiquote_off_tr off_val ctxt antiquoteCur) pos; val var' = lookup_comp ctxt ps var (Bound (off_var var)); val arr' = case pos' of NONE => var' | SOME p => var' $ p; val idxs' = rev (map (antiquote_off_tr off_val ctxt antiquoteCur) idxs); in arr_var_tr ctxt ps var arr' pos' e idxs' end | NONE => raise TERM ("no proper lvalue", [v])) | update_tr _ _ _ _ e t = raise TERM ("update_tr", [t]) fun app_assign_tr f ctxt [v, e] = let fun offset _ = 0; in f $ Abs ("s", dummyT, update_tr ctxt [] offset offset (antiquoteCur_tr ctxt e) v (Bound 0)) end | app_assign_tr _ _ ts = raise TERM ("assign_tr", ts); val assign_tr = app_assign_tr (Syntax.const @{const_syntax Basic}); val raise_tr = app_assign_tr (Syntax.const @{const_syntax raise}); fun basic_assign_tr ctxt (ts as [v, e]) = let fun offset v = 0; in update_tr ctxt [] offset offset (antiquoteCur_tr ctxt e) v end | basic_assign_tr _ ts = raise TERM ("basic_assign_tr", ts); fun basic_assigns_tr ctxt [t] = let fun dest_basic (Const (@{syntax_const "_BAssign"}, _) $ v $ e) = basic_assign_tr ctxt [v,e] | dest_basic _ = raise Match; fun dest_basics (Const (@{syntax_const "_basics"}, _) $ x $ xs) = dest_basic x :: dest_basics xs | dest_basics (t as Const (@{syntax_const "_BAssign"}, _) $_ $ _) = [dest_basic t] | dest_basics _ = [] val upds = dest_basics t; in Abs ("s", dummyT, fold (fn upd => fn s => upd s) upds (Bound 0)) end | basic_assigns_tr _ ts = raise TERM ("basic_assigns_tr", ts); fun basic_tr ctxt [t] = Syntax.const @{const_syntax Basic} $ (Abs ("s", dummyT, antiquote_tr ctxt @{syntax_const "_antiquoteCur"} (Term.incr_boundvars 1 t) $ Bound 0)); fun init_tr ctxt [Const (var,_),comp,value] = let fun dest_set (Const (@{const_syntax Set.empty}, _)) = [] | dest_set (Const (@{const_syntax insert}, _) $ x $ xs) = x :: dest_set xs; fun dest_list (Const (@{const_syntax Nil}, _)) = [] | dest_list (Const (@{const_syntax Cons}, _) $ Free (x, _) $ xs) = x :: dest_list xs; fun dest_val_list (Const (@{const_syntax Nil}, _)) = [] | dest_val_list (Const (@{const_syntax Cons},_) $ x $ xs) = dest_set x :: dest_val_list xs | dest_val_list t = [dest_set t]; val values = (case value of Const (@{const_syntax Cons}, _) $ _ $ _ => map mk_list (transpose (dest_val_list value)) | Const (@{const_syntax insert}, _) $ _ $ _ => dest_set value | _ => raise TERM ("unknown variable initialization", [])) val comps = dest_list comp; fun mk_upd var c v = Syntax.free (suffix Record.updateN (Hoare.varname (suffix ("_" ^ c) var))) $ v; val upds = map2 (mk_upd var) comps values; val app_upds = fold (fn upd => fn s => upd $ s) upds; val upd = if is_global ctxt (Hoare.varname (suffix ("_" ^ hd comps) var)) then Syntax.free (suffix Record.updateN globalsN) $ app_upds (Syntax.free globalsN $ Bound 0) $ Bound 0 else app_upds (Bound 0) in Syntax.const @{const_syntax Basic} $ Abs ("s", dummyT, upd) end | init_tr _ _ = raise Match; fun new_tr ctxt (ts as [var,size,init]) = let fun offset v = 0; fun dest_init (Const (@{syntax_const "_newinit"}, _) $ Const (var, _) $ v) = (var, v) | dest_init _ = raise Match; fun dest_inits (Const (@{syntax_const "_newinits"}, _) $ x $ xs) = dest_init x :: dest_inits xs | dest_inits (t as (Const (@{syntax_const "_newinit"}, _) $_ $ _)) = [dest_init t] | dest_inits _ = raise Match; val g = Syntax.free globalsN $ Bound 0; val alloc = lookup_comp ctxt [] allocN (Bound 0); - val new = Syntax.free "new" $ (Syntax.const @{const_syntax set} $ alloc); (* FIXME new !? *) + val new = Syntax.free "new" $ (Syntax.const @{const_syntax set} $ alloc); (* fixme new !? *) fun mk_upd (var,v) = let val varn = Hoare.varname var; val var' = lookup_comp ctxt [] varn (Bound 0); in - update_comp ctxt [] false false varn + update_comp ctxt NONE [] false false varn (Syntax.const @{const_syntax fun_upd} $ var' $ new $ v) end; val inits = map mk_upd (dest_inits init); val free = lookup_comp ctxt [] freeN (Bound 0); val freetest = Syntax.const @{const_syntax Orderings.less_eq} $ size $ free; val alloc_upd = - update_comp ctxt [] false false allocN + update_comp ctxt NONE [] false false allocN (Syntax.const @{const_syntax Cons} $ new $ alloc); val free_upd = - update_comp ctxt [] false false freeN + update_comp ctxt NONE [] false false freeN (Syntax.const @{const_syntax Groups.minus} $ free $ size); val g' = Syntax.free (suffix Record.updateN globalsN) $ K_rec_syntax (fold (fn upd => fn s => upd s) (alloc_upd :: free_upd :: inits) g) $ Bound 0; val cond = Syntax.const @{const_syntax If} $ freetest $ update_tr ctxt [] offset offset new var g' $ update_tr ctxt [] offset offset Null var (Bound 0); in Syntax.const @{const_syntax Basic} $ Abs ("s", dummyT, cond) end | new_tr _ ts = raise TERM ("new_tr",ts); fun nnew_tr ctxt (ts as [var,size,init]) = let fun offset v = 0; fun dest_init (Const (@{syntax_const "_newinit"}, _) $ Const (var, _) $ v) = (var, v) | dest_init _ = raise Match; fun dest_inits (Const (@{syntax_const "_newinits"}, _) $ x $ xs) = dest_init x :: dest_inits xs | dest_inits (t as (Const (@{syntax_const "_newinit"}, _) $ _ $ _)) = [dest_init t] | dest_inits _ = raise Match; val g = Syntax.free globalsN $ Bound 0; val alloc = lookup_comp ctxt [] allocN (Bound 0); - val new = Syntax.free "new" $ (Syntax.const @{const_syntax set} $ alloc); (* FIXME new !? *) + val new = Syntax.free "new" $ (Syntax.const @{const_syntax set} $ alloc); (* fixme new !? *) fun mk_upd (var,v) = let val varn = Hoare.varname var; val var' = lookup_comp ctxt [] varn (Bound 0); in - update_comp ctxt [] false false varn + update_comp ctxt NONE [] false false varn (Syntax.const @{const_syntax fun_upd} $ var' $ new $ v) end; val inits = map mk_upd (dest_inits init); val free = lookup_comp ctxt [] freeN (Bound 0); val freetest = Syntax.const @{const_syntax Orderings.less_eq} $ size $ free; val alloc_upd = - update_comp ctxt [] false false allocN + update_comp ctxt NONE [] false false allocN (Syntax.const @{const_syntax Cons} $ new $ alloc); val free_upd = - update_comp ctxt [] false false freeN + update_comp ctxt NONE [] false false freeN (Syntax.const @{const_syntax Groups.minus} $ free $ size); val g' = Syntax.free (suffix Record.updateN globalsN) $ K_rec_syntax (fold (fn upd => fn s => upd s) (alloc_upd :: inits @ [free_upd]) g) $ Bound 0; val cond = Syntax.const @{const_syntax if_rel} $ Abs ("s", dummyT, freetest) $ Abs ("s", dummyT, update_tr ctxt [] offset offset new var g') $ Abs ("s", dummyT, update_tr ctxt [] offset offset Null var (Bound 0)) $ Abs ("s", dummyT, update_tr ctxt [] offset offset new var g'); in Syntax.const @{const_syntax Spec} $ cond end | nnew_tr _ ts = raise TERM ("nnew_tr", ts); fun loc_tr ctxt (ts as [init, bdy]) = let fun dest_init (Const (@{syntax_const "_locinit"}, _) $ Const (var,_) $ v) = (var, v) | dest_init (Const (@{syntax_const "_locnoinit"}, _) $ Const (var, _)) = (var, Syntax.const antiquoteCur $ Syntax.free (Hoare.varname var)) - (* FIXME could skip this dummy initialisation v := v s and + (* fixme: could skip this dummy initialisation v := v s and derive non init variables in the print translation from the return function instead the init function *) | dest_init _ = raise Match; fun dest_inits (Const (@{syntax_const "_locinits"}, _) $ x $ xs) = dest_init x :: dest_inits xs | dest_inits (t as (Const (@{syntax_const "_locinit"}, _) $ _ $ _)) = [dest_init t] | dest_inits (t as (Const (@{syntax_const "_locnoinit"}, _) $ _)) = [dest_init t] | dest_inits _ = raise Match; fun mk_init_upd (var, v) = - update_comp ctxt [] true false var (antiquoteCur_tr ctxt v); + update_comp ctxt NONE [] true false var (antiquoteCur_tr ctxt v); fun mk_ret_upd var = - update_comp ctxt [] true false var (lookup_comp ctxt [] var (Bound 1)); + update_comp ctxt NONE [] true false var (lookup_comp ctxt [] var (Bound 1)); val var_vals = map (apfst Hoare.varname) (dest_inits init); val ini = Abs ("s", dummyT, fold mk_init_upd var_vals (Bound 0)); val ret = Abs ("s",dummyT, Abs ("t",dummyT, fold (mk_ret_upd o fst) var_vals (Bound 0))); val c = Abs ("i", dummyT, Abs ("t", dummyT, Syntax.const @{const_syntax Skip})); in Syntax.const @{const_syntax block} $ ini $ bdy $ ret $ c end infixr 9 &; fun (NONE & NONE) = NONE | ((SOME x) & NONE) = SOME x | (NONE & (SOME x)) = SOME x | ((SOME x) & (SOME y)) = SOME (Syntax.const @{const_syntax HOL.conj} $ x $ y); fun mk_imp (l,SOME r) = SOME (HOLogic.mk_imp (l, r)) | mk_imp (l,NONE) = NONE; local fun le l r = Syntax.const @{const_syntax Orderings.less} $ l $ r; -fun in_range t = Syntax.free "in_range" $ t; (* FIXME ?? *) +fun in_range t = Syntax.free "in_range" $ t; (* fixme ?? *) fun not_zero t = Syntax.const @{const_syntax Not} $ (Syntax.const @{const_syntax HOL.eq} $ t $ Syntax.const @{const_syntax Groups.zero}); fun not_Null t = Syntax.const @{const_syntax Not} $ - (Syntax.const @{const_syntax HOL.eq} $ t $ Syntax.free "Simpl_Heap.Null"); (* FIXME ?? *) + (Syntax.const @{const_syntax HOL.eq} $ t $ Syntax.free "Simpl_Heap.Null"); (* fixme ?? *) fun in_length i l = Syntax.const @{const_syntax Orderings.less} $ i $ (Syntax.const @{const_syntax size} $ l); fun is_pos t = Syntax.const @{const_syntax Orderings.less_eq} $ Syntax.const @{const_syntax Groups.zero} $ t; fun infer_type ctxt t = Syntax.check_term ctxt (Exn.release (#2 (Syntax_Phases.decode_term ctxt ([], Exn.Res t)))); (* NOTE: operations on actual terms *) fun is_arr (Const (@{const_name List.nth},_) $ l $ e) = is_arr l | is_arr (Const (a, _) $ Bound 0) = Hoare.is_state_var a | is_arr (Const (a,_) $ (Const (globals,_) $ Bound 0)) = Hoare.is_state_var a | is_arr ((Const (hp,_) $ (Const (globals,_) $ Bound 0)) $ p) = Hoare.is_state_var hp | is_arr _ = false; fun dummyfyT (TFree x) = TFree x | dummyfyT (TVar x) = dummyT | dummyfyT (Type (c, Ts)) = let val Ts' = map dummyfyT Ts; in if exists (fn T => T = dummyT) Ts' then dummyT else Type (c, Ts') end; fun guard ctxt Ts (add as (Const (@{const_name Groups.plus},_) $ l $ r)) = guard ctxt Ts l & guard ctxt Ts r & SOME (in_range add) | guard ctxt Ts (sub as (Const (@{const_name Groups.minus},_) $ l $ r)) = let val g = (if fastype_of1 (Ts,sub) = HOLogic.natT then le r l else in_range sub) handle TERM _ => error ("guard generation, cannot determine type of: " ^ Syntax.string_of_term ctxt sub); in guard ctxt Ts l & guard ctxt Ts r & SOME g end | guard ctxt Ts (mul as (Const (@{const_name Groups.times},_) $ l $r)) = guard ctxt Ts l & guard ctxt Ts r & SOME (in_range mul) | guard ctxt Ts (Const (@{const_name HOL.conj},_) $ l $ r) = guard ctxt Ts l & mk_imp (l,guard ctxt Ts r) | guard ctxt Ts (Const (@{const_name HOL.disj},_) $ l $ r) = guard ctxt Ts l & mk_imp (HOLogic.Not $ l,guard ctxt Ts r) | guard ctxt Ts (dv as (Const (@{const_name Rings.divide},_) $ l $ r)) = - guard ctxt Ts l & guard ctxt Ts r & SOME (not_zero r) & SOME (in_range dv) (* FIXME: Make more concrete guard...*) + guard ctxt Ts l & guard ctxt Ts r & SOME (not_zero r) & SOME (in_range dv) (* fixme: Make more concrete guard...*) | guard ctxt Ts (Const (@{const_name Rings.modulo},_) $ l $ r) = guard ctxt Ts l & guard ctxt Ts r & SOME (not_zero r) | guard ctxt Ts (un as (Const (@{const_name Groups.uminus},_) $ n)) = guard ctxt Ts n & SOME (in_range un) | guard ctxt Ts (Const (@{const_name Int.nat},_) $ i) = guard ctxt Ts i & SOME (is_pos i) | guard ctxt Ts (i as (Const (@{const_abbrev Int.int},_) $ n)) = guard ctxt Ts n & SOME (in_range i) | guard ctxt Ts (Const (@{const_name List.nth},_) $ l $ e) = if is_arr l then guard ctxt Ts l & guard ctxt Ts e & SOME (in_length e l) else NONE (*oder default?*) | guard ctxt Ts (Const (hp,_) $ (Const (globals,_) $ Bound 0) $ p) = if Hoare.is_state_var hp then guard ctxt Ts p & SOME (not_Null p)(*& SOME (in_alloc p)*) else guard ctxt Ts p (* | guard (Const (@{const_name "list_update"},_)$l$i$e) = if is_arr l then guard i & SOME (in_length i l) & guard e else NONE*) (* oder default?*) (* | guard (Const (upd,_)$e$s) = for procedure parameters,like default but left to right if is_some (try (unsuffix updateN) upd) then guard s & guard e else guard e & guard s *) | guard ctxt Ts t = fold_rev (fn l => fn r => guard ctxt Ts l & r) (snd (strip_comb t)) NONE (* default *) | guard _ _ _ = NONE; in fun mk_guard ctxt t = let (* We apply type inference first, so that we can generate different guards, depending on the type, e.g. int vs. nat *) val Abs (_, T, t') = map_types dummyfyT (infer_type ctxt (Abs ("s", dummyT, t))); in guard ctxt [T] t' end; end; -(* FIXME: make guard function a parameter of all parse-translations that need it.*) +(* fixme: make guard function a parameter of all parse-translations that need it.*) val _ = Theory.setup (Context.theory_map (Hoare.install_generate_guard mk_guard)); fun mk_singleton_guard f g = Syntax.const @{const_syntax Cons} $ (Syntax.const @{const_syntax Pair} $ Syntax.const f $ (Syntax.const @{const_syntax Collect} $ Abs ("s", dummyT, g))) $ Syntax.const @{const_syntax Nil}; fun guarded_Assign_tr ctxt (ts as [v, e]) = let val ass = assign_tr ctxt [v, e]; val guard = Hoare.generate_guard ctxt; (* By the artificial "=" between left and right-hand side we get a bigger term and thus more information for type-inference *) in case guard (Syntax.const @{const_syntax HOL.eq} $ antiquoteCur_tr ctxt v $ antiquoteCur_tr ctxt e) of NONE => ass | SOME g => Syntax.const @{const_syntax guards} $ mk_singleton_guard @{const_syntax False} g $ ass end | guarded_Assign_tr _ ts = raise Match; fun guarded_New_tr ctxt (ts as [var, size, init]) = let val new = new_tr ctxt [var, size, init]; val guard = Hoare.generate_guard ctxt; in case guard (antiquoteCur_tr ctxt var) of NONE => new | SOME g => Syntax.const @{const_syntax guards} $ mk_singleton_guard @{const_syntax False} g $ new end | guarded_New_tr _ ts = raise TERM ("guarded_New_tr", ts); fun guarded_NNew_tr ctxt (ts as [var, size, init]) = let val new = nnew_tr ctxt [var, size, init]; val guard = Hoare.generate_guard ctxt; in case guard (antiquoteCur_tr ctxt var) of NONE => new | SOME g => Syntax.const @{const_syntax guards} $ mk_singleton_guard @{const_syntax False} g $ new end | guarded_NNew_tr _ ts = raise TERM ("guarded_NNew_tr", ts); fun guarded_While_tr ctxt (ts as [b,I,V,c]) = let val guard = Hoare.generate_guard ctxt; val cnd as Abs (_, _, b') = quote_tr ctxt antiquoteCur b; val b'' = Syntax.const @{const_syntax Collect} $ cnd; in case guard b' of NONE => Syntax.const @{const_syntax whileAnno} $ b'' $ I $ V $ c | SOME g => Syntax.const @{const_syntax whileAnnoG} $ mk_singleton_guard @{const_syntax False} g $ b'' $ I $ V $ c end | guarded_While_tr _ ts = raise Match; fun guarded_WhileFix_tr ctxt (ts as [b as (_ $ Abs (_, _, b')), I, V, c]) = let val guard = Hoare.generate_guard ctxt; in case guard b' of NONE => Syntax.const @{const_syntax whileAnnoFix} $ b $ I $ V $ c | SOME g => Syntax.const @{const_syntax whileAnnoGFix} $ mk_singleton_guard @{const_syntax False} g $ b $ I $ V $ c end | guarded_WhileFix_tr _ ts = raise Match; fun guarded_Cond_tr ctxt (ts as [b, c, d]) = let val guard = Hoare.generate_guard ctxt; val cnd as Abs (_, _, b') = quote_tr ctxt @{syntax_const "_antiquoteCur"} b; val cond = Syntax.const @{const_syntax Cond} $ (Syntax.const @{const_syntax Collect} $ cnd) $ c $ d; in case guard b' of NONE => cond | SOME g => Syntax.const @{const_syntax guards} $ mk_singleton_guard @{const_syntax False} g $ cond end | guarded_Cond_tr _ ts = raise Match; (* parsing procedure calls *) fun dest_pars (Const (@{syntax_const "_par"}, _) $ p) = [p] | dest_pars (Const (@{syntax_const "_pars"}, _) $ p $ ps) = dest_pars p @ dest_pars ps | dest_pars t = raise TERM ("dest_pars", [t]); fun dest_actuals (Const (@{syntax_const "_actuals_empty"}, _)) = [] | dest_actuals (Const (@{syntax_const "_actuals"}, _) $ pars) = dest_pars pars | dest_actuals t = raise TERM ("dest_actuals", [t]); -fun mk_call_tr ctxt grd Call formals pn pt actuals has_args cont = +fun mk_call_tr ctxt grd Call formals pn pt actuals has_args result_exn cont = let val fcall = cont <> NONE; val state_kind = the_default (Hoare.get_default_state_kind ctxt) (Hoare.get_state_kind pn ctxt); - fun init_par_tr name arg = - update_comp ctxt [] false false name (antiquoteCur_tr ctxt arg); + fun init_par_tr name lense_opt arg = + update_comp ctxt lense_opt [] false false name (antiquoteCur_tr ctxt arg); fun result_par_tr name arg = let fun offset_old n = 2; fun offset n = if is_global ctxt n then 0 else 2; + val lookup = lookup_comp ctxt [] name (Bound 1) in - update_tr ctxt [pn] offset offset_old - (lookup_comp ctxt [] name (Bound 1)) arg + update_tr ctxt [pn] offset offset_old lookup arg end; val _ = if not (Config.get ctxt StateSpace.silent) andalso ((not fcall andalso length formals <> length actuals) orelse (fcall andalso length formals <> length actuals + 1)) then raise TERM ("call_tr: number of formal (" ^ string_of_int (length formals) ^ ") and actual (" ^ string_of_int (length actuals) ^ ") parameters for \"" ^ unsuffix Hoare.proc_deco pn ^ "\" do not match.", []) else (); val globals = [Syntax.const globals_updateN $ (K_rec_syntax (Const (globalsN, dummyT) $ Bound 0))]; val ret = Abs ("s", dummyT, Abs ("t", dummyT, Library.foldr (op $) (globals, Bound 1))); - val val_formals = filter (fn (kind, _) => kind = Hoare.In) formals; - val res_formals = filter (fn (kind, _) => kind = Hoare.Out) formals; + val val_formals = filter (fn (kind, _, _) => kind = Hoare.In) formals; + val res_formals = filter (fn (kind, _, _) => kind = Hoare.Out) formals; val (val_actuals, res_actuals) = chop (length val_formals) actuals; val init_bdy = let val state = (case state_kind of Hoare.Record => Bound 0 - | Hoare.Function => Syntax.const localsN $ Bound 0); - val upds = fold2 (fn (_, name) => init_par_tr name) val_formals val_actuals state; + | _ => Syntax.const localsN $ Bound 0); + val upds = fold2 (fn (_, name, lense_opt) => init_par_tr name lense_opt) val_formals val_actuals state; in (case state_kind of Hoare.Record => upds - | Hoare.Function => Syntax.const locals_updateN $ K_rec_syntax upds $ Bound 0) + | _ => Syntax.const locals_updateN $ K_rec_syntax upds $ Bound 0) end; val init = Abs ("s", dummyT, init_bdy); val call = (case cont of NONE => (* Procedure call *) let val results = - map (fn ((_, name), arg) => result_par_tr name arg) + map (fn ((_, name, _), arg) => result_par_tr name arg) (rev (res_formals ~~ res_actuals)); val res = Abs ("i", dummyT, Abs ("t", dummyT, Syntax.const @{const_syntax Basic} $ Abs ("s", dummyT, fold_rev (fn f => fn s => f s) results (Bound 0)))); - in if has_args then Call $init $ pt $ ret $ res else Call $ pt end + val args = [init, pt, ret] @ result_exn @ [res] + in if has_args then list_comb (Call, args) else Call $ pt end | SOME c => (* Function call *) let val res = (case res_formals of - [(_, n)] => Abs ("s", dummyT, lookup_comp ctxt [] n (Bound 0)) + [(_, n, _)] => Abs ("s", dummyT, lookup_comp ctxt [] n (Bound 0)) | _ => if Config.get ctxt StateSpace.silent then Abs ("s", dummyT, lookup_comp ctxt [] "dummy" (Bound 0)) else raise TERM ("call_tr: function " ^ pn ^ "may only have one result parameter", [])); in Call $ init $ pt $ ret $ res $ c end) val guard = Hoare.generate_guard ctxt; in if grd then (case fold_rev (fn arg => fn g => guard (antiquoteCur_tr ctxt arg) & g) (res_actuals @ val_actuals) NONE of NONE => call | SOME g => Syntax.const @{const_syntax guards} $ mk_singleton_guard @{const_syntax False} g $ call) else call end; -(* FIXME: What is prfx for, maybe unused *) -fun dest_procname ctxt prfx false (Const (p, _)) = - (prfx ^ suffix Hoare.proc_deco p, HOLogic.mk_string p) - | dest_procname ctxt prfx false (t as Free (p, T)) = - (prfx ^ suffix Hoare.proc_deco p, Free (suffix Hoare.proc_deco p, T)) - | dest_procname ctxt prfx false (Const (@{syntax_const "_free"},_) $ Free (p,T)) = - (prfx ^ suffix Hoare.proc_deco p, Free (suffix Hoare.proc_deco p, T)) - | dest_procname ctxt prfx false (t as (Const (@{syntax_const "_antiquoteCur"},_) $ Const (p, _))) = - (prfx ^ Hoare.resuffix Hoare.deco Hoare.proc_deco p, t) - | dest_procname ctxt prfx false (t as (Const (@{syntax_const "_antiquoteCur"}, _) $ Free (p, _))) = - (prfx ^ Hoare.resuffix Hoare.deco Hoare.proc_deco p, t) - | dest_procname ctxt prfx false (t as Const (p, _) $ _) = - (prfx ^ Hoare.resuffix Hoare.deco Hoare.proc_deco p, t) (* antiquoteOld *) - | dest_procname ctxt prfx false (t as Free (p,_)$_) = - (prfx ^ Hoare.resuffix Hoare.deco Hoare.proc_deco p, t) (* antiquoteOld *) - | dest_procname ctxt prfx false (t as Const (@{syntax_const "_antiquoteOld"}, _) $ _ $ Const (p, _)) = - (prfx ^ suffix Hoare.proc_deco p, t) - | dest_procname ctxt prfx false (t as Const (@{syntax_const "_antiquoteOld"}, _) $ _ $ Free (p,_)) = - (prfx ^ suffix Hoare.proc_deco p, t) - (* FIXME StateFun.lookup !? *) - | dest_procname ctxt prfx false (t as Const (@{const_name "StateFun.lookup"}, _) $ _ $ Free (p, _) $ _) = - (prfx ^ suffix Hoare.proc_deco (Hoare.remdeco' p), t) (* antiquoteOld *) +fun dest_procname ctxt false (Const (p, _)) = + (suffix Hoare.proc_deco p, HOLogic.mk_string p) + | dest_procname ctxt false (t as Free (p, T)) = + (case Hoare.get_default_state_space ctxt of + SOME {read_function_name, ...} => (p, read_function_name ctxt p) + | _ => (suffix Hoare.proc_deco p, Free (suffix Hoare.proc_deco p, T))) + | dest_procname ctxt false (Const (@{syntax_const "_free"},_) $ Free (p,T)) = + (suffix Hoare.proc_deco p, Free (suffix Hoare.proc_deco p, T)) + | dest_procname ctxt false (t as (Const (@{syntax_const "_antiquoteCur"},_) $ Const (p, _))) = + (Hoare.resuffix Hoare.deco Hoare.proc_deco p, t) + | dest_procname ctxt false (t as (Const (@{syntax_const "_antiquoteCur"}, _) $ Free (p, _))) = + (Hoare.resuffix Hoare.deco Hoare.proc_deco p, t) + | dest_procname ctxt false (t as Const (p, _) $ _) = + (Hoare.resuffix Hoare.deco Hoare.proc_deco p, t) (* antiquoteOld *) + | dest_procname ctxt false (t as Free (p,_)$_) = + (Hoare.resuffix Hoare.deco Hoare.proc_deco p, t) (* antiquoteOld *) + | dest_procname ctxt false (t as Const (@{syntax_const "_antiquoteOld"}, _) $ _ $ Const (p, _)) = + (suffix Hoare.proc_deco p, t) + | dest_procname ctxt false (t as Const (@{syntax_const "_antiquoteOld"}, _) $ _ $ Free (p,_)) = + (suffix Hoare.proc_deco p, t) + | dest_procname ctxt false (t as Const (@{const_name "StateFun.lookup"}, _) $ _ $ Free (p, _) $ _) = + (suffix Hoare.proc_deco (Hoare.remdeco' ctxt p), t) (* antiquoteOld *) - | dest_procname ctxt prfx false t = (prfx, t) - | dest_procname ctxt prfx true t = + | dest_procname ctxt false t = ("", t) + | dest_procname ctxt true t = let fun quote t = Abs ("s", dummyT, antiquoteCur_tr ctxt t) in (case quote t of (t' as Abs (_, _, Free (p, _) $ Bound 0)) => - (prfx ^ Hoare.resuffix Hoare.deco Hoare.proc_deco p, t') - (* FIXME StateFun.lookup !? *) + (Hoare.resuffix Hoare.deco Hoare.proc_deco p, t') | (t' as Abs (_, _, Const (@{const_name "StateFun.lookup"}, _) $ _ $ Free (p, _) $ (_ $ Bound 0))) => - (prfx ^ suffix Hoare.proc_deco (Hoare.remdeco' p), t') - | t' => (prfx, t')) + (suffix Hoare.proc_deco (Hoare.remdeco' ctxt p), t') + | t' => ("", t')) end -fun gen_call_tr prfx dyn grd ctxt p actuals has_args cont = +fun gen_call_tr dyn grd ctxt p actuals has_args result_exn cont = let - fun Call false true NONE = Const (@{const_syntax call}, dummyT) - | Call false false NONE = Const (@{const_syntax Call}, dummyT) - | Call true true NONE = Const (@{const_syntax dynCall}, dummyT) - | Call false true (SOME c) = Const (@{const_syntax fcall}, dummyT) - | Call _ _ _ = raise TERM ("gen_call_tr: no proper procedure call", []); + fun Call false true [] NONE = Const (@{const_syntax call}, dummyT) + | Call false true _ NONE = Const (@{const_syntax call_exn}, dummyT) + | Call false false [] NONE = Const (@{const_syntax Call}, dummyT) + | Call true true [] NONE = Const (@{const_syntax dynCall}, dummyT) + | Call true true _ NONE = Const (@{const_syntax dynCall_exn}, dummyT) + | Call false true [] (SOME c) = Const (@{const_syntax fcall}, dummyT) + | Call _ _ _ _ = raise TERM ("gen_call_tr: no proper procedure call", []); - val (pn, pt) = dest_procname ctxt prfx dyn (Term_Position.strip_positions p); + val (pn, pt) = dest_procname ctxt dyn (Term_Position.strip_positions p); in (case Hoare.get_params pn ctxt of SOME formals => - mk_call_tr ctxt grd (Call dyn has_args cont) formals pn pt actuals has_args cont + mk_call_tr ctxt grd (Call dyn has_args result_exn cont) formals pn pt actuals has_args result_exn cont | NONE => if Config.get ctxt StateSpace.silent - then mk_call_tr ctxt grd (Call dyn has_args cont) [] pn pt [] has_args cont + then mk_call_tr ctxt grd (Call dyn has_args result_exn cont) [] pn pt [] has_args result_exn cont else raise TERM ("gen_call_tr: procedure " ^ quote pn ^ " not defined", [])) end; -fun call_tr dyn grd ctxt [p, actuals] = - gen_call_tr "" dyn grd ctxt p (dest_actuals actuals) true NONE - | call_tr _ _ _ t = raise TERM ("call_tr", t); +fun call_tr dyn grd result_exn ctxt [p, actuals] = + gen_call_tr dyn grd ctxt p (dest_actuals actuals) true result_exn NONE + | call_tr _ _ _ _ t = raise TERM ("call_tr", t); -fun call_ass_tr dyn grd ctxt [l, p, actuals] = - gen_call_tr "" dyn grd ctxt p (dest_actuals actuals @ [l]) true NONE - | call_ass_tr _ _ _ t = raise TERM ("call_ass_tr", t); +fun call_ass_tr dyn grd result_exn ctxt [l, p, actuals] = + gen_call_tr dyn grd ctxt p (dest_actuals actuals @ [l]) true result_exn NONE + | call_ass_tr _ _ _ _ t = raise TERM ("call_ass_tr", t); fun proc_tr ctxt [p, actuals] = - gen_call_tr "" false false ctxt p (dest_actuals actuals) false NONE + gen_call_tr false false ctxt p (dest_actuals actuals) false [] NONE | proc_tr _ t = raise TERM ("proc_tr", t); fun proc_ass_tr ctxt [l, p, actuals] = - gen_call_tr "" false false ctxt p (dest_actuals actuals @ [l]) false NONE + gen_call_tr false false ctxt p (dest_actuals actuals @ [l]) false [] NONE | proc_ass_tr _ t = raise TERM ("proc_ass_tr", t); fun fcall_tr ctxt [p, actuals, c] = - gen_call_tr "" false false ctxt p (dest_actuals actuals) true (SOME c) + gen_call_tr false false ctxt p (dest_actuals actuals) true [] (SOME c) | fcall_tr _ t = raise TERM ("fcall_tr", t); (* printing procedure calls *) fun upd_tr' ctxt (x_upd, T) = (case try (unsuffix' Record.updateN) x_upd of SOME x => (Hoare.chopsfx Hoare.deco (Hoare.extern ctxt x), if T = dummyT then T else Term.domain_type T) | NONE => (case try (unsuffix Hoare.deco) x_upd of SOME _ => (Hoare.remdeco ctxt x_upd, T) | NONE => raise Match)); fun update_name_tr' ctxt (Free x) = Const (upd_tr' ctxt x) | update_name_tr' ctxt ((c as Const (@{syntax_const "_free"}, _)) $ Free x) = (*c $*) Const (upd_tr' ctxt x) | update_name_tr' ctxt (Const x) = Const (upd_tr' ctxt x) - | update_name_tr' _ _ = raise Match; + | update_name_tr' ctxt t = + (case Hoare.get_default_state_space ctxt of + SOME {update_tr',...} => update_tr' ctxt t + | NONE => raise Match); + fun term_name_eq (Const (x, _)) (Const (y, _)) = (x = y) | term_name_eq (Free (x, _)) (Free (y, _)) = (x = y) | term_name_eq (Var (x, _)) (Var (y, _)) = (x = y) | term_name_eq (a $ b) (c $ d) = term_name_eq a c andalso term_name_eq b d | term_name_eq (Abs (s, _, a)) (Abs (t, _, b)) = (s = t) andalso term_name_eq a b | term_name_eq _ _ = false; fun list_update_tr' l (r as Const (@{const_syntax list_update},_) $ l' $ i $ e) = if term_name_eq l l' then let fun sel_arr a [i] (Const (@{const_syntax nth},_) $ a' $ i') = term_name_eq a a' andalso i = i' | sel_arr a (i::is) (Const (@{const_syntax nth},_) $ sel $ i') = i = i' andalso sel_arr a is sel | sel_arr _ _ _ = false; fun tr' a idxs (e as Const (@{const_syntax list_update}, _) $ sel $ i $ e') = if sel_arr a idxs sel then tr' a (i :: idxs) e' else (idxs, e) | tr' _ idxs e = (idxs, e); val (idxs, e') = tr' l [i] e; val lft = fold_rev (fn i => fn arr => Syntax.const @{const_syntax nth} $ arr $ i) idxs l; in (lft,e') end else (l, r) | list_update_tr' l r = (l, r); fun list_mult_update_tr' l (r as Const (@{const_syntax list_multupd},_) $ var $ idxs $ values) = (Syntax.const @{const_syntax list_multsel} $ var $ idxs, values) | list_mult_update_tr' l r = (l, r); -fun update_tr' l (r as Const (@{const_syntax fun_upd}, _) $ +fun update_tr' ctxt l (r as Const (@{const_syntax fun_upd}, _) $ (hp as (Const (@{syntax_const "_antiquoteCur"}, _) $ _)) $ p $ value) = if term_name_eq l hp then (case value of (Const (@{const_syntax list_update}, _) $ _ $ _ $ _) => list_update_tr' (l $ p) value | (Const (@{const_syntax list_multupd},_) $ _ $ _ $ _) => list_mult_update_tr' (l $ p) value | _ => (l $ p, value)) else (l, r) - | update_tr' l (r as Const (@{const_syntax list_update},_) $ + | update_tr' ctxt l (r as Const (@{const_syntax list_update},_) $ (var as (Const (@{syntax_const "_antiquoteCur"}, _) $ _)) $ i $ value) = if term_name_eq l var then list_update_tr' l r else (l, r) - | update_tr' l (r as Const (@{const_syntax list_multupd}, _) $ + | update_tr' ctxt l (r as Const (@{const_syntax list_multupd}, _) $ (var as (Const (@{syntax_const "_antiquoteCur"}, _) $ _)) $ idxs $ values) = if term_name_eq l var then list_mult_update_tr' l r else (l, r) - | update_tr' l r = (l, r); + | update_tr' ctxt l r = (l, r); -fun dest_K_rec (Abs (_, _, v)) = - if member (op =) (loose_bnos v) 0 then NONE else SOME (incr_boundvars ~1 v) - | dest_K_rec (Abs (_, _, Abs (_, _, v) $ Bound 0)) = (* eta expanded version *) +fun dest_K_rec (Abs (_, _, Abs (_, _, v) $ Bound 0)) = (* eta expanded version *) let val lbv = loose_bnos v; in if member (op =) lbv 0 orelse member (op =) lbv 1 then NONE else SOME (incr_boundvars ~2 v) end + | dest_K_rec (Abs (_, _, v)) = + if member (op =) (loose_bnos v) 0 then NONE else SOME (incr_boundvars ~1 v) + | dest_K_rec (Const (@{const_syntax K_statefun}, _) $ v) = SOME v | dest_K_rec _ = NONE; +fun the_Match (SOME x) = x + | the_Match (NONE) = raise Match + +fun dest_update ctxt (upd' $ dest $ constr $ n $ v $ s) = + (n, v, SOME s) + | dest_update ctxt (upd' $ dest $ constr $ n $ v) = + (n, v, NONE) + | dest_update ctxt t = + case Hoare.get_default_state_space ctxt of + SOME {dest_update_tr', ...} => dest_update_tr' ctxt t + | NONE => raise Match + local -fun uncover (upd,v) = - (case (upd, v) of - (Const (cupd, _), upd' $ dest $ constr $ n $ (Const (@{const_syntax K_statefun}, _) $ v') $ s) => +fun uncover ctxt (upd,v) = (upd, v) |> first_match [ + fn (Const (cupd, _), t) => if member (op =) [globals_updateN, locals_updateN] (Long_Name.base_name cupd) - then (case s of (Const (g, _) $ _) => - if member (op =) [localsN, globalsN] (Long_Name.base_name g) - then (n, v') - else raise Match - | _ => raise Match) - else (upd, v) - | (Const (gupd, _), upd' $ k $ s) => - (case dest_K_rec k of - SOME v' => - if Long_Name.base_name gupd = globals_updateN - then - (case s of - Const (gl, _) $ _ => - if Long_Name.base_name gl = globalsN (* assignment *) - then (upd',v') - else raise Match - | _ => raise Match) + then case dest_update ctxt t of + (n, v', SOME s) => (case s of (Const (g, _) $ _) => + if member (op =) [localsN, globalsN] (Long_Name.base_name g) + then (n, the_Match (dest_K_rec v')) + else raise Match + | _ => raise Match) + | (n, v', NONE) => (n, the_Match (dest_K_rec v')) + else (upd, v), + fn (upd, v ) => + (case (upd, v) of + (Const (gupd, _), t as (upd' $ k $ s)) => + (case dest_K_rec k of + SOME v' => + if Long_Name.base_name gupd = globals_updateN + then + (case s of + Const (gl, _) $ _ => + if Long_Name.base_name gl = globalsN (* assignment *) + then (upd',v') + else raise Match + | _ => raise Match) + else (upd, v) + | _ => (upd, v)) + | (Const (upd_glob, _), upd' $ v') => + if Long_Name.base_name upd_glob = upd_globalsN (* result parameter *) + then (upd', v') + else if Long_Name.base_name upd_glob = globals_updateN + then (case dest_K_rec v' of + SOME v'' => (upd',v'') + | _ => (upd, v)) else (upd, v) - | _ => (upd, v)) - | (Const (upd_glob, _), upd' $ v') => - if Long_Name.base_name upd_glob = upd_globalsN (* result parameter *) - then (upd', v') else (upd, v) - | _ => (upd, v)); + | _ => (upd, v))] in -fun global_upd_tr' upd k = +fun global_upd_tr' ctxt upd k = (case dest_K_rec k of - SOME v => uncover (upd, v) - | NONE => uncover (upd, k)); - + SOME v => uncover ctxt (upd, v) + | NONE => uncover ctxt (upd, k)) end; - -fun dest_updates (t as (upd as Const (u, _)) $ k $ state) = +fun dest_updates ctxt t = t |> first_match [ + fn (t as (upd as Const (u, _)) $ k $ state) => (case dest_K_rec k of SOME value => if member (op =) [globals_updateN, locals_updateN] (Long_Name.base_name u) - then dest_updates value + then dest_updates ctxt value else if can (unsuffix Record.updateN) u orelse Long_Name.base_name u = upd_globalsN - then (upd,value)::dest_updates state + then (upd,value)::dest_updates ctxt state else raise Match - | NONE => raise Match) - | dest_updates (t as (upd as Const (u,_))$k) = + | NONE => raise Match (*dest_updates ctxt k @ dest_updates ctxt state*) (* check for nested update (e.g. locals-stack) *) + (*handle Match => []*)), (* t could be just (globals $ s) *) + fn (t as (upd as Const (u,_))$k) => (case dest_K_rec k of SOME value => if member (op =) [globals_updateN, locals_updateN] (Long_Name.base_name u) - then dest_updates value + then dest_updates ctxt value else if can (unsuffix Record.updateN) u orelse Long_Name.base_name u = upd_globalsN then [(upd,value)] else if Long_Name.base_name u = globalsN then [] else raise Match - | NONE => []) (* t could be just (globals $ s) *) - | dest_updates ((Const (u, _)) $ _ $ _ $ n $ (Const (@{const_syntax K_statefun},_) $ value) $ state) = + | NONE => dest_updates ctxt k (* check for nested update (e.g. locals-stack) *) + handle Match => []), (* t could be just (globals $ s) *) + fn ((Const (u, _)) $ _ $ _ $ n $ k $ state) => if Long_Name.base_name u = Long_Name.base_name StateFun.updateN - then (n, value) :: dest_updates state - else raise Match - | dest_updates t = []; + then case dest_K_rec k of SOME value => (n, value) :: dest_updates ctxt state | _ => raise Match + else raise Match, + fn ((Const (u, _)) $ _ $ _ $ n $ k) => + if Long_Name.base_name u = Long_Name.base_name StateFun.updateN + then case dest_K_rec k of SOME value => [(n, value)] | _ => raise Match + else raise Match, + fn t => + (case Hoare.get_default_state_space ctxt of + SOME {dest_update_tr', ...} => + (case dest_update_tr' ctxt t of + (n, v, SOME s) => (n, the_Match (dest_K_rec v))::dest_updates ctxt s + | (n, v, NONE) => [(n, the_Match (dest_K_rec v))]) + | NONE => raise Match), + fn t => []] -(* FIXME: externalize names properly before removing decoration! *) +fun dest_updates ctxt t = t |> first_match [ + fn (t as (upd as Const (u, _)) $ k $ state) => + if member (op =) [globals_updateN, locals_updateN] (Long_Name.base_name u) then + dest_updates ctxt k @ dest_updates ctxt state + else if can (unsuffix Record.updateN) u orelse Long_Name.base_name u = upd_globalsN then + (upd, the_Match (dest_K_rec k))::dest_updates ctxt state + else raise Match, (* t could be just (globals $ s) *) + fn (t as (upd as Const (u,_))$k) => + if member (op =) [globals_updateN, locals_updateN] (Long_Name.base_name u) then + dest_updates ctxt k + else if can (unsuffix Record.updateN) u orelse Long_Name.base_name u = upd_globalsN then + [(upd, the_Match (dest_K_rec k))] + (*else if Long_Name.base_name u = globalsN then [] *) + else raise Match, + fn ((Const (u, _)) $ _ $ _ $ n $ k $ state) => + if Long_Name.base_name u = Long_Name.base_name StateFun.updateN then + (n, the_Match (dest_K_rec k)) :: dest_updates ctxt state + else raise Match, + fn ((Const (u, _)) $ _ $ _ $ n $ k) => + if Long_Name.base_name u = Long_Name.base_name StateFun.updateN then + [(n, the_Match (dest_K_rec k))] + else raise Match, + fn t => + (case Hoare.get_default_state_space ctxt of + SOME {dest_update_tr', ...} => + (case dest_update_tr' ctxt t of + (n, v, SOME s) => (n, the_Match (dest_K_rec v))::dest_updates ctxt s + | (n, v, NONE) => [(n, the_Match (dest_K_rec v))]) + | NONE => raise Match), + fn t => []] + + +(* fixme: externalize names properly before removing decoration! *) fun init_tr' ctxt [Abs (_,_,t)] = let val upds = - case dest_updates t of + case dest_updates ctxt t of us as [(Const (gupd, _), v)] => if Long_Name.base_name gupd = globals_updateN - then dest_updates v else us + then dest_updates ctxt v else us | us => us; val comps = map (fn (Const (u, _)) => Symbol.explode (unsuffix (Hoare.deco ^ Record.updateN) u)) (map fst upds); val prfx = maxprefixs (op =) comps; fun dest_list (Const (@{const_syntax Nil}, _)) = [] | dest_list (Const (@{const_syntax Cons}, _) $ x $ xs) = x :: dest_list xs | dest_list t = [t]; fun mk_set [] = Syntax.const @{const_syntax Set.empty} | mk_set (x :: xs) = Syntax.const @{const_syntax insert} $ x $ mk_set xs; val l = length prfx; val _ = if l <= 1 then raise Match else (); val comp = mk_list (map (Syntax.const o implode o drop l) comps); val vals = map mk_set (transpose (map (dest_list o snd) upds)); val v = case vals of [v] => v | vs => mk_list vs; in Syntax.const @{syntax_const "_Init"} $ Syntax.const (implode (fst (split_last prfx))) $ comp $ v end; local fun tr' ctxt c (upd,v) = let val l = Syntax.const antiquoteCur $ update_name_tr' ctxt upd; val r = quote_tr' ctxt antiquoteCur (Abs ("s", dummyT, v)); - val (l', r') = update_tr' l r; + val (l', r') = update_tr' ctxt l r; in (c $ l' $ r') end; in fun app_assign_tr' c ctxt (Abs (s, _, upd $ v $ Bound 0) :: ts) = - tr' ctxt c (global_upd_tr' upd v) + tr' ctxt c (global_upd_tr' ctxt upd v) | app_assign_tr' c ctxt ((upd $ v) :: ts) = (case upd of u $ v => raise Match - | _ => tr' ctxt c (global_upd_tr' upd v)) + | _ => tr' ctxt c (global_upd_tr' ctxt upd v)) | app_assign_tr' _ _ _ = raise Match; end; val assign_tr' = app_assign_tr' (Syntax.const @{syntax_const "_Assign"}); val raise_tr' = app_assign_tr' (Syntax.const @{syntax_const "_raise"}); fun split_Let' ((l as Const (@{const_syntax Let'}, _)) $ x $ t) = let val (recomb,t') = split_Let' t in (fn t => l $ x $ recomb t, t') end | split_Let' (Abs (x, T, t)) = let val (recomb, t') = split_Let' t in if t' = t then (I, t') (* Get rid of last abstraction *) else (fn t => Abs (x, T, recomb t), t') end | split_Let' ((s as Const (@{const_syntax case_prod},_)) $ t) = let val (recomb, t') = split_Let' t in (fn t => s $ recomb t, t') end | split_Let' t = (I, t) fun basic_tr' ctxt [Abs (s, T, t)] = let val (has_let, t') = case t of ((t' as (Const (@{const_syntax Let'},_) $ _ $ _)) $ Bound 0) => (true, t') | _ => (false, t); val (recomb, t'') = split_Let' t'; - val upds = dest_updates t''; + val upds = dest_updates ctxt t''; val _ = if length upds <= 1 andalso not has_let then raise Match else (); val ass = map (fn (u, v) => app_assign_tr' (Syntax.const @{syntax_const "_BAssign"}) ctxt [Abs ("s",dummyT,u$v$Bound 0)]) upds; val basics = foldr1 (fn (x, ys) => Syntax.const @{syntax_const "_basics"} $ x $ ys) (rev ass); in Syntax.const @{syntax_const "_Basic"} $ quote_tr' ctxt @{syntax_const "_antiquoteCur"} (Abs (s, T, recomb basics)) end; fun loc_tr' ctxt [init, bdy, return, c] = (let val upds = (case init of - Abs (_, _, t as (upd $ v $ s)) => dest_updates t - | upd $ v => dest_updates (upd $ v $ Bound 0) + Abs (_, _, t as (upd $ v $ s)) => dest_updates ctxt t + | upd $ v => dest_updates ctxt (upd $ v $ Bound 0) | _ => raise Match); fun mk_locinit c v = Syntax.const @{syntax_const "_locinit"} $ Syntax.const c $ quote_tr' ctxt antiquoteCur (Abs ("s", dummyT, v)); fun init_or_not c c' v = if c = c' then Syntax.const @{syntax_const "_locnoinit"} $ Syntax.const (Hoare.remdeco ctxt c') else mk_locinit (Hoare.remdeco ctxt c) v; fun mk_init (Const (c, _), (v as (Const (c', _) $ Bound 0))) = init_or_not (unsuffix' Record.updateN c) c' v | mk_init (Const (c, _), v) = mk_locinit (unsuffix' (Hoare.deco ^ Record.updateN) (Hoare.extern ctxt c)) v | mk_init ((f as Const (@{syntax_const "_free"}, _)) $ Free (c, _), v) = (case v of Const (lookup, _) $ _ $ (Const (@{syntax_const "_free"}, _) $ Free (c', _)) $ (Const (locals,_) $ Bound 0) => if Long_Name.base_name lookup = Long_Name.base_name StateFun.lookupN andalso Long_Name.base_name locals = localsN then init_or_not c c' v - else mk_locinit (Hoare.remdeco' c) v - | _ => mk_locinit (Hoare.remdeco' c) v) + else mk_locinit (Hoare.remdeco' ctxt c) v + | _ => mk_locinit (Hoare.remdeco' ctxt c) v) | mk_init _ = raise Match; val inits = foldr1 (fn (t, u) => Syntax.const @{syntax_const "_locinits"} $ t $ u) (map mk_init (rev upds)); - in Syntax.const @{syntax_const "_Loc"} $ inits $ bdy end handle Fail _ => raise Match) + in Syntax.const @{syntax_const "_Loc"} $ inits $ bdy end handle Fail _ => raise Match | Empty => raise Match) | loc_tr' _ _ = raise Match; fun actuals_tr' acts = (case acts of [] => Syntax.const @{syntax_const "_actuals_empty"} | xs => Syntax.const @{syntax_const "_actuals"} $ foldr1 (fn (l, r) => (Syntax.const @{syntax_const "_pars"} $ l $ r)) xs); fun gen_call_tr' ctxt Call CallAss init p return c = let - fun get_init_updates (Abs (s, _, upds)) = dest_updates upds - | get_init_updates upds = dest_updates upds; + fun get_init_updates (Abs (s, _, upds)) = dest_updates ctxt upds + | get_init_updates upds = dest_updates ctxt upds; fun get_res_updates (Abs (i, _, Abs (t, _, Const (@{const_syntax Basic}, _) $ Abs (s, _, upds)))) = - dest_updates upds + dest_updates ctxt upds | get_res_updates (Abs (i, _, Abs (t, _, Const (@{const_syntax Basic}, _) $ upds))) = - dest_updates upds + dest_updates ctxt upds | get_res_updates _ = raise Match; - fun init_par_tr' par = + val init_upds = rev (get_init_updates init) + fun init_par_tr' par = Syntax.const @{syntax_const "_par"} $ quote_tr' ctxt antiquoteCur (Abs ("s", dummyT, par)); - val init_actuals = - map (fn (_, value) => init_par_tr' value) (rev (get_init_updates init)); + val init_actuals = + map (fn (_, value) => init_par_tr' value) init_upds; - fun tr' c (upd, v) = - let - val l = Syntax.const antiquoteCur $ update_name_tr' ctxt upd; - val r = - quote_tr' ctxt antiquoteCur + fun tr' c (upd, v) = + let + val l = Syntax.const antiquoteCur $ update_name_tr' ctxt upd; + val r = + quote_tr' ctxt antiquoteCur + (quote_tr' ctxt antiquoteCur (quote_tr' ctxt antiquoteCur - (quote_tr' ctxt antiquoteCur - (Abs ("i", dummyT, Abs ("t", dummyT, Abs ("s", dummyT, v)))))); - val (l', _) = update_tr' l r; - in c $ l' end; + (Abs ("i", dummyT, Abs ("t", dummyT, Abs ("s", dummyT, v)))))); + val (l', _) = update_tr' ctxt l r; + in c $ l' end; - fun ret_par_tr' (upd, v) = - tr' (Syntax.const @{syntax_const "_par"}) (global_upd_tr' upd v); + fun ret_par_tr' (upd, v) = + tr' (Syntax.const @{syntax_const "_par"}) (global_upd_tr' ctxt upd v); - val res_updates = rev (get_res_updates c); - val res_actuals = map ret_par_tr' res_updates; - in if Config.get ctxt use_call_tr' then + val res_updates = rev (get_res_updates c); + val res_actuals = map ret_par_tr' res_updates; + in + if Config.get ctxt use_call_tr' then (case res_actuals of [l] => CallAss $ l $ p $ actuals_tr' init_actuals | _ => Call $ p $ actuals_tr' (init_actuals @ res_actuals)) - else raise Match + else raise Match end; fun gen_fcall_tr' ctxt init p return result c = let - fun get_init_updates (Abs (s, _, upds)) = dest_updates upds + fun get_init_updates (Abs (s, _, upds)) = dest_updates ctxt upds | get_init_updates _ = raise Match; fun init_par_tr' par = Syntax.const @{syntax_const "_par"} $ quote_tr' ctxt antiquoteCur (Abs ("s", dummyT, par)); val init_actuals = map (fn (_, value) => init_par_tr' value) (rev (get_init_updates init)); val (v, c') = (case c of Abs abs => Syntax_Trans.atomic_abs_tr' abs | _ => raise Match); in if Config.get ctxt use_call_tr' then Syntax.const @{syntax_const "_FCall"} $ p $ actuals_tr' init_actuals $ v $ c' else raise Match end; fun pname_tr' ctxt ((f as Const (@{syntax_const "_free"}, _)) $ Free (p, T)) = (*f$*) Const (unsuffix' Hoare.proc_deco p, T) | pname_tr' ctxt (Free (p, T)) = Const (unsuffix' Hoare.proc_deco p, T) | pname_tr' ctxt p = let (* from HOL strings to ML strings *) - fun dest_nib c = (* FIXME authentic syntax *) + fun dest_nib c = (* fixme authentic syntax *) (case raw_explode c of ["N", "i", "b", "b", "l", "e", h] => if "0" <= h andalso h <= "9" then ord h - ord "0" else if "A" <= h andalso h <= "F" then ord h - ord "A" + 10 else raise Match | _ => raise Match); fun dest_chr (Const (@{const_syntax Char},_) $ Const (c1, _) $ Const (c2,_)) = let val c = Char.chr (dest_nib c1 * 16 + dest_nib c2) in if Char.isPrint c then c else raise Match end | dest_chr _ = raise Match; fun dest_string (Const (@{const_syntax Nil}, _)) = [] | dest_string (Const (@{const_syntax Cons}, _) $ c $ cs) = dest_chr c :: dest_string cs | dest_string _ = raise Match; in (case try dest_string p of SOME name => Syntax.const (String.implode name) | NONE => antiquote_mult_tr' ctxt (K true) antiquoteCur antiquoteOld p) end; fun call_tr' ctxt [init, p, return, result] = gen_call_tr' ctxt (Const (@{syntax_const "_Call"}, dummyT)) (Const (@{syntax_const "_CallAss"}, dummyT)) init (pname_tr' ctxt p) return result | call_tr' _ _ = raise Match; +fun call_exn_tr' ctxt [init, p, return, result_exn, result] = + gen_call_tr' ctxt + (Const (@{syntax_const "_Call_exn"}, dummyT)) + (Const (@{syntax_const "_CallAss_exn"}, dummyT)) init (pname_tr' ctxt p) return result + | call_exn_tr' _ _ = raise Match; + fun dyn_call_tr' ctxt [init, p, return, result] = let val p' = quote_tr' ctxt antiquoteCur p in gen_call_tr' ctxt (Const (@{syntax_const "_DynCall"}, dummyT)) (Const (@{syntax_const "_DynCallAss"}, dummyT)) init p' return result end | dyn_call_tr' _ _ = raise Match; +fun dyn_call_exn_tr' ctxt [init, p, return, result_exn, result] = + let val p' = quote_tr' ctxt antiquoteCur p + in + gen_call_tr' ctxt + (Const (@{syntax_const "_DynCall_exn"}, dummyT)) + (Const (@{syntax_const "_DynCallAss_exn"}, dummyT)) init p' return result + end + | dyn_call_exn_tr' _ _ = raise Match; + fun proc_tr' ctxt [p] = let val p' = pname_tr' ctxt p; - val pn = fst (dest_procname ctxt "" false p'); + val pn = fst (dest_procname ctxt false p'); val formals = the (Hoare.get_params pn ctxt) handle Option.Option => raise Match; - val val_formals = map_filter (fn (Hoare.In, p) => SOME p | _ => NONE) formals; - val res_formals = map_filter (fn (Hoare.Out, p) => SOME p | _ => NONE) formals; + val val_formals = map_filter (fn (Hoare.In, p, _) => SOME p | _ => NONE) formals; + val res_formals = map_filter (fn (Hoare.Out, p, _) => SOME p | _ => NONE) formals; fun mkpar n = Syntax.const @{syntax_const "_par"} $ (Syntax.const antiquoteCur $ Syntax.const (Hoare.remdeco ctxt n)); in if not (print_mode_active "NoProc") then (case res_formals of [r] => Syntax.const @{syntax_const "_ProcAss"} $ (Syntax.const antiquoteCur $ Syntax.const (Hoare.remdeco ctxt r)) $ p' $ actuals_tr' (map mkpar val_formals) | _ => Syntax.const @{syntax_const "_Proc"} $ p' $ actuals_tr' (map mkpar (val_formals @ res_formals))) else raise Match end | proc_tr' _ _ = raise Match; fun fcall_tr' ctxt [init, p, return, result, c] = gen_fcall_tr' ctxt init (pname_tr' ctxt p) return result c | fcall_tr' _ _ = raise Match; (* misc. print translations *) fun assert_tr' ctxt ((t as Abs (_, _, p)) :: ts) = let fun selector (Const (c, T)) = Hoare.is_state_var c | selector (Const (l, _) $ _ $ _) = Long_Name.base_name l = Long_Name.base_name StateFun.lookupN - | selector t = false; + | selector t = + (case Hoare.get_default_state_space ctxt of + SOME {is_lookup,...} => is_lookup ctxt t + | _ => false) fun fix_state (Const (@{const_syntax HOL.eq},_) $ (Const (@{syntax_const "_free"}, _) $ _)) = true | fix_state (Const (@{const_syntax HOL.eq},_) $ (Const (@{syntax_const "_bound"}, _) $ _)) = true | fix_state (Const (@{const_syntax HOL.eq},_) $ (Const (@{syntax_const "_var"}, _) $ _)) = true | fix_state (Const (@{const_syntax HOL.eq},_) $ Free _) = true | fix_state (Const (@{const_syntax HOL.eq},_) $ Bound _) = true | fix_state (Const (@{const_syntax HOL.eq},_) $ Var _) = true | fix_state _ = false; in if antiquote_applied_only_to (fn t => selector t orelse fix_state t) p andalso not (print_mode_active "NoAssertion") then app_quote_mult_tr' ctxt selector (Syntax.const @{syntax_const "_Assert"}) (t :: ts) else raise Match end | assert_tr' _ _ = raise Match fun bexp_tr' name ctxt ((Const (@{const_syntax Collect}, _) $ t) :: ts) = app_quote_tr' ctxt (Syntax.const name) (t :: ts) | bexp_tr' _ _ _ = raise Match; fun new_tr' ctxt [Abs (s,_, Const (@{const_syntax If}, _) $ (Const (@{const_syntax Orderings.less_eq},_) $ size $ free) $ (upd $ new $ (gupd $ Abs (_, _, inits_free_alloc) $ Bound 0)) $ (upd' $ null $ Bound 0))] = let fun mk_init (Const (upd, _), Const (@{const_syntax fun_upd},_) $ _ $ _ $ v) = let val var = unsuffix' Hoare.deco (unsuffix' Record.updateN (Hoare.extern ctxt upd)) in Syntax.const @{syntax_const "_newinit"} $ Syntax.const var $ v end | mk_init ((f as Const (@{syntax_const "_free"}, _)) $ Free (var, T), Const (@{const_syntax fun_upd},_) $ _ $ _ $ v) = Syntax.const @{syntax_const "_newinit"} $ - (f $ Free (Hoare.remdeco' var, T)) $ v; + (f $ Free (Hoare.remdeco' ctxt var, T)) $ v; - val inits_free_allocs = dest_updates inits_free_alloc; + val inits_free_allocs = dest_updates ctxt inits_free_alloc; val inits = map mk_init (take (length inits_free_allocs - 2) (inits_free_allocs)); val inits' = foldr1 (fn (t1, t2) => Syntax.const @{syntax_const "_newinits"} $ t1 $ t2) (rev inits); fun tr' (upd, v) = let val l = Syntax.const antiquoteCur $ update_name_tr' ctxt upd; val r = quote_tr' ctxt antiquoteCur (Abs (s, dummyT, v)); - val (l', r') = update_tr' l r + val (l', r') = update_tr' ctxt l r in l' end; - val l = tr' (global_upd_tr' upd' null); + val l = tr' (global_upd_tr' ctxt upd' null); in Syntax.const @{syntax_const "_New"} $ l $ size $ inits' end | new_tr' _ _ = raise Match; fun nnew_tr' ctxt [Const (@{const_syntax if_rel},_) $ (Abs (s, _, Const (@{const_syntax Orderings.less_eq}, _) $ size $ free)) $ (Abs (_, _, upd $ new $ (gupd $ (Abs (_, _, free_inits_alloc)) $ Bound 0))) $ (Abs (_, _, (upd' $ null $ Bound 0))) $ _] = let fun mk_init (Const (upd, _), Const (@{const_syntax fun_upd}, _) $ _ $ _ $ v) = let val var = unsuffix' Hoare.deco (unsuffix' Record.updateN (Hoare.extern ctxt upd)) in Syntax.const @{syntax_const "_newinit"} $ Syntax.const var $ v end | mk_init ((f as Const (@{syntax_const "_free"}, _)) $ Free (var, T), Const (@{const_syntax fun_upd}, _) $_ $ _ $ v) = Syntax.const @{syntax_const "_newinit"} $ - (f $ Free (Hoare.remdeco' var, T)) $ v; + (f $ Free (Hoare.remdeco' ctxt var, T)) $ v; - val free_inits_allocs = dest_updates free_inits_alloc; + val free_inits_allocs = dest_updates ctxt free_inits_alloc; val inits = map mk_init (take (length free_inits_allocs - 2) (tl free_inits_allocs)); val inits' = foldr1 (fn (t1, t2) => Syntax.const @{syntax_const "_newinits"} $ t1 $ t2) (rev inits); fun tr' (upd, v) = let val l = Syntax.const antiquoteCur $ update_name_tr' ctxt upd; val r = quote_tr' ctxt antiquoteCur (Abs (s, dummyT, v)); - val (l', r') = update_tr' l r; + val (l', r') = update_tr' ctxt l r; in l' end; - val l = tr' (global_upd_tr' upd' null); + val l = tr' (global_upd_tr' ctxt upd' null); in Syntax.const @{syntax_const "_NNew"} $ l $ size $ inits' end | nnew_tr' _ _ = raise Match; fun switch_tr' ctxt [v, vs] = let fun case_tr' (Const (@{const_syntax Pair}, _) $ V $ c) = Syntax.const @{syntax_const "_switchcase"} $ V $ c | case_tr' _ = raise Match; fun dest_list (Const (@{const_syntax Nil}, _)) = [] | dest_list (Const (@{const_syntax Cons}, _) $ x $ xs) = x :: dest_list xs | dest_list t = raise Match; fun ltr' [] = raise Match | ltr' [Vc] = Syntax.const @{syntax_const "_switchcasesSingle"} $ case_tr' Vc | ltr' (Vc :: xs) = Syntax.const @{syntax_const "_switchcasesCons"} $ case_tr' Vc $ ltr' xs; in app_quote_tr' ctxt (Syntax.const @{syntax_const "_Switch"}) (v :: [ltr' (dest_list vs)]) end; fun bind_tr' ctxt [e, Abs abs] = let val (v, c) = Syntax_Trans.atomic_abs_tr' abs; val e' = case e of Abs a => e | t as Const _ => Abs ("s", dummyT, t $ Bound 0) | _ => raise Match; in app_quote_tr' ctxt (Syntax.const @{syntax_const "_Bind"}) [e', v, c] end | bind_tr' _ _ = raise Match; local fun dest_list (Const (@{const_syntax Nil}, _)) = [] | dest_list (Const (@{const_syntax Cons}, _) $ x $ xs) = x :: dest_list xs | dest_list _ = raise Match; fun guard_tr' fg = let val (flag, g) = HOLogic.dest_prod fg in if flag aconv @{term True} then Syntax.const @{syntax_const "_guarantee"} $ g else if flag aconv @{term False} then g else fg end handle TERM _ => fg; fun guards_lst_tr' [fg] = guard_tr' fg | guards_lst_tr' (t :: ts) = Syntax.const @{syntax_const "_grds"} $ guard_tr' t $ guards_lst_tr' ts | guards_lst_tr' [] = raise Match; fun cond_guards_lst_tr' ctxt ts = if Config.get ctxt hide_guards then Syntax.const @{syntax_const "_hidden_grds"} else guards_lst_tr' ts; in fun guards_tr' ctxt [gs, c] = Syntax.const @{syntax_const "_guards"} $ cond_guards_lst_tr' ctxt (dest_list gs) $ c | guards_tr' _ _ = raise Match; fun whileAnnoG_tr' ctxt [gs, cond as (Const (@{const_syntax Collect}, _) $ b), I, V, c] = let val b' = (case assert_tr' ctxt [b] of Const (@{syntax_const "_Assert"}, _) $ b' => b' | _ => cond) handle Match => cond; in Syntax.const @{syntax_const "_While_guard_inv_var"} $ cond_guards_lst_tr' ctxt (dest_list gs) $ b' $ I $ V $ (Syntax.const @{syntax_const "_DoPre"} $ c) end | whileAnnoG_tr' _ _ = raise Match; fun whileAnnoGFix_tr' ctxt [gs, cond as (Const (@{const_syntax Collect}, _) $ b), I, V, c] = let val b' = (case assert_tr' ctxt [b] of Const (@{syntax_const "_Assert"}, _) $ b' => b' | _ => cond) handle Match => cond; in (case maps strip_abs_vars [I, V, c] of [] => raise Match | ((x, T) :: xs) => let val (x', I') = Syntax_Trans.atomic_abs_tr' (x, T, strip_abs_body I); val (_ , V') = Syntax_Trans.atomic_abs_tr' (x, T, strip_abs_body V); val (_ , c') = Syntax_Trans.atomic_abs_tr' (x, T, strip_abs_body c); in Syntax.const @{syntax_const "_WhileFix_guard_inv_var"} $ cond_guards_lst_tr' ctxt (dest_list gs) $ b' $ x' $ I' $ V' $ (Syntax.const @{syntax_const "_DoPre"} $ c') end) end; end end; diff --git a/web/entries/Catoids.html b/web/entries/Catoids.html --- a/web/entries/Catoids.html +++ b/web/entries/Catoids.html @@ -1,274 +1,274 @@ Catoids, Categories, Groupoids - Archive of Formal Proofs

Catoids, Categories, Groupoids

Georg Struth 📧

August 14, 2023

Abstract

This AFP entry formalises catoids, which are generalisations of single-set categories, and groupoids. More specifically, in catoids, the partial composition of arrows in a category is generalised to a multioperation, which sends pairs of elements to sets of elements, and the definedness condition of arrow composition -- two arrows can be composed if and only the target of the first matches the source of the second -- is relaxed. Beyond a library of basic laws for catoids, single-set categories and groupoids, I formalise the facts that every catoid can be lifted to a modal powerset quantale, that every groupoid can be lifted to a Dedekind quantale and to power set relation algebras, a special case of a famous result of Jónsson and Tarski. Finally, I show that single-set categories are equivalent to a standard axiomatisation of categories based on a set of objects and a set of arrows, and compare catoids with related structures such as multimonoid and relational monoids (monoids in the monoidal category Rel).

License

Topics

Related publications

  • Cranch, J., Doherty, S., & Struth, G. (2020). Convolution and Concurrency (Version 1). arXiv. https://doi.org/10.48550/ARXIV.2002.02321
  • Calk, C., Fahrenberg, U., Johansen, C., Struth, G., & Ziemiański, K. (2021). $$\ell r$$-Multisemigroups, Modal Quantales and the Origin of Locality. Lecture Notes in Computer Science, 90–107. https://doi.org/10.1007/978-3-030-88701-8_6
  • Fahrenberg, U., Johansen, C., Struth, G., & Ziemiański, K. (2023). Catoids and modal convolution algebras. Algebra Universalis, 84(2). https://doi.org/10.1007/s00012-023-00805-9
  • -
  • Calk, C., Malbos, P., Pous, D., & Struth, G. (2023). Higher Catoids, Higher Quantales and their Correspondences (Version 1). arXiv. https://doi.org/10.48550/ARXIV.2307.09253 +
  • Calk, C., Malbos, P., Pous, D., & Struth, G. (2023). Higher Catoids, Higher Quantales and their Correspondences (Version 2). arXiv. https://doi.org/10.48550/ARXIV.2307.09253

Session Catoids

\ No newline at end of file diff --git a/web/entries/Combinatorics_Words.html b/web/entries/Combinatorics_Words.html --- a/web/entries/Combinatorics_Words.html +++ b/web/entries/Combinatorics_Words.html @@ -1,289 +1,289 @@ Combinatorics on Words Basics - Archive of Formal Proofs

Combinatorics on Words Basics

Abstract

We formalize basics of Combinatorics on Words. This is an extension of existing theories on lists. We provide additional properties related to prefix, suffix, factor, length and rotation. The topics include prefix and suffix comparability, mismatch, word power, total and reversed morphisms, border, periods, primitivity and roots. We also formalize basic, mostly folklore results related to word equations: equidivisibility, commutation and conjugation. Slightly advanced properties include the Periodicity lemma (often cited as the Fine and Wilf theorem) and the variant of the Lyndon-Schützenberger theorem for words, including its full parametric solution. We support the algebraic point of view which sees words as generators of submonoids of a free monoid. This leads to the concepts of the (free) hull, the (free) basis (or code). We also provide relevant proof methods and a tool to generate reverse-symmetric claims.

License

History

August 17, 2023
Updated to version v1.10.1.
August 24, 2022
Many updates and additions. New theories: Border_Array, Morphisms, Equations_Basic, and Binary_Code_Morphisms.

Topics

Related publications

Session Combinatorics_Words

\ No newline at end of file diff --git a/web/entries/Efficient_Weighted_Path_Order.html b/web/entries/Efficient_Weighted_Path_Order.html --- a/web/entries/Efficient_Weighted_Path_Order.html +++ b/web/entries/Efficient_Weighted_Path_Order.html @@ -1,260 +1,260 @@ A Verified Efficient Implementation of the Weighted Path Order - Archive of Formal Proofs

A Verified Efficient Implementation of the Weighted Path Order

Abstract

The Weighted Path Order (WPO) of Yamada is a powerful technique for proving termination. In a previous AFP entry, the WPO was defined and properties of WPO have been formally verified. However, the implementation of WPO was naive, leading to an exponential runtime in the worst case.

Therefore, in this AFP entry we provide a poly-time implementation of WPO. The implementation is based on memoization. Since WPO generalizes the recursive path order (RPO), we also easily derive an efficient implementation of RPO.

License

Topics

Related publications

    -
  • Thiemann, R., Schöpf, J., Sternagel, C., & Yamada, A. (2020). Certifying the Weighted Path Order (Invited Talk). Schloss Dagstuhl - Leibniz-Zentrum Für Informatik. https://doi.org/10.4230/LIPICS.FSCD.2020.4 +
  • Thiemann, R., Schöpf, J., Sternagel, C., & Yamada, A. (2020). Certifying the Weighted Path Order (Invited Talk). Schloss Dagstuhl – Leibniz-Zentrum für Informatik. https://doi.org/10.4230/LIPICS.FSCD.2020.4
  • Yamada, A., Kusakari, K., & Sakabe, T. (2015). A unified ordering for termination proving. Science of Computer Programming, 111, 110–134. https://doi.org/10.1016/j.scico.2014.07.009

Session Efficient_Weighted_Path_Order

\ No newline at end of file diff --git a/web/entries/Frequency_Moments.html b/web/entries/Frequency_Moments.html --- a/web/entries/Frequency_Moments.html +++ b/web/entries/Frequency_Moments.html @@ -1,280 +1,280 @@ Formalization of Randomized Approximation Algorithms for Frequency Moments - Archive of Formal Proofs

Formalization of Randomized Approximation Algorithms for Frequency Moments

Emin Karayel 🌐

April 8, 2022

Abstract

In 1999 Alon et. al. introduced the still active research topic of approximating the frequency moments of a data stream using randomized algorithms with minimal space usage. This includes the problem of estimating the cardinality of the stream elements - the zeroth frequency moment. But, also higher-order frequency moments that provide information about the skew of the data stream. (The k-th frequency moment of a data stream is the sum of the k-th powers of the occurrence counts of each element in the stream.) This entry formalizes three randomized algorithms for the approximation of F0, F2 and Fk for k ≥ 3 based on [1, 2] and verifies their expected accuracy, success probability and space usage.

License

Topics

Related publications

    -
  • Karayel, E. (2022). Formalization of Randomized Approximation Algorithms for Frequency Moments. Schloss Dagstuhl - Leibniz-Zentrum Für Informatik. https://doi.org/10.4230/LIPICS.ITP.2022.21 +
  • Karayel, E. (2022). Formalization of Randomized Approximation Algorithms for Frequency Moments. Schloss Dagstuhl – Leibniz-Zentrum für Informatik. https://doi.org/10.4230/LIPICS.ITP.2022.21

Session Frequency_Moments

\ No newline at end of file diff --git a/web/entries/HoareForDivergence.html b/web/entries/HoareForDivergence.html --- a/web/entries/HoareForDivergence.html +++ b/web/entries/HoareForDivergence.html @@ -1,248 +1,248 @@ A Hoare Logic for Diverging Programs - Archive of Formal Proofs

A Hoare Logic for Diverging Programs

Abstract

This submission contains:
  1. a formalisation of a small While language with support for output;
  2. a standard total-correctness Hoare logic that has been proved sound and complete; and
  3. a new Hoare logic for proofs about programs that diverge: this new logic has also been proved sound and complete.

License

Topics

Related publications

    -
  • Åman Pohjola, J., Rostedt, H., & Myreen, M. O. (2019). Characteristic Formulae for Liveness Properties of Non-Terminating CakeML Programs. Schloss Dagstuhl - Leibniz-Zentrum Fuer Informatik GmbH, Wadern/Saarbruecken, Germany. https://doi.org/10.4230/LIPICS.ITP.2019.32 +
  • Åman Pohjola, J., Rostedt, H., & Myreen, M. O. (2019). Characteristic Formulae for Liveness Properties of Non-Terminating CakeML Programs (Version 1.0). Schloss Dagstuhl – Leibniz-Zentrum für Informatik. https://doi.org/10.4230/LIPICS.ITP.2019.32

Session HoareForDivergence

Depends on

\ No newline at end of file diff --git a/web/entries/HyperHoareLogic.html b/web/entries/HyperHoareLogic.html --- a/web/entries/HyperHoareLogic.html +++ b/web/entries/HyperHoareLogic.html @@ -1,233 +1,233 @@ Formalization of Hyper Hoare Logic: A Logic to (Dis-)Prove Program Hyperproperties - Archive of Formal Proofs

Formalization of Hyper Hoare Logic: A Logic to (Dis-)Prove Program Hyperproperties

Thibault Dardinier 📧

April 3, 2023

Abstract

Hoare logics are proof systems that allow one to formally establish properties of computer programs. Traditional Hoare logics prove properties of individual program executions (so-called trace properties, such as functional correctness). On the one hand, Hoare logic has been generalized to prove properties of multiple executions of a program (so-called hyperproperties, such as determinism or non-interference). These program logics prove the absence of (bad combinations of) executions. On the other hand, program logics similar to Hoare logic have been proposed to disprove program properties (e.g., Incorrectness Logic [8]), by proving the existence of (bad combinations of) executions. All of these logics have in common that they specify program properties using assertions over a fixed number of states, for instance, a single pre- and post-state for functional properties or pairs of pre- and post-states for non-interference. In this entry, we formalize Hyper Hoare Logic, a generalization of Hoare logic that lifts assertions to properties of arbitrary sets of states. The resulting logic is simple yet expressive: its judgments can express arbitrary trace- and hyperproperties over the terminating executions of a program. By allowing assertions to reason about sets of states, Hyper Hoare Logic can reason about both the absence and the existence of (combinations of) executions, and, thereby, supports both proving and disproving program (hyper-)properties within the same logic. In fact, we prove that Hyper Hoare Logic subsumes the properties handled by numerous existing correctness and incorrectness logics, and can express hyperproperties that no existing Hoare logic can. We also prove that Hyper Hoare Logic is sound and complete.

License

Topics

Related publications

    -
  • Dardinier, T., & Müller, P. (2023). Hyper Hoare Logic: (Dis-)Proving Program Hyperproperties (extended version) (Version 2). arXiv. https://doi.org/10.48550/ARXIV.2301.10037 +
  • Dardinier, T., & Müller, P. (2023). Hyper Hoare Logic: (Dis-)Proving Program Hyperproperties (extended version) (Version 3). arXiv. https://doi.org/10.48550/ARXIV.2301.10037

Session HyperHoareLogic

\ No newline at end of file diff --git a/web/entries/Hypergraph_Colourings.html b/web/entries/Hypergraph_Colourings.html --- a/web/entries/Hypergraph_Colourings.html +++ b/web/entries/Hypergraph_Colourings.html @@ -1,231 +1,231 @@ Hypergraph Colouring Bounds - Archive of Formal Proofs

Hypergraph Colouring Bounds

Chelsea Edmonds 📧 and Lawrence C. Paulson 📧

September 23, 2023

Abstract

This library includes several example applications of the probabilistic method for combinatorics to establish bounds for hypergraph colourings. This focuses on Property B - the existence of a two-colouring of the vertex set of a hypergraph. A stricter bound was formalised using the Lovász local lemma, which in turn required a surprisingly complex proof of the mutual independence principle for hypergraph edges that is often omitted on paper. The formalisation uncovered several interesting examples of circular intuition on proofs involving independence on paper. The formalisation is based on the textbook proofs from Alon and Spencer's famous textbook, The Probabilistic Method, further supported by notes by Zhao. The mutual independence principle proof is inspired by the less precise proof provided in Molloy and Reed's textbook on graph colourings, as it was omitted in all other sources. Additionally, this library demonstrates how locales can be used to establish a reusable probability space framework, thus minimizing the setup required for future formalisations requiring a probability space on numerous possible properties around an incidence system's vertex set.

License

Topics

Related publications

    -
  • Molloy, M., & Reed, B. (2002). Graph Colouring and the Probabilistic Method. Algorithms and Combinatorics. https://doi.org/10.1007/978-3-642-04016-0 +
  • Molloy, M., & Reed, B. (2002). Graph Colouring and the Probabilistic Method. In Algorithms and Combinatorics. Springer Berlin Heidelberg. https://doi.org/10.1007/978-3-642-04016-0
  • https://dl.acm.org/doi/book/10.5555/3002498
  • https://yufeizhao.com/pm/probmethod_notes.pdf

Session Hypergraph_Colourings

\ No newline at end of file diff --git a/web/entries/List_Update.html b/web/entries/List_Update.html --- a/web/entries/List_Update.html +++ b/web/entries/List_Update.html @@ -1,321 +1,321 @@ Analysis of List Update Algorithms - Archive of Formal Proofs

Analysis of List Update Algorithms

Abstract

These theories formalize the quantitative analysis of a number of classical algorithms for the list update problem: 2-competitiveness of move-to-front, the lower bound of 2 for the competitiveness of deterministic list update algorithms and 1.6-competitiveness of the randomized COMB algorithm, the best randomized list update algorithm known to date. The material is based on the first two chapters of Online Computation and Competitive Analysis by Borodin and El-Yaniv.

For an informal description see the FSTTCS 2016 publication Verified Analysis of List Update Algorithms by Haslbeck and Nipkow.

License

Topics

Related publications

    -
  • Haslbeck, M. P. L., & Nipkow, T. (2016). Verified Analysis of List Update Algorithms. Schloss Dagstuhl - Leibniz-Zentrum Fuer Informatik GmbH, Wadern/Saarbruecken, Germany. https://doi.org/10.4230/LIPICS.FSTTCS.2016.49 +
  • Haslbeck, M. P. L., & Nipkow, T. (2016). Verified Analysis of List Update Algorithms. Schloss Dagstuhl – Leibniz-Zentrum für Informatik. https://doi.org/10.4230/LIPICS.FSTTCS.2016.49

Session List_Update

\ No newline at end of file diff --git a/web/entries/OmegaCatoidsQuantales.html b/web/entries/OmegaCatoidsQuantales.html --- a/web/entries/OmegaCatoidsQuantales.html +++ b/web/entries/OmegaCatoidsQuantales.html @@ -1,249 +1,249 @@ Higher Globular Catoids and Quantales - Archive of Formal Proofs

Higher Globular Catoids and Quantales

Cameron Calk and Georg Struth 📧

January 24, 2024

Abstract

We formalise strict 2-catoids, 2-categories, 2-Kleene algebras and 2-quantales, as well as their omega-variants. Due to strictness, the cells of these higher categories have globular shape. We use a single-set approach, generalised to catoids and based on type classes. The higher Kleene algebras and quantales introduced extend features of modal and concurrent Kleene algebras and quantales. They arise for instance as powerset extensions of higher catoids, and have been used in algebraic confluence proofs in higher-dimensional rewriting. Details are described in two companion articles.

License

Topics

Related publications

  • Calk, C., Goubault, E., Malbos, P., & Struth, G. (2022). Algebraic coherent confluence and higher globular Kleene algebras. Logical Methods in Computer Science, Volume 18, Issue 4. https://doi.org/10.46298/lmcs-18(4:9)2022
  • -
  • Calk, C., Malbos, P., Pous, D., & Struth, G. (2023). Higher Catoids, Higher Quantales and their Correspondences (Version 1). arXiv. https://doi.org/10.48550/ARXIV.2307.09253 +
  • Calk, C., Malbos, P., Pous, D., & Struth, G. (2023). Higher Catoids, Higher Quantales and their Correspondences (Version 2). arXiv. https://doi.org/10.48550/ARXIV.2307.09253

Session OmegaCatoidsQuantales

\ No newline at end of file diff --git a/web/entries/Polygonal_Number_Theorem.html b/web/entries/Polygonal_Number_Theorem.html --- a/web/entries/Polygonal_Number_Theorem.html +++ b/web/entries/Polygonal_Number_Theorem.html @@ -1,246 +1,246 @@ Polygonal Number Theorem - Archive of Formal Proofs

Polygonal Number Theorem

Abstract

We formalize the proofs of Cauchy's and Legendre's Polygonal Number Theorems given in Melvyn B. Nathanson's book "Additive Number Theory: The Classical Bases".

For $m \geq 1$, the $k$-th polygonal number of order $m+2$ is defined to be $p_m(k)=\frac{mk(k-1)}{2}+k$. The theorems state that:

1. If $m \ge 4$ and $N \geq 108m$, then $N$ can be written as the sum of $m+1$ polygonal numbers of order $m+2$, at most four of which are different from $0$ or $1$. If $N \geq 324$, then $N$ can be written as the sum of five pentagonal numbers, at least one of which is $0$ or $1$.

2. Let $m \geq 3$ and $N \geq 28m^3$. If $m$ is odd, then $N$ is the sum of four polygonal numbers of order $m+2$. If $m$ is even, then $N$ is the sum of five polygonal numbers of order $m+2$, at least one of which is $0$ or $1$.

We also formalize the proof of Gauss's theorem which states that every non-negative integer is the sum of three triangular numbers.

License

Topics

Related publications

Session Polygonal_Number_Theorem

\ No newline at end of file diff --git a/web/entries/Prim_Dijkstra_Simple.html b/web/entries/Prim_Dijkstra_Simple.html --- a/web/entries/Prim_Dijkstra_Simple.html +++ b/web/entries/Prim_Dijkstra_Simple.html @@ -1,274 +1,274 @@ Purely Functional, Simple, and Efficient Implementation of Prim and Dijkstra - Archive of Formal Proofs

Purely Functional, Simple, and Efficient Implementation of Prim and Dijkstra

Peter Lammich 🌐 and Tobias Nipkow 🌐

June 25, 2019

Abstract

We verify purely functional, simple and efficient implementations of Prim's and Dijkstra's algorithms. This constitutes the first verification of an executable and even efficient version of Prim's algorithm. This entry formalizes the second part of our ITP-2019 proof pearl Purely Functional, Simple and Efficient Priority Search Trees and Applications to Prim and Dijkstra.

License

Topics

Related publications

    -
  • Lammich, P., & Nipkow, T. (2019). Proof Pearl: Purely Functional, Simple and Efficient Priority Search Trees and Applications to Prim and Dijkstra. Schloss Dagstuhl - Leibniz-Zentrum Fuer Informatik GmbH, Wadern/Saarbruecken, Germany. https://doi.org/10.4230/LIPICS.ITP.2019.23 +
  • Lammich, P., & Nipkow, T. (2019). Proof Pearl: Purely Functional, Simple and Efficient Priority Search Trees and Applications to Prim and Dijkstra (Version 1.0). Schloss Dagstuhl – Leibniz-Zentrum für Informatik. https://doi.org/10.4230/LIPICS.ITP.2019.23

Session Prim_Dijkstra_Simple

\ No newline at end of file diff --git a/web/entries/Propositional_Proof_Systems.html b/web/entries/Propositional_Proof_Systems.html --- a/web/entries/Propositional_Proof_Systems.html +++ b/web/entries/Propositional_Proof_Systems.html @@ -1,333 +1,333 @@ Propositional Proof Systems - Archive of Formal Proofs

Propositional Proof Systems

Abstract

We formalize a range of proof systems for classical propositional logic (sequent calculus, natural deduction, Hilbert systems, resolution) and prove the most important meta-theoretic results about semantics and proofs: compactness, soundness, completeness, translations between proof systems, cut-elimination, interpolation and model existence.

License

Topics

Related publications

    -
  • Michaelis, J., & Nipkow, T. (2018). Formalized Proof Systems for Propositional Logic. Schloss Dagstuhl - Leibniz-Zentrum Fuer Informatik GmbH, Wadern/Saarbruecken, Germany. https://doi.org/10.4230/LIPICS.TYPES.2017.5 +
  • Michaelis, J., & Nipkow, T. (2019). Formalized Proof Systems for Propositional Logic (Version 1.0). Schloss Dagstuhl – Leibniz-Zentrum für Informatik. https://doi.org/10.4230/LIPICS.TYPES.2017.5

Session Propositional_Proof_Systems

\ No newline at end of file diff --git a/web/entries/Quantales_Converse.html b/web/entries/Quantales_Converse.html --- a/web/entries/Quantales_Converse.html +++ b/web/entries/Quantales_Converse.html @@ -1,262 +1,262 @@ Modal quantales, involutive quantales, Dedekind Quantales - Archive of Formal Proofs

Modal quantales, involutive quantales, Dedekind Quantales

Georg Struth 📧 and Cameron Calk

July 25, 2023

Abstract

This AFP entry provides mathematical components for modal quantales, involutive quantales and Dedekind quantales. Modal quantales are simple extensions of modal Kleene algebras useful for the verification of recursive programs. Involutive quantales appear in the study of C*-algebras. Dedekind quantales are relatives of Tarski's relation algebras, hence relevant to program verification and beyond that to higher rewriting. We also provide components for weaker variants such as Kleene algebras with converse and modal Kleene algebras with converse.

License

Topics

Related publications

  • Fahrenberg, U., Johansen, C., Struth, G., & Ziemiański, K. (2023). Catoids and modal convolution algebras. Algebra Universalis, 84(2). https://doi.org/10.1007/s00012-023-00805-9
  • -
  • Calk, C., Malbos, P., Pous, D., & Struth, G. (2023). Higher Catoids, Higher Quantales and their Correspondences (Version 1). arXiv. https://doi.org/10.48550/ARXIV.2307.09253 +
  • Calk, C., Malbos, P., Pous, D., & Struth, G. (2023). Higher Catoids, Higher Quantales and their Correspondences (Version 2). arXiv. https://doi.org/10.48550/ARXIV.2307.09253

Session Quantales_Converse

\ No newline at end of file diff --git a/web/entries/S_Finite_Measure_Monad.html b/web/entries/S_Finite_Measure_Monad.html --- a/web/entries/S_Finite_Measure_Monad.html +++ b/web/entries/S_Finite_Measure_Monad.html @@ -1,244 +1,244 @@ S-Finite Measure Monad on Quasi-Borel Spaces - Archive of Formal Proofs

S-Finite Measure Monad on Quasi-Borel Spaces

Abstract

The s-finite measure monad on quasi-Borel spaces provides a suitable denotational model for higher-order probabilistic programs with conditioning. This entry is a formalization of the s-finite measure monad and related notions, including s-finite measures, s-finite kernels, and a proof automation for quasi-Borel spaces which is an extension of our previous entry Quasi-Borel Spaces. We also implement several examples of probabilistic programs in previous works and prove their property. This work is a part of the work by Hirata, Minamide, and Sato, Semantic Foundations of Higher-Order Probabilistic Programs in Isabelle/HOL presented at the 14th Conference on Interactive Theorem Proving (ITP2023).

License

Topics

Related publications

    -
  • Hirata, M., Minamide, Y., & Sato, T. (2023). Semantic Foundations of Higher-Order Probabilistic Programs in Isabelle/HOL. Schloss Dagstuhl - Leibniz-Zentrum Für Informatik. https://doi.org/10.4230/LIPICS.ITP.2023.18 +
  • Hirata, M., Minamide, Y., & Sato, T. (2023). Semantic Foundations of Higher-Order Probabilistic Programs in Isabelle/HOL. Schloss Dagstuhl – Leibniz-Zentrum für Informatik. https://doi.org/10.4230/LIPICS.ITP.2023.18

Session S_Finite_Measure_Monad

\ No newline at end of file diff --git a/web/entries/Simpl.html b/web/entries/Simpl.html --- a/web/entries/Simpl.html +++ b/web/entries/Simpl.html @@ -1,397 +1,402 @@ A Sequential Imperative Programming Language Syntax, Semantics, Hoare Logics and Verification Environment - Archive of Formal Proofs

A Sequential Imperative Programming Language Syntax, Semantics, Hoare Logics and Verification Environment

Norbert Schirmer

February 29, 2008

Abstract

We present the theory of Simpl, a sequential imperative programming language. We introduce its syntax, its semantics (big and small-step operational semantics) and Hoare logics for both partial as well as total correctness. We prove soundness and completeness of the Hoare logic. We integrate and automate the Hoare logic in Isabelle/HOL to obtain a practically usable verification environment for imperative programs. Simpl is independent of a concrete programming language but expressive enough to cover all common language features: mutually recursive procedures, abrupt termination and exceptions, runtime faults, local and global variables, pointers and heap, expressions with side effects, pointers to procedures, partial application and closures, dynamic method invocation and also unbounded nondeterminism.

License

+

History

+
+
February 18, 2024
+
License change to BSD; more canonical and versatile ML interfaces; some more derived language constructs.
+

Topics

Session Simpl

\ No newline at end of file diff --git a/web/entries/Standard_Borel_Spaces.html b/web/entries/Standard_Borel_Spaces.html --- a/web/entries/Standard_Borel_Spaces.html +++ b/web/entries/Standard_Borel_Spaces.html @@ -1,229 +1,229 @@ Standard Borel Spaces - Archive of Formal Proofs

Standard Borel Spaces

Michikazu Hirata 📧

August 8, 2023

Abstract

This entry includes a formalization of standard Borel spaces and (a variant of) the Borel isomorphism theorem. A separable complete metrizable topological space is called a polish space and a measurable space generated from a polish space is called a standard Borel space. We formalize the notion of standard Borel spaces by establishing set-based metric spaces, and then prove (a variant of) the Borel isomorphism theorem. The theorem states that a standard Borel spaces is either a countable discrete space or isomorphic to $\mathbb{R}$.

License

Topics

Related publications

    -
  • Hirata, M., Minamide, Y., & Sato, T. (2023). Semantic Foundations of Higher-Order Probabilistic Programs in Isabelle/HOL. Schloss Dagstuhl - Leibniz-Zentrum Für Informatik. https://doi.org/10.4230/LIPICS.ITP.2023.18 +
  • Hirata, M., Minamide, Y., & Sato, T. (2023). Semantic Foundations of Higher-Order Probabilistic Programs in Isabelle/HOL. Schloss Dagstuhl – Leibniz-Zentrum für Informatik. https://doi.org/10.4230/LIPICS.ITP.2023.18

Session Standard_Borel_Spaces

\ No newline at end of file diff --git a/web/entries/Transport.html b/web/entries/Transport.html --- a/web/entries/Transport.html +++ b/web/entries/Transport.html @@ -1,350 +1,350 @@ Transport via Partial Galois Connections and Equivalences - Archive of Formal Proofs

Transport via Partial Galois Connections and Equivalences

Kevin Kappelmann 📧

October 11, 2023

Abstract

This entry contains the accompanying formalisation of the paper "Transport via Partial Galois Connections and Equivalences" (APLAS 2023). It contains a theoretical framework to transport programs via equivalences, subsuming the theory of Isabelle's Lifting package. It also contains a prototype to automate transports using this framework in Isabelle/HOL, but this prototype is not yet ready for production. Finally, it contains a library on top of Isabelle/HOL's axioms, including various relativised concepts on orders, functions, binary relations, and Galois connections and equivalences.

License

Topics

Related publications

    -
  • Kappelmann, K. (2023). Transport via Partial Galois Connections and Equivalences (Version 4). arXiv. https://doi.org/10.48550/ARXIV.2303.05244 +
  • Kappelmann, K. (2023). Transport via Partial Galois Connections and Equivalences. ArXiv. https://doi.org/10.48550/ARXIV.2303.05244

Session Transport

\ No newline at end of file diff --git a/web/statistics/index.html b/web/statistics/index.html --- a/web/statistics/index.html +++ b/web/statistics/index.html @@ -1,383 +1,383 @@ Statistics - Archive of Formal Proofs - + - +

Statistics

798 Entries
475 Authors
- ~252,000 + ~252,100 Lemmas
- ~4,087,700 + ~4,088,600 Lines of Code

Most used AFP entries:

Name Used by ? entries
1. List-Index 23
2. Collections 19
3. Show 16
4. Jordan_Normal_Form 15
5. Deriving 14
6. Coinductive 13
7. Polynomial_Factorization 12
8. Regular-Sets 12
9. Landau_Symbols 11
10. Abstract-Rewriting 10
11. Automatic_Refinement 10
12. Native_Word 10

Growth in number of entries:

Growth in lines of code:

Growth in number of authors:

Size of entries:

\ No newline at end of file