# HG changeset patch # User prensani # Date 1015344685 -3600 # Node ID 791e3b4c403957598091e827375a567449b1379e # Parent 98f0a09a33c343206d8710ef246f580a762f98fe HoareParallel Theories diff -r 98f0a09a33c3 -r 791e3b4c4039 src/HOL/HoareParallel/Gar_Coll.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HoareParallel/Gar_Coll.thy Tue Mar 05 17:11:25 2002 +0100 @@ -0,0 +1,880 @@ + +header {* \section{The Single Mutator Case} *} + +theory Gar_Coll = Graph + OG_Syntax: + +text {* Declaration of variables: *} + +record gar_coll_state = + M :: nodes + E :: edges + bc :: "nat set" + obc :: "nat set" + Ma :: nodes + ind :: nat + k :: nat + z :: bool + +subsection {* The Mutator *} + +text {* The mutator first redirects an arbitrary edge @{text "R"} from +an arbitrary accessible node towards an arbitrary accessible node +@{text "T"}. It then colors the new target @{text "T"} black. + +We declare the arbitrarily selected node and edge as constants:*} + +consts R :: nat T :: nat + +text {* \noindent The following predicate states, given a list of +nodes @{text "m"} and a list of edges @{text "e"}, the conditions +under which the selected edge @{text "R"} and node @{text "T"} are +valid: *} + +constdefs + Mut_init :: "gar_coll_state \ bool" + "Mut_init \ \ T \ Reach \E \ R < length \E \ T < length \M \" + +text {* \noindent For the mutator we +consider two modules, one for each action. An auxiliary variable +@{text "\z"} is set to false if the mutator has already redirected an +edge but has not yet colored the new target. *} + +constdefs + Redirect_Edge :: "gar_coll_state ann_com" + "Redirect_Edge \ .{\Mut_init \ \z}. \\E:=\E[R:=(fst(\E!R), T)],, \z:= (\\z)\" + + Color_Target :: "gar_coll_state ann_com" + "Color_Target \ .{\Mut_init \ \\z}. \\M:=\M[T:=Black],, \z:= (\\z)\" + + Mutator :: "gar_coll_state ann_com" + "Mutator \ + .{\Mut_init \ \z}. + WHILE True INV .{\Mut_init \ \z}. + DO Redirect_Edge ;; Color_Target OD" + +subsubsection {* Correctness of the mutator *} + +lemmas mutator_defs = Mut_init_def Redirect_Edge_def Color_Target_def + +lemma Redirect_Edge: + "\ Redirect_Edge pre(Color_Target)" +apply (unfold mutator_defs) +apply annhoare +apply(simp_all) +apply(force elim:Graph2) +done + +lemma Color_Target: + "\ Color_Target .{\Mut_init \ \z}." +apply (unfold mutator_defs) +apply annhoare +apply(simp_all) +done + +lemma Mutator: + "\ Mutator .{False}." +apply(unfold Mutator_def) +apply annhoare +apply(simp_all add:Redirect_Edge Color_Target) +apply(simp add:mutator_defs Redirect_Edge_def) +done + +subsection {* The Collector *} + +text {* \noindent A constant @{text "M_init"} is used to give @{text "\Ma"} a +suitable first value, defined as a list of nodes where only the @{text +"Roots"} are black. *} + +consts M_init :: nodes + +constdefs + Proper_M_init :: "gar_coll_state \ bool" + "Proper_M_init \ \ Blacks M_init=Roots \ length M_init=length \M \" + + Proper :: "gar_coll_state \ bool" + "Proper \ \ Proper_Roots \M \ Proper_Edges(\M, \E) \ \Proper_M_init \" + + Safe :: "gar_coll_state \ bool" + "Safe \ \ Reach \E \ Blacks \M \" + +lemmas collector_defs = Proper_M_init_def Proper_def Safe_def + +subsubsection {* Blackening the roots *} + +constdefs + Blacken_Roots :: " gar_coll_state ann_com" + "Blacken_Roots \ + .{\Proper}. + \ind:=0;; + .{\Proper \ \ind=0}. + WHILE \indM + INV .{\Proper \ (\i<\ind. i \ Roots \ \M!i=Black) \ \ind\length \M}. + DO .{\Proper \ (\i<\ind. i \ Roots \ \M!i=Black) \ \indM}. + IF \ind\Roots THEN + .{\Proper \ (\i<\ind. i \ Roots \ \M!i=Black) \ \indM \ \ind\Roots}. + \M:=\M[\ind:=Black] FI;; + .{\Proper \ (\i<\ind+1. i \ Roots \ \M!i=Black) \ \indM}. + \ind:=\ind+1 + OD" + +lemma Blacken_Roots: + "\ Blacken_Roots .{\Proper \ Roots\Blacks \M}." +apply (unfold Blacken_Roots_def) +apply annhoare +apply(simp_all add:collector_defs Graph_defs) +apply safe +apply(simp_all add:nth_list_update) + apply (erule less_SucE) + apply simp+ + apply (erule less_SucE) + apply simp+ + apply(drule le_imp_less_or_eq) + apply force +apply force +done + +subsubsection {* Propagating black *} + +constdefs + PBInv :: "gar_coll_state \ nat \ bool" + "PBInv \ \ \ind. \obc < Blacks \M \ (\i BtoW (\E!i, \M) \ + (\\z \ i=R \ (snd(\E!R)) = T \ (\r. ind \ r \ r < length \E \ BtoW(\E!r,\M))))\" + +constdefs + Propagate_Black_aux :: "gar_coll_state ann_com" + "Propagate_Black_aux \ + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M}. + \ind:=0;; + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M \ \ind=0}. + WHILE \indE + INV .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv \ind \ \ind\length \E}. + DO .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv \ind \ \indE}. + IF \M!(fst (\E!\ind)) = Black THEN + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv \ind \ \indE \ \M!fst(\E!\ind)=Black}. + \M:=\M[snd(\E!\ind):=Black];; + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv (\ind + 1) \ \indE}. + \ind:=\ind+1 + FI + OD" + +lemma Propagate_Black_aux: + "\ Propagate_Black_aux + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ ( \obc < Blacks \M \ \Safe)}." +apply (unfold Propagate_Black_aux_def PBInv_def collector_defs) +apply annhoare +apply(simp_all add:Graph6 Graph7 Graph8 Graph12) + apply force + apply force + apply force +--{* 4 subgoals left *} +apply clarify +apply(simp add:Proper_Edges_def Proper_Roots_def Graph6 Graph7 Graph8 Graph12) +apply (erule disjE) + apply(rule disjI1) + apply(erule Graph13) + apply force +apply (case_tac "M x ! snd (E x ! ind x)=Black") + apply (simp add: Graph10 BtoW_def) + apply (rule disjI2) + apply clarify + apply (erule less_SucE) + apply (erule_tac x=i in allE , erule (1) notE impE) + apply simp + apply clarify + apply (drule le_imp_less_or_eq) + apply (erule disjE) + apply (subgoal_tac "Suc (ind x)\r") + apply fast + apply arith + apply fast + apply fast +apply(rule disjI1) +apply(erule subset_psubset_trans) +apply(erule Graph11) +apply fast +--{* 3 subgoals left *} +apply force +apply force +--{* last *} +apply clarify +apply simp +apply(subgoal_tac "ind x = length (E x)") + apply (rotate_tac -1) + apply simp + apply(drule Graph1) + apply simp + apply clarify + apply(erule allE, erule impE, assumption) + apply force + apply force +apply arith +done + +subsubsection {* Refining propagating black *} + +constdefs + Auxk :: "gar_coll_state \ bool" + "Auxk \ \\kM \ (\M!\k\Black \ \BtoW(\E!\ind, \M) \ + \obcM \ (\\z \ \ind=R \ snd(\E!R)=T + \ (\r. \ind rE \ BtoW(\E!r, \M))))\" + +constdefs + Propagate_Black :: " gar_coll_state ann_com" + "Propagate_Black \ + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M}. + \ind:=0;; + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M \ \ind=0}. + WHILE \indE + INV .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv \ind \ \ind\length \E}. + DO .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv \ind \ \indE}. + IF (\M!(fst (\E!\ind)))=Black THEN + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv \ind \ \indE \ (\M!fst(\E!\ind))=Black}. + \k:=(snd(\E!\ind));; + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv \ind \ \indE \ (\M!fst(\E!\ind))=Black + \ \Auxk}. + \\M:=\M[\k:=Black],, \ind:=\ind+1\ + ELSE .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv \ind \ \indE}. + \IF (\M!(fst (\E!\ind)))\Black THEN \ind:=\ind+1 FI\ + FI + OD" + +lemma Propagate_Black: + "\ Propagate_Black + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ ( \obc < Blacks \M \ \Safe)}." +apply (unfold Propagate_Black_def PBInv_def Auxk_def collector_defs) +apply annhoare +apply(simp_all add:Graph6 Graph7 Graph8 Graph12) + apply force + apply force + apply force +--{* 5 subgoals left *} +apply clarify +apply(simp add:BtoW_def Proper_Edges_def) +--{* 4 subgoals left *} +apply clarify +apply(simp add:Proper_Edges_def Graph6 Graph7 Graph8 Graph12) +apply (erule disjE) + apply (rule disjI1) + apply (erule psubset_subset_trans) + apply (erule Graph9) +apply (case_tac "M x!k x=Black") + apply (case_tac "M x ! snd (E x ! ind x)=Black") + apply (simp add: Graph10 BtoW_def) + apply (rule disjI2) + apply clarify + apply (erule less_SucE) + apply (erule_tac x=i in allE , erule (1) notE impE) + apply simp + apply clarify + apply (drule le_imp_less_or_eq) + apply (erule disjE) + apply (subgoal_tac "Suc (ind x)\r") + apply fast + apply arith + apply fast + apply fast + apply (simp add: Graph10 BtoW_def) + apply (erule disjE) + apply (erule disjI1) + apply clarify + apply (erule less_SucE) + apply force + apply simp + apply (subgoal_tac "Suc R\r") + apply fast + apply arith +apply(rule disjI1) +apply(erule subset_psubset_trans) +apply(erule Graph11) +apply fast +--{* 3 subgoals left *} +apply force +--{* 2 subgoals left *} +apply clarify +apply(simp add:Proper_Edges_def Graph6 Graph7 Graph8 Graph12) +apply (erule disjE) + apply fast +apply clarify +apply (erule less_SucE) + apply (erule_tac x=i in allE , erule (1) notE impE) + apply simp + apply clarify + apply (drule le_imp_less_or_eq) + apply (erule disjE) + apply (subgoal_tac "Suc (ind x)\r") + apply fast + apply arith + apply (simp add: BtoW_def) +apply (simp add: BtoW_def) +--{* last *} +apply clarify +apply simp +apply(subgoal_tac "ind x = length (E x)") + apply (rotate_tac -1) + apply simp + apply(drule Graph1) + apply simp + apply clarify + apply(erule allE, erule impE, assumption) + apply force + apply force +apply arith +done + +subsubsection {* Counting black nodes *} + +constdefs + CountInv :: "gar_coll_state \ nat \ bool" + "CountInv \ \ \ind. {i. i \Ma!i=Black}\\bc \" + +constdefs + Count :: " gar_coll_state ann_com" + "Count \ + .{\Proper \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ (\obc < Blacks \Ma \ \Safe) \ \bc={}}. + \ind:=0;; + .{\Proper \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ (\obc < Blacks \Ma \ \Safe) \ \bc={} + \ \ind=0}. + WHILE \indM + INV .{\Proper \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ \CountInv \ind + \ ( \obc < Blacks \Ma \ \Safe) \ \ind\length \M}. + DO .{\Proper \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ \CountInv \ind + \ ( \obc < Blacks \Ma \ \Safe) \ \indM}. + IF \M!\ind=Black + THEN .{\Proper \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ \CountInv \ind + \ ( \obc < Blacks \Ma \ \Safe) \ \indM \ \M!\ind=Black}. + \bc:=insert \ind \bc + FI;; + .{\Proper \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ \CountInv (\ind+1) + \ ( \obc < Blacks \Ma \ \Safe) \ \indM}. + \ind:=\ind+1 + OD" + +lemma Count: + "\ Count + .{\Proper \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\\bc \ \bc\Blacks \M \ length \Ma=length \M + \ (\obc < Blacks \Ma \ \Safe)}." +apply(unfold Count_def) +apply annhoare +apply(simp_all add:CountInv_def Graph6 Graph7 Graph8 Graph12 Blacks_def collector_defs) + apply force + apply force + apply force + apply clarify + apply simp + apply(fast elim:less_SucE) + apply clarify + apply simp + apply(fast elim:less_SucE) + apply force +apply force +done + +subsubsection {* Appending garbage nodes to the free list *} + +consts Append_to_free :: "nat \ edges \ edges" + +axioms + Append_to_free0: "length (Append_to_free (i, e)) = length e" + Append_to_free1: "Proper_Edges (m, e) + \ Proper_Edges (m, Append_to_free(i, e))" + Append_to_free2: "i \ Reach e + \ n \ Reach (Append_to_free(i, e)) = ( n = i \ n \ Reach e)" + +constdefs + AppendInv :: "gar_coll_state \ nat \ bool" + "AppendInv \ \\ind. \iM. ind\i \ i\Reach \E \ \M!i=Black\" + +constdefs + Append :: " gar_coll_state ann_com" + "Append \ + .{\Proper \ Roots\Blacks \M \ \Safe}. + \ind:=0;; + .{\Proper \ Roots\Blacks \M \ \Safe \ \ind=0}. + WHILE \indM + INV .{\Proper \ \AppendInv \ind \ \ind\length \M}. + DO .{\Proper \ \AppendInv \ind \ \indM}. + IF \M!\ind=Black THEN + .{\Proper \ \AppendInv \ind \ \indM \ \M!\ind=Black}. + \M:=\M[\ind:=White] + ELSE .{\Proper \ \AppendInv \ind \ \indM \ \ind\Reach \E}. + \E:=Append_to_free(\ind,\E) + FI;; + .{\Proper \ \AppendInv (\ind+1) \ \indM}. + \ind:=\ind+1 + OD" + +lemma Append: + "\ Append .{\Proper}." +apply(unfold Append_def AppendInv_def) +apply annhoare +apply(simp_all add:collector_defs Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) + apply(force simp:Blacks_def nth_list_update) + apply force + apply force + apply(force simp add:Graph_defs) + apply force + apply clarify + apply simp + apply(rule conjI) + apply (erule Append_to_free1) + apply clarify + apply (drule_tac n = "i" in Append_to_free2) + apply force + apply force +apply force +done + +subsubsection {* Correctness of the Collector *} + +constdefs + Collector :: " gar_coll_state ann_com" + "Collector \ +.{\Proper}. + WHILE True INV .{\Proper}. + DO + Blacken_Roots;; + .{\Proper \ Roots\Blacks \M}. + \obc:={};; + .{\Proper \ Roots\Blacks \M \ \obc={}}. + \bc:=Roots;; + .{\Proper \ Roots\Blacks \M \ \obc={} \ \bc=Roots}. + \Ma:=M_init;; + .{\Proper \ Roots\Blacks \M \ \obc={} \ \bc=Roots \ \Ma=M_init}. + WHILE \obc\\bc + INV .{\Proper \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\\bc \ \bc\Blacks \M + \ length \Ma=length \M \ (\obc < Blacks \Ma \ \Safe)}. + DO .{\Proper \ Roots\Blacks \M \ \bc\Blacks \M}. + \obc:=\bc;; + Propagate_Black;; + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ (\obc < Blacks \M \ \Safe)}. + \Ma:=\M;; + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \Ma + \ Blacks \Ma\Blacks \M \ \bc\Blacks \M \ length \Ma=length \M + \ ( \obc < Blacks \Ma \ \Safe)}. + \bc:={};; + Count + OD;; + Append + OD" + +lemma Collector: + "\ Collector .{False}." +apply(unfold Collector_def) +apply annhoare +apply(simp_all add: Blacken_Roots Propagate_Black Count Append) +apply(simp_all add:Blacken_Roots_def Propagate_Black_def Count_def Append_def collector_defs) + apply (force simp add: Proper_Roots_def) + apply force + apply force +apply clarify +apply (erule disjE) +apply(simp add:psubsetI) + apply(force dest:subset_antisym) +apply force +done + +subsection {* Interference Freedom *} + +lemmas modules = Redirect_Edge_def Color_Target_def Blacken_Roots_def + Propagate_Black_def Count_def Append_def +lemmas Invariants = PBInv_def Auxk_def CountInv_def AppendInv_def +lemmas abbrev = collector_defs mutator_defs Invariants + +lemma interfree_Blacken_Roots_Redirect_Edge: + "interfree_aux (Some Blacken_Roots, {}, Some Redirect_Edge)" +apply (unfold modules) +apply interfree_aux +apply safe +apply (simp_all add:Graph6 Graph12 abbrev) +done + +lemma interfree_Redirect_Edge_Blacken_Roots: + "interfree_aux (Some Redirect_Edge, {}, Some Blacken_Roots)" +apply (unfold modules) +apply interfree_aux +apply safe +apply(simp add:abbrev)+ +done + +lemma interfree_Blacken_Roots_Color_Target: + "interfree_aux (Some Blacken_Roots, {}, Some Color_Target)" +apply (unfold modules) +apply interfree_aux +apply safe +apply(simp_all add:Graph7 Graph8 nth_list_update abbrev) +done + +lemma interfree_Color_Target_Blacken_Roots: + "interfree_aux (Some Color_Target, {}, Some Blacken_Roots)" +apply (unfold modules ) +apply interfree_aux +apply safe +apply(simp add:abbrev)+ +done + +lemma interfree_Propagate_Black_Redirect_Edge: + "interfree_aux (Some Propagate_Black, {}, Some Redirect_Edge)" +apply (unfold modules ) +apply interfree_aux +--{* 11 subgoals left *} +apply(clarify, simp add:abbrev Graph6 Graph12) +apply(clarify, simp add:abbrev Graph6 Graph12) +apply(clarify, simp add:abbrev Graph6 Graph12) +apply(clarify, simp add:abbrev Graph6 Graph12) +apply(erule conjE)+ +apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) + apply(erule Graph4) + apply(simp)+ + apply (simp add:BtoW_def) + apply (simp add:BtoW_def) +apply(rule conjI) + apply (force simp add:BtoW_def) +apply(erule Graph4) + apply simp+ + apply (simp add:BtoW_def) + apply force + apply (simp add:BtoW_def) + apply force +apply (simp add:BtoW_def) +apply force +--{* 7 subgoals left *} +apply(clarify, simp add:abbrev Graph6 Graph12) +apply(erule conjE)+ +apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) + apply(erule Graph4) + apply(simp)+ + apply (simp add:BtoW_def) + apply (simp add:BtoW_def) +apply(rule conjI) + apply (force simp add:BtoW_def) +apply(erule Graph4) + apply simp+ + apply (simp add:BtoW_def) + apply force + apply (simp add:BtoW_def) + apply force +apply (simp add:BtoW_def) +apply force +--{* 6 subgoals left *} +apply(clarify, simp add:abbrev Graph6 Graph12) +apply(erule conjE)+ +apply(rule conjI) + apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) + apply(erule Graph4) + apply(simp)+ + apply (simp add:BtoW_def) + apply (simp add:BtoW_def) + apply(rule conjI) + apply (force simp add:BtoW_def) + apply(erule Graph4) + apply simp+ + apply (simp add:BtoW_def) + apply force + apply (simp add:BtoW_def) + apply force + apply (simp add:BtoW_def) + apply force +apply(simp add:BtoW_def nth_list_update) +apply force +--{* 5 subgoals left *} +apply(clarify, simp add:abbrev Graph6 Graph12) +--{* 4 subgoals left *} +apply(clarify, simp add:abbrev Graph6 Graph12) +apply(rule conjI) + apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) + apply(erule Graph4) + apply(simp)+ + apply (simp add:BtoW_def) + apply (simp add:BtoW_def) + apply(rule conjI) + apply (force simp add:BtoW_def) + apply(erule Graph4) + apply simp+ + apply (simp add:BtoW_def) + apply force + apply (simp add:BtoW_def) + apply force + apply (simp add:BtoW_def) + apply force +apply(rule conjI) + apply(simp add:nth_list_update) + apply force +apply(rule impI, rule impI, erule disjE, erule disjI1, case_tac "R = (ind x)" ,case_tac "M x ! T = Black") + apply(force simp add:BtoW_def) + apply(case_tac "M x !snd (E x ! ind x)=Black") + apply(rule disjI2) + apply simp + apply (erule Graph5) + apply simp+ + apply(force simp add:BtoW_def) +apply(force simp add:BtoW_def) +--{* 3 subgoals left *} +apply(clarify, simp add:abbrev Graph6 Graph12) +--{* 2 subgoals left *} +apply(clarify, simp add:abbrev Graph6 Graph12) +apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) + apply clarify + apply(erule Graph4) + apply(simp)+ + apply (simp add:BtoW_def) + apply (simp add:BtoW_def) +apply(rule conjI) + apply (force simp add:BtoW_def) +apply(erule Graph4) + apply simp+ + apply (simp add:BtoW_def) + apply force + apply (simp add:BtoW_def) + apply force +apply (simp add:BtoW_def) +apply force +--{* 1 subgoals left *} +apply(simp add:abbrev) +done + +lemma interfree_Redirect_Edge_Propagate_Black: + "interfree_aux (Some Redirect_Edge, {}, Some Propagate_Black)" +apply (unfold modules ) +apply interfree_aux +apply(clarify, simp add:abbrev)+ +done + +lemma interfree_Propagate_Black_Color_Target: + "interfree_aux (Some Propagate_Black, {}, Some Color_Target)" +apply (unfold modules ) +apply interfree_aux +--{* 11 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)+ +apply(erule conjE)+ +apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, + case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, + erule allE, erule impE, assumption, erule impE, assumption, + simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) +--{* 7 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) +apply(erule conjE)+ +apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, + case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, + erule allE, erule impE, assumption, erule impE, assumption, + simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) +--{* 6 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) +apply clarify +apply (rule conjI) + apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, + case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, + erule allE, erule impE, assumption, erule impE, assumption, + simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) +apply(simp add:nth_list_update) +--{* 5 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) +--{* 4 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) +apply (rule conjI) + apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, + case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, + erule allE, erule impE, assumption, erule impE, assumption, + simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) +apply(rule conjI) +apply(simp add:nth_list_update) +apply(rule impI,rule impI, case_tac "M x!T=Black",rotate_tac -1, force simp add: BtoW_def Graph10, + erule subset_psubset_trans, erule Graph11, force) +--{* 3 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) +--{* 2 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) +apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, + case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, + erule allE, erule impE, assumption, erule impE, assumption, + simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) +--{* 3 subgoals left *} +apply(simp add:abbrev) +done + +lemma interfree_Color_Target_Propagate_Black: + "interfree_aux (Some Color_Target, {}, Some Propagate_Black)" +apply (unfold modules ) +apply interfree_aux +apply(clarify, simp add:abbrev)+ +done + +lemma interfree_Count_Redirect_Edge: + "interfree_aux (Some Count, {}, Some Redirect_Edge)" +apply (unfold modules) +apply interfree_aux +--{* 9 subgoals left *} +apply(simp_all add:abbrev Graph6 Graph12) +--{* 6 subgoals left *} +apply(clarify, simp add:abbrev Graph6 Graph12, + erule disjE,erule disjI1,rule disjI2,rule subset_trans, erule Graph3,force,force)+ +done + +lemma interfree_Redirect_Edge_Count: + "interfree_aux (Some Redirect_Edge, {}, Some Count)" +apply (unfold modules ) +apply interfree_aux +apply(clarify,simp add:abbrev)+ +apply(simp add:abbrev) +done + +lemma interfree_Count_Color_Target: + "interfree_aux (Some Count, {}, Some Color_Target)" +apply (unfold modules ) +apply interfree_aux +--{* 9 subgoals left *} +apply(simp_all add:abbrev Graph7 Graph8 Graph12) +--{* 6 subgoals left *} +apply(clarify,simp add:abbrev Graph7 Graph8 Graph12, + erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9)+ +--{* 2 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) +apply(rule conjI) + apply(erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9) +apply(simp add:nth_list_update) +--{* 1 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12, + erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9) +done + +lemma interfree_Color_Target_Count: + "interfree_aux (Some Color_Target, {}, Some Count)" +apply (unfold modules ) +apply interfree_aux +apply(clarify, simp add:abbrev)+ +apply(simp add:abbrev) +done + +lemma interfree_Append_Redirect_Edge: + "interfree_aux (Some Append, {}, Some Redirect_Edge)" +apply (unfold modules ) +apply interfree_aux +apply( simp_all add:abbrev Graph6 Append_to_free0 Append_to_free1 Graph12) +apply(clarify, simp add:abbrev Graph6 Append_to_free0 Append_to_free1 Graph12, force dest:Graph3)+ +done + +lemma interfree_Redirect_Edge_Append: + "interfree_aux (Some Redirect_Edge, {}, Some Append)" +apply (unfold modules ) +apply interfree_aux +apply(clarify, simp add:abbrev Append_to_free0)+ +apply (force simp add: Append_to_free2) +apply(clarify, simp add:abbrev Append_to_free0)+ +done + +lemma interfree_Append_Color_Target: + "interfree_aux (Some Append, {}, Some Color_Target)" +apply (unfold modules ) +apply interfree_aux +apply(clarify, simp add:abbrev Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12 nth_list_update)+ +apply(simp add:abbrev Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12 nth_list_update) +done + +lemma interfree_Color_Target_Append: + "interfree_aux (Some Color_Target, {}, Some Append)" +apply (unfold modules ) +apply interfree_aux +apply(clarify, simp add:abbrev Append_to_free0)+ +apply (force simp add: Append_to_free2) +apply(clarify,simp add:abbrev Append_to_free0)+ +done + +lemmas collector_mutator_interfree = + interfree_Blacken_Roots_Redirect_Edge interfree_Blacken_Roots_Color_Target + interfree_Propagate_Black_Redirect_Edge interfree_Propagate_Black_Color_Target + interfree_Count_Redirect_Edge interfree_Count_Color_Target + interfree_Append_Redirect_Edge interfree_Append_Color_Target + interfree_Redirect_Edge_Blacken_Roots interfree_Color_Target_Blacken_Roots + interfree_Redirect_Edge_Propagate_Black interfree_Color_Target_Propagate_Black + interfree_Redirect_Edge_Count interfree_Color_Target_Count + interfree_Redirect_Edge_Append interfree_Color_Target_Append + +subsubsection {* Interference freedom Collector-Mutator *} + +lemma interfree_Collector_Mutator: + "interfree_aux (Some Collector, {}, Some Mutator)" +apply(unfold Collector_def Mutator_def) +apply interfree_aux +apply(simp_all add:collector_mutator_interfree) +apply(unfold modules collector_defs mutator_defs) +apply(tactic {* TRYALL (interfree_aux_tac) *}) +--{* 32 subgoals left *} +apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) +--{* 20 subgoals left *} +apply(tactic{* TRYALL Clarify_tac *}) +apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) +apply(tactic {* TRYALL (etac disjE) *}) +apply simp_all +apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac subset_trans,etac (thm "Graph3"),Force_tac, assume_tac]) *}) +apply(tactic {* TRYALL(EVERY'[rtac disjI2,etac subset_trans,rtac (thm "Graph9"),Force_tac]) *}) +apply(tactic {* TRYALL(EVERY'[rtac disjI1,etac psubset_subset_trans,rtac (thm "Graph9"),Force_tac]) *}) +done + +subsubsection {* Interference freedom Mutator-Collector *} + +lemma interfree_Mutator_Collector: + "interfree_aux (Some Mutator, {}, Some Collector)" +apply(unfold Collector_def Mutator_def) +apply interfree_aux +apply(simp_all add:collector_mutator_interfree) +apply(unfold modules collector_defs mutator_defs) +apply(tactic {* TRYALL (interfree_aux_tac) *}) +--{* 64 subgoals left *} +apply(simp_all add:nth_list_update Invariants Append_to_free0)+ +apply(tactic{* TRYALL Clarify_tac *}) +--{* 4 subgoals left *} +apply force +apply(simp add:Append_to_free2) +apply force +apply(simp add:Append_to_free2) +done + +subsubsection {* The Garbage Collection algorithm *} + +text {* In total there are 289 verification conditions. *} + +lemma Gar_Coll: + "\- .{\Proper \ \Mut_init \ \z}. + COBEGIN + Collector + .{False}. + \ + Mutator + .{False}. + COEND + .{False}." +apply oghoare +apply(force simp add: Mutator_def Collector_def modules) +apply(rule Collector) +apply(rule Mutator) +apply(simp add:interfree_Collector_Mutator) +apply(simp add:interfree_Mutator_Collector) +apply force +done + +end diff -r 98f0a09a33c3 -r 791e3b4c4039 src/HOL/HoareParallel/Graph.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HoareParallel/Graph.thy Tue Mar 05 17:11:25 2002 +0100 @@ -0,0 +1,426 @@ + +header {* \chapter{Case Study: Single and Multi-Mutator Garbage Collection Algorithms} + +\section {Formalization of the Memory} *} + +theory Graph = Main: + +datatype node = Black | White + +types + nodes = "node list" + edge = "nat \ nat" + edges = "edge list" + +consts Roots :: "nat set" + +constdefs + Proper_Roots :: "nodes \ bool" + "Proper_Roots M \ Roots\{} \ Roots \ {i. i edges) \ bool" + "Proper_Edges \ (\(M,E). \i snd(E!i) nodes) \ bool" + "BtoW \ (\(e,M). (M!fst e)=Black \ (M!snd e)\Black)" + + Blacks :: "nodes \ nat set" + "Blacks M \ {i. i M!i=Black}" + + Reach :: "edges \ nat set" + "Reach E \ {x. (\path. 1 path!(length path - 1)\Roots \ x=path!0 + \ (\ij x\Roots}" + +text{* Reach: the set of reachable nodes is the set of Roots together with the +nodes reachable from some Root by a path represented by a list of + nodes (at least two since we traverse at least one edge), where two +consecutive nodes correspond to an edge in E. *} + +subsection {* Proofs about Graphs *} + +lemmas Graph_defs= Blacks_def Proper_Roots_def Proper_Edges_def BtoW_def +declare Graph_defs [simp] + +subsubsection{* Graph 1. *} + +lemma Graph1_aux [rule_format]: + "\ Roots\Blacks M; \iBtoW(E!i,M)\ + \ 1< length path \ (path!(length path - 1))\Roots \ + (\ij. j < length E \ E!j=(path!(Suc i), path!i))) + \ M!(path!0) = Black" +apply(induct_tac "path") + apply force +apply clarify +apply simp +apply(case_tac "list") + apply force +apply simp +apply(rotate_tac -1) +apply(erule_tac x = "0" in all_dupE) +apply simp +apply clarify +apply(erule allE , erule (1) notE impE) +apply simp +apply(erule mp) +apply(case_tac "lista") + apply force +apply simp +apply(erule mp) +apply clarify +apply(erule_tac x = "Suc i" in allE) +apply force +done + +lemma Graph1: + "\Roots\Blacks M; Proper_Edges(M, E); \iBtoW(E!i,M) \ + \ Reach E\Blacks M" +apply (unfold Reach_def) +apply simp +apply clarify +apply(erule disjE) + apply clarify + apply(rule conjI) + apply(subgoal_tac "0< length path - Suc 0") + apply(erule allE , erule (1) notE impE) + apply force + apply simp + apply(rule Graph1_aux) +apply auto +done + +subsubsection{* Graph 2. *} + +lemma Ex_first_occurrence [rule_format]: + "P (n::nat) \ (\m. P m \ (\i. i \ P i))"; +apply(rule nat_less_induct) +apply clarify +apply(case_tac "\m. m \ P m") +apply auto +done + +lemma Compl_lemma: "(n::nat)\l \ (\m. m\l \ n=l - m)" +apply(rule_tac x = "l - n" in exI) +apply arith +done + +lemma Ex_last_occurrence: + "\P (n::nat); n\l\ \ (\m. P (l - m) \ (\i. i \P (l - i)))" +apply(drule Compl_lemma) +apply clarify +apply(erule Ex_first_occurrence) +done + +lemma Graph2: + "\T \ Reach E; R \ T \ Reach (E[R:=(fst(E!R), T)])" +apply (unfold Reach_def) +apply clarify +apply simp +apply(case_tac "\zpath!z") + apply(rule_tac x = "path" in exI) + apply simp + apply clarify + apply(erule allE , erule (1) notE impE) + apply clarify + apply(rule_tac x = "j" in exI) + apply(case_tac "j=R") + apply(erule_tac x = "Suc i" in allE) + apply simp + apply arith + apply (force simp add:nth_list_update) +apply simp +apply(erule exE) +apply(subgoal_tac "z \ length path - Suc 0") + prefer 2 apply arith +apply(drule_tac P = "\m. m fst(E!R)=path!m" in Ex_last_occurrence) + apply assumption +apply clarify +apply simp +apply(rule_tac x = "(path!0)#(drop (length path - Suc m) path)" in exI) +apply simp +apply(case_tac "length path - (length path - Suc m)") + apply arith +apply simp +apply(subgoal_tac "(length path - Suc m) + nat \ length path") + prefer 2 apply arith +apply(drule nth_drop) +apply simp +apply(subgoal_tac "length path - Suc m + nat = length path - Suc 0") + prefer 2 apply arith +apply simp +apply clarify +apply(case_tac "i") + apply(force simp add: nth_list_update) +apply simp +apply(subgoal_tac "(length path - Suc m) + nata \ length path") + prefer 2 apply arith +apply simp +apply(subgoal_tac "(length path - Suc m) + (Suc nata) \ length path") + prefer 2 apply arith +apply simp +apply(erule_tac x = "length path - Suc m + nata" in allE) +apply simp +apply clarify +apply(rule_tac x = "j" in exI) +apply(case_tac "R=j") + prefer 2 apply force +apply simp +apply(drule_tac t = "path ! (length path - Suc m)" in sym) +apply simp +apply(case_tac " length path - Suc 0 < m") + apply(subgoal_tac "(length path - Suc m)=0") + prefer 2 apply arith + apply(rotate_tac -1) + apply(simp del: diff_is_0_eq) + apply(subgoal_tac "Suc nata\nat") + prefer 2 apply arith + apply(drule_tac n = "Suc nata" in Compl_lemma) + apply clarify + apply force +apply(drule leI) +apply(subgoal_tac "Suc (length path - Suc m + nata)=(length path - Suc 0) - (m - Suc nata)") + apply(erule_tac x = "m - (Suc nata)" in allE) + apply(case_tac "m") + apply simp + apply simp + apply(subgoal_tac "natb - nata < Suc natb") + prefer 2 apply(erule thin_rl)+ apply arith + apply simp + apply(case_tac "length path") + apply force + apply simp +apply(frule_tac i1 = "length path" and j1 = "length path - Suc 0" and k1 = "m" in diff_diff_right [THEN mp]) +apply(erule_tac V = "length path - Suc m + nat = length path - Suc 0" in thin_rl) +apply simp +apply arith +done + + +subsubsection{* Graph 3. *} + +lemma Graph3: + "\ T\Reach E; R \ Reach(E[R:=(fst(E!R),T)]) \ Reach E" +apply (unfold Reach_def) +apply clarify +apply simp +apply(case_tac "\i(length path\m)") + prefer 2 apply arith + apply(simp add: min_def) + apply(rule conjI) + apply(subgoal_tac "\(m + length patha - 1 < m)") + prefer 2 apply arith + apply(simp add: nth_append min_def) + apply(rule conjI) + apply(case_tac "m") + apply force + apply(case_tac "path") + apply force + apply force + apply clarify + apply(case_tac "Suc i\m") + apply(erule_tac x = "i" in allE) + apply simp + apply clarify + apply(rule_tac x = "j" in exI) + apply(case_tac "Suc i(length path\Suc m)" ) + prefer 2 apply arith + apply(simp add: min_def) + apply clarify + apply(erule_tac x = "i" in allE) + apply simp + apply clarify + apply(case_tac "R=j") + apply(force simp add: nth_list_update) + apply(force simp add: nth_list_update) +--{* the changed edge is not part of the path *} +apply(rule_tac x = "path" in exI) +apply simp +apply clarify +apply(erule_tac x = "i" in allE) +apply clarify +apply(case_tac "R=j") + apply(erule_tac x = "i" in allE) + apply simp +apply(force simp add: nth_list_update) +done + +subsubsection{* Graph 4. *} + +lemma Graph4: + "\T \ Reach E; Roots\Blacks M; I\length E; TiBtoW(E!i,M); RBlack\ \ + (\r. I\r \ r BtoW(E[R:=(fst(E!R),T)]!r,M))" +apply (unfold Reach_def) +apply simp +apply(erule disjE) + prefer 2 apply force +apply clarify +--{* there exist a black node in the path to T *} +apply(case_tac "\m T \ Reach E ; Roots \ Blacks M; \iBtoW(E!i,M); T Black\ + \ (\r. R r BtoW(E[R:=(fst(E!R),T)]!r,M))" +apply (unfold Reach_def) +apply simp +apply(erule disjE) + prefer 2 apply force +apply clarify +--{* there exist a black node in the path to T*} +apply(case_tac "\mR") + apply(drule le_imp_less_or_eq) + apply(erule disjE) + apply(erule allE , erule (1) notE impE) + apply force + apply force + apply(rule_tac x = "j" in exI) + apply(force simp add: nth_list_update) +apply simp +apply(rotate_tac -1) +apply(erule_tac x = "length path - 1" in allE) +apply(case_tac "length path") + apply force +apply force +done + +subsubsection {* Graph 6, 7, 8. *} + +lemma Graph6: + "\Proper_Edges(M,E); R \ Proper_Edges(M,E[R:=(fst(E!R),T)])" +apply (unfold Proper_Edges_def) + apply(force simp add: nth_list_update) +done + +lemma Graph7: + "\Proper_Edges(M,E)\ \ Proper_Edges(M[T:=a],E)" +apply (unfold Proper_Edges_def) +apply force +done + +lemma Graph8: + "\Proper_Roots(M)\ \ Proper_Roots(M[T:=a])" +apply (unfold Proper_Roots_def) +apply force +done + +text{* Some specific lemmata for the verification of garbage collection algorithms. *} + +lemma Graph9: "j Blacks M\Blacks (M[j := Black])" +apply (unfold Blacks_def) + apply(force simp add: nth_list_update) +done + +lemma Graph10 [rule_format (no_asm)]: "\i. M!i=a \M[i:=a]=M" +apply(induct_tac "M") +apply auto +apply(case_tac "i") +apply auto +done + +lemma Graph11 [rule_format (no_asm)]: + "\ M!j\Black;j \ Blacks M \ Blacks (M[j := Black])" +apply (unfold Blacks_def) +apply(rule psubsetI) + apply(force simp add: nth_list_update) +apply safe +apply(erule_tac c = "j" in equalityCE) +apply auto +done + +lemma Graph12: "\a\Blacks M;j \ a\Blacks (M[j := Black])" +apply (unfold Blacks_def) +apply(force simp add: nth_list_update) +done + +lemma Graph13: "\a\ Blacks M;j \ a \ Blacks (M[j := Black])" +apply (unfold Blacks_def) +apply(erule psubset_subset_trans) +apply(force simp add: nth_list_update) +done + +declare Graph_defs [simp del] + +end diff -r 98f0a09a33c3 -r 791e3b4c4039 src/HOL/HoareParallel/Mul_Gar_Coll.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HoareParallel/Mul_Gar_Coll.thy Tue Mar 05 17:11:25 2002 +0100 @@ -0,0 +1,1291 @@ + +header {* \section{The Multi-Mutator Case} *} + +theory Mul_Gar_Coll = Graph + OG_Syntax: + +text {* The full theory takes aprox. 18 minutes. *} + +record mut = + Z :: bool + R :: nat + T :: nat + +text {* Declaration of variables: *} + +record mul_gar_coll_state = + M :: nodes + E :: edges + bc :: "nat set" + obc :: "nat set" + Ma :: nodes + ind :: nat + k :: nat + q :: nat + l :: nat + Muts :: "mut list" + +subsection {* The Mutators *} + +constdefs + Mul_mut_init :: "mul_gar_coll_state \ nat \ bool" + "Mul_mut_init \ \ \n. n=length \Muts \ (\iMuts!i)E + \ T (\Muts!i)M) \" + + Mul_Redirect_Edge :: "nat \ nat \ mul_gar_coll_state ann_com" + "Mul_Redirect_Edge j n \ + .{\Mul_mut_init n \ Z (\Muts!j)}. + \IF T(\Muts!j) \ Reach \E THEN + \E:= \E[R (\Muts!j):= (fst (\E!R(\Muts!j)), T (\Muts!j))] FI,, + \Muts:= \Muts[j:= (\Muts!j) \Z:=False\]\" + + Mul_Color_Target :: "nat \ nat \ mul_gar_coll_state ann_com" + "Mul_Color_Target j n \ + .{\Mul_mut_init n \ \ Z (\Muts!j)}. + \\M:=\M[T (\Muts!j):=Black],, \Muts:=\Muts[j:= (\Muts!j) \Z:=True\]\" + + Mul_Mutator :: "nat \ nat \ mul_gar_coll_state ann_com" + "Mul_Mutator j n \ + .{\Mul_mut_init n \ Z (\Muts!j)}. + WHILE True + INV .{\Mul_mut_init n \ Z (\Muts!j)}. + DO Mul_Redirect_Edge j n ;; + Mul_Color_Target j n + OD" + +lemmas mul_mutator_defs = Mul_mut_init_def Mul_Redirect_Edge_def Mul_Color_Target_def + +subsubsection {* Correctness of the proof outline of one mutator *} + +lemma Mul_Redirect_Edge: "0\j \ j + \ Mul_Redirect_Edge j n + pre(Mul_Color_Target j n)" +apply (unfold mul_mutator_defs) +apply annhoare +apply(simp_all) +apply clarify +apply(simp add:nth_list_update) +done + +lemma Mul_Color_Target: "0\j \ j + \ Mul_Color_Target j n + .{\Mul_mut_init n \ Z (\Muts!j)}." +apply (unfold mul_mutator_defs) +apply annhoare +apply(simp_all) +apply clarify +apply(simp add:nth_list_update) +done + +lemma Mul_Mutator: "0\j \ j + \ Mul_Mutator j n .{False}." +apply(unfold Mul_Mutator_def) +apply annhoare +apply(simp_all add:Mul_Redirect_Edge Mul_Color_Target) +apply(simp add:mul_mutator_defs Mul_Redirect_Edge_def) +done + +subsubsection {* Interference freedom between mutators *} + +lemma Mul_interfree_Redirect_Edge_Redirect_Edge: + "\0\i; ij; jj\ \ + interfree_aux (Some (Mul_Redirect_Edge i n),{}, Some(Mul_Redirect_Edge j n))" +apply (unfold mul_mutator_defs) +apply interfree_aux +apply safe +apply(simp_all add: nth_list_update) +apply force+ +done + +lemma Mul_interfree_Redirect_Edge_Color_Target: + "\0\i; ij; jj\ \ + interfree_aux (Some(Mul_Redirect_Edge i n),{},Some(Mul_Color_Target j n))" +apply (unfold mul_mutator_defs) +apply interfree_aux +apply safe +apply(simp_all add: nth_list_update) +done + +lemma Mul_interfree_Color_Target_Redirect_Edge: + "\0\i; ij; jj\ \ + interfree_aux (Some(Mul_Color_Target i n),{},Some(Mul_Redirect_Edge j n))" +apply (unfold mul_mutator_defs) +apply interfree_aux +apply safe +apply(simp_all add:nth_list_update) +apply (drule not_sym,force)+ +done + +lemma Mul_interfree_Color_Target_Color_Target: + " \0\i; ij; jj\ \ + interfree_aux (Some(Mul_Color_Target i n),{},Some(Mul_Color_Target j n))" +apply (unfold mul_mutator_defs) +apply interfree_aux +apply safe +apply(simp_all add: nth_list_update) +apply (drule not_sym,force) +done + +lemmas mul_mutator_interfree = + Mul_interfree_Redirect_Edge_Redirect_Edge Mul_interfree_Redirect_Edge_Color_Target + Mul_interfree_Color_Target_Redirect_Edge Mul_interfree_Color_Target_Color_Target + +lemma Mul_interfree_Mutator_Mutator: "\i < n; j < n; i \ j\ \ + interfree_aux (Some (Mul_Mutator i n), {}, Some (Mul_Mutator j n))" +apply(unfold Mul_Mutator_def) +apply(interfree_aux) +apply(simp_all add:mul_mutator_interfree) +apply(simp_all add: mul_mutator_defs) +apply(tactic {* TRYALL (interfree_aux_tac) *}) +apply(tactic {* ALLGOALS Clarify_tac *}) +apply (simp_all add:nth_list_update) +apply force+ +done + +subsubsection {* Modular Parameterized Mutators *} + +lemma Mul_Parameterized_Mutators: "0 + \- .{\Mul_mut_init n \ (\iMuts!i))}. + COBEGIN + SCHEME [0\ j< n] + Mul_Mutator j n + .{False}. + COEND + .{False}." +apply oghoare +apply(force simp add:Mul_Mutator_def mul_mutator_defs nth_list_update) +apply(erule Mul_Mutator) +apply(simp add:Mul_interfree_Mutator_Mutator) +apply(force simp add:Mul_Mutator_def mul_mutator_defs nth_list_update) +done + +subsection {* The Collector *} + +constdefs + Queue :: "mul_gar_coll_state \ nat" + "Queue \ \ length (filter (\i. \ Z i \ \M!(T i) \ Black) \Muts) \" + +consts M_init :: nodes + +constdefs + Proper_M_init :: "mul_gar_coll_state \ bool" + "Proper_M_init \ \ Blacks M_init=Roots \ length M_init=length \M \" + + Mul_Proper :: "mul_gar_coll_state \ nat \ bool" + "Mul_Proper \ \ \n. Proper_Roots \M \ Proper_Edges (\M, \E) \ \Proper_M_init \ n=length \Muts \" + + Safe :: "mul_gar_coll_state \ bool" + "Safe \ \ Reach \E \ Blacks \M \" + +lemmas mul_collector_defs = Proper_M_init_def Mul_Proper_def Safe_def + +subsubsection {* Blackening Roots *} + +constdefs + Mul_Blacken_Roots :: "nat \ mul_gar_coll_state ann_com" + "Mul_Blacken_Roots n \ + .{\Mul_Proper n}. + \ind:=0;; + .{\Mul_Proper n \ \ind=0}. + WHILE \indM + INV .{\Mul_Proper n \ (\i<\ind. i\Roots \ \M!i=Black) \ \ind\length \M}. + DO .{\Mul_Proper n \ (\i<\ind. i\Roots \ \M!i=Black) \ \indM}. + IF \ind\Roots THEN + .{\Mul_Proper n \ (\i<\ind. i\Roots \ \M!i=Black) \ \indM \ \ind\Roots}. + \M:=\M[\ind:=Black] FI;; + .{\Mul_Proper n \ (\i<\ind+1. i\Roots \ \M!i=Black) \ \indM}. + \ind:=\ind+1 + OD" + +lemma Mul_Blacken_Roots: + "\ Mul_Blacken_Roots n + .{\Mul_Proper n \ Roots \ Blacks \M}." +apply (unfold Mul_Blacken_Roots_def) +apply annhoare +apply(simp_all add:mul_collector_defs Graph_defs) +apply safe +apply(simp_all add:nth_list_update) + apply (erule less_SucE) + apply simp+ + apply (erule less_SucE) + apply simp+ + apply(drule le_imp_less_or_eq) + apply force +apply force +done + +subsubsection {* Propagating Black *} + +constdefs + Mul_PBInv :: "mul_gar_coll_state \ bool" + "Mul_PBInv \ \\Safe \ \obc\Blacks \M \ \l<\Queue + \ (\i<\ind. \BtoW(\E!i,\M)) \ \l\\Queue\" + + Mul_Auxk :: "mul_gar_coll_state \ bool" + "Mul_Auxk \ \\l<\Queue \ \M!\k\Black \ \BtoW(\E!\ind, \M) \ \obc\Blacks \M\" + +constdefs + Mul_Propagate_Black :: "nat \ mul_gar_coll_state ann_com" + "Mul_Propagate_Black n \ + .{\Mul_Proper n \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ (\Safe \ \l\\Queue \ \obc\Blacks \M)}. + \ind:=0;; + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \M \ Blacks \M\Blacks \M \ \bc\Blacks \M + \ (\Safe \ \l\\Queue \ \obc\Blacks \M) \ \ind=0}. + WHILE \indE + INV .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \M \ \bc\Blacks \M + \ \Mul_PBInv \ \ind\length \E}. + DO .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \M \ \bc\Blacks \M + \ \Mul_PBInv \ \indE}. + IF \M!(fst (\E!\ind))=Black THEN + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \M \ \bc\Blacks \M + \ \Mul_PBInv \ (\M!fst(\E!\ind))=Black \ \indE}. + \k:=snd(\E!\ind);; + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \M \ \bc\Blacks \M + \ (\Safe \ \obc\Blacks \M \ \l<\Queue \ (\i<\ind. \BtoW(\E!i,\M)) + \ \l\\Queue \ \Mul_Auxk ) \ \kM \ \M!fst(\E!\ind)=Black + \ \indE}. + \\M:=\M[\k:=Black],,\ind:=\ind+1\ + ELSE .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \M \ \bc\Blacks \M + \ \Mul_PBInv \ \indE}. + \IF \M!(fst (\E!\ind))\Black THEN \ind:=\ind+1 FI\ FI + OD" + +lemma Mul_Propagate_Black: + "\ Mul_Propagate_Black n + .{\Mul_Proper n \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ (\Safe \ \obc\Blacks \M \ \l<\Queue \ (\l\\Queue \ \obc\Blacks \M))}." +apply(unfold Mul_Propagate_Black_def) +apply annhoare +apply(simp_all add:Mul_PBInv_def mul_collector_defs Mul_Auxk_def Graph6 Graph7 Graph8 Graph12 mul_collector_defs Queue_def) +--{* 8 subgoals left *} +apply force +apply force +apply force +apply(force simp add:BtoW_def Graph_defs) +--{* 4 subgoals left *} +apply clarify +apply(simp add: mul_collector_defs Graph12 Graph6 Graph7 Graph8) +apply(disjE_tac) + apply(simp_all add:Graph12 Graph13) + apply(case_tac "M x! k x=Black") + apply(simp add: Graph10) + apply(rule disjI2, rule disjI1, erule subset_psubset_trans, erule Graph11, force) +apply(case_tac "M x! k x=Black") + apply(simp add: Graph10 BtoW_def) + apply(rule disjI2, clarify, erule less_SucE, force) + apply(case_tac "M x!snd(E x! ind x)=Black") + apply(force) + apply(force) +apply(rule disjI2, rule disjI1, erule subset_psubset_trans, erule Graph11, force) +--{* 3 subgoals left *} +apply force +--{* 2 subgoals left *} +apply clarify +apply(conjI_tac) +apply(disjE_tac) + apply (simp_all) +apply clarify +apply(erule less_SucE) + apply force +apply (simp add:BtoW_def) +--{* 1 subgoals left *} +apply clarify +apply simp +apply(disjE_tac) +apply (simp_all) +apply(rule disjI1 , rule Graph1) + apply simp_all +done + +subsubsection {* Counting Black Nodes *} + +constdefs + Mul_CountInv :: "mul_gar_coll_state \ nat \ bool" + "Mul_CountInv \ \ \ind. {i. i \Ma!i=Black}\\bc \" + + Mul_Count :: "nat \ mul_gar_coll_state ann_com" + "Mul_Count n \ + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M) ) + \ \q \bc={}}. + \ind:=0;; + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M) ) + \ \q \bc={} \ \ind=0}. + WHILE \indM + INV .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ \Mul_CountInv \ind + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) + \ \q \ind\length \M}. + DO .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ \Mul_CountInv \ind + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) + \ \q \indM}. + IF \M!\ind=Black + THEN .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ \Mul_CountInv \ind + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) + \ \q \indM \ \M!\ind=Black}. + \bc:=insert \ind \bc + FI;; + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ \Mul_CountInv (\ind+1) + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) + \ \q \indM}. + \ind:=\ind+1 + OD" + +lemma Mul_Count: + "\ Mul_Count n + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ Blacks \Ma\\bc + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) + \ \q edges \ edges" + +axioms + Append_to_free0: "length (Append_to_free (i, e)) = length e" + Append_to_free1: "Proper_Edges (m, e) + \ Proper_Edges (m, Append_to_free(i, e))" + Append_to_free2: "i \ Reach e + \ n \ Reach (Append_to_free(i, e)) = ( n = i \ n \ Reach e)" + +constdefs + Mul_AppendInv :: "mul_gar_coll_state \ nat \ bool" + "Mul_AppendInv \ \ \ind. (\i. ind\i \ iM \ i\Reach \E \ \M!i=Black)\" + + Mul_Append :: "nat \ mul_gar_coll_state ann_com" + "Mul_Append n \ + .{\Mul_Proper n \ Roots\Blacks \M \ \Safe}. + \ind:=0;; + .{\Mul_Proper n \ Roots\Blacks \M \ \Safe \ \ind=0}. + WHILE \indM + INV .{\Mul_Proper n \ \Mul_AppendInv \ind \ \ind\length \M}. + DO .{\Mul_Proper n \ \Mul_AppendInv \ind \ \indM}. + IF \M!\ind=Black THEN + .{\Mul_Proper n \ \Mul_AppendInv \ind \ \indM \ \M!\ind=Black}. + \M:=\M[\ind:=White] + ELSE + .{\Mul_Proper n \ \Mul_AppendInv \ind \ \indM \ \ind\Reach \E}. + \E:=Append_to_free(\ind,\E) + FI;; + .{\Mul_Proper n \ \Mul_AppendInv (\ind+1) \ \indM}. + \ind:=\ind+1 + OD" + +lemma Mul_Append: + "\ Mul_Append n + .{\Mul_Proper n}." +apply(unfold Mul_Append_def) +apply annhoare +apply(simp_all add: mul_collector_defs Mul_AppendInv_def + Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) +apply(force simp add:Blacks_def) +apply(force simp add:Blacks_def) +apply(force simp add:Blacks_def) +apply(force simp add:Graph_defs) +apply force +apply(force simp add:Append_to_free1 Append_to_free2) +apply force +apply force +done + +subsubsection {* Collector *} + +constdefs + Mul_Collector :: "nat \ mul_gar_coll_state ann_com" + "Mul_Collector n \ +.{\Mul_Proper n}. +WHILE True INV .{\Mul_Proper n}. +DO +Mul_Blacken_Roots n ;; +.{\Mul_Proper n \ Roots\Blacks \M}. + \obc:={};; +.{\Mul_Proper n \ Roots\Blacks \M \ \obc={}}. + \bc:=Roots;; +.{\Mul_Proper n \ Roots\Blacks \M \ \obc={} \ \bc=Roots}. + \l:=0;; +.{\Mul_Proper n \ Roots\Blacks \M \ \obc={} \ \bc=Roots \ \l=0}. + WHILE \lMul_Proper n \ Roots\Blacks \M \ \bc\Blacks \M \ + (\Safe \ (\l\\Queue \ \bc\Blacks \M) \ \lMul_Proper n \ Roots\Blacks \M \ \bc\Blacks \M + \ (\Safe \ \l\\Queue \ \bc\Blacks \M)}. + \obc:=\bc;; + Mul_Propagate_Black n;; + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \M \ \bc\Blacks \M + \ (\Safe \ \obc\Blacks \M \ \l<\Queue + \ (\l\\Queue \ \obc\Blacks \M))}. + \bc:={};; + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \M \ \bc\Blacks \M + \ (\Safe \ \obc\Blacks \M \ \l<\Queue + \ (\l\\Queue \ \obc\Blacks \M)) \ \bc={}}. + \ \Ma:=\M,, \q:=\Queue \;; + Mul_Count n;; + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ Blacks \Ma\\bc + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) + \ \qobc=\bc THEN + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ Blacks \Ma\\bc + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) + \ \q \obc=\bc}. + \l:=\l+1 + ELSE .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ Blacks \Ma\\bc + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) + \ \q \obc\\bc}. + \l:=0 FI + OD;; + Mul_Append n +OD" + +lemmas mul_modules = Mul_Redirect_Edge_def Mul_Color_Target_def + Mul_Blacken_Roots_def Mul_Propagate_Black_def + Mul_Count_def Mul_Append_def + +lemma Mul_Collector: + "\ Mul_Collector n + .{False}." +apply(unfold Mul_Collector_def) +apply annhoare +apply(simp_all only:pre.simps Mul_Blacken_Roots + Mul_Propagate_Black Mul_Count Mul_Append) +apply(simp_all add:mul_modules) +apply(simp_all add:mul_collector_defs Queue_def) +apply force +apply force +apply force +apply (force simp add: less_Suc_eq_le length_filter) +apply force +apply (force dest:subset_antisym) +apply force +apply force +apply force +done + +subsection {* Interference Freedom *} + +lemma le_length_filter_update[rule_format]: + "\i. (\P (list!i) \ P j) \ i length(filter P list) \ length(filter P (list[i:=j]))" +apply(induct_tac "list") + apply(simp) +apply(clarify) +apply(case_tac i) + apply(simp) +apply(simp) +done + +lemma less_length_filter_update [rule_format]: + "\i. P j \ \(P (list!i)) \ i length(filter P list) < length(filter P (list[i:=j]))" +apply(induct_tac "list") + apply(simp) +apply(clarify) +apply(case_tac i) + apply(simp) +apply(simp) +done + +lemma Mul_interfree_Blacken_Roots_Redirect_Edge: "\0\j; j \ + interfree_aux (Some(Mul_Blacken_Roots n),{},Some(Mul_Redirect_Edge j n))" +apply (unfold mul_modules) +apply interfree_aux +apply safe +apply(simp_all add:Graph6 Graph9 Graph12 nth_list_update mul_mutator_defs mul_collector_defs) +done + +lemma Mul_interfree_Redirect_Edge_Blacken_Roots: "\0\j; j\ + interfree_aux (Some(Mul_Redirect_Edge j n ),{},Some (Mul_Blacken_Roots n))" +apply (unfold mul_modules) +apply interfree_aux +apply safe +apply(simp_all add:mul_mutator_defs nth_list_update) +done + +lemma Mul_interfree_Blacken_Roots_Color_Target: "\0\j; j\ + interfree_aux (Some(Mul_Blacken_Roots n),{},Some (Mul_Color_Target j n ))" +apply (unfold mul_modules) +apply interfree_aux +apply safe +apply(simp_all add:mul_mutator_defs mul_collector_defs nth_list_update Graph7 Graph8 Graph9 Graph12) +done + +lemma Mul_interfree_Color_Target_Blacken_Roots: "\0\j; j\ + interfree_aux (Some(Mul_Color_Target j n ),{},Some (Mul_Blacken_Roots n ))" +apply (unfold mul_modules) +apply interfree_aux +apply safe +apply(simp_all add:mul_mutator_defs nth_list_update) +done + +lemma Mul_interfree_Propagate_Black_Redirect_Edge: "\0\j; j\ + interfree_aux (Some(Mul_Propagate_Black n),{},Some (Mul_Redirect_Edge j n ))" +apply (unfold mul_modules) +apply interfree_aux +apply(simp_all add:mul_mutator_defs mul_collector_defs Mul_PBInv_def nth_list_update Graph6) +--{* 7 subgoals left *} +apply clarify +apply(disjE_tac) + apply(simp_all add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) +apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule impI,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +--{* 6 subgoals left *} +apply clarify +apply(disjE_tac) + apply(simp_all add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) +apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule impI,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +--{* 5 subgoals left *} +apply clarify +apply(disjE_tac) + apply(simp_all add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) +apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(erule conjE) +apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule conjI) + apply(rule impI,(rule disjI2)+,rule conjI) + apply clarify + apply(case_tac "R (Muts x! j)=i") + apply (force simp add: nth_list_update BtoW_def) + apply (force simp add: nth_list_update) + apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,(rule disjI2)+, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) + apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) +apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) +apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) +--{* 4 subgoals left *} +apply clarify +apply(disjE_tac) + apply(simp_all add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) +apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(erule conjE) +apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule conjI) + apply(rule impI,(rule disjI2)+,rule conjI) + apply clarify + apply(case_tac "R (Muts x! j)=i") + apply (force simp add: nth_list_update BtoW_def) + apply (force simp add: nth_list_update) + apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,(rule disjI2)+, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) + apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) +apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) +apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) +--{* 3 subgoals left *} +apply clarify +apply(disjE_tac) + apply(simp_all add:Graph6) + apply (rule impI) + apply(rule conjI) + apply(rule disjI1,rule subset_trans,erule Graph3,simp,simp) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule conjI) + apply(rule impI) + apply(rule conjI) + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) + apply(rule impI) + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule conjI) + apply(rule impI) + apply(rule conjI) + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) + apply(rule impI) + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(erule conjE) + apply(rule conjI) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule impI,rule conjI,(rule disjI2)+,rule conjI) + apply clarify + apply(case_tac "R (Muts x! j)=i") + apply (force simp add: nth_list_update BtoW_def) + apply (force simp add: nth_list_update) + apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) + apply(rule impI,rule conjI) + apply(rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) + apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) + apply(case_tac "R (Muts x! j)=ind x") + apply (force simp add: nth_list_update) + apply (force simp add: nth_list_update) +apply(rule impI, (rule disjI2)+, erule le_trans) +apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +--{* 2 subgoals left *} +apply clarify +apply(rule conjI) + apply(disjE_tac) + apply(simp_all add:Mul_Auxk_def Graph6) + apply (rule impI) + apply(rule conjI) + apply(rule disjI1,rule subset_trans,erule Graph3,simp,simp) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule impI) + apply(rule conjI) + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) + apply(rule impI) + apply(rule conjI) + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) +apply(rule impI) +apply(rule conjI) + apply(erule conjE)+ + apply(case_tac "M x!(T (Muts x!j))=Black") + apply((rule disjI2)+,rule conjI) + apply clarify + apply(case_tac "R (Muts x! j)=i") + apply (force simp add: nth_list_update BtoW_def) + apply (force simp add: nth_list_update) + apply(rule conjI) + apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update BtoW_def) + apply (simp add:nth_list_update) + apply(rule impI) + apply simp + apply(disjE_tac) + apply(rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply force + apply(rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) + apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) +apply(disjE_tac) +apply simp_all +apply(conjI_tac) + apply(rule impI) + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(erule conjE)+ +apply(rule impI,(rule disjI2)+,rule conjI) + apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule impI)+ +apply simp +apply(disjE_tac) + apply(rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply force +--{* 1 subgoals left *} +apply clarify +apply(disjE_tac) + apply(simp_all add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) +apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(erule conjE) +apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule conjI) + apply(rule impI,(rule disjI2)+,rule conjI) + apply clarify + apply(case_tac "R (Muts x! j)=i") + apply (force simp add: nth_list_update BtoW_def) + apply (force simp add: nth_list_update) + apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,(rule disjI2)+, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) + apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) +apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) +apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) +done + +lemma Mul_interfree_Redirect_Edge_Propagate_Black: "\0\j; j\ + interfree_aux (Some(Mul_Redirect_Edge j n ),{},Some (Mul_Propagate_Black n))" +apply (unfold mul_modules) +apply interfree_aux +apply safe +apply(simp_all add:mul_mutator_defs nth_list_update) +done + +lemma Mul_interfree_Propagate_Black_Color_Target: "\0\j; j\ + interfree_aux (Some(Mul_Propagate_Black n),{},Some (Mul_Color_Target j n ))" +apply (unfold mul_modules) +apply interfree_aux +apply(simp_all add: mul_collector_defs mul_mutator_defs) +--{* 7 subgoals left *} +apply clarify +apply (simp add:Graph7 Graph8 Graph12) +apply(disjE_tac) + apply(simp add:Graph7 Graph8 Graph12) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI1, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule disjI2)+,erule subset_psubset_trans, erule Graph11, simp) +apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) +--{* 6 subgoals left *} +apply clarify +apply (simp add:Graph7 Graph8 Graph12) +apply(disjE_tac) + apply(simp add:Graph7 Graph8 Graph12) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI1, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule disjI2)+,erule subset_psubset_trans, erule Graph11, simp) +apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) +--{* 5 subgoals left *} +apply clarify +apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12) +apply(disjE_tac) + apply(simp add:Graph7 Graph8 Graph12) + apply(rule disjI2,rule disjI1, erule psubset_subset_trans,simp add:Graph9) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply(rule disjI2,rule disjI1,erule subset_psubset_trans, erule Graph11, simp) +apply(erule conjE) +apply(case_tac "M x!(T (Muts x!j))=Black") + apply((rule disjI2)+) + apply (rule conjI) + apply(simp add:Graph10) + apply(erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) +apply(rule disjI2,rule disjI1,erule subset_psubset_trans, erule Graph11, simp) +--{* 4 subgoals left *} +apply clarify +apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12) +apply(disjE_tac) + apply(simp add:Graph7 Graph8 Graph12) + apply(rule disjI2,rule disjI1, erule psubset_subset_trans,simp add:Graph9) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply(rule disjI2,rule disjI1,erule subset_psubset_trans, erule Graph11, simp) +apply(erule conjE) +apply(case_tac "M x!(T (Muts x!j))=Black") + apply((rule disjI2)+) + apply (rule conjI) + apply(simp add:Graph10) + apply(erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) +apply(rule disjI2,rule disjI1,erule subset_psubset_trans, erule Graph11, simp) +--{* 3 subgoals left *} +apply clarify +apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12) +apply(case_tac "M x!(T (Muts x!j))=Black") + apply(simp add:Graph10) + apply(disjE_tac) + apply simp_all + apply(rule disjI2, rule disjI2, rule disjI1,erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply(erule conjE) + apply((rule disjI2)+,erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) +apply(rule conjI) + apply(rule disjI2,rule disjI1, erule subset_psubset_trans,simp add:Graph11) +apply (force simp add:nth_list_update) +--{* 2 subgoals left *} +apply clarify +apply(simp add:Mul_Auxk_def Graph7 Graph8 Graph12) +apply(case_tac "M x!(T (Muts x!j))=Black") + apply(simp add:Graph10) + apply(disjE_tac) + apply simp_all + apply(rule disjI2, rule disjI2, rule disjI1,erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply(erule conjE)+ + apply((rule disjI2)+,rule conjI, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule impI)+) + apply simp + apply(erule disjE) + apply(rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply force +apply(rule conjI) + apply(rule disjI2,rule disjI1, erule subset_psubset_trans,simp add:Graph11) +apply (force simp add:nth_list_update) +--{* 1 subgoals left *} +apply clarify +apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12) +apply(case_tac "M x!(T (Muts x!j))=Black") + apply(simp add:Graph10) + apply(disjE_tac) + apply simp_all + apply(rule disjI2, rule disjI2, rule disjI1,erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply(erule conjE) + apply((rule disjI2)+,erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) +apply(rule disjI2,rule disjI1, erule subset_psubset_trans,simp add:Graph11) +done + +lemma Mul_interfree_Color_Target_Propagate_Black: "\0\j; j\ + interfree_aux (Some(Mul_Color_Target j n),{},Some(Mul_Propagate_Black n ))" +apply (unfold mul_modules) +apply interfree_aux +apply safe +apply(simp_all add:mul_mutator_defs nth_list_update) +done + +lemma Mul_interfree_Count_Redirect_Edge: "\0\j; j\ + interfree_aux (Some(Mul_Count n ),{},Some(Mul_Redirect_Edge j n))" +apply (unfold mul_modules) +apply interfree_aux +--{* 9 subgoals left *} +apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def Graph6) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) + apply(simp add:Graph6) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(simp add:Graph6) +--{* 8 subgoals left *} +apply(simp add:mul_mutator_defs nth_list_update) +--{* 7 subgoals left *} +apply(simp add:mul_mutator_defs mul_collector_defs) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) + apply(simp add:Graph6) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(simp add:Graph6) +--{* 6 subgoals left *} +apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def) +apply clarify +apply disjE_tac + apply(simp add:Graph6 Queue_def) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) + apply(simp add:Graph6) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(simp add:Graph6) +--{* 5 subgoals left *} +apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) + apply(simp add:Graph6) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(simp add:Graph6) +--{* 4 subgoals left *} +apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) + apply(simp add:Graph6) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(simp add:Graph6) +--{* 3 subgoals left *} +apply(simp add:mul_mutator_defs nth_list_update) +--{* 2 subgoals left *} +apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) + apply(simp add:Graph6) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(simp add:Graph6) +--{* 1 subgoals left *} +apply(simp add:mul_mutator_defs nth_list_update) +done + +lemma Mul_interfree_Redirect_Edge_Count: "\0\j; j\ + interfree_aux (Some(Mul_Redirect_Edge j n),{},Some(Mul_Count n ))" +apply (unfold mul_modules) +apply interfree_aux +apply safe +apply(simp_all add:mul_mutator_defs nth_list_update) +done + +lemma Mul_interfree_Count_Color_Target: "\0\j; j\ + interfree_aux (Some(Mul_Count n ),{},Some(Mul_Color_Target j n))" +apply (unfold mul_modules) +apply interfree_aux +apply(simp_all add:mul_collector_defs mul_mutator_defs Mul_CountInv_def) +--{* 6 subgoals left *} +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply (simp add: Graph7 Graph8 Graph12) +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) +apply (simp add: Graph7 Graph8 Graph12) +apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) +--{* 5 subgoals left *} +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply (simp add: Graph7 Graph8 Graph12) +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) +apply (simp add: Graph7 Graph8 Graph12) +apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) +--{* 4 subgoals left *} +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply (simp add: Graph7 Graph8 Graph12) +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) +apply (simp add: Graph7 Graph8 Graph12) +apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) +--{* 3 subgoals left *} +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply (simp add: Graph7 Graph8 Graph12) +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) +apply (simp add: Graph7 Graph8 Graph12) +apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) +--{* 2 subgoals left *} +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12 nth_list_update) + apply (simp add: Graph7 Graph8 Graph12 nth_list_update) +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply(rule conjI) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) + apply (simp add: nth_list_update) +apply (simp add: Graph7 Graph8 Graph12) +apply(rule conjI) + apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) +apply (simp add: nth_list_update) +--{* 1 subgoals left *} +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply (simp add: Graph7 Graph8 Graph12) +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) +apply (simp add: Graph7 Graph8 Graph12) +apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) +done + +lemma Mul_interfree_Color_Target_Count: "\0\j; j\ + interfree_aux (Some(Mul_Color_Target j n),{}, Some(Mul_Count n ))" +apply (unfold mul_modules) +apply interfree_aux +apply safe +apply(simp_all add:mul_mutator_defs nth_list_update) +done + +lemma Mul_interfree_Append_Redirect_Edge: "\0\j; j\ + interfree_aux (Some(Mul_Append n),{}, Some(Mul_Redirect_Edge j n))" +apply (unfold mul_modules) +apply interfree_aux +apply(tactic {* ALLGOALS Clarify_tac *}) +apply(simp_all add:Graph6 Append_to_free0 Append_to_free1 mul_collector_defs mul_mutator_defs Mul_AppendInv_def) +apply(erule_tac x=j in allE, force dest:Graph3)+ +done + +lemma Mul_interfree_Redirect_Edge_Append: "\0\j; j\ + interfree_aux (Some(Mul_Redirect_Edge j n),{},Some(Mul_Append n))" +apply (unfold mul_modules) +apply interfree_aux +apply(tactic {* ALLGOALS Clarify_tac *}) +apply(simp_all add:mul_collector_defs Append_to_free0 Mul_AppendInv_def mul_mutator_defs nth_list_update) +done + +lemma Mul_interfree_Append_Color_Target: "\0\j; j\ + interfree_aux (Some(Mul_Append n),{}, Some(Mul_Color_Target j n))" +apply (unfold mul_modules) +apply interfree_aux +apply(tactic {* ALLGOALS Clarify_tac *}) +apply(simp_all add:mul_mutator_defs mul_collector_defs Mul_AppendInv_def Graph7 Graph8 Append_to_free0 Append_to_free1 + Graph12 nth_list_update) +done + +lemma Mul_interfree_Color_Target_Append: "\0\j; j\ + interfree_aux (Some(Mul_Color_Target j n),{}, Some(Mul_Append n))" +apply (unfold mul_modules) +apply interfree_aux +apply(tactic {* ALLGOALS Clarify_tac *}) +apply(simp_all add: mul_mutator_defs nth_list_update) +apply(simp add:Mul_AppendInv_def Append_to_free0) +done + +subsubsection {* Interference freedom Collector-Mutator *} + +lemmas mul_collector_mutator_interfree = + Mul_interfree_Blacken_Roots_Redirect_Edge Mul_interfree_Blacken_Roots_Color_Target + Mul_interfree_Propagate_Black_Redirect_Edge Mul_interfree_Propagate_Black_Color_Target + Mul_interfree_Count_Redirect_Edge Mul_interfree_Count_Color_Target + Mul_interfree_Append_Redirect_Edge Mul_interfree_Append_Color_Target + Mul_interfree_Redirect_Edge_Blacken_Roots Mul_interfree_Color_Target_Blacken_Roots + Mul_interfree_Redirect_Edge_Propagate_Black Mul_interfree_Color_Target_Propagate_Black + Mul_interfree_Redirect_Edge_Count Mul_interfree_Color_Target_Count + Mul_interfree_Redirect_Edge_Append Mul_interfree_Color_Target_Append + +lemma Mul_interfree_Collector_Mutator: "j + interfree_aux (Some (Mul_Collector n), {}, Some (Mul_Mutator j n))" +apply(unfold Mul_Collector_def Mul_Mutator_def) +apply interfree_aux +apply(simp_all add:mul_collector_mutator_interfree) +apply(unfold mul_modules mul_collector_defs mul_mutator_defs) +apply(tactic {* TRYALL (interfree_aux_tac) *}) +--{* 42 subgoals left *} +apply (clarify,simp add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)+ +--{* 24 subgoals left *} +apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) +--{* 14 subgoals left *} +apply(tactic {* TRYALL Clarify_tac *}) +apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) +apply(tactic {* TRYALL (rtac conjI) *}) +apply(tactic {* TRYALL (rtac impI) *}) +apply(tactic {* TRYALL (etac disjE) *}) +apply(tactic {* TRYALL (etac conjE) *}) +apply(tactic {* TRYALL (etac disjE) *}) +apply(tactic {* TRYALL (etac disjE) *}) +--{* 72 subgoals left *} +apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) +--{* 35 subgoals left *} +apply(tactic {* TRYALL(EVERY'[rtac disjI1,rtac subset_trans,etac (thm "Graph3"),Force_tac, assume_tac]) *}) +--{* 28 subgoals left *} +apply(tactic {* TRYALL (etac conjE) *}) +apply(tactic {* TRYALL (etac disjE) *}) +--{* 34 subgoals left *} +apply(rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(tactic {* ALLGOALS(case_tac "M x!(T (Muts x ! j))=Black") *}) +apply(simp_all add:Graph10) +--{* 47 subgoals left *} +apply(tactic {* TRYALL(EVERY'[REPEAT o (rtac disjI2),etac subset_psubset_trans,etac (thm "Graph11"),Force_tac]) *}) +--{* 41 subgoals left *} +apply(tactic {* TRYALL(EVERY'[rtac disjI2, rtac disjI1, etac le_trans, force_tac (claset(),simpset() addsimps [thm "Queue_def", less_Suc_eq_le, thm "le_length_filter_update"])]) *}) +--{* 35 subgoals left *} +apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac disjI1,etac psubset_subset_trans,rtac (thm "Graph9"),Force_tac]) *}) +--{* 31 subgoals left *} +apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac disjI1,etac subset_psubset_trans,etac (thm "Graph11"),Force_tac]) *}) +--{* 29 subgoals left *} +apply(tactic {* TRYALL(EVERY'[REPEAT o (rtac disjI2),etac subset_psubset_trans,etac subset_psubset_trans,etac (thm "Graph11"),Force_tac]) *}) +--{* 25 subgoals left *} +apply(tactic {* TRYALL(EVERY'[rtac disjI2, rtac disjI2, rtac disjI1, etac le_trans, force_tac (claset(),simpset() addsimps [thm "Queue_def", less_Suc_eq_le, thm "le_length_filter_update"])]) *}) +--{* 10 subgoals left *} +apply(rule disjI2,rule disjI2,rule conjI,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update, rule disjI1, rule less_imp_le, erule less_le_trans, force simp add:Queue_def less_Suc_eq_le le_length_filter_update)+ +done + +subsubsection {* Interference freedom Mutator-Collector *} + +lemma Mul_interfree_Mutator_Collector: " j < n \ + interfree_aux (Some (Mul_Mutator j n), {}, Some (Mul_Collector n))" +apply(unfold Mul_Collector_def Mul_Mutator_def) +apply interfree_aux +apply(simp_all add:mul_collector_mutator_interfree) +apply(unfold mul_modules mul_collector_defs mul_mutator_defs) +apply(tactic {* TRYALL (interfree_aux_tac) *}) +--{* 76 subgoals left *} +apply (clarify,simp add: nth_list_update)+ +--{* 56 subgoals left *} +apply(clarify,simp add:Mul_AppendInv_def Append_to_free0 nth_list_update)+ +done + +subsubsection {* The Multi-Mutator Garbage Collection Algorithm *} + +text {* The total number of verification conditions is 328 *} + +lemma Mul_Gar_Coll: + "\- .{\Mul_Proper n \ \Mul_mut_init n \ (\iMuts!i))}. + COBEGIN + Mul_Collector n + .{False}. + \ + SCHEME [0\ j< n] + Mul_Mutator j n + .{False}. + COEND + .{False}." +apply oghoare +--{* Strengthening the precondition *} +apply(rule Int_greatest) + apply (case_tac n) + apply(force simp add: Mul_Collector_def mul_mutator_defs mul_collector_defs nth_append) + apply(simp add: Mul_Mutator_def mul_collector_defs mul_mutator_defs nth_append) + apply force +apply clarify +apply(case_tac xa) + apply(simp add:Mul_Collector_def mul_mutator_defs mul_collector_defs nth_append) +apply(simp add: Mul_Mutator_def mul_mutator_defs mul_collector_defs nth_append nth_map_upt) +--{* Collector *} +apply(rule Mul_Collector) +--{* Mutator *} +apply(erule Mul_Mutator) +--{* Interference freedom *} +apply(simp add:Mul_interfree_Collector_Mutator) +apply(simp add:Mul_interfree_Mutator_Collector) +apply(simp add:Mul_interfree_Mutator_Mutator) +--{* Weakening of the postcondition *} +apply(case_tac n) + apply(simp add:Mul_Collector_def mul_mutator_defs mul_collector_defs nth_append) +apply(simp add:Mul_Mutator_def mul_mutator_defs mul_collector_defs nth_append) +done + +end \ No newline at end of file diff -r 98f0a09a33c3 -r 791e3b4c4039 src/HOL/HoareParallel/OG_Com.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HoareParallel/OG_Com.thy Tue Mar 05 17:11:25 2002 +0100 @@ -0,0 +1,55 @@ + +header {* \chapter{The Owicki-Gries Method} + +\section{Abstract Syntax} *} + +theory OG_Com = Main: + +text {* Type abbreviations for boolean expressions and assertions: *} + +types + 'a bexp = "'a set" + 'a assn = "'a set" + +text {* The syntax of commands is defined by two mutually recursive +datatypes: @{text "'a ann_com"} for annotated commands and @{text "'a +com"} for non-annotated commands. *} + +datatype 'a ann_com = + AnnBasic "('a assn)" "('a \ 'a)" + | AnnSeq "('a ann_com)" "('a ann_com)" + | AnnCond1 "('a assn)" "('a bexp)" "('a ann_com)" "('a ann_com)" + | AnnCond2 "('a assn)" "('a bexp)" "('a ann_com)" + | AnnWhile "('a assn)" "('a bexp)" "('a assn)" "('a ann_com)" + | AnnAwait "('a assn)" "('a bexp)" "('a com)" +and 'a com = + Parallel "('a ann_com option \ 'a assn) list" + | Basic "('a \ 'a)" + | Seq "('a com)" "('a com)" + | Cond "('a bexp)" "('a com)" "('a com)" + | While "('a bexp)" "('a assn)" "('a com)" + +text {* The function @{text pre} extracts the precondition of an +annotated command: *} + +consts + pre ::"'a ann_com \ 'a assn" +primrec + "pre (AnnBasic r f) = r" + "pre (AnnSeq c1 c2) = pre c1" + "pre (AnnCond1 r b c1 c2) = r" + "pre (AnnCond2 r b c) = r" + "pre (AnnWhile r b i c) = r" + "pre (AnnAwait r b c) = r" + +text {* Well-formedness predicate for atomic programs: *} + +consts atom_com :: "'a com \ bool" +primrec + "atom_com (Parallel Ts) = False" + "atom_com (Basic f) = True" + "atom_com (Seq c1 c2) = (atom_com c1 \ atom_com c2)" + "atom_com (Cond b c1 c2) = (atom_com c1 \ atom_com c2)" + "atom_com (While b i c) = atom_com c" + +end \ No newline at end of file diff -r 98f0a09a33c3 -r 791e3b4c4039 src/HOL/HoareParallel/OG_Examples.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HoareParallel/OG_Examples.thy Tue Mar 05 17:11:25 2002 +0100 @@ -0,0 +1,564 @@ + +header {* \section{Examples} *} + +theory OG_Examples = OG_Syntax: + +subsection {* Mutual Exclusion *} + +subsubsection {* Peterson's Algorithm I*} + +text {* Eike Best. "Semantics of Sequential and Parallel Programs", page 217. *} + +record Petersons_mutex_1 = + pr1 :: nat + pr2 :: nat + in1 :: bool + in2 :: bool + hold :: nat + +lemma Petersons_mutex_1: + "\- .{\pr1=0 \ \\in1 \ \pr2=0 \ \\in2 }. + COBEGIN .{\pr1=0 \ \\in1}. + WHILE True INV .{\pr1=0 \ \\in1}. + DO + .{\pr1=0 \ \\in1}. \ \in1:=True,,\pr1:=1 \;; + .{\pr1=1 \ \in1}. \ \hold:=1,,\pr1:=2 \;; + .{\pr1=2 \ \in1 \ (\hold=1 \ \hold=2 \ \pr2=2)}. + AWAIT (\\in2 \ \(\hold=1)) THEN \pr1:=3 END;; + .{\pr1=3 \ \in1 \ (\hold=1 \ \hold=2 \ \pr2=2)}. + \\in1:=False,,\pr1:=0\ + OD .{\pr1=0 \ \\in1}. + \ + .{\pr2=0 \ \\in2}. + WHILE True INV .{\pr2=0 \ \\in2}. + DO + .{\pr2=0 \ \\in2}. \ \in2:=True,,\pr2:=1 \;; + .{\pr2=1 \ \in2}. \ \hold:=2,,\pr2:=2 \;; + .{\pr2=2 \ \in2 \ (\hold=2 \ (\hold=1 \ \pr1=2))}. + AWAIT (\\in1 \ \(\hold=2)) THEN \pr2:=3 END;; + .{\pr2=3 \ \in2 \ (\hold=2 \ (\hold=1 \ \pr1=2))}. + \\in2:=False,,\pr2:=0\ + OD .{\pr2=0 \ \\in2}. + COEND + .{\pr1=0 \ \\in1 \ \pr2=0 \ \\in2}." +apply oghoare +--{* 104 verification conditions. *} +apply auto +done + +subsubsection {*Peterson's Algorithm II: A Busy Wait Solution *} + +text {* Apt and Olderog. "Verification of sequential and concurrent Programs", page 282. *} + +record Busy_wait_mutex = + flag1 :: bool + flag2 :: bool + turn :: nat + after1 :: bool + after2 :: bool + +lemma Busy_wait_mutex: + "\- .{True}. + \flag1:=False,, \flag2:=False,, + COBEGIN .{\\flag1}. + WHILE True + INV .{\\flag1}. + DO .{\\flag1}. \ \flag1:=True,,\after1:=False \;; + .{\flag1 \ \\after1}. \ \turn:=1,,\after1:=True \;; + .{\flag1 \ \after1 \ (\turn=1 \ \turn=2)}. + WHILE \(\flag2 \ \turn=2) + INV .{\flag1 \ \after1 \ (\turn=1 \ \turn=2)}. + DO .{\flag1 \ \after1 \ (\turn=1 \ \turn=2)}. SKIP OD;; + .{\flag1 \ \after1 \ (\flag2 \ \after2 \ \turn=2)}. + \flag1:=False + OD + .{False}. + \ + .{\\flag2}. + WHILE True + INV .{\\flag2}. + DO .{\\flag2}. \ \flag2:=True,,\after2:=False \;; + .{\flag2 \ \\after2}. \ \turn:=2,,\after2:=True \;; + .{\flag2 \ \after2 \ (\turn=1 \ \turn=2)}. + WHILE \(\flag1 \ \turn=1) + INV .{\flag2 \ \after2 \ (\turn=1 \ \turn=2)}. + DO .{\flag2 \ \after2 \ (\turn=1 \ \turn=2)}. SKIP OD;; + .{\flag2 \ \after2 \ (\flag1 \ \after1 \ \turn=1)}. + \flag2:=False + OD + .{False}. + COEND + .{False}." +apply oghoare +--{* 122 vc *} +apply auto +done + +subsubsection {* Peterson's Algorithm III: A Solution using Semaphores *} + +record Semaphores_mutex = + out :: bool + who :: nat + +lemma Semaphores_mutex: + "\- .{i\j}. + \out:=True ,, + COBEGIN .{i\j}. + WHILE True INV .{i\j}. + DO .{i\j}. AWAIT \out THEN \out:=False,, \who:=i END;; + .{\\out \ \who=i \ i\j}. \out:=True OD + .{False}. + \ + .{i\j}. + WHILE True INV .{i\j}. + DO .{i\j}. AWAIT \out THEN \out:=False,,\who:=j END;; + .{\\out \ \who=j \ i\j}. \out:=True OD + .{False}. + COEND + .{False}." +apply oghoare +--{* 38 vc *} +apply auto +done + +subsubsection {* Peterson's Algorithm III: Parameterized version: *} + +lemma Semaphores_parameterized_mutex: + "0 \- .{True}. + \out:=True ,, + COBEGIN + SCHEME [0\ i< n] + .{True}. + WHILE True INV .{True}. + DO .{True}. AWAIT \out THEN \out:=False,, \who:=i END;; + .{\\out \ \who=i}. \out:=True OD + .{False}. + COEND + .{False}." +apply oghoare +apply auto +done + +subsubsection{* The Ticket Algorithm *} + +record Ticket_mutex = + num :: nat + nextv :: nat + turn :: "nat list" + index :: nat + +lemma Ticket_mutex: + "\ 0n=length \turn \ 0<\nextv \ (\k l. k l k\l + \ \turn!k < \num \ (\turn!k =0 \ \turn!k\\turn!l))\ \ + \ \- .{n=length \turn}. + \index:= 0,, + WHILE \index < n INV .{n=length \turn \ (\i<\index. \turn!i=0)}. + DO \turn:= \turn[\index:=0],, \index:=\index +1 OD,, + \num:=1 ,, \nextv:=1 ,, + COBEGIN + SCHEME [0\ i< n] + .{\I}. + WHILE True INV .{\I}. + DO .{\I}. \ \turn :=\turn[i:=\num],, \num:=\num+1 \;; + .{\I}. WAIT \turn!i=\nextv END;; + .{\I \ \turn!i=\nextv}. \nextv:=\nextv+1 + OD + .{False}. + COEND + .{False}." +apply oghoare +--{* 35 vc *} +apply simp_all +apply(tactic {* ALLGOALS Clarify_tac *}) +apply simp_all +apply(tactic {* ALLGOALS Clarify_tac *}) +--{* 11 subgoals left *} +apply(erule less_SucE) + apply simp +apply simp +--{* 10 subgoals left *} +apply force +apply(case_tac "i=k") + apply force +apply simp +apply(case_tac "i=l") + apply force +apply force +--{* 8 subgoals left *} +prefer 8 +apply force +apply force +--{* 6 subgoals left *} +prefer 6 +apply(erule_tac x=i in allE) +apply force +--{* 5 subgoals left *} +prefer 5 +apply(rule conjI) + apply clarify +prefer 2 +apply(case_tac "j=i") + apply simp +apply simp +--{* 4 subgoals left *} +apply(tactic {* ALLGOALS (case_tac "j=k") *}) +apply simp_all +apply(erule_tac x=i in allE) +apply force +apply(case_tac "j=l") + apply simp + apply(erule_tac x=k in allE) + apply(erule_tac x=k in allE) + apply(erule_tac x=l in allE) + apply force +apply(erule_tac x=k in allE) +apply(erule_tac x=k in allE) +apply(erule_tac x=l in allE) +apply force +--{* 8 subgoals left *} +apply force +apply(case_tac "j=l") + apply simp +apply(erule_tac x=k in allE) +apply(erule_tac x=l in allE) +apply force +apply force +apply force +apply(erule_tac x=k in allE) +apply(erule_tac x=l in allE) +apply(case_tac "j=l") + apply force +apply force +apply force +apply(erule_tac x=k in allE) +apply(erule_tac x=l in allE) +apply(case_tac "j=l") + apply force +apply force +apply force +apply(erule_tac x=k in allE) +apply(erule_tac x=l in allE) +apply(case_tac "j=l") + apply force +apply force +done + +subsection{* Parallel Zero Search *} + +text {* Synchronized Zero Search. Zero-6 *} + +text {*Apt and Olderog. "Verification of sequential and concurrent Programs" page 294: *} + +record Zero_search = + turn :: nat + found :: bool + x :: nat + y :: nat + +lemma Zero_search: + "\I1= \ a\\x \ (\found \ (a<\x \ f(\x)=0) \ (\y\a \ f(\y)=0)) + \ (\\found \ a<\ x \ f(\x)\0) \ ; + I2= \\y\a+1 \ (\found \ (a<\x \ f(\x)=0) \ (\y\a \ f(\y)=0)) + \ (\\found \ \y\a \ f(\y)\0) \ \ \ + \- .{\ u. f(u)=0}. + \turn:=1,, \found:= False,, + \x:=a,, \y:=a+1 ,, + COBEGIN .{\I1}. + WHILE \\found + INV .{\I1}. + DO .{a\\x \ (\found \ \y\a \ f(\y)=0) \ (a<\x \ f(\x)\0)}. + WAIT \turn=1 END;; + .{a\\x \ (\found \ \y\a \ f(\y)=0) \ (a<\x \ f(\x)\0)}. + \turn:=2;; + .{a\\x \ (\found \ \y\a \ f(\y)=0) \ (a<\x \ f(\x)\0)}. + \ \x:=\x+1,, + IF f(\x)=0 THEN \found:=True ELSE SKIP FI\ + OD;; + .{\I1 \ \found}. + \turn:=2 + .{\I1 \ \found}. + \ + .{\I2}. + WHILE \\found + INV .{\I2}. + DO .{\y\a+1 \ (\found \ a<\x \ f(\x)=0) \ (\y\a \ f(\y)\0)}. + WAIT \turn=2 END;; + .{\y\a+1 \ (\found \ a<\x \ f(\x)=0) \ (\y\a \ f(\y)\0)}. + \turn:=1;; + .{\y\a+1 \ (\found \ a<\x \ f(\x)=0) \ (\y\a \ f(\y)\0)}. + \ \y:=(\y - 1),, + IF f(\y)=0 THEN \found:=True ELSE SKIP FI\ + OD;; + .{\I2 \ \found}. + \turn:=1 + .{\I2 \ \found}. + COEND + .{f(\x)=0 \ f(\y)=0}." +apply oghoare +--{* 98 verification conditions *} +apply auto +--{* auto takes about 3 minutes !! *} +apply arith+ +done + +text {* Easier Version: without AWAIT. Apt and Olderog. page 256: *} + +lemma Zero_Search_2: +"\I1=\ a\\x \ (\found \ (a<\x \ f(\x)=0) \ (\y\a \ f(\y)=0)) + \ (\\found \ a<\x \ f(\x)\0)\; + I2= \\y\a+1 \ (\found \ (a<\x \ f(\x)=0) \ (\y\a \ f(\y)=0)) + \ (\\found \ \y\a \ f(\y)\0)\\ \ + \- .{\u. f(u)=0}. + \found:= False,, + \x:=a,, \y:=a+1,, + COBEGIN .{\I1}. + WHILE \\found + INV .{\I1}. + DO .{a\\x \ (\found \ \y\a \ f(\y)=0) \ (a<\x \ f(\x)\0)}. + \ \x:=\x+1,,IF f(\x)=0 THEN \found:=True ELSE SKIP FI\ + OD + .{\I1 \ \found}. + \ + .{\I2}. + WHILE \\found + INV .{\I2}. + DO .{\y\a+1 \ (\found \ a<\x \ f(\x)=0) \ (\y\a \ f(\y)\0)}. + \ \y:=(\y - 1),,IF f(\y)=0 THEN \found:=True ELSE SKIP FI\ + OD + .{\I2 \ \found}. + COEND + .{f(\x)=0 \ f(\y)=0}." +apply oghoare +--{* 20 vc *} +apply auto +--{* auto takes aprox. 2 minutes. *} +apply arith+ +done + +subsection {* Producer/Consumer *} + +subsubsection {* Previous lemmas *} + +lemma nat_lemma1: "\(c::nat) \ a ;a < b\ \ b - a \ b - c" +by (simp split: nat_diff_split) + +lemma nat_lemma2: "\ b = m*(n::nat) + t; a = s*n + t ; b - a < n \ \ m - s = 0" +proof - + assume "b = m*(n::nat) + t" and "a = s*n + t" + hence "(m - s) * n = b - a" by (simp add: diff_mult_distrib) + also assume "\ < n" + finally have "m - s < 1" by simp + thus ?thesis by arith +qed + +lemma less_imp_diff_is_0: "m < Suc(n) \ m-n = 0" +by arith +lemma mod_lemma: "\ (c::nat) \ a; a < b; b - c < n \ \ b mod n \ a mod n" +apply(subgoal_tac "b=b div n*n + b mod n" ) + prefer 2 apply (simp add: mod_div_equality [symmetric]) +apply(subgoal_tac "a=a div n*n + a mod n") + prefer 2 + apply(simp add: mod_div_equality [symmetric]) +apply(frule nat_lemma1 , assumption) +apply(drule le_less_trans) +back + apply assumption +apply(frule less_not_refl2) +apply(drule less_imp_le) +apply (drule_tac m = "a" in div_le_mono) +apply(drule_tac m = "a div n" in le_imp_less_Suc) +apply(drule less_imp_diff_is_0) +apply(safe) +apply(simp) +apply(frule_tac b = "b" and a = "a" and n = "n" in nat_lemma2 , assumption) +apply assumption +apply(drule diffs0_imp_equal) +apply(simp) +apply(simp) +done + +subsubsection {* Producer/Consumer Algorithm *} + +record Producer_consumer = + ins :: nat + outs :: nat + li :: nat + lj :: nat + vx :: nat + vy :: nat + buffer :: "nat list" + b :: "nat list" + +text {* The whole proof takes aprox. 4 minutes. *} + +lemma Producer_consumer: + "\INIT= \0 0buffer \ length \b=length a\ ; + I= \(\k<\ins. \outs\k \ (a ! k) = \buffer ! (k mod (length \buffer))) \ + \outs\\ins \ \ins-\outs\length \buffer\ ; + I1= \\I \ \li\length a\ ; + p1= \\I1 \ \li=\ins\ ; + I2 = \\I \ (\k<\lj. (a ! k)=(\b ! k)) \ \lj\length a\ ; + p2 = \\I2 \ \lj=\outs\ \ \ + \- .{\INIT}. + \ins:=0,, \outs:=0,, \li:=0,, \lj:=0,, + COBEGIN .{\p1 \ \INIT}. + WHILE \li p1 \ \INIT}. + DO .{\p1 \ \INIT \ \livx:= (a ! \li);; + .{\p1 \ \INIT \ \li \vx=(a ! \li)}. + WAIT \ins-\outs < length \buffer END;; + .{\p1 \ \INIT \ \li \vx=(a ! \li) + \ \ins-\outs < length \buffer}. + \buffer:=(list_update \buffer (\ins mod (length \buffer)) \vx);; + .{\p1 \ \INIT \ \li (a ! \li)=(\buffer ! (\ins mod (length \buffer))) + \ \ins-\outs buffer}. + \ins:=\ins+1;; + .{\I1 \ \INIT \ (\li+1)=\ins \ \lili:=\li+1 + OD + .{\p1 \ \INIT \ \li=length a}. + \ + .{\p2 \ \INIT}. + WHILE \lj < length a + INV .{\p2 \ \INIT}. + DO .{\p2 \ \lj \INIT}. + WAIT \outs<\ins END;; + .{\p2 \ \lj \outs<\ins \ \INIT}. + \vy:=(\buffer ! (\outs mod (length \buffer)));; + .{\p2 \ \lj \outs<\ins \ \vy=(a ! \lj) \ \INIT}. + \outs:=\outs+1;; + .{\I2 \ (\lj+1)=\outs \ \lj \vy=(a ! \lj) \ \INIT}. + \b:=(list_update \b \lj \vy);; + .{\I2 \ (\lj+1)=\outs \ \lj (a ! \lj)=(\b ! \lj) \ \INIT}. + \lj:=\lj+1 + OD + .{\p2 \ \lj=length a \ \INIT}. + COEND + .{ \kb ! k)}." +apply oghoare +--{* 138 vc *} +apply(tactic {* ALLGOALS Clarify_tac *}) +--{* 112 subgoals left *} +apply(simp_all (no_asm)) +apply(tactic {*ALLGOALS (conjI_Tac (K all_tac)) *}) +--{* 860 subgoals left *} +apply(tactic {* ALLGOALS Clarify_tac *}) +apply(simp_all only:length_0_conv [THEN sym]) +--{* 36 subgoals left *} +apply (simp_all del:length_0_conv add: nth_list_update mod_less_divisor mod_lemma) +--{* 32 subgoals left *} +apply(tactic {* ALLGOALS Clarify_tac *}) +apply(tactic {* TRYALL arith_tac *}) +--{* 9 subgoals left *} +apply (force simp add:less_Suc_eq) +apply(drule sym) +apply (force simp add:less_Suc_eq)+ +done + +subsection {* Parameterized Examples *} + +subsubsection {* Set Elements of an Array to Zero *} + +record scheme1_vars = + a :: "nat \ nat" +lemma scheme1: + "\- .{True}. + COBEGIN SCHEME [0\ia:=\a (i:=0) .{\a i=0}. COEND + .{\i < n. \a i = 0}." +apply oghoare +apply simp_all +done + +text {* Same example with lists as auxiliary variables. *} +record scheme1_list_vars = + a :: "nat list" +lemma scheme1_list: + "\- .{n < length \a}. + COBEGIN + SCHEME [0\ia}. \a:=\a[i:=0] .{\a!i=0}. + COEND + .{\i < n. \a!i = 0}." +apply oghoare +apply simp_all + apply force+ +apply clarify +apply (simp add:nth_list_update) +done + +subsubsection {* Increment a Variable in Parallel *} + +text {* First some lemmas about summation properties. Summation is +defined in PreList. *} + +lemma scheme2_lemma1: "!!b. j (\i b j = 0 " +apply(induct n) + apply simp_all +apply(force simp add: less_Suc_eq) +done + +lemma scheme2_lemma2_aux: "!!b. j + (\iii s \ (\iij \ Suc (\i< n. b i)=(\i< n. (b (j := Suc 0)) i)" +apply(frule_tac b="(b (j:=(Suc 0)))" in scheme2_lemma2_aux) +apply(erule_tac t="Summation (b(j := (Suc 0))) n" in ssubst) +apply(frule_tac b=b in scheme2_lemma2_aux) +apply(erule_tac t="Summation b n" in ssubst) +apply(subgoal_tac "Suc (Summation b j + b j + (\iij") + apply(drule_tac b="b" and t="(Suc 0)" in scheme2_lemma2_aux2) +apply(rotate_tac -1) +apply(erule ssubst) +apply simp_all +done + +lemma scheme2_lemma3: "!!b. \i< n. b i = (Suc 0) \ (\i nat" + x :: nat +lemma scheme_2: "0 + \- .{\x=0 \ (\i< n. \c i)=0}. + COBEGIN + SCHEME [0\ix=(\i< n. \c i) \ \c i=0}. + \ \x:=\x+(Suc 0),, \c:=\c (i:=(Suc 0)) \ + .{\x=(\i< n. \c i) \ \c i=(Suc 0)}. + COEND + .{\x=n}." +apply oghoare +apply simp_all +apply (tactic {* ALLGOALS Clarify_tac *}) +apply simp_all + apply(force elim:scheme2_lemma1) + apply(erule scheme2_lemma2) + apply simp + apply(erule scheme2_lemma2) + apply simp + apply(erule scheme2_lemma2) + apply simp +apply(force intro: scheme2_lemma3) +done + +end \ No newline at end of file diff -r 98f0a09a33c3 -r 791e3b4c4039 src/HOL/HoareParallel/OG_Hoare.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HoareParallel/OG_Hoare.thy Tue Mar 05 17:11:25 2002 +0100 @@ -0,0 +1,495 @@ + +header {* \section{The Proof System} *} + +theory OG_Hoare = OG_Tran: + +consts assertions :: "'a ann_com \ ('a assn) set" +primrec + "assertions (AnnBasic r f) = {r}" + "assertions (AnnSeq c1 c2) = assertions c1 \ assertions c2" + "assertions (AnnCond1 r b c1 c2) = {r} \ assertions c1 \ assertions c2" + "assertions (AnnCond2 r b c) = {r} \ assertions c" + "assertions (AnnWhile r b i c) = {r, i} \ assertions c" + "assertions (AnnAwait r b c) = {r}" + +consts atomics :: "'a ann_com \ ('a assn \ 'a com) set" +primrec + "atomics (AnnBasic r f) = {(r, Basic f)}" + "atomics (AnnSeq c1 c2) = atomics c1 \ atomics c2" + "atomics (AnnCond1 r b c1 c2) = atomics c1 \ atomics c2" + "atomics (AnnCond2 r b c) = atomics c" + "atomics (AnnWhile r b i c) = atomics c" + "atomics (AnnAwait r b c) = {(r \ b, c)}" + +consts com :: "'a ann_triple_op \ 'a ann_com_op" +primrec "com (c, q) = c" + +consts post :: "'a ann_triple_op \ 'a assn" +primrec "post (c, q) = q" + +constdefs interfree_aux :: "('a ann_com_op \ 'a assn \ 'a ann_com_op) \ bool" + "interfree_aux \ \(co, q, co'). co'= None \ + (\(r,a) \ atomics (the co'). \= (q \ r) a q \ + (co = None \ (\p \ assertions (the co). \= (p \ r) a p)))" + +constdefs interfree :: "(('a ann_triple_op) list) \ bool" + "interfree Ts \ \i j. i < length Ts \ j < length Ts \ i \ j \ + interfree_aux (com (Ts!i), post (Ts!i), com (Ts!j)) " + +consts ann_hoare :: "('a ann_com \ 'a assn) set" +syntax "_ann_hoare" :: "'a ann_com \ 'a assn \ bool" ("(2\ _// _)" [60,90] 45) +translations "\ c q" \ "(c, q) \ ann_hoare" + +consts oghoare :: "('a assn \ 'a com \ 'a assn) set" +syntax "_oghoare" :: "'a assn \ 'a com \ 'a assn \ bool" ("(3\- _//_//_)" [90,55,90] 50) +translations "\- p c q" \ "(p, c, q) \ oghoare" + +inductive oghoare ann_hoare +intros + AnnBasic: "r \ {s. f s \ q} \ \ (AnnBasic r f) q" + + AnnSeq: "\ \ c0 pre c1; \ c1 q \ \ \ (AnnSeq c0 c1) q" + + AnnCond1: "\ r \ b \ pre c1; \ c1 q; r \ -b \ pre c2; \ c2 q\ + \ \ (AnnCond1 r b c1 c2) q" + AnnCond2: "\ r \ b \ pre c; \ c q; r \ -b \ q \ \ \ (AnnCond2 r b c) q" + + AnnWhile: "\ r \ i; i \ b \ pre c; \ c i; i \ -b \ q \ + \ \ (AnnWhile r b i c) q" + + AnnAwait: "\ atom_com c; \- (r \ b) c q \ \ \ (AnnAwait r b c) q" + + AnnConseq: "\\ c q; q \ q' \ \ \ c q'" + + + Parallel: "\ \ic q. Ts!i = (Some c, q) \ \ c q; interfree Ts \ + \ \- (\i\{i. ii\{i. i- {s. f s \q} (Basic f) q" + + Seq: "\ \- p c1 r; \- r c2 q \ \ \- p (Seq c1 c2) q " + + Cond: "\ \- (p \ b) c1 q; \- (p \ -b) c2 q \ \ \- p (Cond b c1 c2) q" + + While: "\ \- (p \ b) c p \ \ \- p (While b i c) (p \ -b)" + + Conseq: "\ p' \ p; \- p c q ; q \ q' \ \ \- p' c q'" + +section {* Soundness *} +(* In the version Isabelle-10-Sep-1999: HOL: The THEN and ELSE +parts of conditional expressions (if P then x else y) are no longer +simplified. (This allows the simplifier to unfold recursive +functional programs.) To restore the old behaviour, we declare +@{text "lemmas [cong del] = if_weak_cong"}. *) + +lemmas [cong del] = if_weak_cong + +lemmas ann_hoare_induct = oghoare_ann_hoare.induct [THEN conjunct2] +lemmas oghoare_induct = oghoare_ann_hoare.induct [THEN conjunct1] + +lemmas AnnBasic = oghoare_ann_hoare.AnnBasic +lemmas AnnSeq = oghoare_ann_hoare.AnnSeq +lemmas AnnCond1 = oghoare_ann_hoare.AnnCond1 +lemmas AnnCond2 = oghoare_ann_hoare.AnnCond2 +lemmas AnnWhile = oghoare_ann_hoare.AnnWhile +lemmas AnnAwait = oghoare_ann_hoare.AnnAwait +lemmas AnnConseq = oghoare_ann_hoare.AnnConseq + +lemmas Parallel = oghoare_ann_hoare.Parallel +lemmas Basic = oghoare_ann_hoare.Basic +lemmas Seq = oghoare_ann_hoare.Seq +lemmas Cond = oghoare_ann_hoare.Cond +lemmas While = oghoare_ann_hoare.While +lemmas Conseq = oghoare_ann_hoare.Conseq + +subsection {* Soundness of the System for Atomic Programs *} + +lemma Basic_ntran [rule_format]: + "(Basic f, s) -Pn\ (Parallel Ts, t) \ All_None Ts \ t = f s" +apply(induct "n") + apply(simp (no_asm)) +apply(fast dest: rel_pow_Suc_D2 Parallel_empty_lemma elim: transition_cases) +done + +lemma SEM_fwhile: "SEM S (p \ b) \ p \ SEM (fwhile b S k) p \ (p \ -b)" +apply (induct "k") + apply(simp (no_asm) add: L3_5v_lemma3) +apply(simp (no_asm) add: L3_5iv L3_5ii Parallel_empty) +apply(rule Un_least) + apply(rule subset_trans) + prefer 2 apply simp + apply(erule L3_5i) +apply(simp add: SEM_def sem_def id_def) +apply clarify +apply(drule rtrancl_imp_UN_rel_pow) +apply clarify +apply(drule Basic_ntran) + apply fast+ +done + +lemma atom_hoare_sound [rule_format (no_asm)]: + " \- p c q \ atom_com(c) \ \= p c q" +apply (unfold com_validity_def) +apply(rule oghoare_induct) +apply simp_all +--{*Basic*} + apply(simp add: SEM_def sem_def) + apply(fast dest: rtrancl_imp_UN_rel_pow Basic_ntran) +--{* Seq *} + apply(rule impI) + apply(rule subset_trans) + prefer 2 apply simp + apply(simp add: L3_5ii L3_5i) +--{* Cond *} + apply(rule impI) + apply(simp add: L3_5iv) + apply(erule Un_least) + apply assumption +--{* While *} + apply(rule impI) + apply(simp add: L3_5v) + apply(rule UN_least) + apply(drule SEM_fwhile) + apply assumption +--{* Conseq *} +apply(simp add: SEM_def sem_def) +apply force +done + +subsection {* Soundness of the System for Component Programs *} + +inductive_cases ann_transition_cases: + "(None,s) -1\ t" + "(Some (AnnBasic r f),s) -1\ t" + "(Some (AnnSeq c1 c2), s) -1\ t" + "(Some (AnnCond1 r b c1 c2), s) -1\ t" + "(Some (AnnCond2 r b c), s) -1\ t" + "(Some (AnnWhile r b I c), s) -1\ t" + "(Some (AnnAwait r b c),s) -1\ t" + +text {* Strong Soundness for Component Programs:*} + +lemma ann_hoare_case_analysis [rule_format]: "\ C q' \ + ((\r f. C = AnnBasic r f \ (\q. r \ {s. f s \ q} \ q \ q')) \ + (\c0 c1. C = AnnSeq c0 c1 \ (\q. q \ q' \ \ c0 pre c1 \ \ c1 q)) \ + (\r b c1 c2. C = AnnCond1 r b c1 c2 \ (\q. q \ q' \ + r \ b \ pre c1 \ \ c1 q \ r \ -b \ pre c2 \ \ c2 q)) \ + (\r b c. C = AnnCond2 r b c \ + (\q. q \ q' \ r \ b \ pre c \ \ c q \ r \ -b \ q)) \ + (\r i b c. C = AnnWhile r b i c \ + (\q. q \ q' \ r \ i \ i \ b \ pre c \ \ c i \ i \ -b \ q)) \ + (\r b c. C = AnnAwait r b c \ (\q. q \ q' \ \- (r \ b) c q)))" +apply(rule ann_hoare_induct) +apply simp_all + apply(rule_tac x=q in exI,simp)+ +apply(rule conjI,clarify,simp,clarify,rule_tac x=qa in exI,fast)+ +apply(clarify,simp,clarify,rule_tac x=qa in exI,fast) +done + +lemma Help: "(transition \ {(v,v,u). True}) = (transition)" +apply force +done + +lemma Strong_Soundness_aux_aux [rule_format]: + "(co, s) -1\ (co', t) \ (\c. co = Some c \ s\ pre c \ + (\q. \ c q \ (if co' = None then t\q else t \ pre(the co') \ \ (the co') q )))" +apply(rule ann_transition_transition.induct [THEN conjunct1]) +apply simp_all +--{* Basic *} + apply clarify + apply(frule ann_hoare_case_analysis) + apply force +--{* Seq *} + apply clarify + apply(frule ann_hoare_case_analysis,simp) + apply(fast intro: AnnConseq) + apply clarify + apply(frule ann_hoare_case_analysis,simp) + apply clarify + apply(rule conjI) + apply force + apply(rule AnnSeq,simp) + apply(fast intro: AnnConseq) +--{* Cond1 *} + apply clarify + apply(frule ann_hoare_case_analysis,simp) + apply(fast intro: AnnConseq) + apply clarify + apply(frule ann_hoare_case_analysis,simp) + apply(fast intro: AnnConseq) +--{* Cond2 *} + apply clarify + apply(frule ann_hoare_case_analysis,simp) + apply(fast intro: AnnConseq) + apply clarify + apply(frule ann_hoare_case_analysis,simp) + apply(fast intro: AnnConseq) +--{* While *} + apply clarify + apply(frule ann_hoare_case_analysis,simp) + apply force + apply clarify + apply(frule ann_hoare_case_analysis,simp) + apply auto + apply(rule AnnSeq) + apply simp + apply(rule AnnWhile) + apply simp_all + apply(fast) +--{* Await *} +apply(frule ann_hoare_case_analysis,simp) +apply clarify +apply(drule atom_hoare_sound) + apply simp +apply(simp add: com_validity_def SEM_def sem_def) +apply(simp add: Help All_None_def) +apply force +done + +lemma Strong_Soundness_aux: "\ (Some c, s) -*\ (co, t); s \ pre c; \ c q \ + \ if co = None then t \ q else t \ pre (the co) \ \ (the co) q" +apply(erule rtrancl_induct2) + apply simp +apply(case_tac "a") + apply(fast elim: ann_transition_cases) +apply(erule Strong_Soundness_aux_aux) + apply simp +apply simp_all +done + +lemma Strong_Soundness: "\ (Some c, s)-*\(co, t); s \ pre c; \ c q \ + \ if co = None then t\q else t \ pre (the co)" +apply(force dest:Strong_Soundness_aux) +done + +lemma ann_hoare_sound: "\ c q \ \ c q" +apply (unfold ann_com_validity_def ann_SEM_def ann_sem_def) +apply clarify +apply(drule Strong_Soundness) +apply simp_all +done + +subsection {* Soundness of the System for Parallel Programs *} + +lemma Parallel_length_post_P1: "(Parallel Ts,s) -P1\ (R', t) \ + (\Rs. R' = (Parallel Rs) \ (length Rs) = (length Ts) \ + (\i. i post(Rs ! i) = post(Ts ! i)))" +apply(erule transition_cases) +apply simp +apply clarify +apply(case_tac "i=ia") +apply simp+ +done + +lemma Parallel_length_post_PStar: "(Parallel Ts,s) -P*\ (R',t) \ + (\Rs. R' = (Parallel Rs) \ (length Rs) = (length Ts) \ + (\i. i post(Ts ! i) = post(Rs ! i)))" +apply(erule rtrancl_induct2) + apply(simp_all) +apply clarify +apply simp +apply(drule Parallel_length_post_P1) +apply auto +done + +lemma assertions_lemma: "pre c \ assertions c" +apply(rule ann_com_com.induct [THEN conjunct1]) +apply auto +done + +lemma interfree_aux1 [rule_format]: + "(c,s) -1\ (r,t) \ (interfree_aux(c1, q1, c) \ interfree_aux(c1, q1, r))" +apply (rule ann_transition_transition.induct [THEN conjunct1]) +apply(safe) +prefer 13 +apply (rule TrueI) +apply (simp_all add:interfree_aux_def) +apply force+ +done + +lemma interfree_aux2 [rule_format]: + "(c,s) -1\ (r,t) \ (interfree_aux(c, q, a) \ interfree_aux(r, q, a) )" +apply (rule ann_transition_transition.induct [THEN conjunct1]) +apply(force simp add:interfree_aux_def)+ +done + +lemma interfree_lemma: "\ (Some c, s) -1\ (r, t);interfree Ts ; i \ interfree (Ts[i:= (r, q)])" +apply(simp add: interfree_def) +apply clarify +apply(case_tac "i=j") + apply(drule_tac t = "ia" in not_sym) + apply simp_all +apply(force elim: interfree_aux1) +apply(force elim: interfree_aux2 simp add:nth_list_update) +done + +text {* Strong Soundness Theorem for Parallel Programs:*} + +lemma Parallel_Strong_Soundness_Seq_aux: + "\interfree Ts; i + \ interfree (Ts[i:=(Some c0, pre c1)])" +apply(simp add: interfree_def) +apply clarify +apply(case_tac "i=j") + apply(force simp add: nth_list_update interfree_aux_def) +apply(case_tac "i=ia") + apply(erule_tac x=ia in allE) + apply(force simp add:interfree_aux_def assertions_lemma) +apply simp +done + +lemma Parallel_Strong_Soundness_Seq [rule_format (no_asm)]: + "\ \i post(Ts!i) + else b \ pre(the(com(Ts!i))) \ \ the(com(Ts!i)) post(Ts!i)); + com(Ts ! i) = Some(AnnSeq c0 c1); i \ + (\ia post(Ts[i:=(Some c0, pre c1)]! ia) + else b \ pre(the(com(Ts[i:=(Some c0, pre c1)]! ia))) \ + \ the(com(Ts[i:=(Some c0, pre c1)]! ia)) post(Ts[i:=(Some c0, pre c1)]! ia))) + \ interfree (Ts[i:= (Some c0, pre c1)])" +apply(rule conjI) + apply safe + apply(case_tac "i=ia") + apply simp + apply(force dest: ann_hoare_case_analysis) + apply simp +apply(fast elim: Parallel_Strong_Soundness_Seq_aux) +done + +lemma Parallel_Strong_Soundness_aux_aux [rule_format]: + "(Some c, b) -1\ (co, t) \ + (\Ts. i com(Ts ! i) = Some c \ + (\ipost(Ts!i) + else b\pre(the(com(Ts!i))) \ \ the(com(Ts!i)) post(Ts!i))) \ + interfree Ts \ + (\j. j i\j \ (if com(Ts!j) = None then t\post(Ts!j) + else t\pre(the(com(Ts!j))) \ \ the(com(Ts!j)) post(Ts!j))) )" +apply(rule ann_transition_transition.induct [THEN conjunct1]) +apply safe +prefer 11 +apply(rule TrueI) +apply simp_all +--{* Basic *} + apply(erule_tac x = "i" in all_dupE, erule (1) notE impE) + apply(erule_tac x = "j" in allE , erule (1) notE impE) + apply(simp add: interfree_def) + apply(erule_tac x = "j" in allE,simp) + apply(erule_tac x = "i" in allE,simp) + apply(drule_tac t = "i" in not_sym) + apply(case_tac "com(Ts ! j)=None") + apply(force intro: converse_rtrancl_into_rtrancl + simp add: interfree_aux_def com_validity_def SEM_def sem_def All_None_def) + apply(simp add:interfree_aux_def) + apply clarify + apply simp + apply clarify + apply(erule_tac x="pre y" in ballE) + apply(force intro: converse_rtrancl_into_rtrancl + simp add: com_validity_def SEM_def sem_def All_None_def) + apply(simp add:assertions_lemma) +--{* Seqs *} + apply(erule_tac x = "Ts[i:=(Some c0, pre c1)]" in allE) + apply(drule Parallel_Strong_Soundness_Seq,simp+) + apply(erule_tac x = "Ts[i:=(Some c0, pre c1)]" in allE) + apply(drule Parallel_Strong_Soundness_Seq,simp+) +--{* Await *} +apply(rule_tac x = "i" in allE , assumption , erule (1) notE impE) +apply(erule_tac x = "j" in allE , erule (1) notE impE) +apply(simp add: interfree_def) +apply(erule_tac x = "j" in allE,simp) +apply(erule_tac x = "i" in allE,simp) +apply(drule_tac t = "i" in not_sym) +apply(case_tac "com(Ts ! j)=None") + apply(force intro: converse_rtrancl_into_rtrancl simp add: interfree_aux_def + com_validity_def SEM_def sem_def All_None_def Help) +apply(simp add:interfree_aux_def) +apply clarify +apply simp +apply clarify +apply(erule_tac x="pre y" in ballE) + apply(force intro: converse_rtrancl_into_rtrancl + simp add: com_validity_def SEM_def sem_def All_None_def Help) +apply(simp add:assertions_lemma) +done + +lemma Parallel_Strong_Soundness_aux [rule_format]: + "\(Ts',s) -P*\ (Rs',t); Ts' = (Parallel Ts); interfree Ts; + \i. i (\c q. (Ts ! i) = (Some c, q) \ s\(pre c) \ \ c q ) \ \ + \Rs. Rs' = (Parallel Rs) \ (\j. j + (if com(Rs ! j) = None then t\post(Ts ! j) + else t\pre(the(com(Rs ! j))) \ \ the(com(Rs ! j)) post(Ts ! j))) \ interfree Rs" +apply(erule rtrancl_induct2) + apply clarify +--{* Base *} + apply force +--{* Induction step *} +apply clarify +apply(drule Parallel_length_post_PStar) +apply clarify +apply (ind_cases "(Parallel Ts, s) -P1\ (Parallel Rs, t)") +apply(rule conjI) + apply clarify + apply(case_tac "i=j") + apply(simp split del:split_if) + apply(erule Strong_Soundness_aux_aux,simp+) + apply force + apply force + apply(simp split del: split_if) + apply(erule Parallel_Strong_Soundness_aux_aux) + apply(simp_all add: split del:split_if) + apply force +apply(rule interfree_lemma) +apply simp_all +done + +lemma Parallel_Strong_Soundness: + "\(Parallel Ts, s) -P*\ (Parallel Rs, t); interfree Ts; ji. i (\c q. Ts ! i = (Some c, q) \ s\pre c \ \ c q) \ \ + if com(Rs ! j) = None then t\post(Ts ! j) else t\pre (the(com(Rs ! j)))" +apply(drule Parallel_Strong_Soundness_aux) +apply simp+ +done + +lemma oghoare_sound [rule_format (no_asm)]: "\- p c q \ \= p c q" +apply (unfold com_validity_def) +apply(rule oghoare_induct) +apply(rule TrueI)+ +--{* Parallel *} + apply(simp add: SEM_def sem_def) + apply clarify + apply(frule Parallel_length_post_PStar) + apply clarify + apply(drule_tac j=i in Parallel_Strong_Soundness) + apply clarify + apply simp + apply force + apply simp + apply(erule_tac V = "\i. ?P i" in thin_rl) + apply(drule_tac s = "length Rs" in sym) + apply(erule allE, erule impE, assumption) + apply(force dest: nth_mem simp add: All_None_def) +--{* Basic *} + apply(simp add: SEM_def sem_def) + apply(force dest: rtrancl_imp_UN_rel_pow Basic_ntran) +--{* Seq *} + apply(rule subset_trans) + prefer 2 apply assumption + apply(simp add: L3_5ii L3_5i) +--{* Cond *} + apply(simp add: L3_5iv) + apply(erule Un_least) + apply assumption +--{* While *} + apply(simp add: L3_5v) + apply(rule UN_least) + apply(drule SEM_fwhile) + apply assumption +--{* Conseq *} +apply(simp add: SEM_def sem_def) +apply auto +done + +end \ No newline at end of file diff -r 98f0a09a33c3 -r 791e3b4c4039 src/HOL/HoareParallel/OG_Syntax.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HoareParallel/OG_Syntax.thy Tue Mar 05 17:11:25 2002 +0100 @@ -0,0 +1,137 @@ + +header {* \section{Concrete Syntax} *} + +theory OG_Syntax = Quote_Antiquote + OG_Tactics: + +text{* Syntax for commands and for assertions and boolean expressions in + commands @{text com} and annotated commands @{text ann_com}. *} + +syntax + "_Assign" :: "idt \ 'b \ 'a com" ("(\_ :=/ _)" [70, 65] 61) + "_AnnAssign" :: "'a assn \ idt \ 'b \ 'a com" ("(_ \_ :=/ _)" [90,70,65] 61) + +translations + "\\x := a" \ "Basic \\\(_update_name x a)\" + "r \\x := a" \ "AnnBasic r \\\(_update_name x a)\" + +syntax + "_AnnSkip" :: "'a assn \ 'a ann_com" ("_//SKIP" [90] 63) + "_AnnSeq" :: "'a ann_com \ 'a ann_com \ 'a ann_com" ("_;;/ _" [60,61] 60) + + "_AnnCond1" :: "'a assn \ 'a bexp \ 'a ann_com \ 'a ann_com \ 'a ann_com" + ("_ //IF _ /THEN _ /ELSE _ /FI" [90,0,0,0] 61) + "_AnnCond2" :: "'a assn \ 'a bexp \ 'a ann_com \ 'a ann_com" + ("_ //IF _ /THEN _ /FI" [90,0,0] 61) + "_AnnWhile" :: "'a assn \ 'a bexp \ 'a assn \ 'a ann_com \ 'a ann_com" + ("_ //WHILE _ /INV _ //DO _//OD" [90,0,0,0] 61) + "_AnnAwait" :: "'a assn \ 'a bexp \ 'a com \ 'a ann_com" + ("_ //AWAIT _ /THEN /_ /END" [90,0,0] 61) + "_AnnAtom" :: "'a assn \ 'a com \ 'a ann_com" ("_//\_\" [90,0] 61) + "_AnnWait" :: "'a assn \ 'a bexp \ 'a ann_com" ("_//WAIT _ END" [90,0] 61) + + "_Skip" :: "'a com" ("SKIP" 63) + "_Seq" :: "'a com \ 'a com \ 'a com" ("_,,/ _" [55, 56] 55) + "_Cond" :: "'a bexp \ 'a com \ 'a com \ 'a com" + ("(0IF _/ THEN _/ ELSE _/ FI)" [0, 0, 0] 61) + "_Cond2" :: "'a bexp \ 'a com \ 'a com" ("IF _ THEN _ FI" [0,0] 56) + "_While_inv" :: "'a bexp \ 'a assn \ 'a com \ 'a com" + ("(0WHILE _/ INV _ //DO _ /OD)" [0, 0, 0] 61) + "_While" :: "'a bexp \ 'a com \ 'a com" + ("(0WHILE _ //DO _ /OD)" [0, 0] 61) + +translations + "SKIP" \ "Basic id" + "c_1,, c_2" \ "Seq c_1 c_2" + + "IF b THEN c1 ELSE c2 FI" \ "Cond .{b}. c1 c2" + "IF b THEN c FI" \ "IF b THEN c ELSE SKIP FI" + "WHILE b INV i DO c OD" \ "While .{b}. i c" + "WHILE b DO c OD" \ "WHILE b INV arbitrary DO c OD" + + "r SKIP" \ "AnnBasic r id" + "c_1;; c_2" \ "AnnSeq c_1 c_2" + "r IF b THEN c1 ELSE c2 FI" \ "AnnCond1 r .{b}. c1 c2" + "r IF b THEN c FI" \ "AnnCond2 r .{b}. c" + "r WHILE b INV i DO c OD" \ "AnnWhile r .{b}. i c" + "r AWAIT b THEN c END" \ "AnnAwait r .{b}. c" + "r \c\" \ "r AWAIT True THEN c END" + "r WAIT b END" \ "r AWAIT b THEN SKIP END" + +nonterminals + prgs + +syntax + "_PAR" :: "prgs \ 'a" ("COBEGIN//_//COEND" [57] 56) + "_prg" :: "['a, 'a] \ prgs" ("_//_" [60, 90] 57) + "_prgs" :: "['a, 'a, prgs] \ prgs" ("_//_//\//_" [60,90,57] 57) + + "_prg_scheme" :: "['a, 'a, 'a, 'a, 'a] \ prgs" + ("SCHEME [_ \ _ < _] _// _" [0,0,0,60, 90] 57) + +translations + "_prg c q" \ "[(Some c, q)]" + "_prgs c q ps" \ "(Some c, q) # ps" + "_PAR ps" \ "Parallel ps" + + "_prg_scheme j i k c q" \ "(map (\i. (Some c, q)) [j..k(])" + +print_translation {* + let + fun quote_tr' f (t :: ts) = + Term.list_comb (f $ Syntax.quote_tr' "_antiquote" t, ts) + | quote_tr' _ _ = raise Match; + + fun annquote_tr' f (r :: t :: ts) = + Term.list_comb (f $ r $ Syntax.quote_tr' "_antiquote" t, ts) + | annquote_tr' _ _ = raise Match; + + val assert_tr' = quote_tr' (Syntax.const "_Assert"); + + fun bexp_tr' name ((Const ("Collect", _) $ t) :: ts) = + quote_tr' (Syntax.const name) (t :: ts) + | bexp_tr' _ _ = raise Match; + + fun annbexp_tr' name (r :: (Const ("Collect", _) $ t) :: ts) = + annquote_tr' (Syntax.const name) (r :: t :: ts) + | annbexp_tr' _ _ = raise Match; + + fun upd_tr' (x_upd, T) = + (case try (unsuffix RecordPackage.updateN) x_upd of + Some x => (x, if T = dummyT then T else Term.domain_type T) + | None => raise Match); + + fun update_name_tr' (Free x) = Free (upd_tr' x) + | update_name_tr' ((c as Const ("_free", _)) $ Free x) = + c $ Free (upd_tr' x) + | update_name_tr' (Const x) = Const (upd_tr' x) + | update_name_tr' _ = raise Match; + + fun assign_tr' (Abs (x, _, f $ t $ Bound 0) :: ts) = + quote_tr' (Syntax.const "_Assign" $ update_name_tr' f) + (Abs (x, dummyT, t) :: ts) + | assign_tr' _ = raise Match; + + fun annassign_tr' (r :: Abs (x, _, f $ t $ Bound 0) :: ts) = + quote_tr' (Syntax.const "_AnnAssign" $ r $ update_name_tr' f) + (Abs (x, dummyT, t) :: ts) + | annassign_tr' _ = raise Match; + + fun Parallel_PAR [(Const ("Cons",_) $ (Const ("Pair",_) $ (Const ("Some",_) $ t1 ) $ t2) $ Const ("Nil",_))] = + (Syntax.const "_prg" $ t1 $ t2) + | Parallel_PAR [(Const ("Cons",_) $ (Const ("Pair",_) $ (Const ("Some",_) $ t1) $ t2) $ ts)] = + (Syntax.const "_prgs" $ t1 $ t2 $ Parallel_PAR [ts]) + | Parallel_PAR _ = raise Match; + +fun Parallel_tr' ts = Syntax.const "_PAR" $ Parallel_PAR ts; + in + [("Collect", assert_tr'), ("Basic", assign_tr'), + ("Cond", bexp_tr' "_Cond"), ("While", bexp_tr' "_While_inv"), + ("AnnBasic", annassign_tr'), + ("AnnWhile", annbexp_tr' "_AnnWhile"), ("AnnAwait", annbexp_tr' "_AnnAwait"), + ("AnnCond1", annbexp_tr' "_AnnCond1"), ("AnnCond2", annbexp_tr' "_AnnCond2")] + + end + +*} + +end \ No newline at end of file diff -r 98f0a09a33c3 -r 791e3b4c4039 src/HOL/HoareParallel/OG_Tactics.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HoareParallel/OG_Tactics.thy Tue Mar 05 17:11:25 2002 +0100 @@ -0,0 +1,501 @@ + +header {* \section{Generation of Verification Conditions} *} + +theory OG_Tactics = OG_Hoare: + +lemmas ann_hoare_intros=AnnBasic AnnSeq AnnCond1 AnnCond2 AnnWhile AnnAwait AnnConseq +lemmas oghoare_intros=Parallel Basic Seq Cond While Conseq + +lemma ParallelConseqRule: + "\ p \ (\i\{i. i- (\i\{i. ii\{i. ii\{i. i q \ + \ \- p (Parallel Ts) q" +apply (rule Conseq) +prefer 2 + apply fast +apply assumption+ +done + +lemma SkipRule: "p \ q \ \- p (Basic id) q" +apply(rule oghoare_intros) + prefer 2 apply(rule Basic) + prefer 2 apply(rule subset_refl) +apply(simp add:Id_def) +done + +lemma BasicRule: "p \ {s. (f s)\q} \ \- p (Basic f) q" +apply(rule oghoare_intros) + prefer 2 apply(rule oghoare_intros) + prefer 2 apply(rule subset_refl) +apply assumption +done + +lemma SeqRule: "\ \- p c1 r; \- r c2 q \ \ \- p (Seq c1 c2) q" +apply(rule Seq) +apply fast+ +done + +lemma CondRule: + "\ p \ {s. (s\b \ s\w) \ (s\b \ s\w')}; \- w c1 q; \- w' c2 q \ + \ \- p (Cond b c1 c2) q" +apply(rule Cond) + apply(rule Conseq) + prefer 4 apply(rule Conseq) +apply simp_all +apply force+ +done + +lemma WhileRule: "\ p \ i; \- (i \ b) c i ; (i \ (-b)) \ q \ + \ \- p (While b i c) q" +apply(rule Conseq) + prefer 2 apply(rule While) +apply assumption+ +done + +text {* Three new proof rules for special instances of the @{text +AnnBasic} and the @{text AnnAwait} commands when the transformation +performed on the state is the identity, and for an @{text AnnAwait} +command where the boolean condition is @{text "{s. True}"}: *} + +lemma AnnatomRule: + "\ atom_com(c); \- r c q \ \ \ (AnnAwait r {s. True} c) q" +apply(rule AnnAwait) +apply simp_all +done + +lemma AnnskipRule: + "r \ q \ \ (AnnBasic r id) q" +apply(rule AnnBasic) +apply simp +done + +lemma AnnwaitRule: + "\ (r \ b) \ q \ \ \ (AnnAwait r b (Basic id)) q" +apply(rule AnnAwait) + apply simp +apply(rule BasicRule) +apply simp +done + +text {* Lemmata to avoid using the definition of @{text +map_ann_hoare}, @{text interfree_aux}, @{text interfree_swap} and +@{text interfree} by splitting it into different cases: *} + +lemma interfree_aux_rule1: "interfree_aux(co, q, None)" +by(simp add:interfree_aux_def) + +lemma interfree_aux_rule2: + "\(R,r)\(atomics a). \- (q \ R) r q \ interfree_aux(None, q, Some a)" +apply(simp add:interfree_aux_def) +apply(force elim:oghoare_sound) +done + +lemma interfree_aux_rule3: + "(\(R, r)\(atomics a). \- (q \ R) r q \ (\p\(assertions c). \- (p \ R) r p)) + \ interfree_aux(Some c, q, Some a)" +apply(simp add:interfree_aux_def) +apply(force elim:oghoare_sound) +done + +lemma AnnBasic_assertions: + "\interfree_aux(None, r, Some a); interfree_aux(None, q, Some a)\ \ + interfree_aux(Some (AnnBasic r f), q, Some a)" +apply(simp add: interfree_aux_def) +by force + +lemma AnnSeq_assertions: + "\ interfree_aux(Some c1, q, Some a); interfree_aux(Some c2, q, Some a)\\ + interfree_aux(Some (AnnSeq c1 c2), q, Some a)" +apply(simp add: interfree_aux_def) +by force + +lemma AnnCond1_assertions: + "\ interfree_aux(None, r, Some a); interfree_aux(Some c1, q, Some a); + interfree_aux(Some c2, q, Some a)\\ + interfree_aux(Some(AnnCond1 r b c1 c2), q, Some a)" +apply(simp add: interfree_aux_def) +by force + +lemma AnnCond2_assertions: + "\ interfree_aux(None, r, Some a); interfree_aux(Some c, q, Some a)\\ + interfree_aux(Some (AnnCond2 r b c), q, Some a)" +apply(simp add: interfree_aux_def) +by force + +lemma AnnWhile_assertions: + "\ interfree_aux(None, r, Some a); interfree_aux(None, i, Some a); + interfree_aux(Some c, q, Some a)\\ + interfree_aux(Some (AnnWhile r b i c), q, Some a)" +apply(simp add: interfree_aux_def) +by force + +lemma AnnAwait_assertions: + "\ interfree_aux(None, r, Some a); interfree_aux(None, q, Some a)\\ + interfree_aux(Some (AnnAwait r b c), q, Some a)" +apply(simp add: interfree_aux_def) +by force + +lemma AnnBasic_atomics: + "\- (q \ r) (Basic f) q \ interfree_aux(None, q, Some (AnnBasic r f))" +by(simp add: interfree_aux_def oghoare_sound) + +lemma AnnSeq_atomics: + "\ interfree_aux(Any, q, Some a1); interfree_aux(Any, q, Some a2)\\ + interfree_aux(Any, q, Some (AnnSeq a1 a2))" +apply(simp add: interfree_aux_def) +by force + +lemma AnnCond1_atomics: + "\ interfree_aux(Any, q, Some a1); interfree_aux(Any, q, Some a2)\\ + interfree_aux(Any, q, Some (AnnCond1 r b a1 a2))" +apply(simp add: interfree_aux_def) +by force + +lemma AnnCond2_atomics: + "interfree_aux (Any, q, Some a)\ interfree_aux(Any, q, Some (AnnCond2 r b a))" +by(simp add: interfree_aux_def) + +lemma AnnWhile_atomics: "interfree_aux (Any, q, Some a) + \ interfree_aux(Any, q, Some (AnnWhile r b i a))" +by(simp add: interfree_aux_def) + +lemma Annatom_atomics: + "\- (q \ r) a q \ interfree_aux (None, q, Some (AnnAwait r {x. True} a))" +by(simp add: interfree_aux_def oghoare_sound) + +lemma AnnAwait_atomics: + "\- (q \ (r \ b)) a q \ interfree_aux (None, q, Some (AnnAwait r b a))" +by(simp add: interfree_aux_def oghoare_sound) + +constdefs + interfree_swap :: "('a ann_triple_op * ('a ann_triple_op) list) \ bool" + "interfree_swap == \(x, xs). \y\set xs. interfree_aux (com x, post x, com y) + \ interfree_aux(com y, post y, com x)" + +lemma interfree_swap_Empty: "interfree_swap (x, [])" +by(simp add:interfree_swap_def) + +lemma interfree_swap_List: + "\ interfree_aux (com x, post x, com y); + interfree_aux (com y, post y ,com x); interfree_swap (x, xs) \ + \ interfree_swap (x, y#xs)" +by(simp add:interfree_swap_def) + +lemma interfree_swap_Map: "\k. i\k \ k interfree_aux (com x, post x, c k) + \ interfree_aux (c k, Q k, com x) + \ interfree_swap (x, map (\k. (c k, Q k)) [i..j(])" +by(force simp add: interfree_swap_def less_diff_conv) + +lemma interfree_Empty: "interfree []" +by(simp add:interfree_def) + +lemma interfree_List: + "\ interfree_swap(x, xs); interfree xs \ \ interfree (x#xs)" +apply(simp add:interfree_def interfree_swap_def) +apply clarify +apply(case_tac i) + apply(case_tac j) + apply simp_all +apply(case_tac j,simp+) +done + +lemma interfree_Map: + "(\i j. a\i \ i a\j \ j i\j \ interfree_aux (c i, Q i, c j)) + \ interfree (map (\k. (c k, Q k)) [a..b(])" +by(force simp add: interfree_def less_diff_conv) + +constdefs map_ann_hoare :: "(('a ann_com_op * 'a assn) list) \ bool " ("[\] _" [0] 45) + "[\] Ts == (\ic q. Ts!i=(Some c, q) \ \ c q)" + +lemma MapAnnEmpty: "[\] []" +by(simp add:map_ann_hoare_def) + +lemma MapAnnList: "\ \ c q ; [\] xs \ \ [\] (Some c,q)#xs" +apply(simp add:map_ann_hoare_def) +apply clarify +apply(case_tac i,simp+) +done + +lemma MapAnnMap: + "\k. i\k \ k \ (c k) (Q k) \ [\] map (\k. (Some (c k), Q k)) [i..j(]" +apply(simp add: map_ann_hoare_def less_diff_conv) +done + +lemma ParallelRule:"\ [\] Ts ; interfree Ts \ + \ \- (\i\{i. ii\{i. i \k (c k) (Q k); + \k l. k l k\l \ interfree_aux (Some(c k), Q k, Some(c l)) \ + \ \- (\i\{i. iii\{i. i(b x)}" +by fast +lemma list_length: "length []=0 \ length (x#xs) = Suc(length xs)" +by simp +lemma list_lemmas: "length []=0 \ length (x#xs) = Suc(length xs) +\ (x#xs) ! 0=x \ (x#xs) ! Suc n = xs ! n" +by simp +lemma le_Suc_eq_insert: "{i. i a1 \ a2 \ .. \ an"} returning +@{text n} subgoals, one for each conjunct: *} + +ML {* +fun conjI_Tac tac i st = st |> + ( (EVERY [rtac conjI i, + conjI_Tac tac (i+1), + tac i]) ORELSE (tac i) ) +*} + + +subsubsection {* Tactic for the generation of the verification conditions *} + +text {* The tactic basically uses two subtactics: + +\begin{description} + +\item[HoareRuleTac] is called at the level of parallel programs, it + uses the ParallelTac to solve parallel composition of programs. + This verification has two parts, namely, (1) all component programs are + correct and (2) they are interference free. @{text HoareRuleTac} is + also called at the level of atomic regions, i.e. @{text "\ \"} and + @{text "AWAIT b THEN _ END"}, and at each interference freedom test. + +\item[AnnHoareRuleTac] is for component programs which + are annotated programs and so, there are not unknown assertions + (no need to use the parameter precond, see NOTE). + + NOTE: precond(::bool) informs if the subgoal has the form @{text "\- ?p c q"}, + in this case we have precond=False and the generated verification + condition would have the form @{text "?p \ \"} which can be solved by + @{text "rtac subset_refl"}, if True we proceed to simplify it using + the simplification tactics above. + +\end{description} +*} + +ML {* + + fun WlpTac i = (rtac (thm "SeqRule") i) THEN (HoareRuleTac false (i+1)) +and HoareRuleTac precond i st = st |> + ( (WlpTac i THEN HoareRuleTac precond i) + ORELSE + (FIRST[rtac (thm "SkipRule") i, + rtac (thm "BasicRule") i, + EVERY[rtac (thm "ParallelConseqRule") i, + ParallelConseq (i+2), + ParallelTac (i+1), + ParallelConseq i], + EVERY[rtac (thm "CondRule") i, + HoareRuleTac false (i+2), + HoareRuleTac false (i+1)], + EVERY[rtac (thm "WhileRule") i, + HoareRuleTac true (i+1)], + K all_tac i ] + THEN (if precond then (K all_tac i) else (rtac (thm "subset_refl") i)))) + +and AnnWlpTac i = (rtac (thm "AnnSeq") i) THEN (AnnHoareRuleTac (i+1)) +and AnnHoareRuleTac i st = st |> + ( (AnnWlpTac i THEN AnnHoareRuleTac i ) + ORELSE + (FIRST[(rtac (thm "AnnskipRule") i), + EVERY[rtac (thm "AnnatomRule") i, + HoareRuleTac true (i+1)], + (rtac (thm "AnnwaitRule") i), + rtac (thm "AnnBasic") i, + EVERY[rtac (thm "AnnCond1") i, + AnnHoareRuleTac (i+3), + AnnHoareRuleTac (i+1)], + EVERY[rtac (thm "AnnCond2") i, + AnnHoareRuleTac (i+1)], + EVERY[rtac (thm "AnnWhile") i, + AnnHoareRuleTac (i+2)], + EVERY[rtac (thm "AnnAwait") i, + HoareRuleTac true (i+1)], + K all_tac i])) + +and ParallelTac i = EVERY[rtac (thm "ParallelRule") i, + interfree_Tac (i+1), + MapAnn_Tac i] + +and MapAnn_Tac i st = st |> + (FIRST[rtac (thm "MapAnnEmpty") i, + EVERY[rtac (thm "MapAnnList") i, + MapAnn_Tac (i+1), + AnnHoareRuleTac i], + EVERY[rtac (thm "MapAnnMap") i, + rtac (thm "allI") i,rtac (thm "impI") i, + AnnHoareRuleTac i]]) + +and interfree_swap_Tac i st = st |> + (FIRST[rtac (thm "interfree_swap_Empty") i, + EVERY[rtac (thm "interfree_swap_List") i, + interfree_swap_Tac (i+2), + interfree_aux_Tac (i+1), + interfree_aux_Tac i ], + EVERY[rtac (thm "interfree_swap_Map") i, + rtac (thm "allI") i,rtac (thm "impI") i, + conjI_Tac (interfree_aux_Tac) i]]) + +and interfree_Tac i st = st |> + (FIRST[rtac (thm "interfree_Empty") i, + EVERY[rtac (thm "interfree_List") i, + interfree_Tac (i+1), + interfree_swap_Tac i], + EVERY[rtac (thm "interfree_Map") i, + rtac (thm "allI") i,rtac (thm "allI") i,rtac (thm "impI") i, + interfree_aux_Tac i ]]) + +and interfree_aux_Tac i = (before_interfree_simp_tac i ) THEN + (FIRST[rtac (thm "interfree_aux_rule1") i, + dest_assertions_Tac i]) + +and dest_assertions_Tac i st = st |> + (FIRST[EVERY[rtac (thm "AnnBasic_assertions") i, + dest_atomics_Tac (i+1), + dest_atomics_Tac i], + EVERY[rtac (thm "AnnSeq_assertions") i, + dest_assertions_Tac (i+1), + dest_assertions_Tac i], + EVERY[rtac (thm "AnnCond1_assertions") i, + dest_assertions_Tac (i+2), + dest_assertions_Tac (i+1), + dest_atomics_Tac i], + EVERY[rtac (thm "AnnCond2_assertions") i, + dest_assertions_Tac (i+1), + dest_atomics_Tac i], + EVERY[rtac (thm "AnnWhile_assertions") i, + dest_assertions_Tac (i+2), + dest_atomics_Tac (i+1), + dest_atomics_Tac i], + EVERY[rtac (thm "AnnAwait_assertions") i, + dest_atomics_Tac (i+1), + dest_atomics_Tac i], + dest_atomics_Tac i]) + +and dest_atomics_Tac i st = st |> + (FIRST[EVERY[rtac (thm "AnnBasic_atomics") i, + HoareRuleTac true i], + EVERY[rtac (thm "AnnSeq_atomics") i, + dest_atomics_Tac (i+1), + dest_atomics_Tac i], + EVERY[rtac (thm "AnnCond1_atomics") i, + dest_atomics_Tac (i+1), + dest_atomics_Tac i], + EVERY[rtac (thm "AnnCond2_atomics") i, + dest_atomics_Tac i], + EVERY[rtac (thm "AnnWhile_atomics") i, + dest_atomics_Tac i], + EVERY[rtac (thm "Annatom_atomics") i, + HoareRuleTac true i], + EVERY[rtac (thm "AnnAwait_atomics") i, + HoareRuleTac true i], + K all_tac i]) +*} + + +text {* The final tactic is given the name @{text oghoare}: *} + +ML {* +fun oghoare_tac i thm = SUBGOAL (fn (term, _) => + (HoareRuleTac true i)) i thm +*} + +text {* Notice that the tactic for parallel programs @{text +"oghoare_tac"} is initially invoked with the value @{text true} for +the parameter @{text precond}. + +Parts of the tactic can be also individually used to generate the +verification conditions for annotated sequential programs and to +generate verification conditions out of interference freedom tests: *} + +ML {* fun annhoare_tac i thm = SUBGOAL (fn (term, _) => + (AnnHoareRuleTac i)) i thm + +fun interfree_aux_tac i thm = SUBGOAL (fn (term, _) => + (interfree_aux_Tac i)) i thm +*} + +text {* The so defined ML tactics are then ``exported'' to be used in +Isabelle proofs. *} + +method_setup oghoare = {* + Method.no_args + (Method.SIMPLE_METHOD' HEADGOAL (oghoare_tac)) *} + "verification condition generator for the oghoare logic" + +method_setup annhoare = {* + Method.no_args + (Method.SIMPLE_METHOD' HEADGOAL (annhoare_tac)) *} + "verification condition generator for the ann_hoare logic" + +method_setup interfree_aux = {* + Method.no_args + (Method.SIMPLE_METHOD' HEADGOAL (interfree_aux_tac)) *} + "verification condition generator for interference freedom tests" + +text {* Tactics useful for dealing with the generated verification conditions: *} + +method_setup conjI_tac = {* + Method.no_args + (Method.SIMPLE_METHOD' HEADGOAL (conjI_Tac (K all_tac))) *} + "verification condition generator for interference freedom tests" + +ML {* +fun disjE_Tac tac i st = st |> + ( (EVERY [etac disjE i, + disjE_Tac tac (i+1), + tac i]) ORELSE (tac i) ) +*} + +method_setup disjE_tac = {* + Method.no_args + (Method.SIMPLE_METHOD' HEADGOAL (disjE_Tac (K all_tac))) *} + "verification condition generator for interference freedom tests" + +end \ No newline at end of file diff -r 98f0a09a33c3 -r 791e3b4c4039 src/HOL/HoareParallel/OG_Tran.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HoareParallel/OG_Tran.thy Tue Mar 05 17:11:25 2002 +0100 @@ -0,0 +1,308 @@ + +header {* \section{Operational Semantics} *} + +theory OG_Tran = OG_Com: + +types + 'a ann_com_op = "('a ann_com) option" + 'a ann_triple_op = "('a ann_com_op \ 'a assn)" + +consts com :: "'a ann_triple_op \ 'a ann_com_op" +primrec "com (c, q) = c" + +consts post :: "'a ann_triple_op \ 'a assn" +primrec "post (c, q) = q" + +constdefs + All_None :: "'a ann_triple_op list \ bool" + "All_None Ts \ \(c, q) \ set Ts. c = None" + +subsection {* The Transition Relation *} + +consts + ann_transition :: "(('a ann_com_op \ 'a) \ ('a ann_com_op \ 'a)) set" + transition :: "(('a com \ 'a) \ ('a com \ 'a)) set" + +syntax + "_ann_transition" :: "('a ann_com_op \ 'a) \ ('a ann_com_op \ 'a) \ bool" + ("_ -1\ _"[81,81] 100) + "_ann_transition_n" :: "('a ann_com_op \ 'a) \ nat \ ('a ann_com_op \ 'a) + \ bool" ("_ -_\ _"[81,81] 100) + "_ann_transition_*" :: "('a ann_com_op \ 'a) \ ('a ann_com_op \ 'a) \ bool" + ("_ -*\ _"[81,81] 100) + + "_transition" :: "('a com \ 'a) \ ('a com \ 'a) \ bool" ("_ -P1\ _"[81,81] 100) + "_transition_n" :: "('a com \ 'a) \ nat \ ('a com \ 'a) \ bool" + ("_ -P_\ _"[81,81,81] 100) + "_transition_*" :: "('a com \ 'a) \ ('a com \ 'a) \ bool" ("_ -P*\ _"[81,81] 100) + +text {* The corresponding syntax translations are: *} + +translations + "con_0 -1\ con_1" \ "(con_0, con_1) \ ann_transition" + "con_0 -n\ con_1" \ "(con_0, con_1) \ ann_transition^n" + "con_0 -*\ con_1" \ "(con_0, con_1) \ ann_transition\<^sup>*" + + "con_0 -P1\ con_1" \ "(con_0, con_1) \ transition" + "con_0 -Pn\ con_1" \ "(con_0, con_1) \ transition^n" + "con_0 -P*\ con_1" \ "(con_0, con_1) \ transition\<^sup>*" + +inductive ann_transition transition +intros + AnnBasic: "(Some (AnnBasic r f), s) -1\ (None, f s)" + + AnnSeq1: "(Some c0, s) -1\ (None, t) \ + (Some (AnnSeq c0 c1), s) -1\ (Some c1, t)" + AnnSeq2: "(Some c0, s) -1\ (Some c2, t) \ + (Some (AnnSeq c0 c1), s) -1\ (Some (AnnSeq c2 c1), t)" + + AnnCond1T: "s \ b \ (Some (AnnCond1 r b c1 c2), s) -1\ (Some c1, s)" + AnnCond1F: "s \ b \ (Some (AnnCond1 r b c1 c2), s) -1\ (Some c2, s)" + + AnnCond2T: "s \ b \ (Some (AnnCond2 r b c), s) -1\ (Some c, s)" + AnnCond2F: "s \ b \ (Some (AnnCond2 r b c), s) -1\ (None, s)" + + AnnWhileF: "s \ b \ (Some (AnnWhile r b i c), s) -1\ (None, s)" + AnnWhileT: "s \ b \ (Some (AnnWhile r b i c), s) -1\ + (Some (AnnSeq c (AnnWhile i b i c)), s)" + + AnnAwait: "\ s \ b; atom_com c; (c, s) -P*\ (Parallel [], t) \ \ + (Some (AnnAwait r b c), s) -1\ (None, t)" + + Parallel: "\ i (r, t) \ + \ (Parallel Ts, s) -P1\ (Parallel (Ts [i:=(r, q)]), t)" + + Basic: "(Basic f, s) -P1\ (Parallel [], f s)" + + Seq1: "All_None Ts \ (Seq (Parallel Ts) c, s) -P1\ (c, s)" + Seq2: "(c0, s) -P1\ (c2, t) \ (Seq c0 c1, s) -P1\ (Seq c2 c1, t)" + + CondT: "s \ b \ (Cond b c1 c2, s) -P1\ (c1, s)" + CondF: "s \ b \ (Cond b c1 c2, s) -P1\ (c2, s)" + + WhileF: "s \ b \ (While b i c, s) -P1\ (Parallel [], s)" + WhileT: "s \ b \ (While b i c, s) -P1\ (Seq c (While b i c), s)" + +monos "rtrancl_mono" + +subsection {* Definition of Semantics *} + +constdefs + ann_sem :: "'a ann_com \ 'a \ 'a set" + "ann_sem c \ \s. {t. (Some c, s) -*\ (None, t)}" + + ann_SEM :: "'a ann_com \ 'a set \ 'a set" + "ann_SEM c S \ \ann_sem c ` S" + + sem :: "'a com \ 'a \ 'a set" + "sem c \ \s. {t. \Ts. (c, s) -P*\ (Parallel Ts, t) \ All_None Ts}" + + SEM :: "'a com \ 'a set \ 'a set" + "SEM c S \ \sem c ` S " + +syntax "_Omega" :: "'a com" ("\" 63) +translations "\" \ "While UNIV UNIV (Basic id)" + +consts fwhile :: "'a bexp \ 'a com \ nat \ 'a com" +primrec + "fwhile b c 0 = \" + "fwhile b c (Suc n) = Cond b (Seq c (fwhile b c n)) (Basic id)" + +subsubsection {* Proofs *} + +declare ann_transition_transition.intros [intro] +inductive_cases transition_cases: + "(Parallel T,s) -P1\ t" + "(Basic f, s) -P1\ t" + "(Seq c1 c2, s) -P1\ t" + "(Cond b c1 c2, s) -P1\ t" + "(While b i c, s) -P1\ t" + +lemma Parallel_empty_lemma [rule_format (no_asm)]: + "(Parallel [],s) -Pn\ (Parallel Ts,t) \ Ts=[] \ n=0 \ s=t" +apply(induct n) + apply(simp (no_asm)) +apply clarify +apply(drule rel_pow_Suc_D2) +apply(force elim:transition_cases) +done + +lemma Parallel_AllNone_lemma [rule_format (no_asm)]: + "All_None Ss \ (Parallel Ss,s) -Pn\ (Parallel Ts,t) \ Ts=Ss \ n=0 \ s=t" +apply(induct "n") + apply(simp (no_asm)) +apply clarify +apply(drule rel_pow_Suc_D2) +apply clarify +apply(erule transition_cases,simp_all) +apply(force dest:nth_mem simp add:All_None_def) +done + +lemma Parallel_AllNone: "All_None Ts \ (SEM (Parallel Ts) X) = X" +apply (unfold SEM_def sem_def) +apply auto +apply(drule rtrancl_imp_UN_rel_pow) +apply clarify +apply(drule Parallel_AllNone_lemma) +apply auto +done + +lemma Parallel_empty: "Ts=[] \ (SEM (Parallel Ts) X) = X" +apply(rule Parallel_AllNone) +apply(simp add:All_None_def) +done + +text {* Set of lemmas from Apt and Olderog "Verification of sequential +and concurrent programs", page 63. *} + +lemma L3_5i: "X\Y \ SEM c X \ SEM c Y" +apply (unfold SEM_def) +apply force +done + +lemma L3_5ii_lemma1: + "\ (c1, s1) -P*\ (Parallel Ts, s2); All_None Ts; + (c2, s2) -P*\ (Parallel Ss, s3); All_None Ss \ + \ (Seq c1 c2, s1) -P*\ (Parallel Ss, s3)" +apply(erule converse_rtrancl_induct2) +apply(force intro:converse_rtrancl_into_rtrancl)+ +done + +lemma L3_5ii_lemma2 [rule_format (no_asm)]: + "\c1 c2 s t. (Seq c1 c2, s) -Pn\ (Parallel Ts, t) \ + (All_None Ts) \ (\y m Rs. (c1,s) -P*\ (Parallel Rs, y) \ + (All_None Rs) \ (c2, y) -Pm\ (Parallel Ts, t) \ m \ n)" +apply(induct "n") + apply(force) +apply(safe dest!: rel_pow_Suc_D2) +apply(erule transition_cases,simp_all) + apply (fast intro!: le_SucI) +apply (fast intro!: le_SucI elim!: rel_pow_imp_rtrancl converse_rtrancl_into_rtrancl) +done + +lemma L3_5ii_lemma3: + "\(Seq c1 c2,s) -P*\ (Parallel Ts,t); All_None Ts\ \ + (\y Rs. (c1,s) -P*\ (Parallel Rs,y) \ All_None Rs + \ (c2,y) -P*\ (Parallel Ts,t))" +apply(drule rtrancl_imp_UN_rel_pow) +apply(fast dest: L3_5ii_lemma2 rel_pow_imp_rtrancl) +done + +lemma L3_5ii: "SEM (Seq c1 c2) X = SEM c2 (SEM c1 X)" +apply (unfold SEM_def sem_def) +apply auto + apply(fast dest: L3_5ii_lemma3) +apply(fast elim: L3_5ii_lemma1) +done + +lemma L3_5iii: "SEM (Seq (Seq c1 c2) c3) X = SEM (Seq c1 (Seq c2 c3)) X" +apply (simp (no_asm) add: L3_5ii) +done + +lemma L3_5iv: + "SEM (Cond b c1 c2) X = (SEM c1 (X \ b)) Un (SEM c2 (X \ (-b)))" +apply (unfold SEM_def sem_def) +apply auto +apply(erule converse_rtranclE) + prefer 2 + apply (erule transition_cases,simp_all) + apply(fast intro: converse_rtrancl_into_rtrancl elim: transition_cases)+ +done + + +lemma L3_5v_lemma1[rule_format]: + "(S,s) -Pn\ (T,t) \ S=\ \ (\(\Rs. T=(Parallel Rs) \ All_None Rs))" +apply (unfold UNIV_def) +apply(rule nat_less_induct) +apply safe +apply(erule rel_pow_E2) + apply simp_all +apply(erule transition_cases) + apply simp_all +apply(erule rel_pow_E2) + apply(simp add: Id_def) +apply(erule transition_cases,simp_all) +apply clarify +apply(erule transition_cases,simp_all) +apply(erule rel_pow_E2,simp) +apply clarify +apply(erule transition_cases) + apply simp+ + apply clarify + apply(erule transition_cases) +apply simp_all +done + +lemma L3_5v_lemma2: "\(\, s) -P*\ (Parallel Ts, t); All_None Ts \ \ False" +apply(fast dest: rtrancl_imp_UN_rel_pow L3_5v_lemma1) +done + +lemma L3_5v_lemma3: "SEM (\) S = {}" +apply (unfold SEM_def sem_def) +apply(fast dest: L3_5v_lemma2) +done + +lemma L3_5v_lemma4 [rule_format]: + "\s. (While b i c, s) -Pn\ (Parallel Ts, t) \ All_None Ts \ + (\k. (fwhile b c k, s) -P*\ (Parallel Ts, t))" +apply(rule nat_less_induct) +apply safe +apply(erule rel_pow_E2) + apply safe +apply(erule transition_cases,simp_all) + apply (rule_tac x = "1" in exI) + apply(force dest: Parallel_empty_lemma intro: converse_rtrancl_into_rtrancl simp add: Id_def) +apply safe +apply(drule L3_5ii_lemma2) + apply safe +apply(drule le_imp_less_Suc) +apply (erule allE , erule impE,assumption) +apply (erule allE , erule impE, assumption) +apply safe +apply (rule_tac x = "k+1" in exI) +apply(simp (no_asm)) +apply(rule converse_rtrancl_into_rtrancl) + apply fast +apply(fast elim: L3_5ii_lemma1) +done + +lemma L3_5v_lemma5 [rule_format]: + "\s. (fwhile b c k, s) -P*\ (Parallel Ts, t) \ All_None Ts \ + (While b i c, s) -P*\ (Parallel Ts,t)" +apply(induct "k") + apply(force dest: L3_5v_lemma2) +apply safe +apply(erule converse_rtranclE) + apply simp_all +apply(erule transition_cases,simp_all) + apply(rule converse_rtrancl_into_rtrancl) + apply(fast) + apply(fast elim!: L3_5ii_lemma1 dest: L3_5ii_lemma3) +apply(drule rtrancl_imp_UN_rel_pow) +apply clarify +apply(erule rel_pow_E2) + apply simp_all +apply(erule transition_cases,simp_all) +apply(fast dest: Parallel_empty_lemma) +done + +lemma L3_5v: "SEM (While b i c) = (\x. (\k. SEM (fwhile b c k) x))" +apply(rule ext) +apply (simp add: SEM_def sem_def) +apply safe + apply(drule rtrancl_imp_UN_rel_pow,simp) + apply clarify + apply(fast dest:L3_5v_lemma4) +apply(fast intro: L3_5v_lemma5) +done + +section {* Validity of Correctness Formulas *} + +constdefs + com_validity :: "'a assn \ 'a com \ 'a assn \ bool" ("(3\= _// _//_)" [90,55,90] 50) + "\= p c q \ SEM c p \ q" + + ann_com_validity :: "'a ann_com \ 'a assn \ bool" ("\ _ _" [60,90] 45) + "\ c q \ ann_SEM c (pre c) \ q" + +end \ No newline at end of file diff -r 98f0a09a33c3 -r 791e3b4c4039 src/HOL/HoareParallel/Quote_Antiquote.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HoareParallel/Quote_Antiquote.thy Tue Mar 05 17:11:25 2002 +0100 @@ -0,0 +1,24 @@ + +header {* \section{Concrete Syntax} *} + +theory Quote_Antiquote = Main: + +syntax + "_quote" :: "'b \ ('a \ 'b)" ("(\_\)" [0] 1000) + "_antiquote" :: "('a \ 'b) \ 'b" ("\_" [1000] 1000) + "_Assert" :: "'a \ 'a set" ("(.{_}.)" [0] 1000) + +syntax (xsymbols) + "_Assert" :: "'a \ 'a set" ("(\_\)" [0] 1000) + +translations + ".{b}." \ "Collect \b\" + +parse_translation {* + let + fun quote_tr [t] = Syntax.quote_tr "_antiquote" t + | quote_tr ts = raise TERM ("quote_tr", ts); + in [("_quote", quote_tr)] end +*} + +end \ No newline at end of file diff -r 98f0a09a33c3 -r 791e3b4c4039 src/HOL/HoareParallel/RG_Com.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HoareParallel/RG_Com.thy Tue Mar 05 17:11:25 2002 +0100 @@ -0,0 +1,25 @@ + +header {* \chapter{The Rely-Guarantee Method} + +\section {Abstract Syntax} +*} + +theory RG_Com = Main: + +text {* Semantics of assertions and boolean expressions (bexp) as sets +of states. Syntax of commands @{text com} and parallel commands +@{text par_com}. *} + +types + 'a bexp = "'a set" + +datatype 'a com = + Basic "'a \'a" + | Seq "'a com" "'a com" + | Cond "'a bexp" "'a com" "'a com" + | While "'a bexp" "'a com" + | Await "'a bexp" "'a com" + +types 'a par_com = "(('a com) option) list" + +end \ No newline at end of file diff -r 98f0a09a33c3 -r 791e3b4c4039 src/HOL/HoareParallel/RG_Examples.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HoareParallel/RG_Examples.thy Tue Mar 05 17:11:25 2002 +0100 @@ -0,0 +1,408 @@ + +header {* \section{Examples} *} + +theory RG_Examples = RG_Syntax: + +lemmas definitions [simp]= stable_def Pre_def Rely_def Guar_def Post_def Com_def + +subsection {* Set Elements of an Array to Zero *} + +lemma le_less_trans2: "\(j::nat) j\ \ i (a::nat) < c; b\d \ \ a + b < c + d" +by simp + +record Example1 = + A :: "nat list" + +lemma Example1: + "\ COBEGIN + SCHEME [0 \ i < n] + (\A := \A [i := 0], + \ n < length \A \, + \ length \A = length \A \ \A ! i = \A ! i \, + \ length \A = length \A \ (\j j \ \A ! j = \A ! j) \, + \ \A ! i = 0 \) + COEND + SAT [\ n < length \A \, \ \A = \A \, \ True \, \ \i < n. \A ! i = 0 \]" +apply(rule Parallel) + apply simp + apply clarify + apply simp + apply(erule disjE) + apply simp + apply clarify + apply simp + apply auto +apply(rule Basic) +apply auto +done + +lemma Example1_parameterized: +"k < t \ + \ COBEGIN + SCHEME [k*n\i<(Suc k)*n] (\A:=\A[i:=0], + \t*n < length \A\, + \t*n < length \A \ length \A=length \A \ \A!i = \A!i\, + \t*n < length \A \ length \A=length \A \ (\jA . i\j \ \A!j = \A!j)\, + \\A!i=0\) + COEND + SAT [\t*n < length \A\, + \t*n < length \A \ length \A=length \A \ (\iA!(k*n+i)=\A!(k*n+i))\, + \t*n < length \A \ length \A=length \A \ + (\iA . (i \A!i = \A!i) \ ((Suc k)*n \ i\ \A!i = \A!i))\, + \\iA!(k*n+i) = 0\]" +apply(rule Parallel) + apply simp + apply clarify + apply simp + apply(erule disjE) + apply clarify + apply simp + apply clarify + apply simp + apply clarify + apply simp + apply(erule_tac x="k*n +i" in allE) + apply(subgoal_tac "k*n+i COBEGIN + (\ \x:=\x+1;; \c_0:=\c_0 + 1 \, + \\x=\c_0 + \c_1 \ \c_0=0\, + \\c_0 = \c_0 \ + (\x=\c_0 + \c_1 + \ \x = \c_0 + \c_1)\, + \\c_1 = \c_1 \ + (\x=\c_0 + \c_1 + \ \x =\c_0 + \c_1)\, + \\x=\c_0 + \c_1 \ \c_0=1 \) + \ + (\ \x:=\x+1;; \c_1:=\c_1+1 \, + \\x=\c_0 + \c_1 \ \c_1=0 \, + \\c_1 = \c_1 \ + (\x=\c_0 + \c_1 + \ \x = \c_0 + \c_1)\, + \\c_0 = \c_0 \ + (\x=\c_0 + \c_1 + \ \x =\c_0 + \c_1)\, + \\x=\c_0 + \c_1 \ \c_1=1\) + COEND + SAT [\\x=0 \ \c_0=0 \ \c_1=0\, + \\x=\x \ \c_0= \c_0 \ \c_1=\c_1\, + \True\, + \\x=2\]" +apply(rule Parallel) + apply simp_all + apply clarify + apply(case_tac i) + apply simp + apply(erule disjE) + apply clarify + apply simp + apply clarify + apply simp + apply(case_tac j,simp) + apply simp + apply simp + apply(erule disjE) + apply clarify + apply simp + apply clarify + apply simp + apply(case_tac j,simp,simp) + apply clarify + apply(case_tac i,simp,simp) + apply clarify + apply simp + apply(erule_tac x=0 in all_dupE) + apply(erule_tac x=1 in allE,simp) +apply clarify +apply(case_tac i,simp) + apply(rule Await) + apply simp_all + apply(clarify) + apply(rule Seq) + prefer 2 + apply(rule Basic) + apply simp_all + apply(rule subset_refl) + apply(rule Basic) + apply simp_all + apply clarify + apply simp +apply(rule Await) + apply simp_all +apply(clarify) +apply(rule Seq) + prefer 2 + apply(rule Basic) + apply simp_all + apply(rule subset_refl) +apply(rule Basic) +apply simp_all +apply clarify +apply simp +done + +subsubsection {* Parameterized *} + +lemma Example2_lemma1: "j (\i b j = 0 " +apply(induct n) + apply simp_all +apply(force simp add: less_Suc_eq) +done + +lemma Example2_lemma2_aux: + "j (\iii s \ (\iij \ Suc (\i< n. b i)=(\i< n. (b (j:=1)) i)" +apply(frule_tac b="(b (j:=1))" in Example2_lemma2_aux) +apply(erule_tac t="Summation (b(j := 1)) n" in ssubst) +apply(frule_tac b=b in Example2_lemma2_aux) +apply(erule_tac t="Summation b n" in ssubst) +apply(subgoal_tac "Suc (Summation b j + b j + (\iij") + apply(drule_tac b="b" and t=1 in Example2_lemma2_aux2) + apply(rotate_tac -1) + apply(erule ssubst) +apply simp_all +done + +lemma Example2_lemma2_Suc0: "\j \ Suc (\i< n. b i)=(\i< n. (b (j:=Suc 0)) i)" +by(simp add:Example2_lemma2) + +lemma Example2_lemma3: "\i< n. b i = 1 \ (\i nat" + y :: nat + +lemma Example2_parameterized: "0 + \ COBEGIN SCHEME [0\i \y:=\y+1;; \C:=\C (i:=1) \, + \\y=(\iC i) \ \C i=0\, + \\C i = \C i \ + (\y=(\iC i) \ \y =(\iC i))\, + \(\jj \ \C j = \C j) \ + (\y=(\iC i) \ \y =(\iC i))\, + \\y=(\iC i) \ \C i=1\) + COEND + SAT [\\y=0 \ (\iC i)=0 \, \\C=\C \ \y=\y\, \True\, \\y=n\]" +apply(rule Parallel) +apply force +apply force +apply(force elim:Example2_lemma1) +apply clarify +apply simp +apply(force intro:Example2_lemma3) +apply clarify +apply simp +apply(rule Await) +apply simp_all +apply clarify +apply(rule Seq) +prefer 2 +apply(rule Basic) +apply(rule subset_refl) +apply simp+ +apply(rule Basic) +apply simp +apply clarify +apply simp +apply(force elim:Example2_lemma2_Suc0) +apply simp+ +done + +subsection {* Find Least Element *} + +text {* A previous lemma: *} + +lemma mod_aux :"\i < (n::nat); a mod n = i; j < a + n; j mod n = i; a < j\ \ False" +apply(subgoal_tac "a=a div n*n + a mod n" ) + prefer 2 apply (simp (no_asm_use) only: mod_div_equality [symmetric]) +apply(subgoal_tac "j=j div n*n + j mod n") + prefer 2 apply (simp (no_asm_use) only: mod_div_equality [symmetric]) +apply simp +apply(subgoal_tac "a div n*n < j div n*n") +prefer 2 apply arith +apply(subgoal_tac "j div n*n < (a div n + 1)*n") +prefer 2 apply simp +apply (simp only:mult_less_cancel2) +apply arith +done + +record Example3 = + X :: "nat \ nat" + Y :: "nat \ nat" + +lemma Example3: "m mod n=0 \ + \ COBEGIN + SCHEME [0\ijX i < \Y j) DO + IF P(B!(\X i)) THEN \Y:=\Y (i:=\X i) + ELSE \X:= \X (i:=(\X i)+ n) FI + OD, + \(\X i) mod n=i \ (\j<\X i. j mod n=i \ \P(B!j)) \ (\Y i P(B!(\Y i)) \ \Y i\ m+i)\, + \(\jj \ \Y j \ \Y j) \ \X i = \X i \ + \Y i = \Y i\, + \(\jj \ \X j = \X j \ \Y j = \Y j) \ + \Y i \ \Y i\, + \(\X i) mod n=i \ (\j<\X i. j mod n=i \ \P(B!j)) \ (\Y i P(B!(\Y i)) \ \Y i\ m+i) \ (\jY j \ \X i) \) + COEND + SAT [\ \iX i=i \ \Y i=m+i \,\\X=\X \ \Y=\Y\,\True\, + \\iX i) mod n=i \ (\j<\X i. j mod n=i \ \P(B!j)) \ + (\Y i P(B!(\Y i)) \ \Y i\ m+i) \ (\jY j \ \X i)\]" +apply(rule Parallel) +(*5*) +apply force+ +apply clarify +apply simp +apply(rule While) + apply force + apply force + apply force + apply(rule_tac "pre'"="\ \X i mod n = i \ (\j. j<\X i \ j mod n = i \ \P(B!j)) \ (\Y i < n * q \ P (B!(\Y i))) \ \X i<\Y i\" in Conseq) + apply force + apply(rule subset_refl)+ + apply(rule Cond) + apply force + apply(rule Basic) + apply force + apply force + apply force + apply force + apply(rule Basic) + apply simp + apply clarify + apply simp + apply(case_tac "X x (j mod n)\ j") + apply(drule le_imp_less_or_eq) + apply(erule disjE) + apply(drule_tac j=j and n=n and i="j mod n" and a="X x (j mod n)" in mod_aux) + apply assumption+ + apply simp+ + apply(erule_tac x=j in allE) + apply force + apply simp + apply clarify + apply(rule conjI) + apply clarify + apply simp + apply(erule not_sym) + apply force +apply force+ +done + +text {* Same but with a list as auxiliary variable: *} + +record Example3_list = + X :: "nat list" + Y :: "nat list" + +lemma Example3_list: "m mod n=0 \ \ (COBEGIN SCHEME [0\ijX!i < \Y!j) DO + IF P(B!(\X!i)) THEN \Y:=\Y[i:=\X!i] ELSE \X:= \X[i:=(\X!i)+ n] FI + OD, + \nX \ nY \ (\X!i) mod n=i \ (\j<\X!i. j mod n=i \ \P(B!j)) \ (\Y!i P(B!(\Y!i)) \ \Y!i\ m+i)\, + \(\jj \ \Y!j \ \Y!j) \ \X!i = \X!i \ + \Y!i = \Y!i \ length \X = length \X \ length \Y = length \Y\, + \(\jj \ \X!j = \X!j \ \Y!j = \Y!j) \ + \Y!i \ \Y!i \ length \X = length \X \ length \Y = length \Y\, + \(\X!i) mod n=i \ (\j<\X!i. j mod n=i \ \P(B!j)) \ (\Y!i P(B!(\Y!i)) \ \Y!i\ m+i) \ (\jY!j \ \X!i) \) COEND) + SAT [\nX \ nY \ (\iX!i=i \ \Y!i=m+i) \, + \\X=\X \ \Y=\Y\, + \True\, + \\iX!i) mod n=i \ (\j<\X!i. j mod n=i \ \P(B!j)) \ + (\Y!i P(B!(\Y!i)) \ \Y!i\ m+i) \ (\jY!j \ \X!i)\]" +apply(rule Parallel) +(*5*) +apply force+ +apply clarify +apply simp +apply(rule While) + apply force + apply force + apply force + apply(rule_tac "pre'"="\nX \ nY \ \X ! i mod n = i \ (\j. j < \X ! i \ j mod n = i \ \ P (B ! j)) \ (\Y ! i < n * q \ P (B ! (\Y ! i))) \ \X!i<\Y!i\" in Conseq) + apply force + apply(rule subset_refl)+ + apply(rule Cond) + apply force + apply(rule Basic) + apply force + apply force + apply force + apply force + apply(rule Basic) + apply simp + apply clarify + apply simp + apply(rule allI) + apply(rule impI)+ + apply(case_tac "X x ! i\ j") + apply(drule le_imp_less_or_eq) + apply(erule disjE) + apply(drule_tac j=j and n=n and i=i and a="X x ! i" in mod_aux) + apply assumption+ + apply simp +apply force+ +done + +end \ No newline at end of file diff -r 98f0a09a33c3 -r 791e3b4c4039 src/HOL/HoareParallel/RG_Hoare.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HoareParallel/RG_Hoare.thy Tue Mar 05 17:11:25 2002 +0100 @@ -0,0 +1,1495 @@ + +header {* \section{The Proof System} *} + +theory RG_Hoare = RG_Tran: + +subsection {* Proof System for Component Programs *} + +constdefs + stable :: "'a set \ ('a \ 'a) set \ bool" + "stable \ \f g. (\x y. x \ f \ (x, y) \ g \ y \ f)" + +consts rghoare :: "('a rgformula) set" +syntax + "_rghoare" :: "['a com, 'a set, ('a \ 'a) set, ('a \ 'a) set, 'a set] \ bool" + ("\ _ sat [_, _, _, _]" [60,0,0,0,0] 45) +translations + "\ P sat [pre, rely, guar, post]" \ "(P, pre, rely, guar, post) \ rghoare" + +inductive rghoare +intros + Basic: "\ pre \ {s. f s \ post}; {(s,t). s \ pre \ t=f s} \ guar; + stable pre rely; stable post rely \ + \ \ Basic f sat [pre, rely, guar, post]" + + Seq: "\ \ P sat [pre, rely, guar, mid]; \ Q sat [mid, rely, guar, post] \ + \ \ Seq P Q sat [pre, rely, guar, post]" + + Cond: "\ stable pre rely; \ P1 sat [pre \ b, rely, guar, post]; + \ P2 sat [pre \ -b, rely, guar, post]; \s. (s,s)\guar \ + \ \ Cond b P1 P2 sat [pre, rely, guar, post]" + + While: "\ stable pre rely; (pre \ -b) \ post; stable post rely; + \ P sat [pre \ b, rely, guar, pre]; \s. (s,s)\guar \ + \ \ While b P sat [pre, rely, guar, post]" + + Await: "\ stable pre rely; stable post rely; + \V. \ P sat [pre \ b \ {V}, {(s, t). s = t}, UNIV, {s. (V, s) \ guar} \ post] \ + \ \ Await b P sat [pre, rely, guar, post]" + + Conseq: "\ pre \ pre'; rely \ rely'; guar' \ guar; post' \ post; + \ P sat [pre', rely', guar', post'] \ + \ \ P sat [pre, rely, guar, post]" + +constdefs + Pre :: "'a rgformula \ 'a set" + "Pre x \ fst(snd x)" + Post :: "'a rgformula \ 'a set" + "Post x \ snd(snd(snd(snd x)))" + Rely :: "'a rgformula \ ('a \ 'a) set" + "Rely x \ fst(snd(snd x))" + Guar :: "'a rgformula \ ('a \ 'a) set" + "Guar x \ fst(snd(snd(snd x)))" + Com :: "'a rgformula \ 'a com" + "Com x \ fst x" + +subsection {* Proof System for Parallel Programs *} + +types 'a par_rgformula = "('a rgformula) list \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set" + +consts par_rghoare :: "('a par_rgformula) set" +syntax + "_par_rghoare" :: "('a rgformula) list \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set + \ bool" ("\ _ SAT [_, _, _, _]" [60,0,0,0,0] 45) +translations + "\ Ps SAT [pre, rely, guar, post]" \ "(Ps, pre, rely, guar, post) \ par_rghoare" + +inductive par_rghoare +intros + Parallel: + "\ \i (\j\{j. j j\i}. Guar(xs!j)) \ Rely(xs!i); + (\j\{j. j guar; + pre \ (\i\{i. ii\{i. i post; + \i Com(xs!i) sat [Pre(xs!i),Rely(xs!i),Guar(xs!i),Post(xs!i)] \ + \ \ xs SAT [pre, rely, guar, post]" + +section {* Soundness*} + +subsubsection {* Some previous lemmas *} + +lemma tl_of_assum_in_assum: + "(P, s) # (P, t) # xs \ assum (pre, rely) \ stable pre rely + \ (P, t) # xs \ assum (pre, rely)" +apply(simp add:assum_def) +apply clarify +apply(rule conjI) + apply(erule_tac x=0 in allE) + apply(simp (no_asm_use)only:stable_def) + apply(erule allE,erule allE,erule impE,assumption,erule mp) + apply(simp add:Env) +apply clarify +apply(erule_tac x="Suc i" in allE) +apply simp +done + +lemma etran_in_comm: + "(P, t) # xs \ comm(guar, post) \ (P, s) # (P, t) # xs \ comm(guar, post)" +apply(simp add:comm_def) +apply clarify +apply(case_tac i,simp+) +done + +lemma ctran_in_comm: + "\(s, s) \ guar; (Q, s) # xs \ comm(guar, post)\ + \ (P, s) # (Q, s) # xs \ comm(guar, post)" +apply(simp add:comm_def) +apply clarify +apply(case_tac i,simp+) +done + +lemma takecptn_is_cptn [rule_format, elim!]: + "\j. c \ cptn \ take (Suc j) c \ cptn" +apply(induct "c") + apply(force elim: cptn.elims) +apply clarify +apply(case_tac j) + apply simp + apply(rule CptnOne) +apply simp +apply(force intro:cptn.intros elim:cptn.elims) +done + +lemma dropcptn_is_cptn [rule_format,elim!]: + "\j cptn \ drop j c \ cptn" +apply(induct "c") + apply(force elim: cptn.elims) +apply clarify +apply(case_tac j,simp+) +apply(erule cptn.elims) + apply simp + apply force +apply force +done + +lemma takepar_cptn_is_par_cptn [rule_format,elim]: + "\j. c \ par_cptn \ take (Suc j) c \ par_cptn" +apply(induct "c") + apply(force elim: cptn.elims) +apply clarify +apply(case_tac j,simp) + apply(rule ParCptnOne) +apply(force intro:par_cptn.intros elim:par_cptn.elims) +done + +lemma droppar_cptn_is_par_cptn [rule_format]: + "\j par_cptn \ drop j c \ par_cptn" +apply(induct "c") + apply(force elim: par_cptn.elims) +apply clarify +apply(case_tac j,simp+) +apply(erule par_cptn.elims) + apply simp + apply force +apply force +done + +lemma tl_of_cptn_is_cptn: "\x # xs \ cptn; xs \ []\ \ xs \ cptn" +apply(subgoal_tac "1 < length (x # xs)") + apply(drule dropcptn_is_cptn,simp+) +done + +lemma not_ctran_None [rule_format]: + "\s. (None, s)#xs \ cptn \ (\i xs!i)" +apply(induct xs,simp+) +apply clarify +apply(erule cptn.elims,simp) + apply simp + apply(case_tac i,simp) + apply(rule Env) + apply simp +apply(force elim:ctran.elims) +done + +lemma cptn_not_empty [simp]:"[] \ cptn" +apply(force elim:cptn.elims) +done + +lemma etran_or_ctran [rule_format]: + "\m i. x\cptn \ m \ length x + \ (\i. Suc i < m \ \ x!i -c\ x!Suc i) \ Suc i < m \ x!i -e\ x!Suc i" +apply(induct x,simp) +apply clarify +apply(erule cptn.elims,simp) + apply(case_tac i,simp) + apply(rule Env) + apply simp + apply(erule_tac x="m - 1" in allE) + apply(case_tac m,simp,simp) + apply(subgoal_tac "(\i. Suc i < nata \ (((P, t) # xs) ! i, xs ! i) \ ctran)") + apply force + apply clarify + apply(erule_tac x="Suc ia" in allE,simp) +apply(erule_tac x="0" and P="\j. ?H j \ (?J j) \ ctran" in allE,simp) +done + +lemma etran_or_ctran2 [rule_format]: + "\i. Suc i x\cptn \ (x!i -c\ x!Suc i \ \ x!i -e\ x!Suc i) + \ (x!i -e\ x!Suc i \ \ x!i -c\ x!Suc i)" +apply(induct x) + apply simp +apply clarify +apply(erule cptn.elims,simp) + apply(case_tac i,simp+) +apply(case_tac i,simp) + apply(force elim:etran.elims) +apply simp +done + +lemma etran_or_ctran2_disjI1: + "\ x\cptn; Suc i x!Suc i\ \ \ x!i -e\ x!Suc i" +by(drule etran_or_ctran2,simp_all) + +lemma etran_or_ctran2_disjI2: + "\ x\cptn; Suc i x!Suc i\ \ \ x!i -c\ x!Suc i" +by(drule etran_or_ctran2,simp_all) + +lemma not_ctran_None2 [rule_format]: + "\ (None, s) # xs \cptn; i \ \ ((None, s) # xs) ! i -c\ xs ! i" +apply(frule not_ctran_None,simp) +apply(case_tac i,simp) + apply(force elim:etran.elims) +apply simp +apply(rule etran_or_ctran2_disjI2,simp_all) +apply(force intro:tl_of_cptn_is_cptn) +done + +lemma Ex_first_occurrence [rule_format]: "P (n::nat) \ (\m. P m \ (\i P i))"; +apply(rule nat_less_induct) +apply clarify +apply(case_tac "\m. m \ P m") +apply auto +done + +lemma stability [rule_format]: + "\j k. x \ cptn \ stable p rely \ j\k \ k snd(x!j)\p \ + (\i. (Suc i) (x!i -e\ x!(Suc i)) \ (snd(x!i), snd(x!(Suc i))) \ rely) \ + (\i. j\i \ i x!i -e\ x!Suc i) \ snd(x!k)\p \ fst(x!j)=fst(x!k)" +apply(induct x) + apply clarify + apply(force elim:cptn.elims) +apply clarify +apply(erule cptn.elims,simp) + apply simp + apply(case_tac k,simp,simp) + apply(case_tac j,simp) + apply(erule_tac x=0 in allE) + apply(erule_tac x="nat" and P="\j. (0\j) \ (?J j)" in allE,simp) + apply(subgoal_tac "t\p") + apply(subgoal_tac "(\i. i < length xs \ ((P, t) # xs) ! i -e\ xs ! i \ (snd (((P, t) # xs) ! i), snd (xs ! i)) \ rely)") + apply clarify + apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j)\etran" in allE,simp) + apply clarify + apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j) \ (?T j)\rely" in allE,simp) + apply(erule_tac x=0 and P="\j. (?H j) \ (?J j)\etran \ ?T j" in allE,simp) + apply(simp(no_asm_use) only:stable_def) + apply(erule_tac x=s in allE) + apply(erule_tac x=t in allE) + apply simp + apply(erule mp) + apply(erule mp) + apply(rule Env) + apply simp + apply(erule_tac x="nata" in allE) + apply(erule_tac x="nat" and P="\j. (?s\j) \ (?J j)" in allE,simp) + apply(subgoal_tac "(\i. i < length xs \ ((P, t) # xs) ! i -e\ xs ! i \ (snd (((P, t) # xs) ! i), snd (xs ! i)) \ rely)") + apply clarify + apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j)\etran" in allE,simp) + apply clarify + apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j) \ (?T j)\rely" in allE,simp) +apply(case_tac k,simp,simp) +apply(case_tac j) + apply(erule_tac x=0 and P="\j. (?H j) \ (?J j)\etran" in allE,simp) + apply(erule etran.elims,simp) +apply(erule_tac x="nata" in allE) +apply(erule_tac x="nat" and P="\j. (?s\j) \ (?J j)" in allE,simp) +apply(subgoal_tac "(\i. i < length xs \ ((Q, t) # xs) ! i -e\ xs ! i \ (snd (((Q, t) # xs) ! i), snd (xs ! i)) \ rely)") + apply clarify + apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j)\etran" in allE,simp) +apply clarify +apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j) \ (?T j)\rely" in allE,simp) +done + +subsection {* Soundness of the System for Component Programs *} + +subsubsection {* Soundness of the Basic rule *} + +lemma unique_ctran_Basic [rule_format]: + "\s i. x \ cptn \ x ! 0 = (Some (Basic f), s) \ + Suc i x!i -c\ x!Suc i \ (\j. Suc j i\j \ x!j -e\ x!Suc j)" +apply(induct x,simp) +apply simp +apply clarify +apply(erule cptn.elims,simp) + apply(case_tac i,simp+) + apply clarify + apply(case_tac j,simp) + apply(rule Env) + apply simp +apply clarify +apply simp +apply(case_tac i) + apply(case_tac j,simp,simp) + apply(erule ctran.elims,simp_all) + apply(force elim: not_ctran_None) +apply(ind_cases "((Some (Basic f), sa), Q, t) \ ctran") +apply simp +apply(drule_tac i=nat in not_ctran_None,simp) +apply(erule etran.elims,simp) +done + +lemma exists_ctran_Basic_None [rule_format]: + "\s i. x \ cptn \ x ! 0 = (Some (Basic f), s) + \ i fst(x!i)=None \ (\j x!Suc j)" +apply(induct x,simp) +apply simp +apply clarify +apply(erule cptn.elims,simp) + apply(case_tac i,simp,simp) + apply(erule_tac x=nat in allE,simp) + apply clarify + apply(rule_tac x="Suc j" in exI,simp,simp) +apply clarify +apply(case_tac i,simp,simp) +apply(rule_tac x=0 in exI,simp) +done + +lemma Basic_sound: + " \pre \ {s. f s \ post}; {(s, t). s \ pre \ t = f s} \ guar; + stable pre rely; stable post rely\ + \ \ Basic f sat [pre, rely, guar, post]" +apply(unfold com_validity_def) +apply clarify +apply(simp add:comm_def) +apply(rule conjI) + apply clarify + apply(simp add:cp_def assum_def) + apply clarify + apply(frule_tac j=0 and k=i and p=pre in stability) + apply simp_all + apply(erule_tac x=ia in allE,simp) + apply(erule_tac i=i and f=f in unique_ctran_Basic,simp_all) + apply(erule subsetD,simp) + apply(case_tac "x!i") + apply clarify + apply(drule_tac s="Some (Basic f)" in sym,simp) + apply(thin_tac "\j. ?H j") + apply(force elim:ctran.elims) +apply clarify +apply(simp add:cp_def) +apply clarify +apply(frule_tac i="length x - 1" and f=f in exists_ctran_Basic_None,simp+) + apply(case_tac x,simp+) + apply(rule last_fst_esp,simp add:last_length) + apply (case_tac x,simp+) +apply(simp add:assum_def) +apply clarify +apply(frule_tac j=0 and k="j" and p=pre in stability) + apply simp_all + apply arith + apply(erule_tac x=i in allE,simp) + apply(erule_tac i=j and f=f in unique_ctran_Basic,simp_all) + apply arith + apply arith +apply(case_tac "x!j") +apply clarify +apply simp +apply(drule_tac s="Some (Basic f)" in sym,simp) +apply(case_tac "x!Suc j",simp) +apply(rule ctran.elims,simp) +apply(simp_all) +apply(drule_tac c=sa in subsetD,simp) +apply clarify +apply(frule_tac j="Suc j" and k="length x - 1" and p=post in stability,simp_all) + apply(case_tac x,simp+) + apply(erule_tac x=i in allE) +apply(erule_tac i=j and f=f in unique_ctran_Basic,simp_all) + apply arith+ +apply(case_tac x) +apply(simp add:last_length)+ +done + +subsubsection{* Soundness of the Await rule *} + +lemma unique_ctran_Await [rule_format]: + "\s i. x \ cptn \ x ! 0 = (Some (Await b c), s) \ + Suc i x!i -c\ x!Suc i \ (\j. Suc j i\j \ x!j -e\ x!Suc j)" +apply(induct x,simp+) +apply clarify +apply(erule cptn.elims,simp) + apply(case_tac i,simp+) + apply clarify + apply(case_tac j,simp) + apply(rule Env) + apply simp +apply clarify +apply simp +apply(case_tac i) + apply(case_tac j,simp,simp) + apply(erule ctran.elims,simp_all) + apply(force elim: not_ctran_None) +apply(ind_cases "((Some (Await b c), sa), Q, t) \ ctran",simp) +apply(drule_tac i=nat in not_ctran_None,simp) +apply(erule etran.elims,simp) +done + +lemma exists_ctran_Await_None [rule_format]: + "\s i. x \ cptn \ x ! 0 = (Some (Await b c), s) + \ i fst(x!i)=None \ (\j x!Suc j)" +apply(induct x,simp+) +apply clarify +apply(erule cptn.elims,simp) + apply(case_tac i,simp+) + apply(erule_tac x=nat in allE,simp) + apply clarify + apply(rule_tac x="Suc j" in exI,simp,simp) +apply clarify +apply(case_tac i,simp,simp) +apply(rule_tac x=0 in exI,simp) +done + +lemma Star_imp_cptn: + "(P, s) -c*\ (R, t) \ \l \ cp P s. (last l)=(R, t) + \ (\i. Suc i l!i -c\ l!Suc i)" +apply (erule converse_rtrancl_induct2) + apply(rule_tac x="[(R,t)]" in bexI) + apply simp + apply(simp add:cp_def) + apply(rule CptnOne) +apply clarify +apply(rule_tac x="(a, b)#l" in bexI) + apply (rule conjI) + apply(case_tac l,simp add:cp_def) + apply(simp add:last_length) + apply clarify +apply(case_tac i,simp) +apply(simp add:cp_def) +apply force +apply(simp add:cp_def) + apply(case_tac l) + apply(force elim:cptn.elims) +apply simp +apply(erule CptnComp) +apply clarify +done + +lemma Await_sound: + "\stable pre rely; stable post rely; + \V. \ P sat [pre \ b \ {s. s = V}, {(s, t). s = t}, UNIV, {s. (V, s) \ guar} \ post] \ + \ P sat [pre \ b \ {s. s = V}, {(s, t). s = t}, UNIV, {s. (V, s) \ guar} \ post] \ + \ \ Await b P sat [pre, rely, guar, post]" +apply(unfold com_validity_def) +apply clarify +apply(simp add:comm_def) +apply(rule conjI) + apply clarify + apply(simp add:cp_def assum_def) + apply clarify + apply(frule_tac j=0 and k=i and p=pre in stability,simp_all) + apply(erule_tac x=ia in allE,simp) + apply(subgoal_tac "x\ cp (Some(Await b P)) s") + apply(erule_tac i=i in unique_ctran_Await,force,simp_all) + apply(simp add:cp_def) +--{* here starts the different part. *} + apply(erule ctran.elims,simp_all) + apply(drule Star_imp_cptn) + apply clarify + apply(erule_tac x=sa in allE) + apply clarify + apply(erule_tac x=sa in allE) + apply(drule_tac c=l in subsetD) + apply (simp add:cp_def) + apply clarify + apply(erule_tac x=ia and P="\i. ?H i \ (?J i,?I i)\ctran" in allE,simp) + apply(erule etran.elims,simp) + apply simp +apply clarify +apply(simp add:cp_def) +apply clarify +apply(frule_tac i="length x - 1" in exists_ctran_Await_None,force) + apply (case_tac x,simp+) + apply(rule last_fst_esp,simp add:last_length) + apply(case_tac x, (simp add:cptn_not_empty)+) +apply clarify +apply(simp add:assum_def) +apply clarify +apply(frule_tac j=0 and k="j" and p=pre in stability,simp_all) + apply arith + apply(erule_tac x=i in allE,simp) + apply(erule_tac i=j in unique_ctran_Await,force,simp_all) + apply arith + apply arith +apply(case_tac "x!j") +apply clarify +apply simp +apply(drule_tac s="Some (Await b P)" in sym,simp) +apply(case_tac "x!Suc j",simp) +apply(rule ctran.elims,simp) +apply(simp_all) +apply(drule Star_imp_cptn) +apply clarify +apply(erule_tac x=sa in allE) +apply clarify +apply(erule_tac x=sa in allE) +apply(drule_tac c=l in subsetD) + apply (simp add:cp_def) + apply clarify + apply(erule_tac x=i and P="\i. ?H i \ (?J i,?I i)\ctran" in allE,simp) + apply(erule etran.elims,simp) +apply simp +apply clarify +apply(frule_tac j="Suc j" and k="length x - 1" and p=post in stability,simp_all) + apply(case_tac x,simp+) + apply(erule_tac x=i in allE) +apply(erule_tac i=j in unique_ctran_Await,force,simp_all) + apply arith+ +apply(case_tac x) +apply(simp add:last_length)+ +done + +subsubsection{* Soundness of the Conditional rule *} + +lemma last_length2 [rule_format]: "xs\[] \ (last xs)=(xs!(length xs - (Suc 0)))" +apply(induct xs,simp+) +apply(case_tac "length list",simp+) +done + +lemma last_drop: "Suc m last(drop (Suc m) x) = last x" +apply(case_tac "(drop (Suc m) x)\[]") + apply(drule last_length2) + apply(erule ssubst) + apply(simp only:length_drop) + apply(subgoal_tac "Suc m + (length x - Suc m - (Suc 0)) \ length x") + apply(simp only:nth_drop) + apply(case_tac "x\[]") + apply(drule last_length2) + apply(erule ssubst) + apply simp + apply(subgoal_tac "Suc (length x - 2)=(length x - Suc 0)") + apply simp + apply arith + apply simp + apply arith +apply (simp add:length_greater_0_conv [THEN sym]) +done + +lemma Cond_sound: + "\ stable pre rely; \ P1 sat [pre \ b, rely, guar, post]; + \ P2 sat [pre \ - b, rely, guar, post]; \s. (s,s)\guar\ + \ \ (Cond b P1 P2) sat [pre, rely, guar, post]" +apply(unfold com_validity_def) +apply clarify +apply(simp add:cp_def comm_def) +apply(case_tac "\i. Suc i x!i -c\ x!Suc i") + prefer 2 + apply simp + apply clarify + apply(frule_tac j="0" and k="length x - 1" and p=pre in stability,simp+) + apply(case_tac x,simp+) + apply(simp add:assum_def) + apply(simp add:assum_def) + apply(erule_tac m="length x" in etran_or_ctran,simp+) + apply(case_tac x,simp+) + apply(case_tac x, (simp add:last_length)+) +apply(erule exE) +apply(drule_tac n=i and P="\i. ?H i \ (?J i,?I i)\ ctran" in Ex_first_occurrence) +apply clarify +apply (simp add:assum_def) +apply(frule_tac j=0 and k="m" and p=pre in stability,simp+) + apply(erule_tac m="Suc m" in etran_or_ctran,simp+) +apply(erule ctran.elims,simp_all) + apply(erule_tac x="sa" in allE) + apply(drule_tac c="drop (Suc m) x" in subsetD) + apply simp + apply(rule conjI) + apply force + apply clarify + apply(subgoal_tac "(Suc m) + i \ length x") + apply(subgoal_tac "(Suc m) + (Suc i) \ length x") + apply(rotate_tac -2) + apply simp + apply(erule_tac x="Suc (m + i)" and P="\j. ?H j \ ?J j \ ?I j" in allE) + apply(subgoal_tac "Suc (Suc (m + i)) < length x",simp) + apply arith + apply arith + apply arith + apply simp + apply(rule conjI) + apply clarify + apply(case_tac "i\m") + apply(drule le_imp_less_or_eq) + apply(erule disjE) + apply(erule_tac x=i in allE, erule impE, assumption) + apply simp+ + apply(erule_tac x="i - (Suc m)" and P="\j. ?H j \ ?J j \ (?I j)\guar" in allE) + apply(subgoal_tac "(Suc m)+(i - Suc m) \ length x") + apply(subgoal_tac "(Suc m)+Suc (i - Suc m) \ length x") + apply(rotate_tac -2) + apply simp + apply(erule mp) + apply arith + apply arith + apply arith + apply(simp add:last_drop) +apply(case_tac "length (drop (Suc m) x)",simp) +apply(erule_tac x="sa" in allE) +back +apply(drule_tac c="drop (Suc m) x" in subsetD,simp) + apply(rule conjI) + apply force + apply clarify + apply(subgoal_tac "(Suc m) + i \ length x") + apply(subgoal_tac "(Suc m) + (Suc i) \ length x") + apply(rotate_tac -2) + apply simp + apply(erule_tac x="Suc (m + i)" and P="\j. ?H j \ ?J j \ ?I j" in allE) + apply(subgoal_tac "Suc (Suc (m + i)) < length x") + apply simp + apply arith + apply arith + apply arith +apply simp +apply clarify +apply(rule conjI) + apply clarify + apply(case_tac "i\m") + apply(drule le_imp_less_or_eq) + apply(erule disjE) + apply(erule_tac x=i in allE, erule impE, assumption) + apply simp + apply simp + apply(erule_tac x="i - (Suc m)" and P="\j. ?H j \ ?J j \ (?I j)\guar" in allE) + apply(subgoal_tac "(Suc m)+(i - Suc m) \ length x") + apply(subgoal_tac "(Suc m)+Suc (i - Suc m) \ length x") + apply(rotate_tac -2) + apply simp + apply(erule mp) + apply arith + apply arith + apply arith +apply(simp add:last_drop) +done + +subsubsection{* Soundness of the Sequential rule *} + +inductive_cases Seq_cases [elim!]: "(Some (Seq P Q), s) -c\ t" + +lemma last_lift_not_None: "fst ((lift Q) ((x#xs)!(length xs))) \ None" +apply(subgoal_tac "length xs cptn_mod \ \s P. x !0=(Some (Seq P Q), s) \ + (\iSome Q) \ + (\xs\ cp (Some P) s. x=map (lift Q) xs)" +apply(erule cptn_mod.induct) +apply(unfold cp_def) +apply safe +apply simp_all + apply(simp add:lift_def) + apply(rule_tac x="[(Some Pa, sa)]" in exI,simp add:CptnOne) + apply(subgoal_tac "(\i < Suc (length xs). fst (((Some (Seq Pa Q), t) # xs) ! i) \ Some Q)") + apply clarify + apply(case_tac xsa,simp,simp) + apply(rule_tac x="(Some Pa, sa) #(Some Pa, t) # list" in exI,simp) + apply(rule conjI,erule CptnEnv) + apply(simp add:lift_def) + apply clarify + apply(erule_tac x="Suc i" in allE, simp) + apply(ind_cases "((Some (Seq Pa Q), sa), None, t) \ ctran") + apply(rule_tac x="(Some P, sa) # xs" in exI, simp add:cptn_iff_cptn_mod lift_def) +apply(erule_tac x="length xs" in allE, simp) +apply(simp only:Cons_lift_append) +apply(subgoal_tac "length xs < length ((Some P, sa) # xs)") + apply(simp only :nth_append length_map last_length nth_map) + apply(case_tac "last((Some P, sa) # xs)") + apply(simp add:lift_def) +apply simp +done + +lemma Seq_sound2 [rule_format]: + "x \ cptn \ \s P i. x!0=(Some (Seq P Q), s) \ i fst(x!i)=Some Q \ + (\j(Some Q)) \ + (\xs ys. xs \ cp (Some P) s \ length xs=Suc i \ ys \ cp (Some Q) (snd(xs !i)) \ x=(map (lift Q) xs)@tl ys)" +apply(erule cptn.induct) +apply(unfold cp_def) +apply safe +apply simp_all + apply(case_tac i,simp+) + apply(erule allE,erule impE,assumption,simp) + apply clarify + apply(subgoal_tac "(\j < nat. fst (((Some (Seq Pa Q), t) # xs) ! j) \ Some Q)",clarify) + prefer 2 + apply force + apply(case_tac xsa,simp,simp) + apply(rule_tac x="(Some Pa, sa) #(Some Pa, t) # list" in exI,simp) + apply(rule conjI,erule CptnEnv) + apply(simp add:lift_def) + apply(rule_tac x=ys in exI,simp) +apply(ind_cases "((Some (Seq Pa Q), sa), t) \ ctran") + apply simp + apply(rule_tac x="(Some Pa, sa)#[(None, ta)]" in exI,simp) + apply(rule conjI) + apply(drule_tac xs="[]" in CptnComp,force simp add:CptnOne,simp) + apply(case_tac i, simp+) + apply(case_tac nat,simp+) + apply(rule_tac x="(Some Q,ta)#xs" in exI,simp add:lift_def) + apply(case_tac nat,simp+) + apply(force) +apply(case_tac i, simp+) +apply(case_tac nat,simp+) +apply(erule_tac x="Suc nata" in allE,simp) +apply clarify +apply(subgoal_tac "(\j Some Q)",clarify) + prefer 2 + apply clarify + apply force +apply(rule_tac x="(Some Pa, sa)#(Some P2, ta)#(tl xsa)" in exI,simp) +apply(rule conjI,erule CptnComp) +apply(rule nth_tl_if,force,simp+) +apply(rule_tac x=ys in exI,simp) +apply(rule conjI) +apply(rule nth_tl_if,force,simp+) + apply(rule tl_zero,simp+) + apply force +apply(rule conjI,simp add:lift_def) +apply(subgoal_tac "lift Q (Some P2, ta) =(Some (Seq P2 Q), ta)") + apply(simp add:Cons_lift del:map.simps) + apply(rule nth_tl_if) + apply force + apply simp+ +apply(simp add:lift_def) +done +(* +lemma last_lift_not_None3: "fst (last (map (lift Q) (x#xs))) \ None" +apply(simp only:last_length [THEN sym]) +apply(subgoal_tac "length xs None" +apply(simp only:last_length [THEN sym]) +apply(subgoal_tac "length xs None" +apply(subgoal_tac "length xs\ P sat [pre, rely, guar, mid]; \ Q sat [mid, rely, guar, post]\ + \ \ Seq P Q sat [pre, rely, guar, post]" +apply(unfold com_validity_def) +apply clarify +apply(case_tac "\i[]") + apply(drule last_length2) + apply (simp del:map.simps) + apply(simp only:last_lift_not_None) + apply simp +--{* @{text "\i[]") + apply(drule last_length2,simp) + apply(rule conjI) + apply(erule mp) + apply(case_tac "xs!m") + apply(case_tac "fst(xs!m)",simp) + apply(simp add:lift_def nth_append) + apply clarify + apply(erule_tac x="m+i" in allE) + back + back + apply(case_tac ys,(simp add:nth_append)+) + apply (case_tac i, (simp add:snd_lift)+) + apply(erule mp) + apply(case_tac "xs!m") + apply(force elim:etran.elims intro:Env simp add:lift_def) + apply simp +apply simp +apply clarify +apply(rule conjI,clarify) + apply(case_tac "i[]") + apply(drule last_length2) + apply(simp add: snd_lift nth_append) + apply(rule conjI,clarify) + apply(case_tac ys,simp+) + apply clarify + apply(case_tac ys,simp+) + apply(drule last_length2,simp) +apply simp +done + +subsubsection{* Soundness of the While rule *} + +lemma assum_after_body: + "\ \ P sat [pre \ b, rely, guar, pre]; + (Some P, s) # xs \ cptn_mod; fst (((Some P, s) # xs)!length xs) = None; s \ b; + (Some (While b P), s) # (Some (Seq P (While b P)), s) # map (lift (While b P)) xs @ ys \ assum (pre, rely)\ + \ (Some (While b P), snd (((Some P, s) # xs)!length xs)) # ys \ assum (pre, rely)" +apply(simp add:assum_def com_validity_def cp_def cptn_iff_cptn_mod) +apply clarify +apply(erule_tac x=s in allE) +apply(drule_tac c="(Some P, s) # xs" in subsetD,simp) + apply clarify + apply(erule_tac x="Suc i" in allE) + apply simp + apply(simp add:Cons_lift_append nth_append snd_lift del:map.simps) + apply(erule mp) + apply(erule etran.elims,simp) + apply(case_tac "fst(((Some P, s) # xs) ! i)") + apply(force intro:Env simp add:lift_def) + apply(force intro:Env simp add:lift_def) +apply(rule conjI) + apply(simp add:comm_def last_length) +apply clarify +apply(erule_tac x="Suc(length xs + i)" in allE,simp) +apply(case_tac i, simp add:nth_append Cons_lift_append snd_lift del:map.simps) + apply(erule mp) + apply(case_tac "((Some P, s) # xs) ! length xs") + apply(simp add:lift_def) +apply(simp add:Cons_lift_append nth_append snd_lift del:map.simps) +done + +lemma last_append[rule_format]: + "\xs. ys\[] \ ((xs@ys)!(length (xs@ys) - (Suc 0)))=(ys!(length ys - (Suc 0)))" +apply(induct ys) + apply simp +apply clarify +apply (simp add:nth_append length_append) +done + +lemma assum_after_body: + "\ \ P sat [pre \ b, rely, guar, pre]; + (Some P, s) # xs \ cptn_mod; fst (last ((Some P, s) # xs)) = None; s \ b; + (Some (While b P), s) # (Some (Seq P (While b P)), s) # map (lift (While b P)) xs @ ys + \ assum (pre, rely)\ + \ (Some (While b P), snd (last ((Some P, s) # xs))) # ys \ assum (pre, rely)" +apply(simp add:assum_def com_validity_def cp_def cptn_iff_cptn_mod) +apply clarify +apply(erule_tac x=s in allE) +apply(drule_tac c="(Some P, s) # xs" in subsetD,simp) + apply clarify + apply(erule_tac x="Suc i" in allE) + apply simp + apply(simp add:Cons_lift_append nth_append snd_lift del:map.simps) + apply(erule mp) + apply(erule etran.elims,simp) + apply(case_tac "fst(((Some P, s) # xs) ! i)") + apply(force intro:Env simp add:lift_def) + apply(force intro:Env simp add:lift_def) +apply(rule conjI) + apply clarify + apply(simp add:comm_def last_length) +apply clarify +apply(rule conjI) + apply(simp add:comm_def) +apply clarify +apply(erule_tac x="Suc(length xs + i)" in allE,simp) +apply(case_tac i, simp add:nth_append Cons_lift_append snd_lift del:map.simps) + apply(simp add:last_length) + apply(erule mp) + apply(case_tac "last xs") + apply(simp add:lift_def) +apply(simp add:Cons_lift_append nth_append snd_lift del:map.simps) +done + +lemma last_append2:"ys\[] \ last (xs@ys)=(last ys)" +apply(frule last_length2) +apply simp +apply(subgoal_tac "xs@ys\[]") +apply(drule last_length2) +back +apply simp +apply(drule_tac xs=xs in last_append) +apply simp +apply simp +done + +lemma While_sound_aux [rule_format]: + "\ pre \ - b \ post; \ P sat [pre \ b, rely, guar, pre]; \s. (s, s) \ guar; + stable pre rely; stable post rely; x \ cptn_mod \ + \ \s xs. x=(Some(While b P),s)#xs \ x\assum(pre, rely) \ x \ comm (guar, post)" +apply(erule cptn_mod.induct) +apply safe +apply (simp_all del:last.simps) +--{* 5 subgoals left *} +apply(simp add:comm_def) +--{* 4 subgoals left *} +apply(rule etran_in_comm) +apply(erule mp) +apply(erule tl_of_assum_in_assum,simp) +--{* While-None *} +apply(ind_cases "((Some (While b P), s), None, t) \ ctran") +apply(simp add:comm_def) +apply(simp add:cptn_iff_cptn_mod [THEN sym]) +apply(rule conjI,clarify) + apply(force simp add:assum_def) +apply clarify +apply(rule conjI, clarify) + apply(case_tac i,simp,simp) + apply(force simp add:not_ctran_None2) +apply(subgoal_tac "\i. Suc i < length ((None, sa) # xs) \ (((None, sa) # xs) ! i, ((None, sa) # xs) ! Suc i)\ etran") + prefer 2 + apply clarify + apply(rule_tac m="length ((None, s) # xs)" in etran_or_ctran,simp+) + apply(erule not_ctran_None2,simp) + apply simp+ +apply(frule_tac j="0" and k="length ((None, s) # xs) - 1" and p=post in stability,simp+) + apply(force simp add:assum_def subsetD) + apply(simp add:assum_def) + apply clarify + apply(erule_tac x="i" in allE,simp) + apply(erule_tac x="Suc i" in allE,simp) + apply simp +apply clarify +apply (simp add:last_length) +--{* WhileOne *} +apply(thin_tac "P = While b P \ ?Q") +apply(rule ctran_in_comm,simp) +apply(simp add:Cons_lift del:map.simps) +apply(simp add:comm_def del:map.simps) +apply(rule conjI) + apply clarify + apply(case_tac "fst(((Some P, sa) # xs) ! i)") + apply(case_tac "((Some P, sa) # xs) ! i") + apply (simp add:lift_def) + apply(ind_cases "(Some (While b P), ba) -c\ t") + apply simp + apply simp + apply(simp add:snd_lift del:map.simps) + apply(simp only:com_validity_def cp_def cptn_iff_cptn_mod) + apply(erule_tac x=sa in allE) + apply(drule_tac c="(Some P, sa) # xs" in subsetD) + apply (simp add:assum_def del:map.simps) + apply clarify + apply(erule_tac x="Suc ia" in allE,simp add:snd_lift del:map.simps) + apply(erule mp) + apply(case_tac "fst(((Some P, sa) # xs) ! ia)") + apply(erule etran.elims,simp add:lift_def) + apply(rule Env) + apply(erule etran.elims,simp add:lift_def) + apply(rule Env) + apply (simp add:comm_def del:map.simps) + apply clarify + apply(erule allE,erule impE,assumption) + apply(erule mp) + apply(case_tac "((Some P, sa) # xs) ! i") + apply(case_tac "xs!i") + apply(simp add:lift_def) + apply(case_tac "fst(xs!i)") + apply force + apply force +--{* last=None *} +apply clarify +apply(subgoal_tac "(map (lift (While b P)) ((Some P, sa) # xs))\[]") + apply(drule last_length2) + apply (simp del:map.simps) + apply(simp only:last_lift_not_None) +apply simp +--{* WhileMore *} +apply(thin_tac "P = While b P \ ?Q") +apply(rule ctran_in_comm,simp del:last.simps) +--{* metiendo la hipotesis antes de dividir la conclusion. *} +apply(subgoal_tac "(Some (While b P), snd (last ((Some P, sa) # xs))) # ys \ assum (pre, rely)") + apply (simp del:last.simps) + prefer 2 + apply(erule assum_after_body) + apply (simp del:last.simps)+ +--{* lo de antes. *} +apply(simp add:comm_def del:map.simps last.simps) +apply(rule conjI) + apply clarify + apply(simp only:Cons_lift_append) + apply(case_tac "i t") + apply simp + apply simp + apply(simp add:snd_lift del:map.simps last.simps) + apply(thin_tac " \i. i < length ys \ ?P i") + apply(simp only:com_validity_def cp_def cptn_iff_cptn_mod) + apply(erule_tac x=sa in allE) + apply(drule_tac c="(Some P, sa) # xs" in subsetD) + apply (simp add:assum_def del:map.simps last.simps) + apply clarify + apply(erule_tac x="Suc ia" in allE,simp add:nth_append snd_lift del:map.simps last.simps, erule mp) + apply(case_tac "fst(((Some P, sa) # xs) ! ia)") + apply(erule etran.elims,simp add:lift_def) + apply(rule Env) + apply(erule etran.elims,simp add:lift_def) + apply(rule Env) + apply (simp add:comm_def del:map.simps) + apply clarify + apply(erule allE,erule impE,assumption) + apply(erule mp) + apply(case_tac "((Some P, sa) # xs) ! i") + apply(case_tac "xs!i") + apply(simp add:lift_def) + apply(case_tac "fst(xs!i)") + apply force + apply force +--{* @{text "i \ length xs"} *} +apply(subgoal_tac "i-length xs length xs"} *} +apply(case_tac "i-length xs") + apply arith +apply(simp add:nth_append del:map.simps last.simps) +apply(rule conjI,clarify,arith) +apply clarify +apply(subgoal_tac "i- Suc (length xs)=nat") + prefer 2 + apply arith +apply simp +--{* last=None *} +apply clarify +apply(case_tac ys) + apply(simp add:Cons_lift del:map.simps last.simps) + apply(subgoal_tac "(map (lift (While b P)) ((Some P, sa) # xs))\[]") + apply(drule last_length2) + apply (simp del:map.simps) + apply(simp only:last_lift_not_None) + apply simp +apply(subgoal_tac "((Some (Seq P (While b P)), sa) # map (lift (While b P)) xs @ ys)\[]") + apply(drule last_length2) + apply (simp del:map.simps last.simps) + apply(simp add:nth_append del:last.simps) + apply(subgoal_tac "((Some (While b P), snd (last ((Some P, sa) # xs))) # a # list)\[]") + apply(drule last_length2) + apply (simp del:map.simps last.simps) + apply simp +apply simp +done + +lemma While_sound: + "\stable pre rely; pre \ - b \ post; stable post rely; + \ P sat [pre \ b, rely, guar, pre]; \s. (s,s)\guar\ + \ \ While b P sat [pre, rely, guar, post]" +apply(unfold com_validity_def) +apply clarify +apply(erule_tac xs="tl x" in While_sound_aux) + apply(simp add:com_validity_def) + apply force + apply simp_all +apply(simp add:cptn_iff_cptn_mod cp_def) +apply(simp add:cp_def) +apply clarify +apply(rule nth_equalityI) + apply simp_all + apply(case_tac x,simp+) +apply clarify +apply(case_tac i,simp+) +apply(case_tac x,simp+) +done + +subsubsection{* Soundness of the Rule of Consequence *} + +lemma Conseq_sound: + "\pre \ pre'; rely \ rely'; guar' \ guar; post' \ post; + \ P sat [pre', rely', guar', post']\ + \ \ P sat [pre, rely, guar, post]" +apply(simp add:com_validity_def assum_def comm_def) +apply clarify +apply(erule_tac x=s in allE) +apply(drule_tac c=x in subsetD) + apply force +apply force +done + +subsubsection {* Soundness of the system for sequential component programs *} + +theorem rgsound: + "\ P sat [pre, rely, guar, post] \ \ P sat [pre, rely, guar, post]" +apply(erule rghoare.induct) + apply(force elim:Basic_sound) + apply(force elim:Seq_sound) + apply(force elim:Cond_sound) + apply(force elim:While_sound) + apply(force elim:Await_sound) +apply(erule Conseq_sound,simp+) +done + +subsection {* Soundness of the System for Parallel Programs *} + +constdefs + ParallelCom :: "('a rgformula) list \ 'a par_com" + "ParallelCom Ps \ map (Some \ fst) Ps" + +lemma two: + "\ \i (\j\{j. j < length xs \ j \ i}. Guar (xs ! j)) \ Rely (xs ! i); + pre \ (\i\{i. i < length xs}. Pre (xs ! i)); + \i Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)]; + length xs=length clist; x \ par_cp (ParallelCom xs) s; x\par_assum(pre, rely); + \icp (Some(Com(xs!i))) s; x \ clist \ + \ \j i. i Suc j (clist!i!j) -c\ (clist!i!Suc j) + \ (snd(clist!i!j), snd(clist!i!Suc j)) \ Guar(xs!i)" +apply(unfold par_cp_def) +--{* By contradiction: *} +apply(subgoal_tac "True") + prefer 2 + apply simp +apply(erule_tac Q="True" in contrapos_pp) +apply simp +apply(erule exE) +--{* the first c-tran that does not satisfy the guarantee-condition is from @{text "\_i"} at step @{text "m"}. *} +apply(drule_tac n=j and P="\j. \i. ?H i j" in Ex_first_occurrence) +apply(erule exE) +apply clarify +--{* @{text "\_i \ A(pre, rely_1)"} *} +apply(subgoal_tac "take (Suc (Suc m)) (clist!i) \ assum(Pre(xs!i), Rely(xs!i))") +--{* but this contradicts @{text "\ \_i sat [pre_i,rely_i,guar_i,post_i]"} *} + apply(erule_tac x=i and P="\i. ?H i \ \ (?J i) sat [?I i,?K i,?M i,?N i]" in allE,erule impE,assumption) + apply(simp add:com_validity_def) + apply(erule_tac x=s in allE) + apply(simp add:cp_def comm_def) + apply(drule_tac c="take (Suc (Suc m)) (clist ! i)" in subsetD) + apply simp + apply(erule_tac x=i in allE, erule impE, assumption,erule conjE) + apply(erule takecptn_is_cptn) + apply simp + apply clarify + apply(erule_tac x=m and P="\j. ?I j \ ?J j \ ?H j" in allE) + apply (simp add:conjoin_def same_length_def) +apply(simp add:assum_def) +apply(rule conjI) + apply(erule_tac x=i and P="\j. ?H j \ ?I j \cp (?K j) (?J j)" in allE) + apply(simp add:cp_def par_assum_def) + apply(drule_tac c="s" in subsetD,simp) + apply simp +apply clarify +apply(erule_tac x=i and P="\j. ?H j \ ?M \ UNION (?S j) (?T j) \ (?L j)" in allE) +apply simp +apply(erule subsetD) +apply simp +apply(simp add:conjoin_def compat_label_def) +apply clarify +apply(erule_tac x=ia and P="\j. ?H j \ (?P j) \ ?Q j" in allE,simp) +--{* each etran in @{text "\_1[0\m]"} corresponds to *} +apply(erule disjE) +--{* a c-tran in some @{text "\_{ib}"} *} + apply clarify + apply(case_tac "i=ib",simp) + apply(erule etran.elims,simp) + apply(erule_tac x="ib" and P="\i. ?H i \ (?I i) \ (?J i)" in allE,simp) + apply(erule disjE,arith) + apply(case_tac "ia=m",simp) + apply(erule etran.elims,simp) + apply(erule_tac x=ia and P="\j. ?H j \ (\ i. ?P i j)" in allE) + apply(subgoal_tac "ia"}, +therefore it satisfies @{text "rely \ guar_{ib}"} *} +apply (force simp add:par_assum_def same_state_def) +done + +lemma three [rule_format]: + "\ xs\[]; \i (\j\{j. j < length xs \ j \ i}. Guar (xs ! j)) \ Rely (xs ! i); + pre \ (\i\{i. i < length xs}. Pre (xs ! i)); + \i Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)]; + length xs=length clist; x \ par_cp (ParallelCom xs) s; x\par_assum(pre, rely); + \icp (Some(Com(xs!i))) s; x \ clist \ + \ \j i. i Suc j (clist!i!j) -e\ (clist!i!Suc j) + \ (snd(clist!i!j), snd(clist!i!Suc j)) \ rely \ (\j\{j. j < length xs \ j \ i}. Guar (xs ! j))" +apply(drule two) + apply simp_all +apply clarify +apply(simp add:conjoin_def compat_label_def) +apply clarify +apply(erule_tac x=j and P="\j. ?H j \ (?J j \ (\i. ?P i j)) \ ?I j" in allE,simp) +apply(erule disjE) + prefer 2 + apply(force simp add:same_state_def par_assum_def) +apply clarify +apply(case_tac "i=ia",simp) + apply(erule etran.elims,simp) +apply(erule_tac x="ia" and P="\i. ?H i \ (?I i) \ (?J i)" in allE,simp) + apply(erule disjE,arith) +apply(erule_tac x=j and P="\j. \i. ?S j i \ (?I j i, ?H j i)\ ctran \ (?P i j)" in allE) +apply(erule_tac x=ia and P="\j. ?S j \ (?I j, ?H j)\ ctran \ (?P j)" in allE) +apply(simp add:same_state_def) +apply(erule_tac x=i and P="\j. (?T j) \ (\i. (?H j i) \ (snd (?d j i))=(snd (?e j i)))" in all_dupE,simp) +apply(erule_tac x=ia and P="\j. (?T j) \ (\i. (?H j i) \ (snd (?d j i))=(snd (?e j i)))" in allE,simp) +apply(erule_tac x=j and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in all_dupE) +apply(erule_tac x=j and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in all_dupE) +apply(erule_tac x="Suc j" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) +apply(erule_tac x="Suc j" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) +apply simp +done + +lemma four: + "\xs\[]; \i < length xs. rely \ (\j\{j. j < length xs \ j \ i}. Guar (xs ! j)) \ Rely (xs ! i); + (\j\{j. j < length xs}. Guar (xs ! j)) \ guar; pre \ (\i\{i. i < length xs}. Pre (xs ! i)); + \i < length xs. \ Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)]; + x \ par_cp (ParallelCom xs) s; x \ par_assum (pre, rely); Suc i < length x; x ! i -pc\ x ! Suc i\ + \ (snd (x ! i), snd (x ! Suc i)) \ guar" +apply(simp add: ParallelCom_def) +apply(subgoal_tac "(map (Some \ fst) xs)\[]") + prefer 2 + apply simp +apply(frule rev_subsetD) + apply(erule one [THEN equalityD1]) +apply(erule subsetD) +apply simp +apply clarify +apply(drule_tac pre=pre and rely=rely and x=x and s=s and xs=xs and clist=clist in two) +apply(assumption+) + apply(erule sym) + apply(simp add:ParallelCom_def) + apply assumption + apply(simp add:Com_def) + apply assumption +apply(simp add:conjoin_def same_program_def) +apply clarify +apply(erule_tac x=i and P="\j. ?H j \ fst(?I j)=(?J j)" in all_dupE) +apply(erule_tac x="Suc i" and P="\j. ?H j \ fst(?I j)=(?J j)" in allE) +apply(erule par_ctran.elims,simp) +apply(erule_tac x=i and P="\j. \i. ?S j i \ (?I j i, ?H j i)\ ctran \ (?P i j)" in allE) +apply(erule_tac x=ia and P="\j. ?S j \ (?I j, ?H j)\ ctran \ (?P j)" in allE) +apply(rule_tac x=ia in exI) +apply(simp add:same_state_def) +apply(erule_tac x=ia and P="\j. (?T j) \ (\i. (?H j i) \ (snd (?d j i))=(snd (?e j i)))" in all_dupE,simp) +apply(erule_tac x=ia and P="\j. (?T j) \ (\i. (?H j i) \ (snd (?d j i))=(snd (?e j i)))" in allE,simp) +apply(erule_tac x=i and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in all_dupE) +apply(erule_tac x=i and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in all_dupE,simp) +apply(erule_tac x="Suc i" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) +apply(erule mp) +apply(subgoal_tac "r=fst(clist ! ia ! Suc i)",simp) +apply(drule_tac i=ia in list_eq_if) +back +apply simp_all +done + +lemma parcptn_not_empty [simp]:"[] \ par_cptn" +apply(force elim:par_cptn.elims) +done + +lemma five: + "\xs\[]; \i (\j\{j. j < length xs \ j \ i}. Guar (xs ! j)) \ Rely (xs ! i); + pre \ (\i\{i. i < length xs}. Pre (xs ! i)); (\i\{i. i < length xs}. Post (xs ! i)) \ post; + \i < length xs. \ Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)]; + x \ par_cp (ParallelCom xs) s; x \ par_assum (pre, rely); All_None (fst (last x)) \ + \ snd (last x) \ post" +apply(simp add: ParallelCom_def) +apply(subgoal_tac "(map (Some \ fst) xs)\[]") + prefer 2 + apply simp +apply(frule rev_subsetD) + apply(erule one [THEN equalityD1]) +apply(erule subsetD) +apply simp +apply clarify +apply(subgoal_tac "\iassum(Pre(xs!i), Rely(xs!i))") + apply(erule_tac x=i and P="\i. ?H i \ \ (?J i) sat [?I i,?K i,?M i,?N i]" in allE,erule impE,assumption) + apply(simp add:com_validity_def) + apply(erule_tac x=s in allE) + apply(erule_tac x=i and P="\j. ?H j \ (?I j) \ cp (?J j) s" in allE,simp) + apply(drule_tac c="clist!i" in subsetD) + apply (force simp add:Com_def) + apply(simp add:comm_def conjoin_def same_program_def del:last.simps) + apply clarify + apply(erule_tac x="length x - 1" and P="\j. ?H j \ fst(?I j)=(?J j)" in allE) + apply (simp add:All_None_def same_length_def) + apply(erule_tac x=i and P="\j. ?H j \ length(?J j)=(?K j)" in allE) + apply(subgoal_tac "length x - 1 < length x",simp) + apply(case_tac "x\[]") + apply(drule last_length2,simp) + apply(erule_tac x="clist!i" in ballE) + apply(simp add:same_state_def) + apply(subgoal_tac "clist!i\[]") + apply(drule_tac xs="clist!i" in last_length2,simp) + apply(case_tac x) + apply (force simp add:par_cp_def) + apply (force simp add:par_cp_def) + apply force + apply (force simp add:par_cp_def) + apply(case_tac x) + apply (force simp add:par_cp_def) + apply (force simp add:par_cp_def) +apply clarify +apply(simp add:assum_def) +apply(rule conjI) + apply(simp add:conjoin_def same_state_def par_cp_def) + apply clarify + apply(erule_tac x=ia and P="\j. (?T j) \ (\i. (?H j i) \ (snd (?d j i))=(snd (?e j i)))" in allE,simp) + apply(erule_tac x=0 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) + apply(case_tac x,simp+) + apply (simp add:par_assum_def) + apply clarify + apply(drule_tac c="snd (clist ! ia ! 0)" in subsetD) + apply assumption + apply simp +apply clarify +apply(erule_tac x=ia in all_dupE) +apply simp +apply(rule subsetD) + apply simp +apply(erule_tac pre=pre and rely=rely and x=x and s=s in three) + apply(erule_tac x=ic in allE,erule mp) + apply simp_all + apply(simp add:ParallelCom_def) + apply(force simp add:Com_def) +apply(simp add:conjoin_def same_length_def) +done + +lemma ParallelEmpty [rule_format]: + "\i s. x \ par_cp (ParallelCom []) s \ + Suc i < length x \ (x ! i, x ! Suc i) \ par_ctran" +apply(induct_tac x) + apply(simp add:par_cp_def ParallelCom_def) +apply clarify +apply(case_tac list,simp,simp) +apply(case_tac i) + apply(simp add:par_cp_def ParallelCom_def) + apply(erule par_ctran.elims,simp) +apply(simp add:par_cp_def ParallelCom_def) +apply clarify +apply(erule par_cptn.elims,simp) + apply simp +apply(erule par_ctran.elims) +back +apply simp +done + +theorem par_rgsound: + "\ c SAT [pre, rely, guar, post] \ \ (ParallelCom c) SAT [pre, rely, guar, post]" +apply(erule par_rghoare.induct) +apply(case_tac xs,simp) + apply(simp add:par_com_validity_def par_comm_def) + apply clarify + apply(case_tac "post=UNIV",simp) + apply clarify + apply(drule ParallelEmpty) + apply assumption + apply simp + apply clarify + apply simp +apply(subgoal_tac "xs\[]") + prefer 2 + apply simp +apply(thin_tac "xs = a # list") +apply(simp add:par_com_validity_def par_comm_def) +apply clarify +apply(rule conjI) + apply clarify + apply(erule_tac pre=pre and rely=rely and guar=guar and x=x and s=s and xs=xs in four) + apply(assumption+) + apply clarify + apply (erule allE, erule impE, assumption,erule rgsound) + apply(assumption+) +apply clarify +apply(erule_tac pre=pre and rely=rely and post=post and x=x and s=s and xs=xs in five) + apply(assumption+) + apply clarify + apply (erule allE, erule impE, assumption,erule rgsound) + apply(assumption+) +done + +end \ No newline at end of file diff -r 98f0a09a33c3 -r 791e3b4c4039 src/HOL/HoareParallel/RG_Syntax.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HoareParallel/RG_Syntax.thy Tue Mar 05 17:11:25 2002 +0100 @@ -0,0 +1,90 @@ + +header {* \section{Concrete Syntax} *} + +theory RG_Syntax = Quote_Antiquote + RG_Hoare: + +syntax + "_Assign" :: "idt \ 'b \ 'a com" ("(\_ :=/ _)" [70, 65] 61) + "_skip" :: "'a com" ("SKIP") + "_Seq" :: "'a com \ 'a com \ 'a com" ("(_;;/ _)" [60,61] 60) + "_Cond" :: "'a bexp \ 'a com \ 'a com \ 'a com" ("(0IF _/ THEN _/ ELSE _/FI)" [0, 0, 0] 61) + "_Cond2" :: "'a bexp \ 'a com \ 'a com" ("(0IF _ THEN _ FI)" [0,0] 56) + "_While" :: "'a bexp \ 'a com \ 'a com" ("(0WHILE _ /DO _ /OD)" [0, 0] 61) + "_Await" :: "'a bexp \ 'a com \ 'a com" ("(0AWAIT _ /THEN /_ /END)" [0,0] 61) + "_Atom" :: "'a com \ 'a com" ("(\_\)" 61) + "_Wait" :: "'a bexp \ 'a com" ("(0WAIT _ END)" 61) + +translations + "\\x := a" \ "Basic \\\(_update_name x a)\" + "SKIP" \ "Basic id" + "c1;; c2" \ "Seq c1 c2" + "IF b THEN c1 ELSE c2 FI" \ "Cond .{b}. c1 c2" + "IF b THEN c FI" \ "IF b THEN c ELSE SKIP FI" + "WHILE b DO c OD" \ "While .{b}. c" + "AWAIT b THEN c END" \ "Await .{b}. c" + "\c\" \ "AWAIT True THEN c END" + "WAIT b END" \ "AWAIT b THEN SKIP END" + +nonterminals + prgs + +syntax + "_PAR" :: "prgs \ 'a" ("COBEGIN//_//COEND" 60) + "_prg" :: "'a \ prgs" ("_" 57) + "_prgs" :: "['a, prgs] \ prgs" ("_//\//_" [60,57] 57) + +translations + "_prg a" \ "[a]" + "_prgs a ps" \ "a # ps" + "_PAR ps" \ "ps" + +syntax + "_prg_scheme" :: "['a, 'a, 'a, 'a] \ prgs" ("SCHEME [_ \ _ < _] _" [0,0,0,60] 57) + +translations + "_prg_scheme j i k c" \ "(map (\i. c) [j..k(])" + +text {* Translations for variables before and after a transition: *} + +syntax + "_before" :: "id \ 'a" ("\_") + "_after" :: "id \ 'a" ("\_") + +translations + "\x" \ "x \fst" + "\x" \ "x \snd" + +print_translation {* + let + fun quote_tr' f (t :: ts) = + Term.list_comb (f $ Syntax.quote_tr' "_antiquote" t, ts) + | quote_tr' _ _ = raise Match; + + val assert_tr' = quote_tr' (Syntax.const "_Assert"); + + fun bexp_tr' name ((Const ("Collect", _) $ t) :: ts) = + quote_tr' (Syntax.const name) (t :: ts) + | bexp_tr' _ _ = raise Match; + + fun upd_tr' (x_upd, T) = + (case try (unsuffix RecordPackage.updateN) x_upd of + Some x => (x, if T = dummyT then T else Term.domain_type T) + | None => raise Match); + + fun update_name_tr' (Free x) = Free (upd_tr' x) + | update_name_tr' ((c as Const ("_free", _)) $ Free x) = + c $ Free (upd_tr' x) + | update_name_tr' (Const x) = Const (upd_tr' x) + | update_name_tr' _ = raise Match; + + fun assign_tr' (Abs (x, _, f $ t $ Bound 0) :: ts) = + quote_tr' (Syntax.const "_Assign" $ update_name_tr' f) + (Abs (x, dummyT, t) :: ts) + | assign_tr' _ = raise Match; + in + [("Collect", assert_tr'), ("Basic", assign_tr'), + ("Cond", bexp_tr' "_Cond"), ("While", bexp_tr' "_While_inv")] + end +*} + +end \ No newline at end of file diff -r 98f0a09a33c3 -r 791e3b4c4039 src/HOL/HoareParallel/RG_Tran.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HoareParallel/RG_Tran.thy Tue Mar 05 17:11:25 2002 +0100 @@ -0,0 +1,1076 @@ + +header {* \section{Operational Semantics} *} + +theory RG_Tran = RG_Com: + +subsection {* Semantics of Component Programs *} + +subsubsection {* Environment transitions *} + +types 'a conf = "(('a com) option) \ 'a" + +consts etran :: "('a conf \ 'a conf) set" +syntax "_etran" :: "'a conf \ 'a conf \ bool" ("_ -e\ _" [81,81] 80) +translations "P -e\ Q" \ "(P,Q) \ etran" +inductive etran +intros + Env: "(P, s) -e\ (P, t)" + +subsubsection {* Component transitions *} + +consts ctran :: "('a conf \ 'a conf) set" +syntax + "_ctran" :: "'a conf \ 'a conf \ bool" ("_ -c\ _" [81,81] 80) + "_ctran_*":: "'a conf \ 'a conf \ bool" ("_ -c*\ _" [81,81] 80) +translations + "P -c\ Q" \ "(P,Q) \ ctran" + "P -c*\ Q" \ "(P,Q) \ ctran^*" + +inductive ctran +intros + Basic: "(Some(Basic f), s) -c\ (None, f s)" + + Seq1: "(Some P0, s) -c\ (None, t) \ (Some(Seq P0 P1), s) -c\ (Some P1, t)" + + Seq2: "(Some P0, s) -c\ (Some P2, t) \ (Some(Seq P0 P1), s) -c\ (Some(Seq P2 P1), t)" + + CondT: "s\b \ (Some(Cond b P1 P2), s) -c\ (Some P1, s)" + CondF: "s\b \ (Some(Cond b P1 P2), s) -c\ (Some P2, s)" + + WhileF: "s\b \ (Some(While b P), s) -c\ (None, s)" + WhileT: "s\b \ (Some(While b P), s) -c\ (Some(Seq P (While b P)), s)" + + Await: "\s\b; (Some P, s) -c*\ (None, t)\ \ (Some(Await b P), s) -c\ (None, t)" + +monos "rtrancl_mono" + +subsection {* Semantics of Parallel Programs *} + +types 'a par_conf = "('a par_com) \ 'a" +consts + par_etran :: "('a par_conf \ 'a par_conf) set" + par_ctran :: "('a par_conf \ 'a par_conf) set" +syntax + "_par_etran":: "['a par_conf,'a par_conf] \ bool" ("_ -pe\ _" [81,81] 80) + "_par_ctran":: "['a par_conf,'a par_conf] \ bool" ("_ -pc\ _" [81,81] 80) +translations + "P -pe\ Q" \ "(P,Q) \ par_etran" + "P -pc\ Q" \ "(P,Q) \ par_ctran" + +inductive par_etran +intros + ParEnv: "(Ps, s) -pe\ (Ps, t)" + +inductive par_ctran +intros + ParComp: "\i (r, t)\ \ (Ps, s) -pc\ (Ps[i:=r], t)" + +subsection {* Computations *} + +subsubsection {* Sequential computations *} + +types 'a confs = "('a conf) list" +consts cptn :: "('a confs) set" +inductive "cptn" +intros + CptnOne: "[(P,s)] \ cptn" + CptnEnv: "(P, t)#xs \ cptn \ (P,s)#(P,t)#xs \ cptn" + CptnComp: "\(P,s) -c\ (Q,t); (Q, t)#xs \ cptn \ \ (P,s)#(Q,t)#xs \ cptn" + +constdefs + cp :: "('a com) option \ 'a \ ('a confs) set" + "cp P s \ {l. l!0=(P,s) \ l \ cptn}" + +subsubsection {* Parallel computations *} + +types 'a par_confs = "('a par_conf) list" +consts par_cptn :: "('a par_confs) set" +inductive "par_cptn" +intros + ParCptnOne: "[(P,s)] \ par_cptn" + ParCptnEnv: "(P,t)#xs \ par_cptn \ (P,s)#(P,t)#xs \ par_cptn" + ParCptnComp: "\ (P,s) -pc\ (Q,t); (Q,t)#xs \ par_cptn \ \ (P,s)#(Q,t)#xs \ par_cptn" + +constdefs + par_cp :: "'a par_com \ 'a \ ('a par_confs) set" + "par_cp P s \ {l. l!0=(P,s) \ l \ par_cptn}" + +subsection{* Modular Definition of Computation *} + +constdefs + lift :: "'a com \ 'a conf \ 'a conf" + "lift Q \ \(P, s). (if P=None then (Some Q,s) else (Some(Seq (the P) Q), s))" + +consts cptn_mod :: "('a confs) set" +inductive "cptn_mod" +intros + CptnModOne: "[(P, s)] \ cptn_mod" + CptnModEnv: "(P, t)#xs \ cptn_mod \ (P, s)#(P, t)#xs \ cptn_mod" + CptnModNone: "\(Some P, s) -c\ (None, t); (None, t)#xs \ cptn_mod \ \ (Some P,s)#(None, t)#xs \cptn_mod" + CptnModCondT: "\(Some P0, s)#ys \ cptn_mod; s \ b \ \ (Some(Cond b P0 P1), s)#(Some P0, s)#ys \ cptn_mod" + CptnModCondF: "\(Some P1, s)#ys \ cptn_mod; s \ b \ \ (Some(Cond b P0 P1), s)#(Some P1, s)#ys \ cptn_mod" + CptnModSeq1: "\(Some P0, s)#xs \ cptn_mod; zs=map (lift P1) xs \ + \ (Some(Seq P0 P1), s)#zs \ cptn_mod" + CptnModSeq2: + "\(Some P0, s)#xs \ cptn_mod; fst(last ((Some P0, s)#xs)) = None; + (Some P1, snd(last ((Some P0, s)#xs)))#ys \ cptn_mod; + zs=(map (lift P1) xs)@ys \ \ (Some(Seq P0 P1), s)#zs \ cptn_mod" + + CptnModWhile1: + "\ (Some P, s)#xs \ cptn_mod; s \ b; zs=map (lift (While b P)) xs \ + \ (Some(While b P), s)#(Some(Seq P (While b P)), s)#zs \ cptn_mod" + CptnModWhile2: + "\ (Some P, s)#xs \ cptn_mod; fst(last ((Some P, s)#xs))=None; s \ b; + zs=(map (lift (While b P)) xs)@ys; + (Some(While b P), snd(last ((Some P, s)#xs)))#ys \ cptn_mod\ + \ (Some(While b P), s)#(Some(Seq P (While b P)), s)#zs \ cptn_mod" + +subsection {* Equivalence of Both Definitions.*} + +lemma last_length: "((a#xs)!(length xs))=last (a#xs)" +apply simp +apply(induct xs,simp+) +apply(case_tac list) +apply simp_all +done + +lemma div_seq [rule_format]: "list \ cptn_mod \ + (\s P Q zs. list=(Some (Seq P Q), s)#zs \ + (\xs. (Some P, s)#xs \ cptn_mod \ (zs=(map (lift Q) xs) \ + ( fst(((Some P, s)#xs)!length xs)=None \ + (\ys. (Some Q, snd(((Some P, s)#xs)!length xs))#ys \ cptn_mod + \ zs=(map (lift (Q)) xs)@ys)))))" +apply(erule cptn_mod.induct) +apply simp_all + apply clarify + apply(force intro:CptnModOne) + apply clarify + apply(erule_tac x=Pa in allE) + apply(erule_tac x=Q in allE) + apply simp + apply clarify + apply(erule disjE) + apply(rule_tac x="(Some Pa,t)#xsa" in exI) + apply(rule conjI) + apply clarify + apply(erule CptnModEnv) + apply(rule disjI1) + apply(simp add:lift_def) + apply clarify + apply(rule_tac x="(Some Pa,t)#xsa" in exI) + apply(rule conjI) + apply(erule CptnModEnv) + apply(rule disjI2) + apply(rule conjI) + apply(case_tac xsa,simp,simp) + apply(rule_tac x="ys" in exI) + apply(rule conjI) + apply simp + apply(simp add:lift_def) + apply clarify + apply(erule ctran.elims,simp_all) + apply clarify + apply(rule_tac x="xs" in exI) + apply simp + apply clarify +apply(rule_tac x="xs" in exI) +apply(simp add: last_length) +done + +lemma cptn_onlyif_cptn_mod_aux [rule_format]: + "\s Q t xs.((Some a, s), Q, t) \ ctran \ (Q, t) # xs \ cptn_mod + \ (Some a, s) # (Q, t) # xs \ cptn_mod" +apply(induct a) +apply simp_all +--{* basic *} +apply clarify +apply(erule ctran.elims,simp_all) +apply(rule CptnModNone,rule Basic,simp) +apply clarify +apply(erule ctran.elims,simp_all) +--{* Seq1 *} +apply(rule_tac xs="[(None,ta)]" in CptnModSeq2) + apply(erule CptnModNone) + apply(rule CptnModOne) + apply simp +apply simp +apply(simp add:lift_def) +--{* Seq2 *} +apply(erule_tac x=sa in allE) +apply(erule_tac x="Some P2" in allE) +apply(erule allE,erule impE, assumption) +apply(drule div_seq,simp) +apply force +apply clarify +apply(erule disjE) + apply clarify + apply(erule allE,erule impE, assumption) + apply(erule_tac CptnModSeq1) + apply(simp add:lift_def) +apply clarify +apply(erule allE,erule impE, assumption) +apply(erule_tac CptnModSeq2) + apply (simp add:last_length) + apply (simp add:last_length) +apply(simp add:lift_def) +--{* Cond *} +apply clarify +apply(erule ctran.elims,simp_all) +apply(force elim: CptnModCondT) +apply(force elim: CptnModCondF) +--{* While *} +apply clarify +apply(erule ctran.elims,simp_all) +apply(rule CptnModNone,erule WhileF,simp) +apply(drule div_seq,force) +apply clarify +apply (erule disjE) + apply(force elim:CptnModWhile1) +apply clarify +apply(force simp add:last_length elim:CptnModWhile2) +--{* await *} +apply clarify +apply(erule ctran.elims,simp_all) +apply(rule CptnModNone,erule Await,simp+) +done + +lemma cptn_onlyif_cptn_mod [rule_format]: "c \ cptn \ c \ cptn_mod" +apply(erule cptn.induct) + apply(rule CptnModOne) + apply(erule CptnModEnv) +apply(case_tac P) + apply simp + apply(erule ctran.elims,simp_all) +apply(force elim:cptn_onlyif_cptn_mod_aux) +done + +lemma lift_is_cptn: "c\cptn \ map (lift P) c \ cptn" +apply(erule cptn.induct) + apply(force simp add:lift_def CptnOne) + apply(force intro:CptnEnv simp add:lift_def) +apply(force simp add:lift_def intro:CptnComp Seq2 Seq1 elim:ctran.elims) +done + +lemma cptn_append_is_cptn [rule_format]: + "\b a. b#c1\cptn \ a#c2\cptn \ (b#c1)!length c1=a \ b#c1@c2\cptn" +apply(induct c1) + apply simp +apply clarify +apply(erule cptn.elims,simp_all) + apply(force intro:CptnEnv) +apply(force elim:CptnComp) +done + +lemma last_lift: "\xs\[]; fst(xs!(length xs - (Suc 0)))=None\ + \ fst((map (lift P) xs)!(length (map (lift P) xs)- (Suc 0)))=(Some P)" +apply(case_tac "(xs ! (length xs - (Suc 0)))") +apply (simp add:lift_def) +done + +lemma last_fst [rule_format]: "P((a#x)!length x) \ \P a \ P (x!(length x - (Suc 0)))" +apply(induct x,simp+) +done + +lemma last_fst_esp: + "fst(((Some a,s)#xs)!(length xs))=None \ fst(xs!(length xs - (Suc 0)))=None" +apply(erule last_fst) +apply simp +done + +lemma last_snd: "xs\[] \ + snd(((map (lift P) xs))!(length (map (lift P) xs) - (Suc 0)))=snd(xs!(length xs - (Suc 0)))" +apply(case_tac "(xs ! (length xs - (Suc 0)))",simp) +apply (simp add:lift_def) +done + +lemma Cons_lift: "(Some (Seq P Q), s) # (map (lift Q) xs) = map (lift Q) ((Some P, s) # xs)" +by(simp add:lift_def) + +lemma Cons_lift_append: + "(Some (Seq P Q), s) # (map (lift Q) xs) @ ys = map (lift Q) ((Some P, s) # xs)@ ys " +by(simp add:lift_def) + +lemma lift_nth: "i map (lift Q) xs ! i = lift Q (xs! i)" +by (simp add:lift_def) + +lemma snd_lift: "i< length xs \ snd(lift Q (xs ! i))= snd (xs ! i)" +apply(case_tac "xs!i") +apply(simp add:lift_def) +done + +lemma cptn_if_cptn_mod: "c \ cptn_mod \ c \ cptn" +apply(erule cptn_mod.induct) + apply(rule CptnOne) + apply(erule CptnEnv) + apply(erule CptnComp,simp) + apply(rule CptnComp) + apply(erule CondT,simp) + apply(rule CptnComp) + apply(erule CondF,simp) +--{* Seq1 *} +apply(erule cptn.elims,simp_all) + apply(rule CptnOne) + apply clarify + apply(drule_tac P=P1 in lift_is_cptn) + apply(simp add:lift_def) + apply(rule CptnEnv,simp) +apply clarify +apply(simp add:lift_def) +apply(rule conjI) + apply clarify + apply(rule CptnComp) + apply(rule Seq1,simp) + apply(drule_tac P=P1 in lift_is_cptn) + apply(simp add:lift_def) +apply clarify +apply(rule CptnComp) + apply(rule Seq2,simp) +apply(drule_tac P=P1 in lift_is_cptn) +apply(simp add:lift_def) +--{* Seq2 *} +apply(rule cptn_append_is_cptn) + apply(drule_tac P=P1 in lift_is_cptn) + apply(simp add:lift_def) + apply simp +apply(case_tac "xs\[]") + apply(drule_tac P=P1 in last_lift) + apply(rule last_fst_esp) + apply (simp add:last_length) + apply(simp add:Cons_lift del:map.simps) + apply(rule conjI, clarify, simp) + apply(case_tac "(((Some P0, s) # xs) ! length xs)") + apply clarify + apply (simp add:lift_def last_length) +apply (simp add:last_length) +--{* While1 *} +apply(rule CptnComp) +apply(rule WhileT,simp) +apply(drule_tac P="While b P" in lift_is_cptn) +apply(simp add:lift_def) +--{* While2 *} +apply(rule CptnComp) +apply(rule WhileT,simp) +apply(rule cptn_append_is_cptn) +apply(drule_tac P="While b P" in lift_is_cptn) + apply(simp add:lift_def) + apply simp +apply(case_tac "xs\[]") + apply(drule_tac P="While b P" in last_lift) + apply(rule last_fst_esp,simp add:last_length) + apply(simp add:Cons_lift del:map.simps) + apply(rule conjI, clarify, simp) + apply(case_tac "(((Some P, s) # xs) ! length xs)") + apply clarify + apply (simp add:last_length lift_def) +apply simp +done + +theorem cptn_iff_cptn_mod: "(c \ cptn) = (c \ cptn_mod)" +apply(rule iffI) + apply(erule cptn_onlyif_cptn_mod) +apply(erule cptn_if_cptn_mod) +done + +section {* Validity of Correctness Formulas*} + +subsection {* Validity for Component Programs. *} + +types 'a rgformula = "'a com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set" + +constdefs + assum :: "('a set \ ('a \ 'a) set) \ ('a confs) set" + "assum \ \(pre, rely). {c. snd(c!0) \ pre \ (\i. Suc i + c!i -e\ c!(Suc i) \ (snd(c!i), snd(c!Suc i)) \ rely)}" + + comm :: "(('a \ 'a) set \ 'a set) \ ('a confs) set" + "comm \ \(guar, post). {c. (\i. Suc i + c!i -c\ c!(Suc i) \ (snd(c!i), snd(c!Suc i)) \ guar) \ + (fst (last c) = None \ snd (last c) \ post)}" + + com_validity :: "'a com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set \ bool" + ("\ _ sat [_, _, _, _]" [60,0,0,0,0] 45) + "\ P sat [pre, rely, guar, post] \ + \s. cp (Some P) s \ assum(pre, rely) \ comm(guar, post)" + +subsection {* Validity for Parallel Programs. *} + +constdefs + All_None :: "(('a com) option) list \ bool" + "All_None xs \ \c\set xs. c=None" + + par_assum :: "('a set \ ('a \ 'a) set) \ ('a par_confs) set" + "par_assum \ \(pre, rely). {c. snd(c!0) \ pre \ (\i. Suc i + c!i -pe\ c!Suc i \ (snd(c!i), snd(c!Suc i)) \ rely)}" + + par_comm :: "(('a \ 'a) set \ 'a set) \ ('a par_confs) set" + "par_comm \ \(guar, post). {c. (\i. Suc i + c!i -pc\ c!Suc i \ (snd(c!i), snd(c!Suc i)) \ guar) \ + (All_None (fst (last c)) \ snd( last c) \ post)}" + + par_com_validity :: "'a par_com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set + \ bool" ("\ _ SAT [_, _, _, _]" [60,0,0,0,0] 45) + "\ Ps SAT [pre, rely, guar, post] \ + \s. par_cp Ps s \ par_assum(pre, rely) \ par_comm(guar, post)" + +subsection {* Compositionality of the Semantics *} + +subsubsection {* Definition of the conjoin operator *} + +constdefs + same_length :: "'a par_confs \ ('a confs) list \ bool" + "same_length c clist \ (\i ('a confs) list \ bool" + "same_state c clist \ (\i j ('a confs) list \ bool" + "same_program c clist \ (\jx. fst(nth x j)) clist)" + + compat_label :: "'a par_confs \ ('a confs) list \ bool" + "compat_label c clist \ (\j. Suc j + (c!j -pc\ c!Suc j \ (\i (clist!i)! Suc j \ + (\li \ (clist!l)!j -e\ (clist!l)! Suc j))) \ + (c!j -pe\ c!Suc j \ (\i (clist!i)! Suc j)))" + + conjoin :: "'a par_confs \ ('a confs) list \ bool" ("_ \ _" [65,65] 64) + "c \ clist \ (same_length c clist) \ (same_state c clist) \ (same_program c clist) \ (compat_label c clist)" + +subsubsection {* Some previous lemmas *} + +lemma list_eq_if [rule_format]: "\ys. xs=ys \ (length xs = length ys) \ (\i (\i ys!0=a; ys\[] \ \ ys=(a#(tl ys))" +apply(case_tac ys) + apply simp+ +done + +lemma nth_tl_if [rule_format]: "ys\[] \ ys!0=a \ P ys \ P (a#(tl ys))" +apply(induct ys) + apply simp+ +done + +lemma nth_tl_onlyif [rule_format]: "ys\[] \ ys!0=a \ P (a#(tl ys)) \ P ys" +apply(induct ys) + apply simp+ +done + +lemma seq_not_eq1: "Seq c1 c2\c1" +apply(rule com.induct) +apply simp_all +apply clarify +done + +lemma seq_not_eq2: "Seq c1 c2\c2" +apply(rule com.induct) +apply simp_all +apply clarify +done + +lemma if_not_eq1: "Cond b c1 c2 \c1" +apply(rule com.induct) +apply simp_all +apply clarify +done + +lemma if_not_eq2: "Cond b c1 c2\c2" +apply(rule com.induct) +apply simp_all +apply clarify +done + +lemmas seq_and_if_not_eq [simp] = seq_not_eq1 seq_not_eq2 +seq_not_eq1 [THEN not_sym] seq_not_eq2 [THEN not_sym] +if_not_eq1 if_not_eq2 if_not_eq1 [THEN not_sym] if_not_eq2 [THEN not_sym] + +lemma prog_not_eq_in_ctran_aux [rule_format]: "(P,s) -c\ (Q,t) \ (P\Q)" +apply(erule ctran.induct) +apply simp_all +done + +lemma prog_not_eq_in_ctran [simp]: "\ (P,s) -c\ (P,t)" +apply clarify +apply(drule prog_not_eq_in_ctran_aux) +apply simp +done + +lemma prog_not_eq_in_par_ctran_aux [rule_format]: "(P,s) -pc\ (Q,t) \ (P\Q)" +apply(erule par_ctran.induct) +apply(drule prog_not_eq_in_ctran_aux) +apply clarify +apply(drule list_eq_if) + apply simp_all +apply force +done + +lemma prog_not_eq_in_par_ctran [simp]: "\ (P,s) -pc\ (P,t)" +apply clarify +apply(drule prog_not_eq_in_par_ctran_aux) +apply simp +done + +lemma tl_in_cptn: "\ a#xs \cptn; xs\[] \ \ xs\cptn" +apply(force elim:cptn.elims) +done + +lemma tl_zero[rule_format]: "P (ys!Suc j) \ Suc j ys\[] \ P (tl(ys)!j)" +apply(induct ys) + apply simp_all +done + +subsection {* The Semantics is Compositional *} + +lemma aux_if [rule_format]: + "\xs s clist. (length clist = length xs \ (\i cptn) + \ ((xs, s)#ys \ map (\i. (fst i,s)#snd i) (zip xs clist)) + \ (xs, s)#ys \ par_cptn)" +apply(induct ys) + apply(clarify) + apply(rule ParCptnOne) +apply(clarify) +apply(simp add:conjoin_def compat_label_def) +apply clarify +apply(erule_tac x="0" and P="\j. ?H j \ (?P j \ ?Q j)" in all_dupE,simp) +apply(erule disjE) +--{* first step is a Component step *} + apply clarify + apply simp + apply(subgoal_tac "a=(xs[i:=(fst(clist!i!0))])") + apply(subgoal_tac "b=snd(clist!i!0)",simp) + prefer 2 + apply(simp add: same_state_def) + apply(erule_tac x=i in allE,erule impE,assumption, + erule_tac x=1 and P="\j. (?H j) \ (snd (?d j))=(snd (?e j))" in allE,simp) + prefer 2 + apply(simp add:same_program_def) + apply(erule_tac x=1 and P="\j. ?H j \ (fst (?s j))=(?t j)" in allE,simp) + apply(rule nth_equalityI,simp) + apply clarify + apply(case_tac "i=ia",simp,simp) + apply(erule_tac x=ia and P="\j. ?H j \ ?I j \ ?J j" in allE) + apply(drule_tac t=i in not_sym,simp) + apply(erule etran.elims,simp) + apply(rule ParCptnComp) + apply(erule ParComp,simp) +--{* applying the induction hypothesis *} + apply(erule_tac x="xs[i := fst (clist ! i ! 0)]" in allE) + apply(erule_tac x="snd (clist ! i ! 0)" in allE) + apply(erule mp) + apply(rule_tac x="map tl clist" in exI,simp) + apply(rule conjI,clarify) + apply(case_tac "i=ia",simp) + apply(rule nth_tl_if) + apply(force simp add:same_length_def length_Suc_conv) + apply simp + apply(erule allE,erule impE,assumption,erule tl_in_cptn) + apply(force simp add:same_length_def length_Suc_conv) + apply(rule nth_tl_if) + apply(force simp add:same_length_def length_Suc_conv) + apply(simp add:same_state_def) + apply(erule_tac x=ia in allE, erule impE, assumption, + erule_tac x=1 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) + apply(erule_tac x=ia and P="\j. ?H j \ ?I j \ ?J j" in allE) + apply(drule_tac t=i in not_sym,simp) + apply(erule etran.elims,simp) + apply(erule allE,erule impE,assumption,erule tl_in_cptn) + apply(force simp add:same_length_def length_Suc_conv) + apply(simp add:same_length_def same_state_def) + apply(rule conjI) + apply clarify + apply(case_tac j,simp,simp) + apply(erule_tac x=ia in allE, erule impE, assumption, + erule_tac x="Suc(Suc nat)" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) + apply(force simp add:same_length_def length_Suc_conv) + apply(rule conjI) + apply(simp add:same_program_def) + apply clarify + apply(case_tac j,simp) + apply(rule nth_equalityI,simp) + apply clarify + apply(case_tac "i=ia",simp,simp) + apply(erule_tac x="Suc(Suc nat)" and P="\j. ?H j \ (fst (?s j))=(?t j)" in allE,simp) + apply(rule nth_equalityI,simp,simp) + apply(force simp add:length_Suc_conv) + apply(rule allI,rule impI) + apply(erule_tac x="Suc j" and P="\j. ?H j \ (?I j \ ?J j)" in allE,simp) + apply(erule disjE) + apply clarify + apply(rule_tac x=ia in exI,simp) + apply(case_tac "i=ia",simp) + apply(rule conjI) + apply(force simp add: length_Suc_conv) + apply clarify + apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE,erule impE,assumption) + apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE,erule impE,assumption) + apply simp + apply(case_tac j,simp) + apply(rule tl_zero) + apply(drule_tac t=l in not_sym,simp) + apply(erule_tac x=l in allE, erule impE, assumption, + erule_tac x=1 and P="\j. (?H j) \ (snd (?d j))=(snd (?e j))" in allE,simp) + apply(force elim:etran.elims intro:Env) + apply force + apply force + apply simp + apply(rule tl_zero) + apply(erule tl_zero) + apply force + apply force + apply force + apply force + apply(rule conjI,simp) + apply(rule nth_tl_if) + apply force + apply(erule_tac x=ia in allE, erule impE, assumption, + erule_tac x=1 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) + apply(erule_tac x=ia and P="\j. ?H j \ ?I j \ ?J j" in allE) + apply(drule_tac t=i in not_sym,simp) + apply(erule etran.elims,simp) + apply(erule tl_zero) + apply force + apply force + apply clarify + apply(case_tac "i=l",simp) + apply(rule nth_tl_if) + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply simp + apply(erule_tac P="\j. ?H j \ ?I j \ ?J j" in allE,erule impE,assumption,erule impE,assumption) + apply(erule tl_zero,force) + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply(rule nth_tl_if) + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply(erule_tac x=l in allE, erule impE, assumption, + erule_tac x=1 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) + apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE,erule impE, assumption,simp) + apply(drule not_sym,erule impE,assumption ) + apply(erule etran.elims,simp) + apply(rule tl_zero) + apply force + apply force + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply(rule disjI2) + apply(case_tac j,simp) + apply clarify + apply(rule tl_zero) + apply(erule_tac x=ia and P="\j. ?H j \ ?I j\etran" in allE,erule impE, assumption) + apply(case_tac "i=ia",simp,simp) + apply(erule_tac x=ia in allE, erule impE, assumption, + erule_tac x=1 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) + apply(erule_tac x=ia and P="\j. ?H j \ ?I j \ ?J j" in allE,erule impE, assumption,simp) + apply(drule not_sym,erule impE,assumption) + apply(force elim:etran.elims intro:Env) + apply force + apply(erule_tac x=ia and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply simp + apply clarify + apply(rule tl_zero) + apply(rule tl_zero,force) + apply force + apply(erule_tac x=ia and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply force + apply(erule_tac x=ia and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) +--{* first step is an environmental step *} +apply clarify +apply(erule par_etran.elims) +apply simp +apply(rule ParCptnEnv) +apply(erule_tac x="Ps" in allE) +apply(erule_tac x="t" in allE) +apply(erule mp) +apply(rule_tac x="map tl clist" in exI,simp) +apply(rule conjI) + apply clarify + apply(erule_tac x=i and P="\j. ?H j \ (?I ?s j) \ cptn" in allE,simp) + apply(erule cptn.elims) + apply(simp add:same_length_def) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply(simp add:same_state_def) + apply(erule_tac x=i in allE, erule impE, assumption, + erule_tac x=1 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) + apply(erule_tac x=i and P="\j. ?H j \ ?J j \etran" in allE,simp) + apply(erule etran.elims,simp) +apply(simp add:same_state_def same_length_def) +apply(rule conjI,clarify) + apply(case_tac j,simp,simp) + apply(erule_tac x=i in allE, erule impE, assumption, + erule_tac x="Suc(Suc nat)" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) + apply(rule tl_zero) + apply(simp) + apply force + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) +apply(rule conjI) + apply(simp add:same_program_def) + apply clarify + apply(case_tac j,simp) + apply(rule nth_equalityI,simp) + apply clarify + apply simp + apply(erule_tac x="Suc(Suc nat)" and P="\j. ?H j \ (fst (?s j))=(?t j)" in allE,simp) + apply(rule nth_equalityI,simp,simp) + apply(force simp add:length_Suc_conv) +apply(rule allI,rule impI) +apply(erule_tac x="Suc j" and P="\j. ?H j \ (?I j \ ?J j)" in allE,simp) +apply(erule disjE) + apply clarify + apply(rule_tac x=i in exI,simp) + apply(rule conjI) + apply(erule_tac x=i and P="\i. ?H i \ ?J i \etran" in allE, erule impE, assumption) + apply(erule etran.elims,simp) + apply(erule_tac x=i in allE, erule impE, assumption, + erule_tac x=1 and P="\j. (?H j) \ (snd (?d j))=(snd (?e j))" in allE,simp) + apply(rule nth_tl_if) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply simp + apply(erule tl_zero,force) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply clarify + apply(erule_tac x=l and P="\i. ?H i \ ?J i \etran" in allE, erule impE, assumption) + apply(erule etran.elims,simp) + apply(erule_tac x=l in allE, erule impE, assumption, + erule_tac x=1 and P="\j. (?H j) \ (snd (?d j))=(snd (?e j))" in allE,simp) + apply(rule nth_tl_if) + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply simp + apply(rule tl_zero,force) + apply force + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) +apply(rule disjI2) +apply simp +apply clarify +apply(case_tac j,simp) + apply(rule tl_zero) + apply(erule_tac x=i and P="\i. ?H i \ ?J i \etran" in allE, erule impE, assumption) + apply(erule_tac x=i and P="\i. ?H i \ ?J i \etran" in allE, erule impE, assumption) + apply(force elim:etran.elims intro:Env) + apply force + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) +apply simp +apply(rule tl_zero) + apply(rule tl_zero,force) + apply force + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply force +apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) +done + +lemma less_Suc_0 [iff]: "(n < Suc 0) = (n = 0)" +by auto + +lemma aux_onlyif [rule_format]: "\xs s. (xs, s)#ys \ par_cptn \ + (\clist. (length clist = length xs) \ + (xs, s)#ys \ map (\i. (fst i,s)#(snd i)) (zip xs clist) \ + (\i cptn))" +apply(induct ys) + apply(clarify) + apply(rule_tac x="map (\i. []) [0..length xs(]" in exI) + apply(simp add: conjoin_def same_length_def same_state_def same_program_def compat_label_def) + apply(rule conjI) + apply(rule nth_equalityI,simp,simp) + apply(force intro: cptn.intros) +apply(clarify) +apply(erule par_cptn.elims,simp) + apply simp + apply(erule_tac x="xs" in allE) + apply(erule_tac x="t" in allE,simp) + apply clarify + apply(rule_tac x="(map (\j. (P!j, t)#(clist!j)) [0..length P(])" in exI,simp) + apply(rule conjI) + prefer 2 + apply clarify + apply(rule CptnEnv,simp) + apply(simp add:conjoin_def same_length_def same_state_def) + apply (rule conjI) + apply clarify + apply(case_tac j,simp,simp) + apply(rule conjI) + apply(simp add:same_program_def) + apply clarify + apply(case_tac j,simp) + apply(rule nth_equalityI,simp,simp) + apply simp + apply(rule nth_equalityI,simp,simp) + apply(simp add:compat_label_def) + apply clarify + apply(case_tac j,simp) + apply(simp add:ParEnv) + apply clarify + apply(simp add:Env) + apply simp + apply(erule_tac x=nat in allE,erule impE, assumption) + apply(erule disjE,simp) + apply clarify + apply(rule_tac x=i in exI,simp) + apply force +apply(erule par_ctran.elims,simp) +apply(erule_tac x="Ps[i:=r]" in allE) +apply(erule_tac x="ta" in allE,simp) +apply clarify +apply(rule_tac x="(map (\j. (Ps!j, ta)#(clist!j)) [0..length Ps(]) [i:=((r, ta)#(clist!i))]" in exI,simp) +apply(rule conjI) + prefer 2 + apply clarify + apply(case_tac "i=ia",simp) + apply(erule CptnComp) + apply(erule_tac x=ia and P="\j. ?H j \ (?I j \ cptn)" in allE,simp) + apply simp + apply(erule_tac x=ia in allE) + apply(rule CptnEnv,simp) +apply(simp add:conjoin_def) +apply (rule conjI) + apply(simp add:same_length_def) + apply clarify + apply(case_tac "i=ia",simp,simp) +apply(rule conjI) + apply(simp add:same_state_def) + apply clarify + apply(case_tac j,simp,simp) + apply(case_tac "i=ia",simp,simp) +apply(rule conjI) + apply(simp add:same_program_def) + apply clarify + apply(case_tac j,simp) + apply(rule nth_equalityI,simp,simp) + apply simp + apply(rule nth_equalityI,simp,simp) + apply(erule_tac x=nat and P="\j. ?H j \ (fst (?a j))=((?b j))" in allE) + apply(case_tac nat) + apply clarify + apply(case_tac "i=ia",simp,simp) + apply clarify + apply(case_tac "i=ia",simp,simp) +apply(simp add:compat_label_def) +apply clarify +apply(case_tac j) + apply(rule conjI,simp) + apply(erule ParComp,assumption) + apply clarify + apply(rule_tac x=i in exI,simp) + apply clarify + apply(drule not_sym,simp) + apply(rule Env) +apply simp +apply(erule_tac x=nat and P="\j. ?H j \ (?P j \ ?Q j)" in allE,simp) +apply(erule disjE) + apply clarify + apply(rule_tac x=ia in exI,simp) + apply(rule conjI) + apply(case_tac "i=ia",simp,simp) + apply(rotate_tac -1,simp) + apply clarify + apply(case_tac "i=l",simp) + apply(case_tac "l=ia",simp,simp) + apply(erule_tac x=l in allE,erule impE,assumption,erule impE, assumption,simp) + apply simp + apply(erule_tac x=l in allE,erule impE,assumption,erule impE, assumption,simp) +apply clarify +apply(erule_tac x=ia and P="\j. ?H j \ (?P j)\etran" in allE, erule impE, assumption) +apply(case_tac "i=ia",simp) +apply(rotate_tac -1,simp) +done + +lemma one_iff_aux: "xs\[] \ (\ys. ((xs, s)#ys \ par_cptn) = + (\clist. length clist= length xs \ + ((xs, s)#ys \ map (\i. (fst i,s)#(snd i)) (zip xs clist)) \ + (\i cptn))) = + (par_cp (xs) s = {c. \clist. (length clist)=(length xs) \ + (\i cp(xs!i) s) \ c \ clist})" +apply (rule iffI) + apply(rule subset_antisym) + apply(rule subsetI) + apply(clarify) + apply(simp add:par_cp_def cp_def) + apply(case_tac x) + apply(force elim:par_cptn.elims) + apply simp + apply(erule_tac x="list" in allE) + apply clarify + apply simp + apply(rule_tac x="map (\i. (fst i, s) # snd i) (zip xs clist)" in exI,simp) + apply(rule subsetI) + apply(clarify) + apply(case_tac x) + apply(erule_tac x=0 in allE) + apply(simp add:cp_def conjoin_def same_length_def same_program_def same_state_def compat_label_def) + apply clarify + apply(erule cptn.elims,force,force,force) + apply(simp add:par_cp_def conjoin_def same_length_def same_program_def same_state_def compat_label_def) + apply clarify + apply(erule_tac x=0 and P="\j. ?H j \ (length (?s j) = ?t)" in all_dupE) + apply(subgoal_tac "a = xs") + apply(subgoal_tac "b = s",simp) + prefer 3 + apply(erule_tac x=0 and P="\j. ?H j \ (fst (?s j))=((?t j))" in allE) + apply (simp add:cp_def) + apply(rule nth_equalityI,simp,simp) + prefer 2 + apply(erule_tac x=0 in allE) + apply (simp add:cp_def) + apply(erule_tac x=0 and P="\j. ?H j \ (\i. ?T i \ (snd (?d j i))=(snd (?e j i)))" in allE,simp) + apply(erule_tac x=0 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) + apply(erule_tac x=list in allE) + apply(rule_tac x="map tl clist" in exI,simp) + apply(rule conjI) + apply clarify + apply(case_tac j,simp) + apply(erule_tac x=i in allE, erule impE, assumption, + erule_tac x="0" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) + apply(erule_tac x=i in allE, erule impE, assumption, + erule_tac x="Suc nat" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply(rule conjI) + apply clarify + apply(rule nth_equalityI,simp,simp) + apply(case_tac j) + apply clarify + apply(erule_tac x=i in allE) + apply(simp add:cp_def) + apply clarify + apply simp + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply(thin_tac "?H = (\i. ?J i)") + apply(rule conjI) + apply clarify + apply(erule_tac x=j in allE,erule impE, assumption,erule disjE) + apply clarify + apply(rule_tac x=i in exI,simp) + apply(case_tac j,simp) + apply(rule conjI) + apply(erule_tac x=i in allE) + apply(simp add:cp_def) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply clarify + apply(erule_tac x=l in allE) + apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE) + apply clarify + apply(simp add:cp_def) + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!l",simp,simp) + apply simp + apply(rule conjI) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply clarify + apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE) + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!l",simp,simp) + apply clarify + apply(erule_tac x=i in allE) + apply(simp add:cp_def) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp) + apply(rule nth_tl_if,simp,simp) + apply(erule_tac x=i and P="\j. ?H j \ (?P j)\etran" in allE, erule impE, assumption,simp) + apply(simp add:cp_def) + apply clarify + apply(rule nth_tl_if) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply force + apply force +apply clarify +apply(rule iffI) + apply(simp add:par_cp_def) + apply(erule_tac c="(xs, s) # ys" in equalityCE) + apply simp + apply clarify + apply(rule_tac x="map tl clist" in exI) + apply simp + apply (rule conjI) + apply(simp add:conjoin_def cp_def) + apply(rule conjI) + apply clarify + apply(unfold same_length_def) + apply clarify + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,simp) + apply(rule conjI) + apply(simp add:same_state_def) + apply clarify + apply(erule_tac x=i in allE, erule impE, assumption, + erule_tac x=j and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) + apply(case_tac j,simp) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply(rule conjI) + apply(simp add:same_program_def) + apply clarify + apply(rule nth_equalityI,simp,simp) + apply(case_tac j,simp) + apply clarify + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply clarify + apply(simp add:compat_label_def) + apply(rule allI,rule impI) + apply(erule_tac x=j in allE,erule impE, assumption) + apply(erule disjE) + apply clarify + apply(rule_tac x=i in exI,simp) + apply(rule conjI) + apply(erule_tac x=i in allE) + apply(case_tac j,simp) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply clarify + apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE) + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!l",simp,simp) + apply(erule_tac x=l in allE,simp) + apply(rule disjI2) + apply clarify + apply(rule tl_zero) + apply(case_tac j,simp,simp) + apply(rule tl_zero,force) + apply force + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply force + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply clarify + apply(erule_tac x=i in allE) + apply(simp add:cp_def) + apply(rule nth_tl_if) + apply(simp add:conjoin_def) + apply clarify + apply(simp add:same_length_def) + apply(erule_tac x=i in allE,simp) + apply simp + apply simp + apply simp +apply clarify +apply(erule_tac c="(xs, s) # ys" in equalityCE) + apply(simp add:par_cp_def) +apply simp +apply(erule_tac x="map (\i. (fst i, s) # snd i) (zip xs clist)" in allE) +apply simp +apply clarify +apply(simp add:cp_def) +done + +theorem one: "xs\[] \ + par_cp xs s = {c. \clist. (length clist)=(length xs) \ + (\i cp(xs!i) s) \ c \ clist}" +apply(frule one_iff_aux) +apply(drule sym) +apply(erule iffD2) +apply clarify +apply(rule iffI) + apply(erule aux_onlyif) +apply clarify +apply(force intro:aux_if) +done + +end \ No newline at end of file diff -r 98f0a09a33c3 -r 791e3b4c4039 src/HOL/HoareParallel/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/HoareParallel/ROOT.ML Tue Mar 05 17:11:25 2002 +0100 @@ -0,0 +1,5 @@ + +time_use_thy "OG_Examples"; +time_use_thy "Gar_Coll"; +time_use_thy "Mul_Gar_Coll"; +time_use_thy "RG_Examples";