merged
authorboehmes
Mon, 21 Sep 2009 11:15:55 +0200
changeset 32623 d84b1b0077ae
parent 32622 8ed38c7bd21a (current diff)
parent 32621 a073cb249a06 (diff)
child 32625 f270520df7de
merged
src/HOL/HoareParallel/Gar_Coll.thy
src/HOL/HoareParallel/Graph.thy
src/HOL/HoareParallel/Mul_Gar_Coll.thy
src/HOL/HoareParallel/OG_Com.thy
src/HOL/HoareParallel/OG_Examples.thy
src/HOL/HoareParallel/OG_Hoare.thy
src/HOL/HoareParallel/OG_Syntax.thy
src/HOL/HoareParallel/OG_Tactics.thy
src/HOL/HoareParallel/OG_Tran.thy
src/HOL/HoareParallel/Quote_Antiquote.thy
src/HOL/HoareParallel/RG_Com.thy
src/HOL/HoareParallel/RG_Examples.thy
src/HOL/HoareParallel/RG_Hoare.thy
src/HOL/HoareParallel/RG_Syntax.thy
src/HOL/HoareParallel/RG_Tran.thy
src/HOL/HoareParallel/ROOT.ML
src/HOL/HoareParallel/document/root.bib
src/HOL/HoareParallel/document/root.tex
--- a/src/HOL/HoareParallel/Gar_Coll.thy	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,846 +0,0 @@
-
-header {* \section{The Single Mutator Case} *}
-
-theory Gar_Coll imports Graph OG_Syntax begin
-
-declare psubsetE [rule del]
-
-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 \<Rightarrow> bool"
-  "Mut_init \<equiv> \<guillemotleft> T \<in> Reach \<acute>E \<and> R < length \<acute>E \<and> T < length \<acute>M \<guillemotright>"
-
-text {* \noindent For the mutator we
-consider two modules, one for each action.  An auxiliary variable
-@{text "\<acute>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 \<equiv> .{\<acute>Mut_init \<and> \<acute>z}. \<langle>\<acute>E:=\<acute>E[R:=(fst(\<acute>E!R), T)],, \<acute>z:= (\<not>\<acute>z)\<rangle>"
-
-  Color_Target :: "gar_coll_state ann_com"
-  "Color_Target \<equiv> .{\<acute>Mut_init \<and> \<not>\<acute>z}. \<langle>\<acute>M:=\<acute>M[T:=Black],, \<acute>z:= (\<not>\<acute>z)\<rangle>"
-
-  Mutator :: "gar_coll_state ann_com"
-  "Mutator \<equiv>
-  .{\<acute>Mut_init \<and> \<acute>z}. 
-  WHILE True INV .{\<acute>Mut_init \<and> \<acute>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: 
-  "\<turnstile> Redirect_Edge pre(Color_Target)"
-apply (unfold mutator_defs)
-apply annhoare
-apply(simp_all)
-apply(force elim:Graph2)
-done
-
-lemma Color_Target:
-  "\<turnstile> Color_Target .{\<acute>Mut_init \<and> \<acute>z}."
-apply (unfold mutator_defs)
-apply annhoare
-apply(simp_all)
-done
-
-lemma Mutator: 
- "\<turnstile> 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 "\<acute>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 \<Rightarrow> bool"
-  "Proper_M_init \<equiv>  \<guillemotleft> Blacks M_init=Roots \<and> length M_init=length \<acute>M \<guillemotright>"
- 
-  Proper :: "gar_coll_state \<Rightarrow> bool"
-  "Proper \<equiv> \<guillemotleft> Proper_Roots \<acute>M \<and> Proper_Edges(\<acute>M, \<acute>E) \<and> \<acute>Proper_M_init \<guillemotright>"
-
-  Safe :: "gar_coll_state \<Rightarrow> bool"
-  "Safe \<equiv> \<guillemotleft> Reach \<acute>E \<subseteq> Blacks \<acute>M \<guillemotright>"
-
-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 \<equiv> 
-  .{\<acute>Proper}.
-  \<acute>ind:=0;;
-  .{\<acute>Proper \<and> \<acute>ind=0}.
-  WHILE \<acute>ind<length \<acute>M 
-   INV .{\<acute>Proper \<and> (\<forall>i<\<acute>ind. i \<in> Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind\<le>length \<acute>M}.
-  DO .{\<acute>Proper \<and> (\<forall>i<\<acute>ind. i \<in> Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M}.
-   IF \<acute>ind\<in>Roots THEN 
-   .{\<acute>Proper \<and> (\<forall>i<\<acute>ind. i \<in> Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M \<and> \<acute>ind\<in>Roots}. 
-    \<acute>M:=\<acute>M[\<acute>ind:=Black] FI;;
-   .{\<acute>Proper \<and> (\<forall>i<\<acute>ind+1. i \<in> Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M}.
-    \<acute>ind:=\<acute>ind+1 
-  OD"
-
-lemma Blacken_Roots: 
- "\<turnstile> Blacken_Roots .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>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 force
-apply force
-done
-
-subsubsection {* Propagating black *}
-
-constdefs
-  PBInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
-  "PBInv \<equiv> \<guillemotleft> \<lambda>ind. \<acute>obc < Blacks \<acute>M \<or> (\<forall>i <ind. \<not>BtoW (\<acute>E!i, \<acute>M) \<or>
-   (\<not>\<acute>z \<and> i=R \<and> (snd(\<acute>E!R)) = T \<and> (\<exists>r. ind \<le> r \<and> r < length \<acute>E \<and> BtoW(\<acute>E!r,\<acute>M))))\<guillemotright>"
-
-constdefs  
-  Propagate_Black_aux :: "gar_coll_state ann_com"
-  "Propagate_Black_aux \<equiv>
-  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M}.
-  \<acute>ind:=0;;
-  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M \<and> \<acute>ind=0}. 
-  WHILE \<acute>ind<length \<acute>E 
-   INV .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-         \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind\<le>length \<acute>E}.
-  DO .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-       \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E}. 
-   IF \<acute>M!(fst (\<acute>E!\<acute>ind)) = Black THEN 
-    .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-       \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E \<and> \<acute>M!fst(\<acute>E!\<acute>ind)=Black}.
-     \<acute>M:=\<acute>M[snd(\<acute>E!\<acute>ind):=Black];;
-    .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-       \<and> \<acute>PBInv (\<acute>ind + 1) \<and> \<acute>ind<length \<acute>E}.
-     \<acute>ind:=\<acute>ind+1
-   FI
-  OD"
-
-lemma Propagate_Black_aux: 
-  "\<turnstile>  Propagate_Black_aux
-  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-    \<and> ( \<acute>obc < Blacks \<acute>M \<or> \<acute>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_tac y = r in le_imp_less_or_eq)
-  apply (erule disjE)
-   apply (subgoal_tac "Suc (ind x)\<le>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 (asm_lr))
- 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 \<Rightarrow> bool"
-  "Auxk \<equiv> \<guillemotleft>\<acute>k<length \<acute>M \<and> (\<acute>M!\<acute>k\<noteq>Black \<or> \<not>BtoW(\<acute>E!\<acute>ind, \<acute>M) \<or> 
-          \<acute>obc<Blacks \<acute>M \<or> (\<not>\<acute>z \<and> \<acute>ind=R \<and> snd(\<acute>E!R)=T  
-          \<and> (\<exists>r. \<acute>ind<r \<and> r<length \<acute>E \<and> BtoW(\<acute>E!r, \<acute>M))))\<guillemotright>"
-
-constdefs  
-  Propagate_Black :: " gar_coll_state ann_com"
-  "Propagate_Black \<equiv>
-  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M}.
-  \<acute>ind:=0;;
-  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M \<and> \<acute>ind=0}.
-  WHILE \<acute>ind<length \<acute>E 
-   INV .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-         \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind\<le>length \<acute>E}.
-  DO .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-       \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E}. 
-   IF (\<acute>M!(fst (\<acute>E!\<acute>ind)))=Black THEN 
-    .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-      \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E \<and> (\<acute>M!fst(\<acute>E!\<acute>ind))=Black}.
-     \<acute>k:=(snd(\<acute>E!\<acute>ind));;
-    .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-      \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E \<and> (\<acute>M!fst(\<acute>E!\<acute>ind))=Black 
-      \<and> \<acute>Auxk}.
-     \<langle>\<acute>M:=\<acute>M[\<acute>k:=Black],, \<acute>ind:=\<acute>ind+1\<rangle>
-   ELSE .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-          \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E}. 
-         \<langle>IF (\<acute>M!(fst (\<acute>E!\<acute>ind)))\<noteq>Black THEN \<acute>ind:=\<acute>ind+1 FI\<rangle> 
-   FI
-  OD"
-
-lemma Propagate_Black: 
-  "\<turnstile>  Propagate_Black
-  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-    \<and> ( \<acute>obc < Blacks \<acute>M \<or> \<acute>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_tac y = r in le_imp_less_or_eq)
-   apply (erule disjE)
-    apply (subgoal_tac "Suc (ind x)\<le>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\<le>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_tac y = r in le_imp_less_or_eq)
- apply (erule disjE)
-  apply (subgoal_tac "Suc (ind x)\<le>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 (asm_lr))
- 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 \<Rightarrow> nat \<Rightarrow> bool"
-  "CountInv \<equiv> \<guillemotleft> \<lambda>ind. {i. i<ind \<and> \<acute>Ma!i=Black}\<subseteq>\<acute>bc \<guillemotright>"
-
-constdefs
-  Count :: " gar_coll_state ann_com"
-  "Count \<equiv>
-  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
-    \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-    \<and> length \<acute>Ma=length \<acute>M \<and> (\<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>bc={}}.
-  \<acute>ind:=0;;
-  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
-    \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-   \<and> length \<acute>Ma=length \<acute>M \<and> (\<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>bc={} 
-   \<and> \<acute>ind=0}.
-   WHILE \<acute>ind<length \<acute>M 
-     INV .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
-           \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-           \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>CountInv \<acute>ind
-           \<and> ( \<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>ind\<le>length \<acute>M}.
-   DO .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
-         \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-         \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>CountInv \<acute>ind 
-         \<and> ( \<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>ind<length \<acute>M}. 
-       IF \<acute>M!\<acute>ind=Black 
-          THEN .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
-                 \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-                 \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>CountInv \<acute>ind
-                 \<and> ( \<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>ind<length \<acute>M \<and> \<acute>M!\<acute>ind=Black}.
-          \<acute>bc:=insert \<acute>ind \<acute>bc
-       FI;;
-      .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
-        \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-        \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>CountInv (\<acute>ind+1)
-        \<and> ( \<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>ind<length \<acute>M}.
-      \<acute>ind:=\<acute>ind+1
-   OD"
-
-lemma Count: 
-  "\<turnstile> Count 
-  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
-   \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>\<acute>bc \<and> \<acute>bc\<subseteq>Blacks \<acute>M \<and> length \<acute>Ma=length \<acute>M
-   \<and> (\<acute>obc < Blacks \<acute>Ma \<or> \<acute>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 \<times> edges \<Rightarrow> edges"
-
-axioms
-  Append_to_free0: "length (Append_to_free (i, e)) = length e"
-  Append_to_free1: "Proper_Edges (m, e) 
-                   \<Longrightarrow> Proper_Edges (m, Append_to_free(i, e))"
-  Append_to_free2: "i \<notin> Reach e 
-     \<Longrightarrow> n \<in> Reach (Append_to_free(i, e)) = ( n = i \<or> n \<in> Reach e)"
-
-constdefs
-  AppendInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
-  "AppendInv \<equiv> \<guillemotleft>\<lambda>ind. \<forall>i<length \<acute>M. ind\<le>i \<longrightarrow> i\<in>Reach \<acute>E \<longrightarrow> \<acute>M!i=Black\<guillemotright>"
-
-constdefs
-  Append :: " gar_coll_state ann_com"
-   "Append \<equiv>
-  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe}.
-  \<acute>ind:=0;;
-  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe \<and> \<acute>ind=0}.
-    WHILE \<acute>ind<length \<acute>M 
-      INV .{\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind\<le>length \<acute>M}.
-    DO .{\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M}.
-       IF \<acute>M!\<acute>ind=Black THEN 
-          .{\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>M!\<acute>ind=Black}. 
-          \<acute>M:=\<acute>M[\<acute>ind:=White] 
-       ELSE .{\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>ind\<notin>Reach \<acute>E}.
-              \<acute>E:=Append_to_free(\<acute>ind,\<acute>E)
-       FI;;
-     .{\<acute>Proper \<and> \<acute>AppendInv (\<acute>ind+1) \<and> \<acute>ind<length \<acute>M}. 
-       \<acute>ind:=\<acute>ind+1
-    OD"
-
-lemma Append: 
-  "\<turnstile> Append .{\<acute>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 \<equiv>
-.{\<acute>Proper}.  
- WHILE True INV .{\<acute>Proper}. 
- DO  
-  Blacken_Roots;; 
-  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M}.  
-   \<acute>obc:={};; 
-  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={}}. 
-   \<acute>bc:=Roots;; 
-  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots}. 
-   \<acute>Ma:=M_init;;  
-  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots \<and> \<acute>Ma=M_init}. 
-   WHILE \<acute>obc\<noteq>\<acute>bc  
-     INV .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
-           \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>\<acute>bc \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-           \<and> length \<acute>Ma=length \<acute>M \<and> (\<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe)}. 
-   DO .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M}.
-       \<acute>obc:=\<acute>bc;;
-       Propagate_Black;; 
-      .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-        \<and> (\<acute>obc < Blacks \<acute>M \<or> \<acute>Safe)}. 
-       \<acute>Ma:=\<acute>M;;
-      .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma 
-        \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M \<and> length \<acute>Ma=length \<acute>M 
-        \<and> ( \<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe)}.
-       \<acute>bc:={};;
-       Count 
-   OD;; 
-  Append  
- OD"
-
-lemma Collector: 
-  "\<turnstile> 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)
-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+
---{* 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+
---{* 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 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(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+
-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 subgoal 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 @{claset}) *})
-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 @{clasimpset}, assume_tac]) *})
-apply(tactic {* TRYALL(EVERY'[rtac disjI2,etac subset_trans,rtac @{thm Graph9},force_tac @{clasimpset}]) *})
-apply(tactic {* TRYALL(EVERY'[rtac disjI1,etac @{thm psubset_subset_trans},rtac @{thm Graph9},force_tac @{clasimpset}]) *})
-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 @{claset}) *})
---{* 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: 
-  "\<parallel>- .{\<acute>Proper \<and> \<acute>Mut_init \<and> \<acute>z}.  
-  COBEGIN  
-   Collector
-  .{False}.
- \<parallel>  
-   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
--- a/src/HOL/HoareParallel/Graph.thy	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,410 +0,0 @@
-header {* \chapter{Case Study: Single and Multi-Mutator Garbage Collection Algorithms}
-
-\section {Formalization of the Memory} *}
-
-theory Graph imports Main begin
-
-datatype node = Black | White
-
-types 
-  nodes = "node list"
-  edge  = "nat \<times> nat"
-  edges = "edge list"
-
-consts Roots :: "nat set"
-
-constdefs
-  Proper_Roots :: "nodes \<Rightarrow> bool"
-  "Proper_Roots M \<equiv> Roots\<noteq>{} \<and> Roots \<subseteq> {i. i<length M}"
-
-  Proper_Edges :: "(nodes \<times> edges) \<Rightarrow> bool"
-  "Proper_Edges \<equiv> (\<lambda>(M,E). \<forall>i<length E. fst(E!i)<length M \<and> snd(E!i)<length M)"
-
-  BtoW :: "(edge \<times> nodes) \<Rightarrow> bool"
-  "BtoW \<equiv> (\<lambda>(e,M). (M!fst e)=Black \<and> (M!snd e)\<noteq>Black)"
-
-  Blacks :: "nodes \<Rightarrow> nat set"
-  "Blacks M \<equiv> {i. i<length M \<and> M!i=Black}"
-
-  Reach :: "edges \<Rightarrow> nat set"
-  "Reach E \<equiv> {x. (\<exists>path. 1<length path \<and> path!(length path - 1)\<in>Roots \<and> x=path!0
-              \<and> (\<forall>i<length path - 1. (\<exists>j<length E. E!j=(path!(i+1), path!i))))
-	      \<or> x\<in>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]: 
-  "\<lbrakk> Roots\<subseteq>Blacks M; \<forall>i<length E. \<not>BtoW(E!i,M)\<rbrakk>
-  \<Longrightarrow> 1< length path \<longrightarrow> (path!(length path - 1))\<in>Roots \<longrightarrow>  
-  (\<forall>i<length path - 1. (\<exists>j. j < length E \<and> E!j=(path!(Suc i), path!i))) 
-  \<longrightarrow> 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 -2)
-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: 
-  "\<lbrakk>Roots\<subseteq>Blacks M; Proper_Edges(M, E); \<forall>i<length E. \<not>BtoW(E!i,M) \<rbrakk> 
-  \<Longrightarrow> Reach E\<subseteq>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) \<longrightarrow> (\<exists>m. P m \<and> (\<forall>i. i<m \<longrightarrow> \<not> P i))";
-apply(rule nat_less_induct)
-apply clarify
-apply(case_tac "\<forall>m. m<n \<longrightarrow> \<not> P m")
-apply auto
-done
-
-lemma Compl_lemma: "(n::nat)\<le>l \<Longrightarrow> (\<exists>m. m\<le>l \<and> n=l - m)"
-apply(rule_tac x = "l - n" in exI)
-apply arith
-done
-
-lemma Ex_last_occurrence: 
-  "\<lbrakk>P (n::nat); n\<le>l\<rbrakk> \<Longrightarrow> (\<exists>m. P (l - m) \<and> (\<forall>i. i<m \<longrightarrow> \<not>P (l - i)))"
-apply(drule Compl_lemma)
-apply clarify
-apply(erule Ex_first_occurrence)
-done
-
-lemma Graph2: 
-  "\<lbrakk>T \<in> Reach E; R<length E\<rbrakk> \<Longrightarrow> T \<in> Reach (E[R:=(fst(E!R), T)])"
-apply (unfold Reach_def)
-apply clarify
-apply simp
-apply(case_tac "\<forall>z<length path. fst(E!R)\<noteq>path!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 (force simp add:nth_list_update)
-apply simp
-apply(erule exE)
-apply(subgoal_tac "z \<le> length path - Suc 0")
- prefer 2 apply arith
-apply(drule_tac P = "\<lambda>m. m<length path \<and> 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 \<le> 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 \<le> length path")
- prefer 2 apply arith
-apply(subgoal_tac "(length path - Suc m) + (Suc nata) \<le> 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(simp del: diff_is_0_eq)
- apply(subgoal_tac "Suc nata\<le>nat")
- prefer 2 apply arith
- apply(drule_tac n = "Suc nata" in Compl_lemma)
- apply clarify
- using [[linarith_split_limit = 0]]
- apply force
- using [[linarith_split_limit = 9]]
-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 simp
-done
-
-
-subsubsection{* Graph 3 *}
-
-lemma Graph3: 
-  "\<lbrakk> T\<in>Reach E; R<length E \<rbrakk> \<Longrightarrow> Reach(E[R:=(fst(E!R),T)]) \<subseteq> Reach E"
-apply (unfold Reach_def)
-apply clarify
-apply simp
-apply(case_tac "\<exists>i<length path - 1. (fst(E!R),T)=(path!(Suc i),path!i)")
---{* the changed edge is part of the path *}
- apply(erule exE)
- apply(drule_tac P = "\<lambda>i. i<length path - 1 \<and> (fst(E!R),T)=(path!Suc i,path!i)" in Ex_first_occurrence)
- apply clarify
- apply(erule disjE)
---{* T is NOT a root *}
-  apply clarify
-  apply(rule_tac x = "(take m path)@patha" in exI)
-  apply(subgoal_tac "\<not>(length path\<le>m)")
-   prefer 2 apply arith
-  apply(simp)
-  apply(rule conjI)
-   apply(subgoal_tac "\<not>(m + length patha - 1 < m)")
-    prefer 2 apply arith
-   apply(simp add: nth_append)
-  apply(rule conjI)
-   apply(case_tac "m")
-    apply force
-   apply(case_tac "path")
-    apply force
-   apply force
-  apply clarify
-  apply(case_tac "Suc i\<le>m")
-   apply(erule_tac x = "i" in allE)
-   apply simp
-   apply clarify
-   apply(rule_tac x = "j" in exI)
-   apply(case_tac "Suc i<m")
-    apply(simp add: nth_append)
-    apply(case_tac "R=j")
-     apply(simp add: nth_list_update)
-     apply(case_tac "i=m")
-      apply force
-     apply(erule_tac x = "i" in allE)
-     apply force
-    apply(force simp add: nth_list_update)
-   apply(simp add: nth_append)
-   apply(subgoal_tac "i=m - 1")
-    prefer 2 apply arith
-   apply(case_tac "R=j")
-    apply(erule_tac x = "m - 1" in allE)
-    apply(simp add: nth_list_update)
-   apply(force simp add: nth_list_update)
-  apply(simp add: nth_append)
-  apply(rotate_tac -4)
-  apply(erule_tac x = "i - m" in allE)
-  apply(subgoal_tac "Suc (i - m)=(Suc i - m)" )
-    prefer 2 apply arith
-   apply simp
---{* T is a root *}
- apply(case_tac "m=0")
-  apply force
- apply(rule_tac x = "take (Suc m) path" in exI)
- apply(subgoal_tac "\<not>(length path\<le>Suc m)" )
-  prefer 2 apply arith
- apply clarsimp
- 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: 
-  "\<lbrakk>T \<in> Reach E; Roots\<subseteq>Blacks M; I\<le>length E; T<length M; R<length E; 
-  \<forall>i<I. \<not>BtoW(E!i,M); R<I; M!fst(E!R)=Black; M!T\<noteq>Black\<rbrakk> \<Longrightarrow> 
-  (\<exists>r. I\<le>r \<and> r<length E \<and> 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 "\<exists>m<length path. M!(path!m)=Black")
- apply(erule exE)
- apply(drule_tac P = "\<lambda>m. m<length path \<and> M!(path!m)=Black" in Ex_first_occurrence)
- apply clarify
- apply(case_tac "ma")
-  apply force
- apply simp
- apply(case_tac "length path")
-  apply force
- apply simp
- apply(erule_tac P = "\<lambda>i. i < nata \<longrightarrow> ?P i" and x = "nat" in allE)
- apply simp
- apply clarify
- apply(erule_tac P = "\<lambda>i. i < Suc nat \<longrightarrow> ?P i" and x = "nat" in allE)
- apply simp
- apply(case_tac "j<I")
-  apply(erule_tac x = "j" in allE)
-  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 5 *}
-
-lemma Graph5: 
-  "\<lbrakk> T \<in> Reach E ; Roots \<subseteq> Blacks M; \<forall>i<R. \<not>BtoW(E!i,M); T<length M; 
-    R<length E; M!fst(E!R)=Black; M!snd(E!R)=Black; M!T \<noteq> Black\<rbrakk> 
-   \<Longrightarrow> (\<exists>r. R<r \<and> r<length E \<and> 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 "\<exists>m<length path. M!(path!m)=Black")
- apply(erule exE)
- apply(drule_tac P = "\<lambda>m. m<length path \<and> M!(path!m)=Black" in Ex_first_occurrence)
- apply clarify
- apply(case_tac "ma")
-  apply force
- apply simp
- apply(case_tac "length path")
-  apply force
- apply simp
- apply(erule_tac P = "\<lambda>i. i < nata \<longrightarrow> ?P i" and x = "nat" in allE)
- apply simp
- apply clarify
- apply(erule_tac P = "\<lambda>i. i < Suc nat \<longrightarrow> ?P i" and x = "nat" in allE)
- apply simp
- apply(case_tac "j\<le>R")
-  apply(drule le_imp_less_or_eq [of _ R])
-  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 {* Other lemmas about graphs *}
-
-lemma Graph6: 
- "\<lbrakk>Proper_Edges(M,E); R<length E ; T<length M\<rbrakk> \<Longrightarrow> Proper_Edges(M,E[R:=(fst(E!R),T)])"
-apply (unfold Proper_Edges_def)
- apply(force  simp add: nth_list_update)
-done
-
-lemma Graph7: 
- "\<lbrakk>Proper_Edges(M,E)\<rbrakk> \<Longrightarrow> Proper_Edges(M[T:=a],E)"
-apply (unfold Proper_Edges_def)
-apply force
-done
-
-lemma Graph8: 
- "\<lbrakk>Proper_Roots(M)\<rbrakk> \<Longrightarrow> 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<length M \<Longrightarrow> Blacks M\<subseteq>Blacks (M[j := Black])"
-apply (unfold Blacks_def)
- apply(force simp add: nth_list_update)
-done
-
-lemma Graph10 [rule_format (no_asm)]: "\<forall>i. M!i=a \<longrightarrow>M[i:=a]=M"
-apply(induct_tac "M")
-apply auto
-apply(case_tac "i")
-apply auto
-done
-
-lemma Graph11 [rule_format (no_asm)]: 
-  "\<lbrakk> M!j\<noteq>Black;j<length M\<rbrakk> \<Longrightarrow> Blacks M \<subset> 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: "\<lbrakk>a\<subseteq>Blacks M;j<length M\<rbrakk> \<Longrightarrow> a\<subseteq>Blacks (M[j := Black])"
-apply (unfold Blacks_def)
-apply(force simp add: nth_list_update)
-done
-
-lemma Graph13: "\<lbrakk>a\<subset> Blacks M;j<length M\<rbrakk> \<Longrightarrow> a \<subset> 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
--- a/src/HOL/HoareParallel/Mul_Gar_Coll.thy	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1283 +0,0 @@
-
-header {* \section{The Multi-Mutator Case} *}
-
-theory Mul_Gar_Coll imports Graph OG_Syntax begin
-
-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 \<Rightarrow> nat \<Rightarrow> bool"
-  "Mul_mut_init \<equiv> \<guillemotleft> \<lambda>n. n=length \<acute>Muts \<and> (\<forall>i<n. R (\<acute>Muts!i)<length \<acute>E 
-                          \<and> T (\<acute>Muts!i)<length \<acute>M) \<guillemotright>"
-
-  Mul_Redirect_Edge  :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com"
-  "Mul_Redirect_Edge j n \<equiv>
-  .{\<acute>Mul_mut_init n \<and> Z (\<acute>Muts!j)}.
-  \<langle>IF T(\<acute>Muts!j) \<in> Reach \<acute>E THEN  
-  \<acute>E:= \<acute>E[R (\<acute>Muts!j):= (fst (\<acute>E!R(\<acute>Muts!j)), T (\<acute>Muts!j))] FI,, 
-  \<acute>Muts:= \<acute>Muts[j:= (\<acute>Muts!j) \<lparr>Z:=False\<rparr>]\<rangle>"
-
-  Mul_Color_Target :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com"
-  "Mul_Color_Target j n \<equiv>
-  .{\<acute>Mul_mut_init n \<and> \<not> Z (\<acute>Muts!j)}. 
-  \<langle>\<acute>M:=\<acute>M[T (\<acute>Muts!j):=Black],, \<acute>Muts:=\<acute>Muts[j:= (\<acute>Muts!j) \<lparr>Z:=True\<rparr>]\<rangle>"
-
-  Mul_Mutator :: "nat \<Rightarrow> nat \<Rightarrow>  mul_gar_coll_state ann_com"
-  "Mul_Mutator j n \<equiv>
-  .{\<acute>Mul_mut_init n \<and> Z (\<acute>Muts!j)}.  
-  WHILE True  
-    INV .{\<acute>Mul_mut_init n \<and> Z (\<acute>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\<le>j \<and> j<n \<Longrightarrow> 
-  \<turnstile> 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\<le>j \<and> j<n \<Longrightarrow> 
-  \<turnstile>  Mul_Color_Target j n  
-    .{\<acute>Mul_mut_init n \<and> Z (\<acute>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\<le>j \<and> j<n \<Longrightarrow>  
- \<turnstile> 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: 
-  "\<lbrakk>0\<le>i; i<n; 0\<le>j; j<n; i\<noteq>j\<rbrakk> \<Longrightarrow>  
-  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)
-done
-
-lemma Mul_interfree_Redirect_Edge_Color_Target: 
-  "\<lbrakk>0\<le>i; i<n; 0\<le>j; j<n; i\<noteq>j\<rbrakk> \<Longrightarrow>  
-  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: 
-  "\<lbrakk>0\<le>i; i<n; 0\<le>j; j<n; i\<noteq>j\<rbrakk> \<Longrightarrow> 
-  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)
-done
-
-lemma Mul_interfree_Color_Target_Color_Target: 
-  " \<lbrakk>0\<le>i; i<n; 0\<le>j; j<n; i\<noteq>j\<rbrakk> \<Longrightarrow> 
-  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)
-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: "\<lbrakk>i < n; j < n; i \<noteq> j\<rbrakk> \<Longrightarrow> 
-  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 @{claset}) *})
-apply (simp_all add:nth_list_update)
-done
-
-subsubsection {* Modular Parameterized Mutators *}
-
-lemma Mul_Parameterized_Mutators: "0<n \<Longrightarrow>
- \<parallel>- .{\<acute>Mul_mut_init n \<and> (\<forall>i<n. Z (\<acute>Muts!i))}.
- COBEGIN
- SCHEME  [0\<le> 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 \<Rightarrow> nat"
- "Queue \<equiv> \<guillemotleft> length (filter (\<lambda>i. \<not> Z i \<and> \<acute>M!(T i) \<noteq> Black) \<acute>Muts) \<guillemotright>"
-
-consts  M_init :: nodes
-
-constdefs
-  Proper_M_init :: "mul_gar_coll_state \<Rightarrow> bool"
-  "Proper_M_init \<equiv> \<guillemotleft> Blacks M_init=Roots \<and> length M_init=length \<acute>M \<guillemotright>"
-
-  Mul_Proper :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
-  "Mul_Proper \<equiv> \<guillemotleft> \<lambda>n. Proper_Roots \<acute>M \<and> Proper_Edges (\<acute>M, \<acute>E) \<and> \<acute>Proper_M_init \<and> n=length \<acute>Muts \<guillemotright>"
-
-  Safe :: "mul_gar_coll_state \<Rightarrow> bool"
-  "Safe \<equiv> \<guillemotleft> Reach \<acute>E \<subseteq> Blacks \<acute>M \<guillemotright>"
-
-lemmas mul_collector_defs = Proper_M_init_def Mul_Proper_def Safe_def
-
-subsubsection {* Blackening Roots *}
-
-constdefs
-  Mul_Blacken_Roots :: "nat \<Rightarrow>  mul_gar_coll_state ann_com"
-  "Mul_Blacken_Roots n \<equiv>
-  .{\<acute>Mul_Proper n}.
-  \<acute>ind:=0;;
-  .{\<acute>Mul_Proper n \<and> \<acute>ind=0}.
-  WHILE \<acute>ind<length \<acute>M 
-    INV .{\<acute>Mul_Proper n \<and> (\<forall>i<\<acute>ind. i\<in>Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind\<le>length \<acute>M}.
-  DO .{\<acute>Mul_Proper n \<and> (\<forall>i<\<acute>ind. i\<in>Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M}.
-       IF \<acute>ind\<in>Roots THEN 
-     .{\<acute>Mul_Proper n \<and> (\<forall>i<\<acute>ind. i\<in>Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M \<and> \<acute>ind\<in>Roots}. 
-       \<acute>M:=\<acute>M[\<acute>ind:=Black] FI;;
-     .{\<acute>Mul_Proper n \<and> (\<forall>i<\<acute>ind+1. i\<in>Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M}.
-       \<acute>ind:=\<acute>ind+1 
-  OD"
-
-lemma Mul_Blacken_Roots: 
-  "\<turnstile> Mul_Blacken_Roots n  
-  .{\<acute>Mul_Proper n \<and> Roots \<subseteq> Blacks \<acute>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 force
-apply force
-done
-
-subsubsection {* Propagating Black *} 
-
-constdefs
-  Mul_PBInv :: "mul_gar_coll_state \<Rightarrow> bool"
-  "Mul_PBInv \<equiv>  \<guillemotleft>\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>M \<or> \<acute>l<\<acute>Queue 
-                 \<or> (\<forall>i<\<acute>ind. \<not>BtoW(\<acute>E!i,\<acute>M)) \<and> \<acute>l\<le>\<acute>Queue\<guillemotright>"
-
-  Mul_Auxk :: "mul_gar_coll_state \<Rightarrow> bool"
-  "Mul_Auxk \<equiv> \<guillemotleft>\<acute>l<\<acute>Queue \<or> \<acute>M!\<acute>k\<noteq>Black \<or> \<not>BtoW(\<acute>E!\<acute>ind, \<acute>M) \<or> \<acute>obc\<subset>Blacks \<acute>M\<guillemotright>"
-
-constdefs
-  Mul_Propagate_Black :: "nat \<Rightarrow>  mul_gar_coll_state ann_com"
-  "Mul_Propagate_Black n \<equiv>
- .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-  \<and> (\<acute>Safe \<or> \<acute>l\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M)}. 
- \<acute>ind:=0;;
- .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-   \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> Blacks \<acute>M\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-   \<and> (\<acute>Safe \<or> \<acute>l\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M) \<and> \<acute>ind=0}. 
- WHILE \<acute>ind<length \<acute>E 
-  INV .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-        \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-        \<and> \<acute>Mul_PBInv \<and> \<acute>ind\<le>length \<acute>E}.
- DO .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-     \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-     \<and> \<acute>Mul_PBInv \<and> \<acute>ind<length \<acute>E}.
-   IF \<acute>M!(fst (\<acute>E!\<acute>ind))=Black THEN 
-   .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-     \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-     \<and> \<acute>Mul_PBInv \<and> (\<acute>M!fst(\<acute>E!\<acute>ind))=Black \<and> \<acute>ind<length \<acute>E}.
-    \<acute>k:=snd(\<acute>E!\<acute>ind);;
-   .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-     \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-     \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>M \<or> \<acute>l<\<acute>Queue \<or> (\<forall>i<\<acute>ind. \<not>BtoW(\<acute>E!i,\<acute>M)) 
-        \<and> \<acute>l\<le>\<acute>Queue \<and> \<acute>Mul_Auxk ) \<and> \<acute>k<length \<acute>M \<and> \<acute>M!fst(\<acute>E!\<acute>ind)=Black 
-     \<and> \<acute>ind<length \<acute>E}.
-   \<langle>\<acute>M:=\<acute>M[\<acute>k:=Black],,\<acute>ind:=\<acute>ind+1\<rangle>
-   ELSE .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-         \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-         \<and> \<acute>Mul_PBInv \<and> \<acute>ind<length \<acute>E}.
-	 \<langle>IF \<acute>M!(fst (\<acute>E!\<acute>ind))\<noteq>Black THEN \<acute>ind:=\<acute>ind+1 FI\<rangle> FI
- OD"
-
-lemma Mul_Propagate_Black: 
-  "\<turnstile> Mul_Propagate_Black n  
-   .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-     \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>M \<or> \<acute>l<\<acute>Queue \<and> (\<acute>l\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>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 subgoal 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 \<Rightarrow> nat \<Rightarrow> bool"
- "Mul_CountInv \<equiv> \<guillemotleft> \<lambda>ind. {i. i<ind \<and> \<acute>Ma!i=Black}\<subseteq>\<acute>bc \<guillemotright>"
-
-  Mul_Count :: "nat \<Rightarrow>  mul_gar_coll_state ann_com"
-  "Mul_Count n \<equiv> 
-  .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-    \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-    \<and> length \<acute>Ma=length \<acute>M 
-    \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M) ) 
-    \<and> \<acute>q<n+1 \<and> \<acute>bc={}}.
-  \<acute>ind:=0;;
-  .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-    \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-    \<and> length \<acute>Ma=length \<acute>M 
-    \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M) ) 
-    \<and> \<acute>q<n+1 \<and> \<acute>bc={} \<and> \<acute>ind=0}.
-  WHILE \<acute>ind<length \<acute>M 
-     INV .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-          \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M  
-          \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>Mul_CountInv \<acute>ind 
-          \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M))
-	  \<and> \<acute>q<n+1 \<and> \<acute>ind\<le>length \<acute>M}.
-  DO .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-       \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-       \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>Mul_CountInv \<acute>ind 
-       \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M))
-       \<and> \<acute>q<n+1 \<and> \<acute>ind<length \<acute>M}. 
-     IF \<acute>M!\<acute>ind=Black 
-     THEN .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-            \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M  
-            \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>Mul_CountInv \<acute>ind 
-            \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M))
-            \<and> \<acute>q<n+1 \<and> \<acute>ind<length \<acute>M \<and> \<acute>M!\<acute>ind=Black}.
-          \<acute>bc:=insert \<acute>ind \<acute>bc
-     FI;;
-  .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-    \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-    \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>Mul_CountInv (\<acute>ind+1) 
-    \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M))
-    \<and> \<acute>q<n+1 \<and> \<acute>ind<length \<acute>M}.
-  \<acute>ind:=\<acute>ind+1
-  OD"
- 
-lemma Mul_Count: 
-  "\<turnstile> Mul_Count n  
-  .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-    \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-    \<and> length \<acute>Ma=length \<acute>M \<and> Blacks \<acute>Ma\<subseteq>\<acute>bc 
-    \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M)) 
-    \<and> \<acute>q<n+1}."
-apply (unfold Mul_Count_def)
-apply annhoare
-apply(simp_all add:Mul_CountInv_def mul_collector_defs Mul_Auxk_def Graph6 Graph7 Graph8 Graph12 mul_collector_defs Queue_def)
---{* 7 subgoals left *}
-apply force
-apply force
-apply force
---{* 4 subgoals left *}
-apply clarify
-apply(conjI_tac)
-apply(disjE_tac)
- apply simp_all
-apply(simp add:Blacks_def)
-apply clarify
-apply(erule less_SucE)
- back
- apply force
-apply force
---{* 3 subgoals left *}
-apply clarify
-apply(conjI_tac)
-apply(disjE_tac)
- apply simp_all
-apply clarify
-apply(erule less_SucE)
- back
- apply force
-apply simp
-apply(rotate_tac -1)
-apply (force simp add:Blacks_def)
---{* 2 subgoals left *}
-apply force
---{* 1 subgoal left *}
-apply clarify
-apply(drule_tac x = "ind x" in le_imp_less_or_eq)
-apply (simp_all add:Blacks_def)
-done
-
-subsubsection {* Appending garbage nodes to the free list *}
-
-consts  Append_to_free :: "nat \<times> edges \<Rightarrow> edges"
-
-axioms
-  Append_to_free0: "length (Append_to_free (i, e)) = length e"
-  Append_to_free1: "Proper_Edges (m, e) 
-                    \<Longrightarrow> Proper_Edges (m, Append_to_free(i, e))"
-  Append_to_free2: "i \<notin> Reach e 
-           \<Longrightarrow> n \<in> Reach (Append_to_free(i, e)) = ( n = i \<or> n \<in> Reach e)"
-
-constdefs
-  Mul_AppendInv :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
-  "Mul_AppendInv \<equiv> \<guillemotleft> \<lambda>ind. (\<forall>i. ind\<le>i \<longrightarrow> i<length \<acute>M \<longrightarrow> i\<in>Reach \<acute>E \<longrightarrow> \<acute>M!i=Black)\<guillemotright>"
-
-  Mul_Append :: "nat \<Rightarrow>  mul_gar_coll_state ann_com"
-  "Mul_Append n \<equiv> 
-  .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe}.
-  \<acute>ind:=0;;
-  .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe \<and> \<acute>ind=0}.
-  WHILE \<acute>ind<length \<acute>M 
-    INV .{\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv \<acute>ind \<and> \<acute>ind\<le>length \<acute>M}.
-  DO .{\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M}.
-      IF \<acute>M!\<acute>ind=Black THEN 
-     .{\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>M!\<acute>ind=Black}. 
-      \<acute>M:=\<acute>M[\<acute>ind:=White] 
-      ELSE 
-     .{\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>ind\<notin>Reach \<acute>E}. 
-      \<acute>E:=Append_to_free(\<acute>ind,\<acute>E)
-      FI;;
-  .{\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv (\<acute>ind+1) \<and> \<acute>ind<length \<acute>M}. 
-   \<acute>ind:=\<acute>ind+1
-  OD"
-
-lemma Mul_Append: 
-  "\<turnstile> Mul_Append n  
-     .{\<acute>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 \<Rightarrow>  mul_gar_coll_state ann_com"
-  "Mul_Collector n \<equiv>
-.{\<acute>Mul_Proper n}.  
-WHILE True INV .{\<acute>Mul_Proper n}. 
-DO  
-Mul_Blacken_Roots n ;; 
-.{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M}.  
- \<acute>obc:={};; 
-.{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={}}.  
- \<acute>bc:=Roots;; 
-.{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots}. 
- \<acute>l:=0;; 
-.{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots \<and> \<acute>l=0}. 
- WHILE \<acute>l<n+1  
-   INV .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M \<and>  
-         (\<acute>Safe \<or> (\<acute>l\<le>\<acute>Queue \<or> \<acute>bc\<subset>Blacks \<acute>M) \<and> \<acute>l<n+1)}. 
- DO .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-      \<and> (\<acute>Safe \<or> \<acute>l\<le>\<acute>Queue \<or> \<acute>bc\<subset>Blacks \<acute>M)}.
-    \<acute>obc:=\<acute>bc;;
-    Mul_Propagate_Black n;; 
-    .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-      \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-      \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>M \<or> \<acute>l<\<acute>Queue 
-      \<and> (\<acute>l\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M))}. 
-    \<acute>bc:={};;
-    .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-      \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-      \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>M \<or> \<acute>l<\<acute>Queue 
-      \<and> (\<acute>l\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M)) \<and> \<acute>bc={}}. 
-       \<langle> \<acute>Ma:=\<acute>M,, \<acute>q:=\<acute>Queue \<rangle>;;
-    Mul_Count n;; 
-    .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-      \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-      \<and> length \<acute>Ma=length \<acute>M \<and> Blacks \<acute>Ma\<subseteq>\<acute>bc 
-      \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M)) 
-      \<and> \<acute>q<n+1}. 
-    IF \<acute>obc=\<acute>bc THEN
-    .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-      \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-      \<and> length \<acute>Ma=length \<acute>M \<and> Blacks \<acute>Ma\<subseteq>\<acute>bc 
-      \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M)) 
-      \<and> \<acute>q<n+1 \<and> \<acute>obc=\<acute>bc}.  
-    \<acute>l:=\<acute>l+1  
-    ELSE .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
-          \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
-          \<and> length \<acute>Ma=length \<acute>M \<and> Blacks \<acute>Ma\<subseteq>\<acute>bc 
-          \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M)) 
-          \<and> \<acute>q<n+1 \<and> \<acute>obc\<noteq>\<acute>bc}.  
-        \<acute>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:
-  "\<turnstile> 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)
-apply force
-apply (force dest:subset_antisym)
-apply force
-apply force
-apply force
-done
-
-subsection {* Interference Freedom *}
-
-lemma le_length_filter_update[rule_format]: 
- "\<forall>i. (\<not>P (list!i) \<or> P j) \<and> i<length list 
- \<longrightarrow> length(filter P list) \<le> 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]: 
- "\<forall>i. P j \<and> \<not>(P (list!i)) \<and> i<length list 
- \<longrightarrow> 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: "\<lbrakk>0\<le>j; j<n\<rbrakk> \<Longrightarrow>  
-  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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow> 
-  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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
-  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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
-  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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
-  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 subgoal 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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
-  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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
-  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 subgoal 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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
-  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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
-  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 subgoal left *}
-apply(simp add:mul_mutator_defs nth_list_update)
-done
-
-lemma Mul_interfree_Redirect_Edge_Count: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
-  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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
-  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 subgoal 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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
-  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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
-  interfree_aux (Some(Mul_Append n),{}, Some(Mul_Redirect_Edge j n))"
-apply (unfold mul_modules)
-apply interfree_aux
-apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
-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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
-  interfree_aux (Some(Mul_Redirect_Edge j n),{},Some(Mul_Append n))"
-apply (unfold mul_modules)
-apply interfree_aux
-apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
-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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
-  interfree_aux (Some(Mul_Append n),{}, Some(Mul_Color_Target j n))"
-apply (unfold mul_modules)
-apply interfree_aux
-apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
-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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
-  interfree_aux (Some(Mul_Color_Target j n),{}, Some(Mul_Append n))"
-apply (unfold mul_modules)
-apply interfree_aux
-apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
-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<n  \<Longrightarrow> 
-  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 @{claset}) *})
-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 @{clasimpset}, 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(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 (thm "subset_psubset_trans"),etac (thm "Graph11"),force_tac @{clasimpset}]) *})
---{* 41 subgoals left *}
-apply(tactic {* TRYALL(EVERY'[rtac disjI2, rtac disjI1, etac @{thm le_trans}, force_tac (@{claset},@{simpset} addsimps [@{thm Queue_def}, @{thm less_Suc_eq_le}, @{thm le_length_filter_update}])]) *})
---{* 35 subgoals left *}
-apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac disjI1,etac (thm "psubset_subset_trans"),rtac (thm "Graph9"),force_tac @{clasimpset}]) *})
---{* 31 subgoals left *}
-apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac disjI1,etac (thm "subset_psubset_trans"),etac (thm "Graph11"),force_tac @{clasimpset}]) *})
---{* 29 subgoals left *}
-apply(tactic {* TRYALL(EVERY'[REPEAT o (rtac disjI2),etac (thm "subset_psubset_trans"),etac (thm "subset_psubset_trans"),etac (thm "Graph11"),force_tac @{clasimpset}]) *})
---{* 25 subgoals left *}
-apply(tactic {* TRYALL(EVERY'[rtac disjI2, rtac disjI2, rtac disjI1, etac @{thm le_trans}, force_tac (@{claset},@{simpset} addsimps [@{thm Queue_def}, @{thm 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 \<Longrightarrow> 
-  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: 
- "\<parallel>- .{\<acute>Mul_Proper n \<and> \<acute>Mul_mut_init n \<and> (\<forall>i<n. Z (\<acute>Muts!i))}.  
- COBEGIN  
-  Mul_Collector n
- .{False}.
- \<parallel>  
- SCHEME  [0\<le> 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 i)
- 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
--- a/src/HOL/HoareParallel/OG_Com.thy	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,55 +0,0 @@
-
-header {* \chapter{The Owicki-Gries Method} 
-
-\section{Abstract Syntax} *} 
-
-theory OG_Com imports Main begin
-
-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 \<Rightarrow> '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 \<times> 'a assn) list"
-   | Basic "('a \<Rightarrow> '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 \<Rightarrow> '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 \<Rightarrow> bool"
-primrec  
-  "atom_com (Parallel Ts) = False"
-  "atom_com (Basic f) = True"
-  "atom_com (Seq c1 c2) = (atom_com c1 \<and> atom_com c2)"
-  "atom_com (Cond b c1 c2) = (atom_com c1 \<and> atom_com c2)"
-  "atom_com (While b i c) = atom_com c"
-  
-end
\ No newline at end of file
--- a/src/HOL/HoareParallel/OG_Examples.thy	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,549 +0,0 @@
-
-header {* \section{Examples} *}
-
-theory OG_Examples imports OG_Syntax begin
-
-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: 
-  "\<parallel>- .{\<acute>pr1=0 \<and> \<not>\<acute>in1 \<and> \<acute>pr2=0 \<and> \<not>\<acute>in2 }.  
-  COBEGIN .{\<acute>pr1=0 \<and> \<not>\<acute>in1}.  
-  WHILE True INV .{\<acute>pr1=0 \<and> \<not>\<acute>in1}.  
-  DO  
-  .{\<acute>pr1=0 \<and> \<not>\<acute>in1}. \<langle> \<acute>in1:=True,,\<acute>pr1:=1 \<rangle>;;  
-  .{\<acute>pr1=1 \<and> \<acute>in1}.  \<langle> \<acute>hold:=1,,\<acute>pr1:=2 \<rangle>;;  
-  .{\<acute>pr1=2 \<and> \<acute>in1 \<and> (\<acute>hold=1 \<or> \<acute>hold=2 \<and> \<acute>pr2=2)}.  
-  AWAIT (\<not>\<acute>in2 \<or> \<not>(\<acute>hold=1)) THEN \<acute>pr1:=3 END;;    
-  .{\<acute>pr1=3 \<and> \<acute>in1 \<and> (\<acute>hold=1 \<or> \<acute>hold=2 \<and> \<acute>pr2=2)}. 
-   \<langle>\<acute>in1:=False,,\<acute>pr1:=0\<rangle> 
-  OD .{\<acute>pr1=0 \<and> \<not>\<acute>in1}.  
-  \<parallel>  
-  .{\<acute>pr2=0 \<and> \<not>\<acute>in2}.  
-  WHILE True INV .{\<acute>pr2=0 \<and> \<not>\<acute>in2}.  
-  DO  
-  .{\<acute>pr2=0 \<and> \<not>\<acute>in2}. \<langle> \<acute>in2:=True,,\<acute>pr2:=1 \<rangle>;;  
-  .{\<acute>pr2=1 \<and> \<acute>in2}. \<langle>  \<acute>hold:=2,,\<acute>pr2:=2 \<rangle>;;  
-  .{\<acute>pr2=2 \<and> \<acute>in2 \<and> (\<acute>hold=2 \<or> (\<acute>hold=1 \<and> \<acute>pr1=2))}.  
-  AWAIT (\<not>\<acute>in1 \<or> \<not>(\<acute>hold=2)) THEN \<acute>pr2:=3  END;;    
-  .{\<acute>pr2=3 \<and> \<acute>in2 \<and> (\<acute>hold=2 \<or> (\<acute>hold=1 \<and> \<acute>pr1=2))}. 
-    \<langle>\<acute>in2:=False,,\<acute>pr2:=0\<rangle> 
-  OD .{\<acute>pr2=0 \<and> \<not>\<acute>in2}.  
-  COEND  
-  .{\<acute>pr1=0 \<and> \<not>\<acute>in1 \<and> \<acute>pr2=0 \<and> \<not>\<acute>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: 
- "\<parallel>-  .{True}.  
-  \<acute>flag1:=False,, \<acute>flag2:=False,,  
-  COBEGIN .{\<not>\<acute>flag1}.  
-        WHILE True  
-        INV .{\<not>\<acute>flag1}.  
-        DO .{\<not>\<acute>flag1}. \<langle> \<acute>flag1:=True,,\<acute>after1:=False \<rangle>;;  
-           .{\<acute>flag1 \<and> \<not>\<acute>after1}. \<langle> \<acute>turn:=1,,\<acute>after1:=True \<rangle>;;  
-           .{\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}.  
-            WHILE \<not>(\<acute>flag2 \<longrightarrow> \<acute>turn=2)  
-            INV .{\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}.  
-            DO .{\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}. SKIP OD;; 
-           .{\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>flag2 \<and> \<acute>after2 \<longrightarrow> \<acute>turn=2)}.
-            \<acute>flag1:=False  
-        OD  
-       .{False}.  
-  \<parallel>  
-     .{\<not>\<acute>flag2}.  
-        WHILE True  
-        INV .{\<not>\<acute>flag2}.  
-        DO .{\<not>\<acute>flag2}. \<langle> \<acute>flag2:=True,,\<acute>after2:=False \<rangle>;;  
-           .{\<acute>flag2 \<and> \<not>\<acute>after2}. \<langle> \<acute>turn:=2,,\<acute>after2:=True \<rangle>;;  
-           .{\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}.  
-            WHILE \<not>(\<acute>flag1 \<longrightarrow> \<acute>turn=1)  
-            INV .{\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}.  
-            DO .{\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}. SKIP OD;;  
-           .{\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>flag1 \<and> \<acute>after1 \<longrightarrow> \<acute>turn=1)}. 
-            \<acute>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: 
- "\<parallel>- .{i\<noteq>j}.  
-  \<acute>out:=True ,,  
-  COBEGIN .{i\<noteq>j}.  
-       WHILE True INV .{i\<noteq>j}.  
-       DO .{i\<noteq>j}. AWAIT \<acute>out THEN  \<acute>out:=False,, \<acute>who:=i END;;  
-          .{\<not>\<acute>out \<and> \<acute>who=i \<and> i\<noteq>j}. \<acute>out:=True OD  
-       .{False}.  
-  \<parallel>  
-       .{i\<noteq>j}.  
-       WHILE True INV .{i\<noteq>j}.  
-       DO .{i\<noteq>j}. AWAIT \<acute>out THEN  \<acute>out:=False,,\<acute>who:=j END;;  
-          .{\<not>\<acute>out \<and> \<acute>who=j \<and> i\<noteq>j}. \<acute>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<n \<Longrightarrow> \<parallel>- .{True}.  
-  \<acute>out:=True ,,  
- COBEGIN
-  SCHEME [0\<le> i< n]
-    .{True}.  
-     WHILE True INV .{True}.  
-      DO .{True}. AWAIT \<acute>out THEN  \<acute>out:=False,, \<acute>who:=i END;;  
-         .{\<not>\<acute>out \<and> \<acute>who=i}. \<acute>out:=True OD
-    .{False}. 
- COEND
-  .{False}." 
-apply oghoare
---{* 20 vc *}
-apply auto
-done
-
-subsubsection{* The Ticket Algorithm *}
-
-record Ticket_mutex =
- num :: nat
- nextv :: nat
- turn :: "nat list"
- index :: nat 
-
-lemma Ticket_mutex: 
- "\<lbrakk> 0<n; I=\<guillemotleft>n=length \<acute>turn \<and> 0<\<acute>nextv \<and> (\<forall>k l. k<n \<and> l<n \<and> k\<noteq>l 
-    \<longrightarrow> \<acute>turn!k < \<acute>num \<and> (\<acute>turn!k =0 \<or> \<acute>turn!k\<noteq>\<acute>turn!l))\<guillemotright> \<rbrakk>
-   \<Longrightarrow> \<parallel>- .{n=length \<acute>turn}.  
-   \<acute>index:= 0,,
-   WHILE \<acute>index < n INV .{n=length \<acute>turn \<and> (\<forall>i<\<acute>index. \<acute>turn!i=0)}. 
-    DO \<acute>turn:= \<acute>turn[\<acute>index:=0],, \<acute>index:=\<acute>index +1 OD,,
-  \<acute>num:=1 ,, \<acute>nextv:=1 ,, 
- COBEGIN
-  SCHEME [0\<le> i< n]
-    .{\<acute>I}.  
-     WHILE True INV .{\<acute>I}.  
-      DO .{\<acute>I}. \<langle> \<acute>turn :=\<acute>turn[i:=\<acute>num],, \<acute>num:=\<acute>num+1 \<rangle>;;  
-         .{\<acute>I}. WAIT \<acute>turn!i=\<acute>nextv END;;
-         .{\<acute>I \<and> \<acute>turn!i=\<acute>nextv}. \<acute>nextv:=\<acute>nextv+1
-      OD
-    .{False}. 
- COEND
-  .{False}." 
-apply oghoare
---{* 35 vc *}
-apply simp_all
---{* 21 vc *}
-apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
---{* 11 vc *}
-apply simp_all
-apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
---{* 10 subgoals left *}
-apply(erule less_SucE)
- apply simp
-apply simp
---{* 9 subgoals left *}
-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 fastsimp
---{* 5 subgoals left *}
-prefer 5
-apply(case_tac [!] "j=k")
---{* 10 subgoals left *}
-apply simp_all
-apply(erule_tac x=k in allE)
-apply force
---{* 9 subgoals left *}
-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
---{* 5 subgoals left *}
-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
---{* 3 subgoals left *}
-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
---{* 1 subgoals left *}
-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: 
-  "\<lbrakk>I1= \<guillemotleft> a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> (a<\<acute>x \<and> f(\<acute>x)=0) \<or> (\<acute>y\<le>a \<and> f(\<acute>y)=0)) 
-      \<and> (\<not>\<acute>found \<and> a<\<acute> x \<longrightarrow> f(\<acute>x)\<noteq>0) \<guillemotright> ;  
-    I2= \<guillemotleft>\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> (a<\<acute>x \<and> f(\<acute>x)=0) \<or> (\<acute>y\<le>a \<and> f(\<acute>y)=0)) 
-      \<and> (\<not>\<acute>found \<and> \<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0) \<guillemotright> \<rbrakk> \<Longrightarrow>  
-  \<parallel>- .{\<exists> u. f(u)=0}.  
-  \<acute>turn:=1,, \<acute>found:= False,,  
-  \<acute>x:=a,, \<acute>y:=a+1 ,,  
-  COBEGIN .{\<acute>I1}.  
-       WHILE \<not>\<acute>found  
-       INV .{\<acute>I1}.  
-       DO .{a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> \<acute>y\<le>a \<and> f(\<acute>y)=0) \<and> (a<\<acute>x \<longrightarrow> f(\<acute>x)\<noteq>0)}.  
-          WAIT \<acute>turn=1 END;;  
-          .{a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> \<acute>y\<le>a \<and> f(\<acute>y)=0) \<and> (a<\<acute>x \<longrightarrow> f(\<acute>x)\<noteq>0)}.  
-          \<acute>turn:=2;;  
-          .{a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> \<acute>y\<le>a \<and> f(\<acute>y)=0) \<and> (a<\<acute>x \<longrightarrow> f(\<acute>x)\<noteq>0)}.    
-          \<langle> \<acute>x:=\<acute>x+1,,  
-            IF f(\<acute>x)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>  
-       OD;;  
-       .{\<acute>I1  \<and> \<acute>found}.  
-       \<acute>turn:=2  
-       .{\<acute>I1 \<and> \<acute>found}.  
-  \<parallel>  
-      .{\<acute>I2}.  
-       WHILE \<not>\<acute>found  
-       INV .{\<acute>I2}.  
-       DO .{\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> a<\<acute>x \<and> f(\<acute>x)=0) \<and> (\<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0)}.  
-          WAIT \<acute>turn=2 END;;  
-          .{\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> a<\<acute>x \<and> f(\<acute>x)=0) \<and> (\<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0)}.  
-          \<acute>turn:=1;;  
-          .{\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> a<\<acute>x \<and> f(\<acute>x)=0) \<and> (\<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0)}.  
-          \<langle> \<acute>y:=(\<acute>y - 1),,  
-            IF f(\<acute>y)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>  
-       OD;;  
-       .{\<acute>I2 \<and> \<acute>found}.  
-       \<acute>turn:=1  
-       .{\<acute>I2 \<and> \<acute>found}.  
-  COEND  
-  .{f(\<acute>x)=0 \<or> f(\<acute>y)=0}."
-apply oghoare
---{* 98 verification conditions *}
-apply auto 
---{* auto takes about 3 minutes !! *}
-done
-
-text {* Easier Version: without AWAIT.  Apt and Olderog. page 256: *}
-
-lemma Zero_Search_2: 
-"\<lbrakk>I1=\<guillemotleft> a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> (a<\<acute>x \<and> f(\<acute>x)=0) \<or> (\<acute>y\<le>a \<and> f(\<acute>y)=0)) 
-    \<and> (\<not>\<acute>found \<and> a<\<acute>x \<longrightarrow> f(\<acute>x)\<noteq>0)\<guillemotright>;  
- I2= \<guillemotleft>\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> (a<\<acute>x \<and> f(\<acute>x)=0) \<or> (\<acute>y\<le>a \<and> f(\<acute>y)=0)) 
-    \<and> (\<not>\<acute>found \<and> \<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0)\<guillemotright>\<rbrakk> \<Longrightarrow>  
-  \<parallel>- .{\<exists>u. f(u)=0}.  
-  \<acute>found:= False,,  
-  \<acute>x:=a,, \<acute>y:=a+1,,  
-  COBEGIN .{\<acute>I1}.  
-       WHILE \<not>\<acute>found  
-       INV .{\<acute>I1}.  
-       DO .{a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> \<acute>y\<le>a \<and> f(\<acute>y)=0) \<and> (a<\<acute>x \<longrightarrow> f(\<acute>x)\<noteq>0)}.  
-          \<langle> \<acute>x:=\<acute>x+1,,IF f(\<acute>x)=0 THEN  \<acute>found:=True ELSE  SKIP FI\<rangle>  
-       OD  
-       .{\<acute>I1 \<and> \<acute>found}.  
-  \<parallel>  
-      .{\<acute>I2}.  
-       WHILE \<not>\<acute>found  
-       INV .{\<acute>I2}.  
-       DO .{\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> a<\<acute>x \<and> f(\<acute>x)=0) \<and> (\<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0)}.  
-          \<langle> \<acute>y:=(\<acute>y - 1),,IF f(\<acute>y)=0 THEN  \<acute>found:=True ELSE  SKIP FI\<rangle>  
-       OD  
-       .{\<acute>I2 \<and> \<acute>found}.  
-  COEND  
-  .{f(\<acute>x)=0 \<or> f(\<acute>y)=0}."
-apply oghoare
---{* 20 vc *}
-apply auto
---{* auto takes aprox. 2 minutes. *}
-done
-
-subsection {* Producer/Consumer *}
-
-subsubsection {* Previous lemmas *}
-
-lemma nat_lemma2: "\<lbrakk> b = m*(n::nat) + t; a = s*n + u; t=u; b-a < n \<rbrakk> \<Longrightarrow> m \<le> s"
-proof -
-  assume "b = m*(n::nat) + t" "a = s*n + u" "t=u"
-  hence "(m - s) * n = b - a" by (simp add: diff_mult_distrib)
-  also assume "\<dots> < n"
-  finally have "m - s < 1" by simp
-  thus ?thesis by arith
-qed
-
-lemma mod_lemma: "\<lbrakk> (c::nat) \<le> a; a < b; b - c < n \<rbrakk> \<Longrightarrow> b mod n \<noteq> 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(subgoal_tac "b - a \<le> b - c")
- prefer 2 apply arith
-apply(drule le_less_trans)
-back
- apply assumption
-apply(frule less_not_refl2)
-apply(drule less_imp_le)
-apply (drule_tac m = "a" and k = n in div_le_mono)
-apply(safe)
-apply(frule_tac b = "b" and a = "a" and n = "n" in nat_lemma2, assumption, assumption)
-apply assumption
-apply(drule order_antisym, assumption)
-apply(rotate_tac -3)
-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: 
-  "\<lbrakk>INIT= \<guillemotleft>0<length a \<and> 0<length \<acute>buffer \<and> length \<acute>b=length a\<guillemotright> ;  
-    I= \<guillemotleft>(\<forall>k<\<acute>ins. \<acute>outs\<le>k \<longrightarrow> (a ! k) = \<acute>buffer ! (k mod (length \<acute>buffer))) \<and>  
-            \<acute>outs\<le>\<acute>ins \<and> \<acute>ins-\<acute>outs\<le>length \<acute>buffer\<guillemotright> ;  
-    I1= \<guillemotleft>\<acute>I \<and> \<acute>li\<le>length a\<guillemotright> ;  
-    p1= \<guillemotleft>\<acute>I1 \<and> \<acute>li=\<acute>ins\<guillemotright> ;  
-    I2 = \<guillemotleft>\<acute>I \<and> (\<forall>k<\<acute>lj. (a ! k)=(\<acute>b ! k)) \<and> \<acute>lj\<le>length a\<guillemotright> ;
-    p2 = \<guillemotleft>\<acute>I2 \<and> \<acute>lj=\<acute>outs\<guillemotright> \<rbrakk> \<Longrightarrow>   
-  \<parallel>- .{\<acute>INIT}.  
- \<acute>ins:=0,, \<acute>outs:=0,, \<acute>li:=0,, \<acute>lj:=0,,
- COBEGIN .{\<acute>p1 \<and> \<acute>INIT}. 
-   WHILE \<acute>li <length a 
-     INV .{\<acute>p1 \<and> \<acute>INIT}.   
-   DO .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a}.  
-       \<acute>vx:= (a ! \<acute>li);;  
-      .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a \<and> \<acute>vx=(a ! \<acute>li)}. 
-        WAIT \<acute>ins-\<acute>outs < length \<acute>buffer END;; 
-      .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a \<and> \<acute>vx=(a ! \<acute>li) 
-         \<and> \<acute>ins-\<acute>outs < length \<acute>buffer}. 
-       \<acute>buffer:=(list_update \<acute>buffer (\<acute>ins mod (length \<acute>buffer)) \<acute>vx);; 
-      .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a 
-         \<and> (a ! \<acute>li)=(\<acute>buffer ! (\<acute>ins mod (length \<acute>buffer))) 
-         \<and> \<acute>ins-\<acute>outs <length \<acute>buffer}.  
-       \<acute>ins:=\<acute>ins+1;; 
-      .{\<acute>I1 \<and> \<acute>INIT \<and> (\<acute>li+1)=\<acute>ins \<and> \<acute>li<length a}.  
-       \<acute>li:=\<acute>li+1  
-   OD  
-  .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li=length a}.  
-  \<parallel>  
-  .{\<acute>p2 \<and> \<acute>INIT}.  
-   WHILE \<acute>lj < length a  
-     INV .{\<acute>p2 \<and> \<acute>INIT}.  
-   DO .{\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>INIT}.  
-        WAIT \<acute>outs<\<acute>ins END;; 
-      .{\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>outs<\<acute>ins \<and> \<acute>INIT}.  
-       \<acute>vy:=(\<acute>buffer ! (\<acute>outs mod (length \<acute>buffer)));; 
-      .{\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>outs<\<acute>ins \<and> \<acute>vy=(a ! \<acute>lj) \<and> \<acute>INIT}.  
-       \<acute>outs:=\<acute>outs+1;;  
-      .{\<acute>I2 \<and> (\<acute>lj+1)=\<acute>outs \<and> \<acute>lj<length a \<and> \<acute>vy=(a ! \<acute>lj) \<and> \<acute>INIT}.  
-       \<acute>b:=(list_update \<acute>b \<acute>lj \<acute>vy);; 
-      .{\<acute>I2 \<and> (\<acute>lj+1)=\<acute>outs \<and> \<acute>lj<length a \<and> (a ! \<acute>lj)=(\<acute>b ! \<acute>lj) \<and> \<acute>INIT}.  
-       \<acute>lj:=\<acute>lj+1  
-   OD  
-  .{\<acute>p2 \<and> \<acute>lj=length a \<and> \<acute>INIT}.  
- COEND  
- .{ \<forall>k<length a. (a ! k)=(\<acute>b ! k)}."
-apply oghoare
---{* 138 vc  *}
-apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
---{* 112 subgoals left *}
-apply(simp_all (no_asm))
-apply(tactic {*ALLGOALS (conjI_Tac (K all_tac)) *})
---{* 930 subgoals left *}
-apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
-apply(simp_all (asm_lr) only:length_0_conv [THEN sym])
---{* 44 subgoals left *}
-apply (simp_all (asm_lr) del:length_0_conv add: neq0_conv nth_list_update mod_less_divisor mod_lemma)
---{* 32 subgoals left *}
-apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
-
-apply(tactic {* TRYALL (Lin_Arith.tac @{context}) *})
---{* 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 Example1 =
-  a :: "nat \<Rightarrow> nat"
-
-lemma Example1: 
- "\<parallel>- .{True}.
-   COBEGIN SCHEME [0\<le>i<n] .{True}. \<acute>a:=\<acute>a (i:=0) .{\<acute>a i=0}. COEND 
-  .{\<forall>i < n. \<acute>a i = 0}."
-apply oghoare
-apply simp_all
-done
-
-text {* Same example with lists as auxiliary variables. *}
-record Example1_list =
-  A :: "nat list"
-lemma Example1_list: 
- "\<parallel>- .{n < length \<acute>A}. 
-   COBEGIN 
-     SCHEME [0\<le>i<n] .{n < length \<acute>A}. \<acute>A:=\<acute>A[i:=0] .{\<acute>A!i=0}. 
-   COEND 
-    .{\<forall>i < n. \<acute>A!i = 0}."
-apply oghoare
-apply force+
-done
-
-subsubsection {* Increment a Variable in Parallel *}
-
-text {* First some lemmas about summation properties. *}
-(*
-lemma Example2_lemma1: "!!b. j<n \<Longrightarrow> (\<Sum>i::nat<n. b i) = (0::nat) \<Longrightarrow> b j = 0 "
-apply(induct n)
- apply simp_all
-apply(force simp add: less_Suc_eq)
-done
-*)
-lemma Example2_lemma2_aux: "!!b. j<n \<Longrightarrow> 
- (\<Sum>i=0..<n. (b i::nat)) =
- (\<Sum>i=0..<j. b i) + b j + (\<Sum>i=0..<n-(Suc j) . b (Suc j + i))"
-apply(induct n)
- apply simp_all
-apply(simp add:less_Suc_eq)
- apply(auto)
-apply(subgoal_tac "n - j = Suc(n- Suc j)")
-  apply simp
-apply arith
-done
-
-lemma Example2_lemma2_aux2: 
-  "!!b. j\<le> s \<Longrightarrow> (\<Sum>i::nat=0..<j. (b (s:=t)) i) = (\<Sum>i=0..<j. b i)"
-apply(induct j) 
- apply simp_all
-done
-
-lemma Example2_lemma2: 
- "!!b. \<lbrakk>j<n; b j=0\<rbrakk> \<Longrightarrow> Suc (\<Sum>i::nat=0..<n. b i)=(\<Sum>i=0..<n. (b (j := Suc 0)) i)"
-apply(frule_tac b="(b (j:=(Suc 0)))" in Example2_lemma2_aux)
-apply(erule_tac  t="setsum (b(j := (Suc 0))) {0..<n}" in ssubst)
-apply(frule_tac b=b in Example2_lemma2_aux)
-apply(erule_tac  t="setsum b {0..<n}" in ssubst)
-apply(subgoal_tac "Suc (setsum b {0..<j} + b j + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))=(setsum b {0..<j} + Suc (b j) + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))")
-apply(rotate_tac -1)
-apply(erule ssubst)
-apply(subgoal_tac "j\<le>j")
- apply(drule_tac b="b" and t="(Suc 0)" in Example2_lemma2_aux2)
-apply(rotate_tac -1)
-apply(erule ssubst)
-apply simp_all
-done
-
-
-record Example2 = 
- c :: "nat \<Rightarrow> nat" 
- x :: nat
-
-lemma Example_2: "0<n \<Longrightarrow> 
- \<parallel>- .{\<acute>x=0 \<and> (\<Sum>i=0..<n. \<acute>c i)=0}.  
- COBEGIN 
-   SCHEME [0\<le>i<n] 
-  .{\<acute>x=(\<Sum>i=0..<n. \<acute>c i) \<and> \<acute>c i=0}. 
-   \<langle> \<acute>x:=\<acute>x+(Suc 0),, \<acute>c:=\<acute>c (i:=(Suc 0)) \<rangle>
-  .{\<acute>x=(\<Sum>i=0..<n. \<acute>c i) \<and> \<acute>c i=(Suc 0)}.
- COEND 
- .{\<acute>x=n}."
-apply oghoare
-apply (simp_all cong del: strong_setsum_cong)
-apply (tactic {* ALLGOALS (clarify_tac @{claset}) *})
-apply (simp_all cong del: strong_setsum_cong)
-   apply(erule (1) Example2_lemma2)
-  apply(erule (1) Example2_lemma2)
- apply(erule (1) Example2_lemma2)
-apply(simp)
-done
-
-end
--- a/src/HOL/HoareParallel/OG_Hoare.thy	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,469 +0,0 @@
-
-header {* \section{The Proof System} *}
-
-theory OG_Hoare imports OG_Tran begin
-
-consts assertions :: "'a ann_com \<Rightarrow> ('a assn) set"
-primrec
-  "assertions (AnnBasic r f) = {r}"
-  "assertions (AnnSeq c1 c2) = assertions c1 \<union> assertions c2"
-  "assertions (AnnCond1 r b c1 c2) = {r} \<union> assertions c1 \<union> assertions c2"
-  "assertions (AnnCond2 r b c) = {r} \<union> assertions c"
-  "assertions (AnnWhile r b i c) = {r, i} \<union> assertions c"
-  "assertions (AnnAwait r b c) = {r}" 
-
-consts atomics :: "'a ann_com \<Rightarrow> ('a assn \<times> 'a com) set"       
-primrec
-  "atomics (AnnBasic r f) = {(r, Basic f)}"
-  "atomics (AnnSeq c1 c2) = atomics c1 \<union> atomics c2"
-  "atomics (AnnCond1 r b c1 c2) = atomics c1 \<union> atomics c2"
-  "atomics (AnnCond2 r b c) = atomics c"
-  "atomics (AnnWhile r b i c) = atomics c" 
-  "atomics (AnnAwait r b c) = {(r \<inter> b, c)}"
-
-consts com :: "'a ann_triple_op \<Rightarrow> 'a ann_com_op"
-primrec "com (c, q) = c"
-
-consts post :: "'a ann_triple_op \<Rightarrow> 'a assn"
-primrec "post (c, q) = q"
-
-constdefs  interfree_aux :: "('a ann_com_op \<times> 'a assn \<times> 'a ann_com_op) \<Rightarrow> bool"
-  "interfree_aux \<equiv> \<lambda>(co, q, co'). co'= None \<or>  
-                    (\<forall>(r,a) \<in> atomics (the co'). \<parallel>= (q \<inter> r) a q \<and>
-                    (co = None \<or> (\<forall>p \<in> assertions (the co). \<parallel>= (p \<inter> r) a p)))"
-
-constdefs interfree :: "(('a ann_triple_op) list) \<Rightarrow> bool" 
-  "interfree Ts \<equiv> \<forall>i j. i < length Ts \<and> j < length Ts \<and> i \<noteq> j \<longrightarrow> 
-                         interfree_aux (com (Ts!i), post (Ts!i), com (Ts!j)) "
-
-inductive
-  oghoare :: "'a assn \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool"  ("(3\<parallel>- _//_//_)" [90,55,90] 50)
-  and ann_hoare :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool"  ("(2\<turnstile> _// _)" [60,90] 45)
-where
-  AnnBasic: "r \<subseteq> {s. f s \<in> q} \<Longrightarrow> \<turnstile> (AnnBasic r f) q"
-
-| AnnSeq:   "\<lbrakk> \<turnstile> c0 pre c1; \<turnstile> c1 q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnSeq c0 c1) q"
-  
-| AnnCond1: "\<lbrakk> r \<inter> b \<subseteq> pre c1; \<turnstile> c1 q; r \<inter> -b \<subseteq> pre c2; \<turnstile> c2 q\<rbrakk> 
-              \<Longrightarrow> \<turnstile> (AnnCond1 r b c1 c2) q"
-| AnnCond2: "\<lbrakk> r \<inter> b \<subseteq> pre c; \<turnstile> c q; r \<inter> -b \<subseteq> q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnCond2 r b c) q"
-  
-| AnnWhile: "\<lbrakk> r \<subseteq> i; i \<inter> b \<subseteq> pre c; \<turnstile> c i; i \<inter> -b \<subseteq> q \<rbrakk> 
-              \<Longrightarrow> \<turnstile> (AnnWhile r b i c) q"
-  
-| AnnAwait:  "\<lbrakk> atom_com c; \<parallel>- (r \<inter> b) c q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnAwait r b c) q"
-  
-| AnnConseq: "\<lbrakk>\<turnstile> c q; q \<subseteq> q' \<rbrakk> \<Longrightarrow> \<turnstile> c q'"
-
-
-| Parallel: "\<lbrakk> \<forall>i<length Ts. \<exists>c q. Ts!i = (Some c, q) \<and> \<turnstile> c q; interfree Ts \<rbrakk>
-	   \<Longrightarrow> \<parallel>- (\<Inter>i\<in>{i. i<length Ts}. pre(the(com(Ts!i)))) 
-                     Parallel Ts 
-                  (\<Inter>i\<in>{i. i<length Ts}. post(Ts!i))"
-
-| Basic:   "\<parallel>- {s. f s \<in>q} (Basic f) q"
-  
-| Seq:    "\<lbrakk> \<parallel>- p c1 r; \<parallel>- r c2 q \<rbrakk> \<Longrightarrow> \<parallel>- p (Seq c1 c2) q "
-
-| Cond:   "\<lbrakk> \<parallel>- (p \<inter> b) c1 q; \<parallel>- (p \<inter> -b) c2 q \<rbrakk> \<Longrightarrow> \<parallel>- p (Cond b c1 c2) q"
-
-| While:  "\<lbrakk> \<parallel>- (p \<inter> b) c p \<rbrakk> \<Longrightarrow> \<parallel>- p (While b i c) (p \<inter> -b)"
-
-| Conseq: "\<lbrakk> p' \<subseteq> p; \<parallel>- p c q ; q \<subseteq> q' \<rbrakk> \<Longrightarrow> \<parallel>- 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\<rightarrow> (Parallel Ts, t) \<longrightarrow> All_None Ts \<longrightarrow> 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 \<inter> b) \<subseteq> p \<Longrightarrow> SEM (fwhile b S k) p \<subseteq> (p \<inter> -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 conjI)
- apply (blast dest: L3_5i) 
-apply(simp add: SEM_def sem_def id_def)
-apply (blast dest: Basic_ntran rtrancl_imp_UN_rel_pow) 
-done
-
-lemma atom_hoare_sound [rule_format]: 
- " \<parallel>- p c q \<longrightarrow> atom_com(c) \<longrightarrow> \<parallel>= 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(simp add: L3_5iv)
---{* While *}
- apply (force simp add: L3_5v dest: SEM_fwhile) 
---{* Conseq *}
-apply(force simp add: SEM_def sem_def)
-done
-    
-subsection {* Soundness of the System for Component Programs *}
-
-inductive_cases ann_transition_cases:
-    "(None,s) -1\<rightarrow> (c', s')"
-    "(Some (AnnBasic r f),s) -1\<rightarrow> (c', s')"
-    "(Some (AnnSeq c1 c2), s) -1\<rightarrow> (c', s')"
-    "(Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> (c', s')"
-    "(Some (AnnCond2 r b c), s) -1\<rightarrow> (c', s')"
-    "(Some (AnnWhile r b I c), s) -1\<rightarrow> (c', s')"
-    "(Some (AnnAwait r b c),s) -1\<rightarrow> (c', s')"
-
-text {* Strong Soundness for Component Programs:*}
-
-lemma ann_hoare_case_analysis [rule_format]: 
-  defines I: "I \<equiv> \<lambda>C q'.
-  ((\<forall>r f. C = AnnBasic r f \<longrightarrow> (\<exists>q. r \<subseteq> {s. f s \<in> q} \<and> q \<subseteq> q')) \<and>  
-  (\<forall>c0 c1. C = AnnSeq c0 c1 \<longrightarrow> (\<exists>q. q \<subseteq> q' \<and> \<turnstile> c0 pre c1 \<and> \<turnstile> c1 q)) \<and>  
-  (\<forall>r b c1 c2. C = AnnCond1 r b c1 c2 \<longrightarrow> (\<exists>q. q \<subseteq> q' \<and>  
-  r \<inter> b \<subseteq> pre c1 \<and> \<turnstile> c1 q \<and> r \<inter> -b \<subseteq> pre c2 \<and> \<turnstile> c2 q)) \<and>  
-  (\<forall>r b c. C = AnnCond2 r b c \<longrightarrow> 
-  (\<exists>q. q \<subseteq> q' \<and> r \<inter> b \<subseteq> pre c  \<and> \<turnstile> c q \<and> r \<inter> -b \<subseteq> q)) \<and>  
-  (\<forall>r i b c. C = AnnWhile r b i c \<longrightarrow>  
-  (\<exists>q. q \<subseteq> q' \<and> r \<subseteq> i \<and> i \<inter> b \<subseteq> pre c \<and> \<turnstile> c i \<and> i \<inter> -b \<subseteq> q)) \<and>  
-  (\<forall>r b c. C = AnnAwait r b c \<longrightarrow> (\<exists>q. q \<subseteq> q' \<and> \<parallel>- (r \<inter> b) c q)))"
-  shows "\<turnstile> C q' \<longrightarrow> I C q'"
-apply(rule ann_hoare_induct)
-apply (simp_all add: I)
- 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 \<inter> {(x,y). True}) = (transition)"
-apply force
-done
-
-lemma Strong_Soundness_aux_aux [rule_format]: 
- "(co, s) -1\<rightarrow> (co', t) \<longrightarrow> (\<forall>c. co = Some c \<longrightarrow> s\<in> pre c \<longrightarrow> 
- (\<forall>q. \<turnstile> c q \<longrightarrow> (if co' = None then t\<in>q else t \<in> pre(the co') \<and> \<turnstile> (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
---{* 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: "\<lbrakk> (Some c, s) -*\<rightarrow> (co, t); s \<in> pre c; \<turnstile> c q \<rbrakk>  
-  \<Longrightarrow> if co = None then t \<in> q else t \<in> pre (the co) \<and> \<turnstile> (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: "\<lbrakk> (Some c, s)-*\<rightarrow>(co, t); s \<in> pre c; \<turnstile> c q \<rbrakk>  
-  \<Longrightarrow> if co = None then t\<in>q else t \<in> pre (the co)"
-apply(force dest:Strong_Soundness_aux)
-done
-
-lemma ann_hoare_sound: "\<turnstile> c q  \<Longrightarrow> \<Turnstile> 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\<rightarrow> (R', t) \<Longrightarrow>  
-  (\<exists>Rs. R' = (Parallel Rs) \<and> (length Rs) = (length Ts) \<and>
-  (\<forall>i. i<length Ts \<longrightarrow> 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*\<rightarrow> (R',t) \<Longrightarrow>   
-  (\<exists>Rs. R' = (Parallel Rs) \<and> (length Rs) = (length Ts) \<and>  
-  (\<forall>i. i<length Ts \<longrightarrow> 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 \<in> assertions c"
-apply(rule ann_com_com.induct [THEN conjunct1])
-apply auto
-done
-
-lemma interfree_aux1 [rule_format]: 
-  "(c,s) -1\<rightarrow> (r,t)  \<longrightarrow> (interfree_aux(c1, q1, c) \<longrightarrow> 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\<rightarrow> (r,t) \<longrightarrow> (interfree_aux(c, q, a)  \<longrightarrow> interfree_aux(r, q, a) )"
-apply (rule ann_transition_transition.induct [THEN conjunct1])
-apply(force simp add:interfree_aux_def)+
-done
-
-lemma interfree_lemma: "\<lbrakk> (Some c, s) -1\<rightarrow> (r, t);interfree Ts ; i<length Ts;  
-           Ts!i = (Some c, q) \<rbrakk> \<Longrightarrow> 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: 
-  "\<lbrakk>interfree Ts; i<length Ts; com(Ts ! i) = Some(AnnSeq c0 c1) \<rbrakk> 
-  \<Longrightarrow>  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)]: 
- "\<lbrakk> \<forall>i<length Ts. (if com(Ts!i) = None then b \<in> post(Ts!i) 
-  else b \<in> pre(the(com(Ts!i))) \<and> \<turnstile> the(com(Ts!i)) post(Ts!i));  
-  com(Ts ! i) = Some(AnnSeq c0 c1); i<length Ts; interfree Ts \<rbrakk> \<Longrightarrow> 
- (\<forall>ia<length Ts. (if com(Ts[i:=(Some c0, pre c1)]! ia) = None  
-  then b \<in> post(Ts[i:=(Some c0, pre c1)]! ia) 
- else b \<in> pre(the(com(Ts[i:=(Some c0, pre c1)]! ia))) \<and>  
- \<turnstile> the(com(Ts[i:=(Some c0, pre c1)]! ia)) post(Ts[i:=(Some c0, pre c1)]! ia))) 
-  \<and> 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\<rightarrow> (co, t) \<longrightarrow>  
-  (\<forall>Ts. i<length Ts \<longrightarrow> com(Ts ! i) = Some c \<longrightarrow>  
-  (\<forall>i<length Ts. (if com(Ts ! i) = None then b\<in>post(Ts!i)  
-  else b\<in>pre(the(com(Ts!i))) \<and> \<turnstile> the(com(Ts!i)) post(Ts!i))) \<longrightarrow>  
- interfree Ts \<longrightarrow>  
-  (\<forall>j. j<length Ts \<and> i\<noteq>j \<longrightarrow> (if com(Ts!j) = None then t\<in>post(Ts!j)  
-  else t\<in>pre(the(com(Ts!j))) \<and> \<turnstile> 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(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(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]: 
- "\<lbrakk>(Ts',s) -P*\<rightarrow> (Rs',t);  Ts' = (Parallel Ts); interfree Ts;
- \<forall>i. i<length Ts \<longrightarrow> (\<exists>c q. (Ts ! i) = (Some c, q) \<and> s\<in>(pre c) \<and> \<turnstile> c q ) \<rbrakk> \<Longrightarrow>  
-  \<forall>Rs. Rs' = (Parallel Rs) \<longrightarrow> (\<forall>j. j<length Rs \<longrightarrow> 
-  (if com(Rs ! j) = None then t\<in>post(Ts ! j) 
-  else t\<in>pre(the(com(Rs ! j))) \<and> \<turnstile> the(com(Rs ! j)) post(Ts ! j))) \<and> 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\<rightarrow> (Parallel Rs, t)" for Ts s 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: 
- "\<lbrakk>(Parallel Ts, s) -P*\<rightarrow> (Parallel Rs, t); interfree Ts; j<length Rs; 
-  \<forall>i. i<length Ts \<longrightarrow> (\<exists>c q. Ts ! i = (Some c, q) \<and> s\<in>pre c \<and> \<turnstile> c q) \<rbrakk> \<Longrightarrow>  
-  if com(Rs ! j) = None then t\<in>post(Ts ! j) else t\<in>pre (the(com(Rs ! j)))"
-apply(drule  Parallel_Strong_Soundness_aux)
-apply simp+
-done
-
-lemma oghoare_sound [rule_format]: "\<parallel>- p c q \<longrightarrow> \<parallel>= 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=xb in Parallel_Strong_Soundness)
-         apply clarify
-        apply simp
-       apply force
-      apply simp
-      apply(erule_tac V = "\<forall>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)
---{* While *}
- apply(simp add: L3_5v)
- apply (blast dest: SEM_fwhile) 
---{* Conseq *}
-apply(auto simp add: SEM_def sem_def)
-done
-
-end
\ No newline at end of file
--- a/src/HOL/HoareParallel/OG_Syntax.thy	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,140 +0,0 @@
-theory OG_Syntax
-imports OG_Tactics Quote_Antiquote
-begin
-
-text{* Syntax for commands and for assertions and boolean expressions in 
- commands @{text com} and annotated commands @{text ann_com}. *}
-
-syntax
-  "_Assign"      :: "idt \<Rightarrow> 'b \<Rightarrow> 'a com"    ("(\<acute>_ :=/ _)" [70, 65] 61)
-  "_AnnAssign"   :: "'a assn \<Rightarrow> idt \<Rightarrow> 'b \<Rightarrow> 'a com"    ("(_ \<acute>_ :=/ _)" [90,70,65] 61)
-
-translations
-  "\<acute>\<spacespace>x := a" \<rightharpoonup> "Basic \<guillemotleft>\<acute>\<spacespace>(_update_name x (\<lambda>_. a))\<guillemotright>"
-  "r \<acute>\<spacespace>x := a" \<rightharpoonup> "AnnBasic r \<guillemotleft>\<acute>\<spacespace>(_update_name x (\<lambda>_. a))\<guillemotright>"
-
-syntax
-  "_AnnSkip"     :: "'a assn \<Rightarrow> 'a ann_com"              ("_//SKIP" [90] 63)
-  "_AnnSeq"      :: "'a ann_com \<Rightarrow> 'a ann_com \<Rightarrow> 'a ann_com"  ("_;;/ _" [60,61] 60)
-  
-  "_AnnCond1"    :: "'a assn \<Rightarrow> 'a bexp  \<Rightarrow> 'a ann_com  \<Rightarrow> 'a ann_com \<Rightarrow> 'a ann_com"
-                    ("_ //IF _ /THEN _ /ELSE _ /FI"  [90,0,0,0] 61)
-  "_AnnCond2"    :: "'a assn \<Rightarrow> 'a bexp  \<Rightarrow> 'a ann_com \<Rightarrow> 'a ann_com"
-                    ("_ //IF _ /THEN _ /FI"  [90,0,0] 61)
-  "_AnnWhile"    :: "'a assn \<Rightarrow> 'a bexp  \<Rightarrow> 'a assn \<Rightarrow> 'a ann_com \<Rightarrow> 'a ann_com" 
-                    ("_ //WHILE _ /INV _ //DO _//OD"  [90,0,0,0] 61)
-  "_AnnAwait"    :: "'a assn \<Rightarrow> 'a bexp  \<Rightarrow> 'a com \<Rightarrow> 'a ann_com"
-                    ("_ //AWAIT _ /THEN /_ /END"  [90,0,0] 61)
-  "_AnnAtom"     :: "'a assn  \<Rightarrow> 'a com \<Rightarrow> 'a ann_com"   ("_//\<langle>_\<rangle>" [90,0] 61)
-  "_AnnWait"     :: "'a assn \<Rightarrow> 'a bexp \<Rightarrow> 'a ann_com"   ("_//WAIT _ END" [90,0] 61)
-
-  "_Skip"        :: "'a com"                 ("SKIP" 63)
-  "_Seq"         :: "'a com \<Rightarrow> 'a com \<Rightarrow> 'a com" ("_,,/ _" [55, 56] 55)
-  "_Cond"        :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com \<Rightarrow> 'a com" 
-                                  ("(0IF _/ THEN _/ ELSE _/ FI)" [0, 0, 0] 61)
-  "_Cond2"       :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com"   ("IF _ THEN _ FI" [0,0] 56)
-  "_While_inv"   :: "'a bexp \<Rightarrow> 'a assn \<Rightarrow> 'a com \<Rightarrow> 'a com"
-                    ("(0WHILE _/ INV _ //DO _ /OD)"  [0, 0, 0] 61)
-  "_While"       :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com"
-                    ("(0WHILE _ //DO _ /OD)"  [0, 0] 61)
-
-translations
-  "SKIP" \<rightleftharpoons> "Basic id"
-  "c_1,, c_2" \<rightleftharpoons> "Seq c_1 c_2"
-
-  "IF b THEN c1 ELSE c2 FI" \<rightharpoonup> "Cond .{b}. c1 c2"
-  "IF b THEN c FI" \<rightleftharpoons> "IF b THEN c ELSE SKIP FI"
-  "WHILE b INV i DO c OD" \<rightharpoonup> "While .{b}. i c"
-  "WHILE b DO c OD" \<rightleftharpoons> "WHILE b INV CONST undefined DO c OD"
-
-  "r SKIP" \<rightleftharpoons> "AnnBasic r id"
-  "c_1;; c_2" \<rightleftharpoons> "AnnSeq c_1 c_2" 
-  "r IF b THEN c1 ELSE c2 FI" \<rightharpoonup> "AnnCond1 r .{b}. c1 c2"
-  "r IF b THEN c FI" \<rightharpoonup> "AnnCond2 r .{b}. c"
-  "r WHILE b INV i DO c OD" \<rightharpoonup> "AnnWhile r .{b}. i c"
-  "r AWAIT b THEN c END" \<rightharpoonup> "AnnAwait r .{b}. c"
-  "r \<langle>c\<rangle>" \<rightleftharpoons> "r AWAIT True THEN c END"
-  "r WAIT b END" \<rightleftharpoons> "r AWAIT b THEN SKIP END"
- 
-nonterminals
-  prgs
-
-syntax
-  "_PAR" :: "prgs \<Rightarrow> 'a"              ("COBEGIN//_//COEND" [57] 56)
-  "_prg" :: "['a, 'a] \<Rightarrow> prgs"        ("_//_" [60, 90] 57)
-  "_prgs" :: "['a, 'a, prgs] \<Rightarrow> prgs"  ("_//_//\<parallel>//_" [60,90,57] 57)
-
-  "_prg_scheme" :: "['a, 'a, 'a, 'a, 'a] \<Rightarrow> prgs"  
-                  ("SCHEME [_ \<le> _ < _] _// _" [0,0,0,60, 90] 57)
-
-translations
-  "_prg c q" \<rightleftharpoons> "[(Some c, q)]"
-  "_prgs c q ps" \<rightleftharpoons> "(Some c, q) # ps"
-  "_PAR ps" \<rightleftharpoons> "Parallel ps"
-
-  "_prg_scheme j i k c q" \<rightleftharpoons> "map (\<lambda>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 Record.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 K_tr' (Abs (_,_,t)) = if null (loose_bnos t) then t else raise Match
-      | K_tr' (Abs (_,_,Abs (_,_,t)$Bound 0)) = if null (loose_bnos t) then t else raise Match
-      | K_tr' _ = raise Match;
-
-    fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) =
-          quote_tr' (Syntax.const "_Assign" $ update_name_tr' f)
-            (Abs (x, dummyT, K_tr' k) :: ts)
-      | assign_tr' _ = raise Match;
-
-    fun annassign_tr' (r :: Abs (x, _, f $ k $ Bound 0) :: ts) =
-          quote_tr' (Syntax.const "_AnnAssign" $ r $ update_name_tr' f)
-            (Abs (x, dummyT, K_tr' k) :: 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
--- a/src/HOL/HoareParallel/OG_Tactics.thy	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,496 +0,0 @@
-header {* \section{Generation of Verification Conditions} *}
-
-theory OG_Tactics
-imports OG_Hoare
-begin
-
-lemmas ann_hoare_intros=AnnBasic AnnSeq AnnCond1 AnnCond2 AnnWhile AnnAwait AnnConseq
-lemmas oghoare_intros=Parallel Basic Seq Cond While Conseq
-
-lemma ParallelConseqRule: 
- "\<lbrakk> p \<subseteq> (\<Inter>i\<in>{i. i<length Ts}. pre(the(com(Ts ! i))));  
-  \<parallel>- (\<Inter>i\<in>{i. i<length Ts}. pre(the(com(Ts ! i)))) 
-      (Parallel Ts) 
-     (\<Inter>i\<in>{i. i<length Ts}. post(Ts ! i));  
-  (\<Inter>i\<in>{i. i<length Ts}. post(Ts ! i)) \<subseteq> q \<rbrakk>  
-  \<Longrightarrow> \<parallel>- p (Parallel Ts) q"
-apply (rule Conseq)
-prefer 2 
- apply fast
-apply assumption+
-done
-
-lemma SkipRule: "p \<subseteq> q \<Longrightarrow> \<parallel>- 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 \<subseteq> {s. (f s)\<in>q} \<Longrightarrow> \<parallel>- 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: "\<lbrakk> \<parallel>- p c1 r; \<parallel>- r c2 q \<rbrakk> \<Longrightarrow> \<parallel>- p (Seq c1 c2) q"
-apply(rule Seq)
-apply fast+
-done
-
-lemma CondRule: 
- "\<lbrakk> p \<subseteq> {s. (s\<in>b \<longrightarrow> s\<in>w) \<and> (s\<notin>b \<longrightarrow> s\<in>w')}; \<parallel>- w c1 q; \<parallel>- w' c2 q \<rbrakk> 
-  \<Longrightarrow> \<parallel>- 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: "\<lbrakk> p \<subseteq> i; \<parallel>- (i \<inter> b) c i ; (i \<inter> (-b)) \<subseteq> q \<rbrakk>  
-        \<Longrightarrow> \<parallel>- 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:
-  "\<lbrakk> atom_com(c); \<parallel>- r c q \<rbrakk>  \<Longrightarrow> \<turnstile> (AnnAwait r {s. True} c) q"
-apply(rule AnnAwait)
-apply simp_all
-done
-
-lemma AnnskipRule:
-  "r \<subseteq> q \<Longrightarrow> \<turnstile> (AnnBasic r id) q"
-apply(rule AnnBasic)
-apply simp
-done
-
-lemma AnnwaitRule:
-  "\<lbrakk> (r \<inter> b) \<subseteq> q \<rbrakk> \<Longrightarrow> \<turnstile> (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: 
-  "\<forall>(R,r)\<in>(atomics a). \<parallel>- (q \<inter> R) r q \<Longrightarrow> interfree_aux(None, q, Some a)"
-apply(simp add:interfree_aux_def)
-apply(force elim:oghoare_sound)
-done
-
-lemma interfree_aux_rule3: 
-  "(\<forall>(R, r)\<in>(atomics a). \<parallel>- (q \<inter> R) r q \<and> (\<forall>p\<in>(assertions c). \<parallel>- (p \<inter> R) r p))
-  \<Longrightarrow> interfree_aux(Some c, q, Some a)"
-apply(simp add:interfree_aux_def)
-apply(force elim:oghoare_sound)
-done
-
-lemma AnnBasic_assertions: 
-  "\<lbrakk>interfree_aux(None, r, Some a); interfree_aux(None, q, Some a)\<rbrakk> \<Longrightarrow> 
-    interfree_aux(Some (AnnBasic r f), q, Some a)"
-apply(simp add: interfree_aux_def)
-by force
-
-lemma AnnSeq_assertions: 
-  "\<lbrakk> interfree_aux(Some c1, q, Some a); interfree_aux(Some c2, q, Some a)\<rbrakk>\<Longrightarrow> 
-   interfree_aux(Some (AnnSeq c1 c2), q, Some a)"
-apply(simp add: interfree_aux_def)
-by force
-
-lemma AnnCond1_assertions: 
-  "\<lbrakk> interfree_aux(None, r, Some a); interfree_aux(Some c1, q, Some a); 
-  interfree_aux(Some c2, q, Some a)\<rbrakk>\<Longrightarrow> 
-  interfree_aux(Some(AnnCond1 r b c1 c2), q, Some a)"
-apply(simp add: interfree_aux_def)
-by force
-
-lemma AnnCond2_assertions: 
-  "\<lbrakk> interfree_aux(None, r, Some a); interfree_aux(Some c, q, Some a)\<rbrakk>\<Longrightarrow> 
-  interfree_aux(Some (AnnCond2 r b c), q, Some a)"
-apply(simp add: interfree_aux_def)
-by force
-
-lemma AnnWhile_assertions: 
-  "\<lbrakk> interfree_aux(None, r, Some a); interfree_aux(None, i, Some a); 
-  interfree_aux(Some c, q, Some a)\<rbrakk>\<Longrightarrow> 
-  interfree_aux(Some (AnnWhile r b i c), q, Some a)"
-apply(simp add: interfree_aux_def)
-by force
- 
-lemma AnnAwait_assertions: 
-  "\<lbrakk> interfree_aux(None, r, Some a); interfree_aux(None, q, Some a)\<rbrakk>\<Longrightarrow> 
-  interfree_aux(Some (AnnAwait r b c), q, Some a)"
-apply(simp add: interfree_aux_def)
-by force
- 
-lemma AnnBasic_atomics: 
-  "\<parallel>- (q \<inter> r) (Basic f) q \<Longrightarrow> interfree_aux(None, q, Some (AnnBasic r f))"
-by(simp add: interfree_aux_def oghoare_sound)
-
-lemma AnnSeq_atomics: 
-  "\<lbrakk> interfree_aux(Any, q, Some a1); interfree_aux(Any, q, Some a2)\<rbrakk>\<Longrightarrow> 
-  interfree_aux(Any, q, Some (AnnSeq a1 a2))"
-apply(simp add: interfree_aux_def)
-by force
-
-lemma AnnCond1_atomics:
-  "\<lbrakk> interfree_aux(Any, q, Some a1); interfree_aux(Any, q, Some a2)\<rbrakk>\<Longrightarrow> 
-   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)\<Longrightarrow> interfree_aux(Any, q, Some (AnnCond2 r b a))"
-by(simp add: interfree_aux_def)
-
-lemma AnnWhile_atomics: "interfree_aux (Any, q, Some a) 
-     \<Longrightarrow> interfree_aux(Any, q, Some (AnnWhile r b i a))"
-by(simp add: interfree_aux_def)
-
-lemma Annatom_atomics: 
-  "\<parallel>- (q \<inter> r) a q \<Longrightarrow> interfree_aux (None, q, Some (AnnAwait r {x. True} a))"
-by(simp add: interfree_aux_def oghoare_sound) 
-
-lemma AnnAwait_atomics: 
-  "\<parallel>- (q \<inter> (r \<inter> b)) a q \<Longrightarrow> 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) \<Rightarrow> bool"
-  "interfree_swap == \<lambda>(x, xs). \<forall>y\<in>set xs. interfree_aux (com x, post x, com y)
-  \<and> 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:  
-  "\<lbrakk> interfree_aux (com x, post x, com y); 
-  interfree_aux (com y, post y ,com x); interfree_swap (x, xs) \<rbrakk> 
-  \<Longrightarrow> interfree_swap (x, y#xs)"
-by(simp add:interfree_swap_def)
-
-lemma interfree_swap_Map: "\<forall>k. i\<le>k \<and> k<j \<longrightarrow> interfree_aux (com x, post x, c k) 
- \<and> interfree_aux (c k, Q k, com x)   
- \<Longrightarrow> interfree_swap (x, map (\<lambda>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: 
-  "\<lbrakk> interfree_swap(x, xs); interfree xs \<rbrakk> \<Longrightarrow> 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: 
-  "(\<forall>i j. a\<le>i \<and> i<b \<and> a\<le>j \<and> j<b  \<and> i\<noteq>j \<longrightarrow> interfree_aux (c i, Q i, c j))  
-  \<Longrightarrow> interfree (map (\<lambda>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) \<Rightarrow> bool " ("[\<turnstile>] _" [0] 45)
-  "[\<turnstile>] Ts == (\<forall>i<length Ts. \<exists>c q. Ts!i=(Some c, q) \<and> \<turnstile> c q)"
-
-lemma MapAnnEmpty: "[\<turnstile>] []"
-by(simp add:map_ann_hoare_def)
-
-lemma MapAnnList: "\<lbrakk> \<turnstile> c q ; [\<turnstile>] xs \<rbrakk> \<Longrightarrow> [\<turnstile>] (Some c,q)#xs"
-apply(simp add:map_ann_hoare_def)
-apply clarify
-apply(case_tac i,simp+)
-done
-
-lemma MapAnnMap: 
-  "\<forall>k. i\<le>k \<and> k<j \<longrightarrow> \<turnstile> (c k) (Q k) \<Longrightarrow> [\<turnstile>] map (\<lambda>k. (Some (c k), Q k)) [i..<j]"
-apply(simp add: map_ann_hoare_def less_diff_conv)
-done
-
-lemma ParallelRule:"\<lbrakk> [\<turnstile>] Ts ; interfree Ts \<rbrakk>
-  \<Longrightarrow> \<parallel>- (\<Inter>i\<in>{i. i<length Ts}. pre(the(com(Ts!i)))) 
-          Parallel Ts 
-        (\<Inter>i\<in>{i. i<length Ts}. post(Ts!i))"
-apply(rule Parallel)
- apply(simp add:map_ann_hoare_def)
-apply simp
-done
-(*
-lemma ParamParallelRule:
- "\<lbrakk> \<forall>k<n. \<turnstile> (c k) (Q k); 
-   \<forall>k l. k<n \<and> l<n  \<and> k\<noteq>l \<longrightarrow> interfree_aux (Some(c k), Q k, Some(c l)) \<rbrakk>
-  \<Longrightarrow> \<parallel>- (\<Inter>i\<in>{i. i<n} . pre(c i)) COBEGIN SCHEME [0\<le>i<n] (c i) (Q i) COEND  (\<Inter>i\<in>{i. i<n} . Q i )"
-apply(rule ParallelConseqRule)
-  apply simp
-  apply clarify
-  apply force
- apply(rule ParallelRule)
-  apply(rule MapAnnMap)
-  apply simp
- apply(rule interfree_Map)
- apply simp
-apply simp
-apply clarify
-apply force
-done
-*)
-
-text {* The following are some useful lemmas and simplification
-tactics to control which theorems are used to simplify at each moment,
-so that the original input does not suffer any unexpected
-transformation. *}
-
-lemma Compl_Collect: "-(Collect b) = {x. \<not>(b x)}"
-by fast
-lemma list_length: "length []=0 \<and> length (x#xs) = Suc(length xs)"
-by simp
-lemma list_lemmas: "length []=0 \<and> length (x#xs) = Suc(length xs) 
-\<and> (x#xs) ! 0=x \<and> (x#xs) ! Suc n = xs ! n"
-by simp
-lemma le_Suc_eq_insert: "{i. i <Suc n} = insert n {i. i< n}"
-by auto
-lemmas primrecdef_list = "pre.simps" "assertions.simps" "atomics.simps" "atom_com.simps"
-lemmas my_simp_list = list_lemmas fst_conv snd_conv
-not_less0 refl le_Suc_eq_insert Suc_not_Zero Zero_not_Suc nat.inject
-Collect_mem_eq ball_simps option.simps primrecdef_list
-lemmas ParallelConseq_list = INTER_def Collect_conj_eq length_map length_upt length_append list_length
-
-ML {*
-val before_interfree_simp_tac = (simp_tac (HOL_basic_ss addsimps [thm "com.simps", thm "post.simps"]))
-
-val  interfree_simp_tac = (asm_simp_tac (HOL_ss addsimps [thm "split", thm "ball_Un", thm "ball_empty"]@(thms "my_simp_list")))
-
-val ParallelConseq = (simp_tac (HOL_basic_ss addsimps (thms "ParallelConseq_list")@(thms "my_simp_list")))
-*}
-
-text {* The following tactic applies @{text tac} to each conjunct in a
-subgoal of the form @{text "A \<Longrightarrow> a1 \<and> a2 \<and> .. \<and> 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 "\<langle> \<rangle>"} 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 "\<parallel>- ?p c q"},
- in this case we have precond=False and the generated  verification     
- condition would have the form @{text "?p \<subseteq> \<dots>"} 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 {* 
-val oghoare_tac = SUBGOAL (fn (_, i) =>
-   (HoareRuleTac true i))
-*}
-
-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 {* val annhoare_tac = SUBGOAL (fn (_, i) =>
-  (AnnHoareRuleTac i))
-
-val interfree_aux_tac = SUBGOAL (fn (_, i) =>
-   (interfree_aux_Tac i))
-*}
-
-text {* The so defined ML tactics are then ``exported'' to be used in
-Isabelle proofs. *}
-
-method_setup oghoare = {*
-  Scan.succeed (K (SIMPLE_METHOD' oghoare_tac)) *}
-  "verification condition generator for the oghoare logic"
-
-method_setup annhoare = {*
-  Scan.succeed (K (SIMPLE_METHOD' annhoare_tac)) *}
-  "verification condition generator for the ann_hoare logic"
-
-method_setup interfree_aux = {*
-  Scan.succeed (K (SIMPLE_METHOD' interfree_aux_tac)) *}
-  "verification condition generator for interference freedom tests"
-
-text {* Tactics useful for dealing with the generated verification conditions: *}
-
-method_setup conjI_tac = {*
-  Scan.succeed (K (SIMPLE_METHOD' (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 = {*
-  Scan.succeed (K (SIMPLE_METHOD' (disjE_Tac (K all_tac)))) *}
-  "verification condition generator for interference freedom tests"
-
-end
--- a/src/HOL/HoareParallel/OG_Tran.thy	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,309 +0,0 @@
-
-header {* \section{Operational Semantics} *}
-
-theory OG_Tran imports OG_Com begin
-
-types
-  'a ann_com_op = "('a ann_com) option"
-  'a ann_triple_op = "('a ann_com_op \<times> 'a assn)"
-  
-consts com :: "'a ann_triple_op \<Rightarrow> 'a ann_com_op"
-primrec "com (c, q) = c"
-
-consts post :: "'a ann_triple_op \<Rightarrow> 'a assn"
-primrec "post (c, q) = q"
-
-constdefs
-  All_None :: "'a ann_triple_op list \<Rightarrow> bool"
-  "All_None Ts \<equiv> \<forall>(c, q) \<in> set Ts. c = None"
-
-subsection {* The Transition Relation *}
-
-inductive_set
-  ann_transition :: "(('a ann_com_op \<times> 'a) \<times> ('a ann_com_op \<times> 'a)) set"        
-  and transition :: "(('a com \<times> 'a) \<times> ('a com \<times> 'a)) set"
-  and ann_transition' :: "('a ann_com_op \<times> 'a) \<Rightarrow> ('a ann_com_op \<times> 'a) \<Rightarrow> bool"
-    ("_ -1\<rightarrow> _"[81,81] 100)
-  and transition' :: "('a com \<times> 'a) \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"
-    ("_ -P1\<rightarrow> _"[81,81] 100)
-  and transitions :: "('a com \<times> 'a) \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"
-    ("_ -P*\<rightarrow> _"[81,81] 100)
-where
-  "con_0 -1\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> ann_transition"
-| "con_0 -P1\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> transition"
-| "con_0 -P*\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> transition\<^sup>*"
-
-| AnnBasic:  "(Some (AnnBasic r f), s) -1\<rightarrow> (None, f s)"
-
-| AnnSeq1: "(Some c0, s) -1\<rightarrow> (None, t) \<Longrightarrow> 
-               (Some (AnnSeq c0 c1), s) -1\<rightarrow> (Some c1, t)"
-| AnnSeq2: "(Some c0, s) -1\<rightarrow> (Some c2, t) \<Longrightarrow> 
-               (Some (AnnSeq c0 c1), s) -1\<rightarrow> (Some (AnnSeq c2 c1), t)"
-
-| AnnCond1T: "s \<in> b  \<Longrightarrow> (Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> (Some c1, s)"
-| AnnCond1F: "s \<notin> b \<Longrightarrow> (Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> (Some c2, s)"
-
-| AnnCond2T: "s \<in> b  \<Longrightarrow> (Some (AnnCond2 r b c), s) -1\<rightarrow> (Some c, s)"
-| AnnCond2F: "s \<notin> b \<Longrightarrow> (Some (AnnCond2 r b c), s) -1\<rightarrow> (None, s)"
-
-| AnnWhileF: "s \<notin> b \<Longrightarrow> (Some (AnnWhile r b i c), s) -1\<rightarrow> (None, s)"
-| AnnWhileT: "s \<in> b  \<Longrightarrow> (Some (AnnWhile r b i c), s) -1\<rightarrow> 
-                         (Some (AnnSeq c (AnnWhile i b i c)), s)"
-
-| AnnAwait: "\<lbrakk> s \<in> b; atom_com c; (c, s) -P*\<rightarrow> (Parallel [], t) \<rbrakk> \<Longrightarrow>
-	           (Some (AnnAwait r b c), s) -1\<rightarrow> (None, t)" 
-
-| Parallel: "\<lbrakk> i<length Ts; Ts!i = (Some c, q); (Some c, s) -1\<rightarrow> (r, t) \<rbrakk>
-              \<Longrightarrow> (Parallel Ts, s) -P1\<rightarrow> (Parallel (Ts [i:=(r, q)]), t)"
-
-| Basic:  "(Basic f, s) -P1\<rightarrow> (Parallel [], f s)"
-
-| Seq1:   "All_None Ts \<Longrightarrow> (Seq (Parallel Ts) c, s) -P1\<rightarrow> (c, s)"
-| Seq2:   "(c0, s) -P1\<rightarrow> (c2, t) \<Longrightarrow> (Seq c0 c1, s) -P1\<rightarrow> (Seq c2 c1, t)"
-
-| CondT: "s \<in> b \<Longrightarrow> (Cond b c1 c2, s) -P1\<rightarrow> (c1, s)"
-| CondF: "s \<notin> b \<Longrightarrow> (Cond b c1 c2, s) -P1\<rightarrow> (c2, s)"
-
-| WhileF: "s \<notin> b \<Longrightarrow> (While b i c, s) -P1\<rightarrow> (Parallel [], s)"
-| WhileT: "s \<in> b \<Longrightarrow> (While b i c, s) -P1\<rightarrow> (Seq c (While b i c), s)"
-
-monos "rtrancl_mono"
-
-text {* The corresponding syntax translations are: *}
-
-abbreviation
-  ann_transition_n :: "('a ann_com_op \<times> 'a) \<Rightarrow> nat \<Rightarrow> ('a ann_com_op \<times> 'a) 
-                           \<Rightarrow> bool"  ("_ -_\<rightarrow> _"[81,81] 100)  where
-  "con_0 -n\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> ann_transition ^^ n"
-
-abbreviation
-  ann_transitions :: "('a ann_com_op \<times> 'a) \<Rightarrow> ('a ann_com_op \<times> 'a) \<Rightarrow> bool"
-                           ("_ -*\<rightarrow> _"[81,81] 100)  where
-  "con_0 -*\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> ann_transition\<^sup>*"
-
-abbreviation
-  transition_n :: "('a com \<times> 'a) \<Rightarrow> nat \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"  
-                          ("_ -P_\<rightarrow> _"[81,81,81] 100)  where
-  "con_0 -Pn\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> transition ^^ n"
-
-subsection {* Definition of Semantics *}
-
-constdefs
-  ann_sem :: "'a ann_com \<Rightarrow> 'a \<Rightarrow> 'a set"
-  "ann_sem c \<equiv> \<lambda>s. {t. (Some c, s) -*\<rightarrow> (None, t)}"
-
-  ann_SEM :: "'a ann_com \<Rightarrow> 'a set \<Rightarrow> 'a set"
-  "ann_SEM c S \<equiv> \<Union>ann_sem c ` S"  
-
-  sem :: "'a com \<Rightarrow> 'a \<Rightarrow> 'a set"
-  "sem c \<equiv> \<lambda>s. {t. \<exists>Ts. (c, s) -P*\<rightarrow> (Parallel Ts, t) \<and> All_None Ts}"
-
-  SEM :: "'a com \<Rightarrow> 'a set \<Rightarrow> 'a set"
-  "SEM c S \<equiv> \<Union>sem c ` S "
-
-syntax "_Omega" :: "'a com"    ("\<Omega>" 63)
-translations  "\<Omega>" \<rightleftharpoons> "While CONST UNIV CONST UNIV (Basic id)"
-
-consts fwhile :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> nat \<Rightarrow> 'a com"
-primrec 
-   "fwhile b c 0 = \<Omega>"
-   "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\<rightarrow> t"  
-    "(Basic f, s) -P1\<rightarrow> t"
-    "(Seq c1 c2, s) -P1\<rightarrow> t" 
-    "(Cond b c1 c2, s) -P1\<rightarrow> t"
-    "(While b i c, s) -P1\<rightarrow> t"
-
-lemma Parallel_empty_lemma [rule_format (no_asm)]: 
-  "(Parallel [],s) -Pn\<rightarrow> (Parallel Ts,t) \<longrightarrow> Ts=[] \<and> n=0 \<and> 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 \<longrightarrow> (Parallel Ss,s) -Pn\<rightarrow> (Parallel Ts,t) \<longrightarrow> Ts=Ss \<and> n=0 \<and> 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 \<Longrightarrow> (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=[] \<Longrightarrow> (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\<subseteq>Y \<Longrightarrow> SEM c X \<subseteq> SEM c Y" 
-apply (unfold SEM_def)
-apply force
-done
-
-lemma L3_5ii_lemma1: 
- "\<lbrakk> (c1, s1) -P*\<rightarrow> (Parallel Ts, s2); All_None Ts;  
-  (c2, s2) -P*\<rightarrow> (Parallel Ss, s3); All_None Ss \<rbrakk> 
- \<Longrightarrow> (Seq c1 c2, s1) -P*\<rightarrow> (Parallel Ss, s3)"
-apply(erule converse_rtrancl_induct2)
-apply(force intro:converse_rtrancl_into_rtrancl)+
-done
-
-lemma L3_5ii_lemma2 [rule_format (no_asm)]: 
- "\<forall>c1 c2 s t. (Seq c1 c2, s) -Pn\<rightarrow> (Parallel Ts, t) \<longrightarrow>  
-  (All_None Ts) \<longrightarrow> (\<exists>y m Rs. (c1,s) -P*\<rightarrow> (Parallel Rs, y) \<and> 
-  (All_None Rs) \<and> (c2, y) -Pm\<rightarrow> (Parallel Ts, t) \<and>  m \<le> 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: 
- "\<lbrakk>(Seq c1 c2,s) -P*\<rightarrow> (Parallel Ts,t); All_None Ts\<rbrakk> \<Longrightarrow> 
-    (\<exists>y Rs. (c1,s) -P*\<rightarrow> (Parallel Rs,y) \<and> All_None Rs 
-   \<and> (c2,y) -P*\<rightarrow> (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 \<inter> b)) Un (SEM c2 (X \<inter> (-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\<rightarrow> (T,t) \<longrightarrow> S=\<Omega> \<longrightarrow> (\<not>(\<exists>Rs. T=(Parallel Rs) \<and> 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: "\<lbrakk>(\<Omega>, s) -P*\<rightarrow> (Parallel Ts, t); All_None Ts \<rbrakk> \<Longrightarrow> False"
-apply(fast dest: rtrancl_imp_UN_rel_pow L3_5v_lemma1)
-done
-
-lemma L3_5v_lemma3: "SEM (\<Omega>) S = {}"
-apply (unfold SEM_def sem_def)
-apply(fast dest: L3_5v_lemma2)
-done
-
-lemma L3_5v_lemma4 [rule_format]: 
- "\<forall>s. (While b i c, s) -Pn\<rightarrow> (Parallel Ts, t) \<longrightarrow> All_None Ts \<longrightarrow>  
-  (\<exists>k. (fwhile b c k, s) -P*\<rightarrow> (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]: 
- "\<forall>s. (fwhile b c k, s) -P*\<rightarrow> (Parallel Ts, t) \<longrightarrow> All_None Ts \<longrightarrow>  
-  (While b i c, s) -P*\<rightarrow> (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) = (\<lambda>x. (\<Union>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 \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool"  ("(3\<parallel>= _// _//_)" [90,55,90] 50)
-  "\<parallel>= p c q \<equiv> SEM c p \<subseteq> q"
-
-  ann_com_validity :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool"   ("\<Turnstile> _ _" [60,90] 45)
-  "\<Turnstile> c q \<equiv> ann_SEM c (pre c) \<subseteq> q"
-
-end
\ No newline at end of file
--- a/src/HOL/HoareParallel/Quote_Antiquote.thy	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,24 +0,0 @@
-
-header {* \section{Concrete Syntax} *}
-
-theory Quote_Antiquote imports Main begin
-
-syntax
-  "_quote"     :: "'b \<Rightarrow> ('a \<Rightarrow> 'b)"                ("(\<guillemotleft>_\<guillemotright>)" [0] 1000)
-  "_antiquote" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b"                ("\<acute>_" [1000] 1000)
-  "_Assert"    :: "'a \<Rightarrow> 'a set"                    ("(.{_}.)" [0] 1000)
-
-syntax (xsymbols)
-  "_Assert"    :: "'a \<Rightarrow> 'a set"            ("(\<lbrace>_\<rbrace>)" [0] 1000)
-
-translations
-  ".{b}." \<rightharpoonup> "Collect \<guillemotleft>b\<guillemotright>"
-
-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
--- a/src/HOL/HoareParallel/RG_Com.thy	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,25 +0,0 @@
-
-header {* \chapter{The Rely-Guarantee Method} 
-
-\section {Abstract Syntax}
-*}
-
-theory RG_Com imports Main begin
-
-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 \<Rightarrow>'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
--- a/src/HOL/HoareParallel/RG_Examples.thy	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,359 +0,0 @@
-header {* \section{Examples} *}
-
-theory RG_Examples
-imports RG_Syntax
-begin
-
-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: "\<lbrakk>(j::nat)<k; i\<le> j\<rbrakk> \<Longrightarrow> i<k"
-by simp
-
-lemma add_le_less_mono: "\<lbrakk> (a::nat) < c; b\<le>d \<rbrakk> \<Longrightarrow> a + b < c + d"
-by simp
-
-record Example1 =
-  A :: "nat list"
-
-lemma Example1: 
- "\<turnstile> COBEGIN
-      SCHEME [0 \<le> i < n]
-     (\<acute>A := \<acute>A [i := 0], 
-     \<lbrace> n < length \<acute>A \<rbrace>, 
-     \<lbrace> length \<ordmasculine>A = length \<ordfeminine>A \<and> \<ordmasculine>A ! i = \<ordfeminine>A ! i \<rbrace>, 
-     \<lbrace> length \<ordmasculine>A = length \<ordfeminine>A \<and> (\<forall>j<n. i \<noteq> j \<longrightarrow> \<ordmasculine>A ! j = \<ordfeminine>A ! j) \<rbrace>, 
-     \<lbrace> \<acute>A ! i = 0 \<rbrace>) 
-    COEND
- SAT [\<lbrace> n < length \<acute>A \<rbrace>, \<lbrace> \<ordmasculine>A = \<ordfeminine>A \<rbrace>, \<lbrace> True \<rbrace>, \<lbrace> \<forall>i < n. \<acute>A ! i = 0 \<rbrace>]"
-apply(rule Parallel)
-apply (auto intro!: Basic) 
-done
-
-lemma Example1_parameterized: 
-"k < t \<Longrightarrow>
-  \<turnstile> COBEGIN 
-    SCHEME [k*n\<le>i<(Suc k)*n] (\<acute>A:=\<acute>A[i:=0], 
-   \<lbrace>t*n < length \<acute>A\<rbrace>, 
-   \<lbrace>t*n < length \<ordmasculine>A \<and> length \<ordmasculine>A=length \<ordfeminine>A \<and> \<ordmasculine>A!i = \<ordfeminine>A!i\<rbrace>, 
-   \<lbrace>t*n < length \<ordmasculine>A \<and> length \<ordmasculine>A=length \<ordfeminine>A \<and> (\<forall>j<length \<ordmasculine>A . i\<noteq>j \<longrightarrow> \<ordmasculine>A!j = \<ordfeminine>A!j)\<rbrace>, 
-   \<lbrace>\<acute>A!i=0\<rbrace>) 
-   COEND  
- SAT [\<lbrace>t*n < length \<acute>A\<rbrace>, 
-      \<lbrace>t*n < length \<ordmasculine>A \<and> length \<ordmasculine>A=length \<ordfeminine>A \<and> (\<forall>i<n. \<ordmasculine>A!(k*n+i)=\<ordfeminine>A!(k*n+i))\<rbrace>, 
-      \<lbrace>t*n < length \<ordmasculine>A \<and> length \<ordmasculine>A=length \<ordfeminine>A \<and> 
-      (\<forall>i<length \<ordmasculine>A . (i<k*n \<longrightarrow> \<ordmasculine>A!i = \<ordfeminine>A!i) \<and> ((Suc k)*n \<le> i\<longrightarrow> \<ordmasculine>A!i = \<ordfeminine>A!i))\<rbrace>, 
-      \<lbrace>\<forall>i<n. \<acute>A!(k*n+i) = 0\<rbrace>]"
-apply(rule Parallel)
-    apply auto
-  apply(erule_tac x="k*n +i" in allE)
-  apply(subgoal_tac "k*n+i <length (A b)")
-   apply force
-  apply(erule le_less_trans2) 
-  apply(case_tac t,simp+)
-  apply (simp add:add_commute)
-  apply(simp add: add_le_mono)
-apply(rule Basic)
-   apply simp
-   apply clarify
-   apply (subgoal_tac "k*n+i< length (A x)")
-    apply simp
-   apply(erule le_less_trans2)
-   apply(case_tac t,simp+)
-   apply (simp add:add_commute)
-   apply(rule add_le_mono, auto)
-done
-
-
-subsection {* Increment a Variable in Parallel *}
-
-subsubsection {* Two components *}
-
-record Example2 =
-  x  :: nat
-  c_0 :: nat
-  c_1 :: nat
-
-lemma Example2: 
- "\<turnstile>  COBEGIN
-    (\<langle> \<acute>x:=\<acute>x+1;; \<acute>c_0:=\<acute>c_0 + 1 \<rangle>, 
-     \<lbrace>\<acute>x=\<acute>c_0 + \<acute>c_1  \<and> \<acute>c_0=0\<rbrace>, 
-     \<lbrace>\<ordmasculine>c_0 = \<ordfeminine>c_0 \<and> 
-        (\<ordmasculine>x=\<ordmasculine>c_0 + \<ordmasculine>c_1 
-        \<longrightarrow> \<ordfeminine>x = \<ordfeminine>c_0 + \<ordfeminine>c_1)\<rbrace>,  
-     \<lbrace>\<ordmasculine>c_1 = \<ordfeminine>c_1 \<and> 
-         (\<ordmasculine>x=\<ordmasculine>c_0 + \<ordmasculine>c_1 
-         \<longrightarrow> \<ordfeminine>x =\<ordfeminine>c_0 + \<ordfeminine>c_1)\<rbrace>,
-     \<lbrace>\<acute>x=\<acute>c_0 + \<acute>c_1 \<and> \<acute>c_0=1 \<rbrace>)
-  \<parallel>
-      (\<langle> \<acute>x:=\<acute>x+1;; \<acute>c_1:=\<acute>c_1+1 \<rangle>, 
-     \<lbrace>\<acute>x=\<acute>c_0 + \<acute>c_1 \<and> \<acute>c_1=0 \<rbrace>, 
-     \<lbrace>\<ordmasculine>c_1 = \<ordfeminine>c_1 \<and> 
-        (\<ordmasculine>x=\<ordmasculine>c_0 + \<ordmasculine>c_1 
-        \<longrightarrow> \<ordfeminine>x = \<ordfeminine>c_0 + \<ordfeminine>c_1)\<rbrace>,  
-     \<lbrace>\<ordmasculine>c_0 = \<ordfeminine>c_0 \<and> 
-         (\<ordmasculine>x=\<ordmasculine>c_0 + \<ordmasculine>c_1 
-        \<longrightarrow> \<ordfeminine>x =\<ordfeminine>c_0 + \<ordfeminine>c_1)\<rbrace>,
-     \<lbrace>\<acute>x=\<acute>c_0 + \<acute>c_1 \<and> \<acute>c_1=1\<rbrace>)
- COEND
- SAT [\<lbrace>\<acute>x=0 \<and> \<acute>c_0=0 \<and> \<acute>c_1=0\<rbrace>, 
-      \<lbrace>\<ordmasculine>x=\<ordfeminine>x \<and>  \<ordmasculine>c_0= \<ordfeminine>c_0 \<and> \<ordmasculine>c_1=\<ordfeminine>c_1\<rbrace>,
-      \<lbrace>True\<rbrace>,
-      \<lbrace>\<acute>x=2\<rbrace>]"
-apply(rule Parallel)
-   apply simp_all
-   apply clarify
-   apply(case_tac i)
-    apply simp
-    apply(rule conjI)
-     apply clarify
-     apply simp
-    apply clarify
-    apply simp
-    apply(case_tac j,simp)
-    apply simp
-   apply simp
-   apply(rule conjI)
-    apply clarify
-    apply simp
-   apply clarify
-   apply simp
-   apply(subgoal_tac "j=0")
-    apply (rotate_tac -1)
-    apply (simp (asm_lr))
-   apply arith
-  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(auto intro!: Basic)
-done
-
-subsubsection {* Parameterized *}
-
-lemma Example2_lemma2_aux: "j<n \<Longrightarrow> 
- (\<Sum>i=0..<n. (b i::nat)) =
- (\<Sum>i=0..<j. b i) + b j + (\<Sum>i=0..<n-(Suc j) . b (Suc j + i))"
-apply(induct n)
- apply simp_all
-apply(simp add:less_Suc_eq)
- apply(auto)
-apply(subgoal_tac "n - j = Suc(n- Suc j)")
-  apply simp
-apply arith
-done
-
-lemma Example2_lemma2_aux2: 
-  "j\<le> s \<Longrightarrow> (\<Sum>i::nat=0..<j. (b (s:=t)) i) = (\<Sum>i=0..<j. b i)"
-apply(induct j)
- apply (simp_all cong:setsum_cong)
-done
-
-lemma Example2_lemma2: 
- "\<lbrakk>j<n; b j=0\<rbrakk> \<Longrightarrow> Suc (\<Sum>i::nat=0..<n. b i)=(\<Sum>i=0..<n. (b (j := Suc 0)) i)"
-apply(frule_tac b="(b (j:=(Suc 0)))" in Example2_lemma2_aux)
-apply(erule_tac  t="setsum (b(j := (Suc 0))) {0..<n}" in ssubst)
-apply(frule_tac b=b in Example2_lemma2_aux)
-apply(erule_tac  t="setsum b {0..<n}" in ssubst)
-apply(subgoal_tac "Suc (setsum b {0..<j} + b j + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))=(setsum b {0..<j} + Suc (b j) + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))")
-apply(rotate_tac -1)
-apply(erule ssubst)
-apply(subgoal_tac "j\<le>j")
- apply(drule_tac b="b" and t="(Suc 0)" in Example2_lemma2_aux2)
-apply(rotate_tac -1)
-apply(erule ssubst)
-apply simp_all
-done
-
-lemma Example2_lemma2_Suc0: "\<lbrakk>j<n; b j=0\<rbrakk> \<Longrightarrow>
- Suc (\<Sum>i::nat=0..< n. b i)=(\<Sum>i=0..< n. (b (j:=Suc 0)) i)"
-by(simp add:Example2_lemma2)
-
-record Example2_parameterized =   
-  C :: "nat \<Rightarrow> nat"
-  y  :: nat
-
-lemma Example2_parameterized: "0<n \<Longrightarrow> 
-  \<turnstile> COBEGIN SCHEME  [0\<le>i<n]
-     (\<langle> \<acute>y:=\<acute>y+1;; \<acute>C:=\<acute>C (i:=1) \<rangle>, 
-     \<lbrace>\<acute>y=(\<Sum>i=0..<n. \<acute>C i) \<and> \<acute>C i=0\<rbrace>, 
-     \<lbrace>\<ordmasculine>C i = \<ordfeminine>C i \<and> 
-      (\<ordmasculine>y=(\<Sum>i=0..<n. \<ordmasculine>C i) \<longrightarrow> \<ordfeminine>y =(\<Sum>i=0..<n. \<ordfeminine>C i))\<rbrace>,  
-     \<lbrace>(\<forall>j<n. i\<noteq>j \<longrightarrow> \<ordmasculine>C j = \<ordfeminine>C j) \<and> 
-       (\<ordmasculine>y=(\<Sum>i=0..<n. \<ordmasculine>C i) \<longrightarrow> \<ordfeminine>y =(\<Sum>i=0..<n. \<ordfeminine>C i))\<rbrace>,
-     \<lbrace>\<acute>y=(\<Sum>i=0..<n. \<acute>C i) \<and> \<acute>C i=1\<rbrace>) 
-    COEND
- SAT [\<lbrace>\<acute>y=0 \<and> (\<Sum>i=0..<n. \<acute>C i)=0 \<rbrace>, \<lbrace>\<ordmasculine>C=\<ordfeminine>C \<and> \<ordmasculine>y=\<ordfeminine>y\<rbrace>, \<lbrace>True\<rbrace>, \<lbrace>\<acute>y=n\<rbrace>]"
-apply(rule Parallel)
-apply force
-apply force
-apply(force)
-apply clarify
-apply simp
-apply(simp cong:setsum_ivl_cong)
-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(simp add:Example2_lemma2_Suc0 cong:if_cong)
-apply simp+
-done
-
-subsection {* Find Least Element *}
-
-text {* A previous lemma: *}
-
-lemma mod_aux :"\<lbrakk>i < (n::nat); a mod n = i;  j < a + n; j mod n = i; a < j\<rbrakk> \<Longrightarrow> False"
-apply(subgoal_tac "a=a div n*n + a mod n" )
- prefer 2 apply (simp (no_asm_use))
-apply(subgoal_tac "j=j div n*n + j mod n")
- prefer 2 apply (simp (no_asm_use))
-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 \<Rightarrow> nat"
-  Y :: "nat \<Rightarrow> nat"
-
-lemma Example3: "m mod n=0 \<Longrightarrow> 
- \<turnstile> COBEGIN 
- SCHEME [0\<le>i<n]
- (WHILE (\<forall>j<n. \<acute>X i < \<acute>Y j)  DO 
-   IF P(B!(\<acute>X i)) THEN \<acute>Y:=\<acute>Y (i:=\<acute>X i) 
-   ELSE \<acute>X:= \<acute>X (i:=(\<acute>X i)+ n) FI 
-  OD,
- \<lbrace>(\<acute>X i) mod n=i \<and> (\<forall>j<\<acute>X i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and> (\<acute>Y i<m \<longrightarrow> P(B!(\<acute>Y i)) \<and> \<acute>Y i\<le> m+i)\<rbrace>,
- \<lbrace>(\<forall>j<n. i\<noteq>j \<longrightarrow> \<ordfeminine>Y j \<le> \<ordmasculine>Y j) \<and> \<ordmasculine>X i = \<ordfeminine>X i \<and> 
-   \<ordmasculine>Y i = \<ordfeminine>Y i\<rbrace>,
- \<lbrace>(\<forall>j<n. i\<noteq>j \<longrightarrow> \<ordmasculine>X j = \<ordfeminine>X j \<and> \<ordmasculine>Y j = \<ordfeminine>Y j) \<and>   
-   \<ordfeminine>Y i \<le> \<ordmasculine>Y i\<rbrace>,
- \<lbrace>(\<acute>X i) mod n=i \<and> (\<forall>j<\<acute>X i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and> (\<acute>Y i<m \<longrightarrow> P(B!(\<acute>Y i)) \<and> \<acute>Y i\<le> m+i) \<and> (\<exists>j<n. \<acute>Y j \<le> \<acute>X i) \<rbrace>) 
- COEND
- SAT [\<lbrace> \<forall>i<n. \<acute>X i=i \<and> \<acute>Y i=m+i \<rbrace>,\<lbrace>\<ordmasculine>X=\<ordfeminine>X \<and> \<ordmasculine>Y=\<ordfeminine>Y\<rbrace>,\<lbrace>True\<rbrace>,
-  \<lbrace>\<forall>i<n. (\<acute>X i) mod n=i \<and> (\<forall>j<\<acute>X i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and> 
-    (\<acute>Y i<m \<longrightarrow> P(B!(\<acute>Y i)) \<and> \<acute>Y i\<le> m+i) \<and> (\<exists>j<n. \<acute>Y j \<le> \<acute>X i)\<rbrace>]"
-apply(rule Parallel)
---{*5 subgoals left *}
-apply force+
-apply clarify
-apply simp
-apply(rule While)
-    apply force
-   apply force
-  apply force
- apply(rule_tac pre'="\<lbrace> \<acute>X i mod n = i \<and> (\<forall>j. j<\<acute>X i \<longrightarrow> j mod n = i \<longrightarrow> \<not>P(B!j)) \<and> (\<acute>Y i < n * q \<longrightarrow> P (B!(\<acute>Y i))) \<and> \<acute>X i<\<acute>Y i\<rbrace>" in Conseq)
-     apply force
-    apply(rule subset_refl)+
- apply(rule Cond)
-    apply force
-   apply(rule Basic)
-      apply force
-     apply fastsimp
-    apply force
-   apply force
-  apply(rule Basic)
-     apply simp
-     apply clarify
-     apply simp
-     apply (case_tac "X x (j mod n) \<le> 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 auto
-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 \<Longrightarrow> \<turnstile> (COBEGIN SCHEME [0\<le>i<n]
- (WHILE (\<forall>j<n. \<acute>X!i < \<acute>Y!j)  DO 
-     IF P(B!(\<acute>X!i)) THEN \<acute>Y:=\<acute>Y[i:=\<acute>X!i] ELSE \<acute>X:= \<acute>X[i:=(\<acute>X!i)+ n] FI 
-  OD,
- \<lbrace>n<length \<acute>X \<and> n<length \<acute>Y \<and> (\<acute>X!i) mod n=i \<and> (\<forall>j<\<acute>X!i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and> (\<acute>Y!i<m \<longrightarrow> P(B!(\<acute>Y!i)) \<and> \<acute>Y!i\<le> m+i)\<rbrace>,
- \<lbrace>(\<forall>j<n. i\<noteq>j \<longrightarrow> \<ordfeminine>Y!j \<le> \<ordmasculine>Y!j) \<and> \<ordmasculine>X!i = \<ordfeminine>X!i \<and> 
-   \<ordmasculine>Y!i = \<ordfeminine>Y!i \<and> length \<ordmasculine>X = length \<ordfeminine>X \<and> length \<ordmasculine>Y = length \<ordfeminine>Y\<rbrace>,
- \<lbrace>(\<forall>j<n. i\<noteq>j \<longrightarrow> \<ordmasculine>X!j = \<ordfeminine>X!j \<and> \<ordmasculine>Y!j = \<ordfeminine>Y!j) \<and>   
-   \<ordfeminine>Y!i \<le> \<ordmasculine>Y!i \<and> length \<ordmasculine>X = length \<ordfeminine>X \<and> length \<ordmasculine>Y = length \<ordfeminine>Y\<rbrace>,
- \<lbrace>(\<acute>X!i) mod n=i \<and> (\<forall>j<\<acute>X!i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and> (\<acute>Y!i<m \<longrightarrow> P(B!(\<acute>Y!i)) \<and> \<acute>Y!i\<le> m+i) \<and> (\<exists>j<n. \<acute>Y!j \<le> \<acute>X!i) \<rbrace>) COEND)
- SAT [\<lbrace>n<length \<acute>X \<and> n<length \<acute>Y \<and> (\<forall>i<n. \<acute>X!i=i \<and> \<acute>Y!i=m+i) \<rbrace>,
-      \<lbrace>\<ordmasculine>X=\<ordfeminine>X \<and> \<ordmasculine>Y=\<ordfeminine>Y\<rbrace>,
-      \<lbrace>True\<rbrace>,
-      \<lbrace>\<forall>i<n. (\<acute>X!i) mod n=i \<and> (\<forall>j<\<acute>X!i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and> 
-        (\<acute>Y!i<m \<longrightarrow> P(B!(\<acute>Y!i)) \<and> \<acute>Y!i\<le> m+i) \<and> (\<exists>j<n. \<acute>Y!j \<le> \<acute>X!i)\<rbrace>]"
-apply(rule Parallel)
---{* 5 subgoals left *}
-apply force+
-apply clarify
-apply simp
-apply(rule While)
-    apply force
-   apply force
-  apply force
- apply(rule_tac pre'="\<lbrace>n<length \<acute>X \<and> n<length \<acute>Y \<and> \<acute>X ! i mod n = i \<and> (\<forall>j. j < \<acute>X ! i \<longrightarrow> j mod n = i \<longrightarrow> \<not> P (B ! j)) \<and> (\<acute>Y ! i < n * q \<longrightarrow> P (B ! (\<acute>Y ! i))) \<and> \<acute>X!i<\<acute>Y!i\<rbrace>" 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\<le> 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 auto
-done
-
-end
--- a/src/HOL/HoareParallel/RG_Hoare.thy	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1375 +0,0 @@
-header {* \section{The Proof System} *}
-
-theory RG_Hoare imports RG_Tran begin
-
-subsection {* Proof System for Component Programs *}
-
-declare Un_subset_iff [iff del]
-declare Cons_eq_map_conv[iff]
-
-constdefs
-  stable :: "'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> bool"  
-  "stable \<equiv> \<lambda>f g. (\<forall>x y. x \<in> f \<longrightarrow> (x, y) \<in> g \<longrightarrow> y \<in> f)" 
-
-inductive
-  rghoare :: "['a com, 'a set, ('a \<times> 'a) set, ('a \<times> 'a) set, 'a set] \<Rightarrow> bool"  
-    ("\<turnstile> _ sat [_, _, _, _]" [60,0,0,0,0] 45)
-where
-  Basic: "\<lbrakk> pre \<subseteq> {s. f s \<in> post}; {(s,t). s \<in> pre \<and> (t=f s \<or> t=s)} \<subseteq> guar; 
-            stable pre rely; stable post rely \<rbrakk> 
-           \<Longrightarrow> \<turnstile> Basic f sat [pre, rely, guar, post]"
-
-| Seq: "\<lbrakk> \<turnstile> P sat [pre, rely, guar, mid]; \<turnstile> Q sat [mid, rely, guar, post] \<rbrakk> 
-           \<Longrightarrow> \<turnstile> Seq P Q sat [pre, rely, guar, post]"
-
-| Cond: "\<lbrakk> stable pre rely; \<turnstile> P1 sat [pre \<inter> b, rely, guar, post];
-           \<turnstile> P2 sat [pre \<inter> -b, rely, guar, post]; \<forall>s. (s,s)\<in>guar \<rbrakk>
-          \<Longrightarrow> \<turnstile> Cond b P1 P2 sat [pre, rely, guar, post]"
-
-| While: "\<lbrakk> stable pre rely; (pre \<inter> -b) \<subseteq> post; stable post rely;
-            \<turnstile> P sat [pre \<inter> b, rely, guar, pre]; \<forall>s. (s,s)\<in>guar \<rbrakk>
-          \<Longrightarrow> \<turnstile> While b P sat [pre, rely, guar, post]"
-
-| Await: "\<lbrakk> stable pre rely; stable post rely; 
-            \<forall>V. \<turnstile> P sat [pre \<inter> b \<inter> {V}, {(s, t). s = t}, 
-                UNIV, {s. (V, s) \<in> guar} \<inter> post] \<rbrakk>
-           \<Longrightarrow> \<turnstile> Await b P sat [pre, rely, guar, post]"
-  
-| Conseq: "\<lbrakk> pre \<subseteq> pre'; rely \<subseteq> rely'; guar' \<subseteq> guar; post' \<subseteq> post;
-             \<turnstile> P sat [pre', rely', guar', post'] \<rbrakk>
-            \<Longrightarrow> \<turnstile> P sat [pre, rely, guar, post]"
-
-constdefs 
-  Pre :: "'a rgformula \<Rightarrow> 'a set"
-  "Pre x \<equiv> fst(snd x)"
-  Post :: "'a rgformula \<Rightarrow> 'a set"
-  "Post x \<equiv> snd(snd(snd(snd x)))"
-  Rely :: "'a rgformula \<Rightarrow> ('a \<times> 'a) set"
-  "Rely x \<equiv> fst(snd(snd x))"
-  Guar :: "'a rgformula \<Rightarrow> ('a \<times> 'a) set"
-  "Guar x \<equiv> fst(snd(snd(snd x)))"
-  Com :: "'a rgformula \<Rightarrow> 'a com"
-  "Com x \<equiv> fst x"
-
-subsection {* Proof System for Parallel Programs *}
-
-types 'a par_rgformula = "('a rgformula) list \<times> 'a set \<times> ('a \<times> 'a) set \<times> ('a \<times> 'a) set \<times> 'a set"
-
-inductive
-  par_rghoare :: "('a rgformula) list \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set \<Rightarrow> bool"
-    ("\<turnstile> _ SAT [_, _, _, _]" [60,0,0,0,0] 45)
-where
-  Parallel: 
-  "\<lbrakk> \<forall>i<length xs. rely \<union> (\<Union>j\<in>{j. j<length xs \<and> j\<noteq>i}. Guar(xs!j)) \<subseteq> Rely(xs!i);
-    (\<Union>j\<in>{j. j<length xs}. Guar(xs!j)) \<subseteq> guar;
-     pre \<subseteq> (\<Inter>i\<in>{i. i<length xs}. Pre(xs!i)); 
-    (\<Inter>i\<in>{i. i<length xs}. Post(xs!i)) \<subseteq> post;
-    \<forall>i<length xs. \<turnstile> Com(xs!i) sat [Pre(xs!i),Rely(xs!i),Guar(xs!i),Post(xs!i)] \<rbrakk>
-   \<Longrightarrow>  \<turnstile> xs SAT [pre, rely, guar, post]"
-
-section {* Soundness*}
-
-subsubsection {* Some previous lemmas *}
-
-lemma tl_of_assum_in_assum: 
-  "(P, s) # (P, t) # xs \<in> assum (pre, rely) \<Longrightarrow> stable pre rely 
-  \<Longrightarrow> (P, t) # xs \<in> 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 \<in> comm(guar, post) \<Longrightarrow> (P, s) # (P, t) # xs \<in> comm(guar, post)"
-apply(simp add:comm_def)
-apply clarify
-apply(case_tac i,simp+)
-done
-
-lemma ctran_in_comm: 
-  "\<lbrakk>(s, s) \<in> guar; (Q, s) # xs \<in> comm(guar, post)\<rbrakk> 
-  \<Longrightarrow> (P, s) # (Q, s) # xs \<in> comm(guar, post)"
-apply(simp add:comm_def)
-apply clarify
-apply(case_tac i,simp+)
-done
-
-lemma takecptn_is_cptn [rule_format, elim!]: 
-  "\<forall>j. c \<in> cptn \<longrightarrow> take (Suc j) c \<in> cptn"
-apply(induct "c")
- apply(force elim: cptn.cases)
-apply clarify
-apply(case_tac j) 
- apply simp
- apply(rule CptnOne)
-apply simp
-apply(force intro:cptn.intros elim:cptn.cases)
-done
-
-lemma dropcptn_is_cptn [rule_format,elim!]: 
-  "\<forall>j<length c. c \<in> cptn \<longrightarrow> drop j c \<in> cptn"
-apply(induct "c")
- apply(force elim: cptn.cases)
-apply clarify
-apply(case_tac j,simp+) 
-apply(erule cptn.cases)
-  apply simp
- apply force
-apply force
-done
-
-lemma takepar_cptn_is_par_cptn [rule_format,elim]: 
-  "\<forall>j. c \<in> par_cptn \<longrightarrow> take (Suc j) c \<in> par_cptn"
-apply(induct "c")
- apply(force elim: cptn.cases)
-apply clarify
-apply(case_tac j,simp) 
- apply(rule ParCptnOne)
-apply(force intro:par_cptn.intros elim:par_cptn.cases)
-done
-
-lemma droppar_cptn_is_par_cptn [rule_format]:
-  "\<forall>j<length c. c \<in> par_cptn \<longrightarrow> drop j c \<in> par_cptn"
-apply(induct "c")
- apply(force elim: par_cptn.cases)
-apply clarify
-apply(case_tac j,simp+) 
-apply(erule par_cptn.cases)
-  apply simp
- apply force
-apply force
-done
-
-lemma tl_of_cptn_is_cptn: "\<lbrakk>x # xs \<in> cptn; xs \<noteq> []\<rbrakk> \<Longrightarrow> xs  \<in> cptn"
-apply(subgoal_tac "1 < length (x # xs)") 
- apply(drule dropcptn_is_cptn,simp+)
-done
-
-lemma not_ctran_None [rule_format]: 
-  "\<forall>s. (None, s)#xs \<in> cptn \<longrightarrow> (\<forall>i<length xs. ((None, s)#xs)!i -e\<rightarrow> xs!i)"
-apply(induct xs,simp+)
-apply clarify
-apply(erule cptn.cases,simp)
- apply simp
- apply(case_tac i,simp)
-  apply(rule Env)
- apply simp
-apply(force elim:ctran.cases)
-done
-
-lemma cptn_not_empty [simp]:"[] \<notin> cptn"
-apply(force elim:cptn.cases)
-done
-
-lemma etran_or_ctran [rule_format]: 
-  "\<forall>m i. x\<in>cptn \<longrightarrow> m \<le> length x 
-   \<longrightarrow> (\<forall>i. Suc i < m \<longrightarrow> \<not> x!i -c\<rightarrow> x!Suc i) \<longrightarrow> Suc i < m 
-   \<longrightarrow> x!i -e\<rightarrow> x!Suc i"
-apply(induct x,simp)
-apply clarify
-apply(erule cptn.cases,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 "(\<forall>i. Suc i < nata \<longrightarrow> (((P, t) # xs) ! i, xs ! i) \<notin> ctran)")
-  apply force
- apply clarify
- apply(erule_tac x="Suc ia" in allE,simp)
-apply(erule_tac x="0" and P="\<lambda>j. ?H j \<longrightarrow> (?J j) \<notin> ctran" in allE,simp)
-done
-
-lemma etran_or_ctran2 [rule_format]: 
-  "\<forall>i. Suc i<length x \<longrightarrow> x\<in>cptn \<longrightarrow> (x!i -c\<rightarrow> x!Suc i \<longrightarrow> \<not> x!i -e\<rightarrow> x!Suc i)
-  \<or> (x!i -e\<rightarrow> x!Suc i \<longrightarrow> \<not> x!i -c\<rightarrow> x!Suc i)"
-apply(induct x)
- apply simp
-apply clarify
-apply(erule cptn.cases,simp)
- apply(case_tac i,simp+)
-apply(case_tac i,simp)
- apply(force elim:etran.cases)
-apply simp
-done
-
-lemma etran_or_ctran2_disjI1: 
-  "\<lbrakk> x\<in>cptn; Suc i<length x; x!i -c\<rightarrow> x!Suc i\<rbrakk> \<Longrightarrow> \<not> x!i -e\<rightarrow> x!Suc i"
-by(drule etran_or_ctran2,simp_all)
-
-lemma etran_or_ctran2_disjI2: 
-  "\<lbrakk> x\<in>cptn; Suc i<length x; x!i -e\<rightarrow> x!Suc i\<rbrakk> \<Longrightarrow> \<not> x!i -c\<rightarrow> x!Suc i"
-by(drule etran_or_ctran2,simp_all)
-
-lemma not_ctran_None2 [rule_format]: 
-  "\<lbrakk> (None, s) # xs \<in>cptn; i<length xs\<rbrakk> \<Longrightarrow> \<not> ((None, s) # xs) ! i -c\<rightarrow> xs ! i"
-apply(frule not_ctran_None,simp)
-apply(case_tac i,simp)
- apply(force elim:etranE)
-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) \<longrightarrow> (\<exists>m. P m \<and> (\<forall>i<m. \<not> P i))";
-apply(rule nat_less_induct)
-apply clarify
-apply(case_tac "\<forall>m. m<n \<longrightarrow> \<not> P m")
-apply auto
-done
- 
-lemma stability [rule_format]: 
-  "\<forall>j k. x \<in> cptn \<longrightarrow> stable p rely \<longrightarrow> j\<le>k \<longrightarrow> k<length x \<longrightarrow> snd(x!j)\<in>p  \<longrightarrow>
-  (\<forall>i. (Suc i)<length x \<longrightarrow> 
-          (x!i -e\<rightarrow> x!(Suc i)) \<longrightarrow> (snd(x!i), snd(x!(Suc i))) \<in> rely) \<longrightarrow> 
-  (\<forall>i. j\<le>i \<and> i<k \<longrightarrow> x!i -e\<rightarrow> x!Suc i) \<longrightarrow> snd(x!k)\<in>p \<and> fst(x!j)=fst(x!k)"
-apply(induct x)
- apply clarify
- apply(force elim:cptn.cases)
-apply clarify
-apply(erule cptn.cases,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="\<lambda>j. (0\<le>j) \<longrightarrow> (?J j)" in allE,simp)
-  apply(subgoal_tac "t\<in>p")
-   apply(subgoal_tac "(\<forall>i. i < length xs \<longrightarrow> ((P, t) # xs) ! i -e\<rightarrow> xs ! i \<longrightarrow> (snd (((P, t) # xs) ! i), snd (xs ! i)) \<in> rely)")
-    apply clarify
-    apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j)\<in>etran" in allE,simp)
-   apply clarify
-   apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j) \<longrightarrow> (?T j)\<in>rely" in allE,simp)
-  apply(erule_tac x=0 and P="\<lambda>j. (?H j) \<longrightarrow> (?J j)\<in>etran \<longrightarrow> ?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="\<lambda>j. (?s\<le>j) \<longrightarrow> (?J j)" in allE,simp)
- apply(subgoal_tac "(\<forall>i. i < length xs \<longrightarrow> ((P, t) # xs) ! i -e\<rightarrow> xs ! i \<longrightarrow> (snd (((P, t) # xs) ! i), snd (xs ! i)) \<in> rely)")
-  apply clarify
-  apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j)\<in>etran" in allE,simp)
- apply clarify
- apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j) \<longrightarrow> (?T j)\<in>rely" in allE,simp)
-apply(case_tac k,simp,simp)
-apply(case_tac j)
- apply(erule_tac x=0 and P="\<lambda>j. (?H j) \<longrightarrow> (?J j)\<in>etran" in allE,simp)
- apply(erule etran.cases,simp)
-apply(erule_tac x="nata" in allE)
-apply(erule_tac x="nat" and P="\<lambda>j. (?s\<le>j) \<longrightarrow> (?J j)" in allE,simp)
-apply(subgoal_tac "(\<forall>i. i < length xs \<longrightarrow> ((Q, t) # xs) ! i -e\<rightarrow> xs ! i \<longrightarrow> (snd (((Q, t) # xs) ! i), snd (xs ! i)) \<in> rely)")
- apply clarify
- apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j)\<in>etran" in allE,simp)
-apply clarify
-apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j) \<longrightarrow> (?T j)\<in>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]: 
-  "\<forall>s i. x \<in> cptn \<longrightarrow> x ! 0 = (Some (Basic f), s) \<longrightarrow> 
-  Suc i<length x \<longrightarrow> x!i -c\<rightarrow> x!Suc i \<longrightarrow> 
-  (\<forall>j. Suc j<length x \<longrightarrow> i\<noteq>j \<longrightarrow> x!j -e\<rightarrow> x!Suc j)"
-apply(induct x,simp)
-apply simp
-apply clarify
-apply(erule cptn.cases,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.cases,simp_all)
- apply(force elim: not_ctran_None)
-apply(ind_cases "((Some (Basic f), sa), Q, t) \<in> ctran" for sa Q t)
-apply simp
-apply(drule_tac i=nat in not_ctran_None,simp)
-apply(erule etranE,simp)
-done
-
-lemma exists_ctran_Basic_None [rule_format]: 
-  "\<forall>s i. x \<in> cptn \<longrightarrow> x ! 0 = (Some (Basic f), s) 
-  \<longrightarrow> i<length x \<longrightarrow> fst(x!i)=None \<longrightarrow> (\<exists>j<i. x!j -c\<rightarrow> x!Suc j)"
-apply(induct x,simp)
-apply simp
-apply clarify
-apply(erule cptn.cases,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: 
-  " \<lbrakk>pre \<subseteq> {s. f s \<in> post}; {(s, t). s \<in> pre \<and> t = f s} \<subseteq> guar; 
-  stable pre rely; stable post rely\<rbrakk>
-  \<Longrightarrow> \<Turnstile> 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 "\<forall>j. ?H j")
- apply(force elim:ctran.cases)
-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(erule_tac x=i in allE,simp)
- apply(erule_tac i=j and f=f in unique_ctran_Basic,simp_all)
-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.cases,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]: 
-  "\<forall>s i. x \<in> cptn \<longrightarrow> x ! 0 = (Some (Await b c), s) \<longrightarrow> 
-  Suc i<length x \<longrightarrow> x!i -c\<rightarrow> x!Suc i \<longrightarrow> 
-  (\<forall>j. Suc j<length x \<longrightarrow> i\<noteq>j \<longrightarrow> x!j -e\<rightarrow> x!Suc j)"
-apply(induct x,simp+)
-apply clarify
-apply(erule cptn.cases,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.cases,simp_all)
- apply(force elim: not_ctran_None)
-apply(ind_cases "((Some (Await b c), sa), Q, t) \<in> ctran" for sa Q t,simp)
-apply(drule_tac i=nat in not_ctran_None,simp)
-apply(erule etranE,simp)
-done
-
-lemma exists_ctran_Await_None [rule_format]: 
-  "\<forall>s i.  x \<in> cptn \<longrightarrow> x ! 0 = (Some (Await b c), s) 
-  \<longrightarrow> i<length x \<longrightarrow> fst(x!i)=None \<longrightarrow> (\<exists>j<i. x!j -c\<rightarrow> x!Suc j)"
-apply(induct x,simp+)
-apply clarify
-apply(erule cptn.cases,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*\<rightarrow> (R, t) \<Longrightarrow> \<exists>l \<in> cp P s. (last l)=(R, t)
-  \<and> (\<forall>i. Suc i<length l \<longrightarrow> l!i -c\<rightarrow> 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.cases)
-apply simp
-apply(erule CptnComp)
-apply clarify
-done
- 
-lemma Await_sound: 
-  "\<lbrakk>stable pre rely; stable post rely;
-  \<forall>V. \<turnstile> P sat [pre \<inter> b \<inter> {s. s = V}, {(s, t). s = t}, 
-                 UNIV, {s. (V, s) \<in> guar} \<inter> post] \<and>
-  \<Turnstile> P sat [pre \<inter> b \<inter> {s. s = V}, {(s, t). s = t}, 
-                 UNIV, {s. (V, s) \<in> guar} \<inter> post] \<rbrakk>
-  \<Longrightarrow> \<Turnstile> 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\<in> 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.cases,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="\<lambda>i. ?H i \<longrightarrow> (?J i,?I i)\<in>ctran" in allE,simp)
-  apply(erule etranE,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(erule_tac x=i in allE,simp)
- apply(erule_tac i=j in unique_ctran_Await,force,simp_all)
-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.cases,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="\<lambda>i. ?H i \<longrightarrow> (?J i,?I i)\<in>ctran" in allE,simp)
- apply(erule etranE,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 Cond_sound: 
-  "\<lbrakk> stable pre rely; \<Turnstile> P1 sat [pre \<inter> b, rely, guar, post]; 
-  \<Turnstile> P2 sat [pre \<inter> - b, rely, guar, post]; \<forall>s. (s,s)\<in>guar\<rbrakk>
-  \<Longrightarrow> \<Turnstile> (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 "\<exists>i. Suc i<length x \<and> x!i -c\<rightarrow> 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 add:last_length)+)
-apply(erule exE)
-apply(drule_tac n=i and P="\<lambda>i. ?H i \<and> (?J i,?I i)\<in> 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.cases,simp_all)
- apply(erule_tac x="sa" in allE)
- apply(drule_tac c="drop (Suc m) x" in subsetD)
-  apply simp
-  apply clarify
- apply simp
- apply clarify
- apply(case_tac "i\<le>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="\<lambda>j. ?H j \<longrightarrow> ?J j \<longrightarrow> (?I j)\<in>guar" in allE)
- apply(subgoal_tac "(Suc m)+(i - Suc m) \<le> length x")
-  apply(subgoal_tac "(Suc m)+Suc (i - Suc m) \<le> length x")
-   apply(rotate_tac -2)
-   apply simp
-  apply arith
- apply arith
-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 clarify
-apply simp
-apply clarify
-apply(case_tac "i\<le>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="\<lambda>j. ?H j \<longrightarrow> ?J j \<longrightarrow> (?I j)\<in>guar" in allE)
-apply(subgoal_tac "(Suc m)+(i - Suc m) \<le> length x")
- apply(subgoal_tac "(Suc m)+Suc (i - Suc m) \<le> length x")
-  apply(rotate_tac -2)
-  apply simp
- apply arith
-apply arith
-done
-
-subsubsection{* Soundness of the Sequential rule *}
-
-inductive_cases Seq_cases [elim!]: "(Some (Seq P Q), s) -c\<rightarrow> t"
-
-lemma last_lift_not_None: "fst ((lift Q) ((x#xs)!(length xs))) \<noteq> None"
-apply(subgoal_tac "length xs<length (x # xs)")
- apply(drule_tac Q=Q in  lift_nth)
- apply(erule ssubst)
- apply (simp add:lift_def)
- apply(case_tac "(x # xs) ! length xs",simp)
-apply simp
-done
-
-declare map_eq_Cons_conv [simp del] Cons_eq_map_conv [simp del]
-lemma Seq_sound1 [rule_format]: 
-  "x\<in> cptn_mod \<Longrightarrow> \<forall>s P. x !0=(Some (Seq P Q), s) \<longrightarrow> 
-  (\<forall>i<length x. fst(x!i)\<noteq>Some Q) \<longrightarrow> 
-  (\<exists>xs\<in> 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 "(\<forall>i < Suc (length xs). fst (((Some (Seq Pa Q), t) # xs) ! i) \<noteq> Some Q)")
-    apply clarify
-    apply(rule_tac x="(Some Pa, sa) #(Some Pa, t) # zs" in exI,simp)
-    apply(rule conjI,erule CptnEnv)
-    apply(simp (no_asm_use) add:lift_def)
-   apply clarify
-   apply(erule_tac x="Suc i" in allE, simp)
-  apply(ind_cases "((Some (Seq Pa Q), sa), None, t) \<in> ctran" for Pa sa t)
- 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
-declare map_eq_Cons_conv [simp del] Cons_eq_map_conv [simp del]
-
-lemma Seq_sound2 [rule_format]: 
-  "x \<in> cptn \<Longrightarrow> \<forall>s P i. x!0=(Some (Seq P Q), s) \<longrightarrow> i<length x 
-  \<longrightarrow> fst(x!i)=Some Q \<longrightarrow>
-  (\<forall>j<i. fst(x!j)\<noteq>(Some Q)) \<longrightarrow>
-  (\<exists>xs ys. xs \<in> cp (Some P) s \<and> length xs=Suc i 
-   \<and> ys \<in> cp (Some Q) (snd(xs !i)) \<and> 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 "(\<forall>j < nat. fst (((Some (Seq Pa Q), t) # xs) ! j) \<noteq> 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 (no_asm_use) add:lift_def)
- apply(rule_tac x=ys in exI,simp)
-apply(ind_cases "((Some (Seq Pa Q), sa), t) \<in> ctran" for Pa sa t)
- 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 "(\<forall>j<Suc nata. fst (((Some (Seq P2 Q), ta) # xs) ! j) \<noteq> 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))) \<noteq> None"
-apply(simp only:last_length [THEN sym])
-apply(subgoal_tac "length xs<length (x # xs)")
- apply(drule_tac Q=Q in  lift_nth)
- apply(erule ssubst)
- apply (simp add:lift_def)
- apply(case_tac "(x # xs) ! length xs",simp)
-apply simp
-done
-*)
-
-lemma last_lift_not_None2: "fst ((lift Q) (last (x#xs))) \<noteq> None"
-apply(simp only:last_length [THEN sym])
-apply(subgoal_tac "length xs<length (x # xs)")
- apply(drule_tac Q=Q in  lift_nth)
- apply(erule ssubst)
- apply (simp add:lift_def)
- apply(case_tac "(x # xs) ! length xs",simp)
-apply simp
-done
-
-lemma Seq_sound: 
-  "\<lbrakk>\<Turnstile> P sat [pre, rely, guar, mid]; \<Turnstile> Q sat [mid, rely, guar, post]\<rbrakk>
-  \<Longrightarrow> \<Turnstile> Seq P Q sat [pre, rely, guar, post]"
-apply(unfold com_validity_def)
-apply clarify
-apply(case_tac "\<exists>i<length x. fst(x!i)=Some Q")
- prefer 2
- apply (simp add:cp_def cptn_iff_cptn_mod)
- apply clarify
- apply(frule_tac Seq_sound1,force)
-  apply force
- apply clarify
- apply(erule_tac x=s in allE,simp)
- apply(drule_tac c=xs in subsetD,simp add:cp_def cptn_iff_cptn_mod)
-  apply(simp add:assum_def)
-  apply clarify
-  apply(erule_tac P="\<lambda>j. ?H j \<longrightarrow> ?J j \<longrightarrow> ?I j" in allE,erule impE, assumption)
-  apply(simp add:snd_lift)
-  apply(erule mp)
-  apply(force elim:etranE intro:Env simp add:lift_def)
- apply(simp add:comm_def)
- apply(rule conjI)
-  apply clarify
-  apply(erule_tac P="\<lambda>j. ?H j \<longrightarrow> ?J j \<longrightarrow> ?I j" in allE,erule impE, assumption)
-  apply(simp add:snd_lift)
-  apply(erule mp)
-  apply(case_tac "(xs!i)")
-  apply(case_tac "(xs! Suc i)")
-  apply(case_tac "fst(xs!i)")
-   apply(erule_tac x=i in allE, simp add:lift_def)
-  apply(case_tac "fst(xs!Suc i)")
-   apply(force simp add:lift_def)
-  apply(force simp add:lift_def)
- apply clarify
- apply(case_tac xs,simp add:cp_def)
- apply clarify
- apply (simp del:map.simps)
- apply(subgoal_tac "(map (lift Q) ((a, b) # list))\<noteq>[]")
-  apply(drule last_conv_nth)
-  apply (simp del:map.simps)
-  apply(simp only:last_lift_not_None)
- apply simp
---{* @{text "\<exists>i<length x. fst (x ! i) = Some Q"} *}
-apply(erule exE)
-apply(drule_tac n=i and P="\<lambda>i. i < length x \<and> fst (x ! i) = Some Q" in Ex_first_occurrence)
-apply clarify
-apply (simp add:cp_def)
- apply clarify
- apply(frule_tac i=m in Seq_sound2,force)
-  apply simp+
-apply clarify
-apply(simp add:comm_def)
-apply(erule_tac x=s in allE)
-apply(drule_tac c=xs in subsetD,simp)
- apply(case_tac "xs=[]",simp)
- apply(simp add:cp_def assum_def nth_append)
- apply clarify
- apply(erule_tac x=i in allE)
-  back 
- apply(simp add:snd_lift)
- apply(erule mp)
- apply(force elim:etranE intro:Env simp add:lift_def)
-apply simp
-apply clarify
-apply(erule_tac x="snd(xs!m)" in allE)
-apply(drule_tac c=ys in subsetD,simp add:cp_def assum_def)
- apply(case_tac "xs\<noteq>[]")
- apply(drule last_conv_nth,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.cases intro:Env simp add:lift_def)
- apply simp 
-apply simp
-apply clarify
-apply(rule conjI,clarify)
- apply(case_tac "i<m",simp add:nth_append)
-  apply(simp add:snd_lift)
-  apply(erule allE, erule impE, assumption, erule mp)
-  apply(case_tac "(xs ! i)")
-  apply(case_tac "(xs ! Suc i)")   
-  apply(case_tac "fst(xs ! i)",force simp add:lift_def)   
-  apply(case_tac "fst(xs ! Suc i)")
-   apply (force simp add:lift_def)
-  apply (force simp add:lift_def)
- apply(erule_tac x="i-m" in allE) 
- back
- back
- apply(subgoal_tac "Suc (i - m) < length ys",simp)
-  prefer 2
-  apply arith
- apply(simp add:nth_append snd_lift)
- apply(rule conjI,clarify)
-  apply(subgoal_tac "i=m")
-   prefer 2
-   apply arith
-  apply clarify
-  apply(simp add:cp_def)
-  apply(rule tl_zero)
-    apply(erule mp)
-    apply(case_tac "lift Q (xs!m)",simp add:snd_lift)
-    apply(case_tac "xs!m",case_tac "fst(xs!m)",simp add:lift_def snd_lift)
-     apply(case_tac ys,simp+)
-    apply(simp add:lift_def)
-   apply simp 
-  apply force
- apply clarify
- apply(rule tl_zero)
-   apply(rule tl_zero)
-     apply (subgoal_tac "i-m=Suc(i-Suc m)")
-      apply simp
-      apply(erule mp)
-      apply(case_tac ys,simp+)
-   apply force
-  apply arith
- apply force
-apply clarify
-apply(case_tac "(map (lift Q) xs @ tl ys)\<noteq>[]")
- apply(drule last_conv_nth)
- apply(simp add: snd_lift nth_append)
- apply(rule conjI,clarify)
-  apply(case_tac ys,simp+)
- apply clarify
- apply(case_tac ys,simp+)
-done
-
-subsubsection{* Soundness of the While rule *}
-
-lemma last_append[rule_format]:
-  "\<forall>xs. ys\<noteq>[] \<longrightarrow> ((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: 
-  "\<lbrakk> \<Turnstile> P sat [pre \<inter> b, rely, guar, pre]; 
-  (Some P, s) # xs \<in> cptn_mod; fst (last ((Some P, s) # xs)) = None; s \<in> b;
-  (Some (While b P), s) # (Some (Seq P (While b P)), s) # 
-   map (lift (While b P)) xs @ ys \<in> assum (pre, rely)\<rbrakk>
-  \<Longrightarrow> (Some (While b P), snd (last ((Some P, s) # xs))) # ys \<in> 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 etranE,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 While_sound_aux [rule_format]: 
-  "\<lbrakk> pre \<inter> - b \<subseteq> post; \<Turnstile> P sat [pre \<inter> b, rely, guar, pre]; \<forall>s. (s, s) \<in> guar;
-   stable pre rely;  stable post rely; x \<in> cptn_mod \<rbrakk> 
-  \<Longrightarrow>  \<forall>s xs. x=(Some(While b P),s)#xs \<longrightarrow> x\<in>assum(pre, rely) \<longrightarrow> x \<in> 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) \<in> ctran" for s t)
-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 "\<forall>i. Suc i < length ((None, t) # xs) \<longrightarrow> (((None, t) # xs) ! i, ((None, t) # xs) ! Suc i)\<in> 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 \<longrightarrow> ?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\<rightarrow> t" for ba 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 etranE,simp add:lift_def)
-   apply(rule Env)
-  apply(erule etranE,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))\<noteq>[]")
- apply(drule last_conv_nth)
- apply (simp del:map.simps)
- apply(simp only:last_lift_not_None)
-apply simp
---{* WhileMore *}
-apply(thin_tac "P = While b P \<longrightarrow> ?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 \<in> 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<length xs")
-  apply(simp add:nth_append del:map.simps last.simps)
-  apply(case_tac "fst(((Some P, sa) # xs) ! i)")
-   apply(case_tac "((Some P, sa) # xs) ! i")
-   apply (simp add:lift_def del:last.simps)
-   apply(ind_cases "(Some (While b P), ba) -c\<rightarrow> t" for ba t)
-    apply simp
-   apply simp
-  apply(simp add:snd_lift del:map.simps last.simps)
-  apply(thin_tac " \<forall>i. i < length ys \<longrightarrow> ?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 etranE,simp add:lift_def)
-    apply(rule Env)
-   apply(erule etranE,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 \<ge> length xs"} *}
-apply(subgoal_tac "i-length xs <length ys") 
- prefer 2
- apply arith
-apply(erule_tac x="i-length xs" in allE,clarify)
-apply(case_tac "i=length xs")
- apply (simp add:nth_append snd_lift del:map.simps last.simps)
- apply(simp add:last_length del:last.simps)
- apply(erule mp)
- apply(case_tac "last((Some P, sa) # xs)")
- apply(simp add:lift_def del:last.simps)
---{* @{text "i>length xs"} *} 
-apply(case_tac "i-length xs")
- apply arith
-apply(simp add:nth_append del:map.simps last.simps)
-apply(rotate_tac -3)
-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))\<noteq>[]")
-  apply(drule last_conv_nth)
-  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)\<noteq>[]")
- apply(drule last_conv_nth)
- 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)\<noteq>[]")
-  apply(drule last_conv_nth)
-  apply (simp del:map.simps last.simps)
- apply simp
-apply simp
-done
-
-lemma While_sound: 
-  "\<lbrakk>stable pre rely; pre \<inter> - b \<subseteq> post; stable post rely;
-    \<Turnstile> P sat [pre \<inter> b, rely, guar, pre]; \<forall>s. (s,s)\<in>guar\<rbrakk>
-  \<Longrightarrow> \<Turnstile> 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: 
-  "\<lbrakk>pre \<subseteq> pre'; rely \<subseteq> rely'; guar' \<subseteq> guar; post' \<subseteq> post; 
-  \<Turnstile> P sat [pre', rely', guar', post']\<rbrakk>
-  \<Longrightarrow> \<Turnstile> 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: 
-  "\<turnstile> P sat [pre, rely, guar, post] \<Longrightarrow> \<Turnstile> 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 \<Rightarrow> 'a par_com"
-  "ParallelCom Ps \<equiv> map (Some \<circ> fst) Ps" 
-
-lemma two: 
-  "\<lbrakk> \<forall>i<length xs. rely \<union> (\<Union>j\<in>{j. j < length xs \<and> j \<noteq> i}. Guar (xs ! j)) 
-     \<subseteq> Rely (xs ! i);
-   pre \<subseteq> (\<Inter>i\<in>{i. i < length xs}. Pre (xs ! i));
-   \<forall>i<length xs. 
-   \<Turnstile> Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)];
-   length xs=length clist; x \<in> par_cp (ParallelCom xs) s; x\<in>par_assum(pre, rely);
-  \<forall>i<length clist. clist!i\<in>cp (Some(Com(xs!i))) s; x \<propto> clist \<rbrakk>
-  \<Longrightarrow> \<forall>j i. i<length clist \<and> Suc j<length x \<longrightarrow> (clist!i!j) -c\<rightarrow> (clist!i!Suc j) 
-  \<longrightarrow> (snd(clist!i!j), snd(clist!i!Suc j)) \<in> Guar(xs!i)"
-apply(unfold par_cp_def)
-apply (rule ccontr) 
---{* By contradiction: *}
-apply (simp del: Un_subset_iff)
-apply(erule exE)
---{* the first c-tran that does not satisfy the guarantee-condition is from @{text "\<sigma>_i"} at step @{text "m"}. *}
-apply(drule_tac n=j and P="\<lambda>j. \<exists>i. ?H i j" in Ex_first_occurrence)
-apply(erule exE)
-apply clarify
---{* @{text "\<sigma>_i \<in> A(pre, rely_1)"} *}
-apply(subgoal_tac "take (Suc (Suc m)) (clist!i) \<in> assum(Pre(xs!i), Rely(xs!i))")
---{* but this contradicts @{text "\<Turnstile> \<sigma>_i sat [pre_i,rely_i,guar_i,post_i]"} *}
- apply(erule_tac x=i and P="\<lambda>i. ?H i \<longrightarrow> \<Turnstile> (?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 (blast intro: takecptn_is_cptn) 
- apply simp
- apply clarify
- apply(erule_tac x=m and P="\<lambda>j. ?I j \<and> ?J j \<longrightarrow> ?H j" in allE)
- apply (simp add:conjoin_def same_length_def)
-apply(simp add:assum_def del: Un_subset_iff)
-apply(rule conjI)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow>  ?I j \<in>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="\<lambda>j. ?H j \<longrightarrow> ?M \<union> UNION (?S j) (?T j) \<subseteq>  (?L j)" in allE)
-apply(simp del: Un_subset_iff)
-apply(erule subsetD)
-apply simp
-apply(simp add:conjoin_def compat_label_def)
-apply clarify
-apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (?P j) \<or> ?Q j" in allE,simp)
---{* each etran in @{text "\<sigma>_1[0\<dots>m]"} corresponds to  *}
-apply(erule disjE)
---{* a c-tran in some @{text "\<sigma>_{ib}"}  *}
- apply clarify
- apply(case_tac "i=ib",simp)
-  apply(erule etranE,simp)
- apply(erule_tac x="ib" and P="\<lambda>i. ?H i \<longrightarrow> (?I i) \<or> (?J i)" in allE)
- apply (erule etranE)
- apply(case_tac "ia=m",simp)
- apply simp
- apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (\<forall> i. ?P i j)" in allE)
- apply(subgoal_tac "ia<m",simp)
-  prefer 2
-  apply arith
- apply(erule_tac x=ib and P="\<lambda>j. (?I j, ?H j)\<in> ctran \<longrightarrow> (?P i j)" in allE,simp)
- apply(simp add:same_state_def)
- apply(erule_tac x=i and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in all_dupE)
- apply(erule_tac x=ib and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in allE,simp)
---{* or an e-tran in @{text "\<sigma>"}, 
-therefore it satisfies @{text "rely \<or> guar_{ib}"} *}
-apply (force simp add:par_assum_def same_state_def)
-done
-
-
-lemma three [rule_format]: 
-  "\<lbrakk> xs\<noteq>[]; \<forall>i<length xs. rely \<union> (\<Union>j\<in>{j. j < length xs \<and> j \<noteq> i}. Guar (xs ! j)) 
-   \<subseteq> Rely (xs ! i);
-   pre \<subseteq> (\<Inter>i\<in>{i. i < length xs}. Pre (xs ! i));
-   \<forall>i<length xs. 
-    \<Turnstile> Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)];
-   length xs=length clist; x \<in> par_cp (ParallelCom xs) s; x \<in> par_assum(pre, rely);
-  \<forall>i<length clist. clist!i\<in>cp (Some(Com(xs!i))) s; x \<propto> clist \<rbrakk>
-  \<Longrightarrow> \<forall>j i. i<length clist \<and> Suc j<length x \<longrightarrow> (clist!i!j) -e\<rightarrow> (clist!i!Suc j) 
-  \<longrightarrow> (snd(clist!i!j), snd(clist!i!Suc j)) \<in> rely \<union> (\<Union>j\<in>{j. j < length xs \<and> j \<noteq> 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="\<lambda>j. ?H j \<longrightarrow> (?J j \<and> (\<exists>i. ?P i j)) \<or> ?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 etranE,simp)
-apply(erule_tac x="ia" and P="\<lambda>i. ?H i \<longrightarrow> (?I i) \<or> (?J i)" in allE,simp)
-apply(erule_tac x=j and P="\<lambda>j. \<forall>i. ?S j i \<longrightarrow> (?I j i, ?H j i)\<in> ctran \<longrightarrow> (?P i j)" in allE)
-apply(erule_tac x=ia and P="\<lambda>j. ?S j \<longrightarrow> (?I j, ?H j)\<in> ctran \<longrightarrow> (?P j)" in allE)
-apply(simp add:same_state_def)
-apply(erule_tac x=i and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in all_dupE)
-apply(erule_tac x=ia and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in allE,simp)
-done
-
-lemma four: 
-  "\<lbrakk>xs\<noteq>[]; \<forall>i < length xs. rely \<union> (\<Union>j\<in>{j. j < length xs \<and> j \<noteq> i}. Guar (xs ! j)) 
-    \<subseteq> Rely (xs ! i);
-   (\<Union>j\<in>{j. j < length xs}. Guar (xs ! j)) \<subseteq> guar; 
-   pre \<subseteq> (\<Inter>i\<in>{i. i < length xs}. Pre (xs ! i));
-   \<forall>i < length xs.
-    \<Turnstile> Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)];
-   x \<in> par_cp (ParallelCom xs) s; x \<in> par_assum (pre, rely); Suc i < length x; 
-   x ! i -pc\<rightarrow> x ! Suc i\<rbrakk>
-  \<Longrightarrow> (snd (x ! i), snd (x ! Suc i)) \<in> guar"
-apply(simp add: ParallelCom_def del: Un_subset_iff)
-apply(subgoal_tac "(map (Some \<circ> fst) xs)\<noteq>[]")
- prefer 2
- apply simp
-apply(frule rev_subsetD)
- apply(erule one [THEN equalityD1])
-apply(erule subsetD)
-apply (simp del: Un_subset_iff)
-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="\<lambda>j. ?H j \<longrightarrow> fst(?I j)=(?J j)" in all_dupE)
-apply(erule_tac x="Suc i" and P="\<lambda>j. ?H j \<longrightarrow> fst(?I j)=(?J j)" in allE)
-apply(erule par_ctranE,simp)
-apply(erule_tac x=i and P="\<lambda>j. \<forall>i. ?S j i \<longrightarrow> (?I j i, ?H j i)\<in> ctran \<longrightarrow> (?P i j)" in allE)
-apply(erule_tac x=ia and P="\<lambda>j. ?S j \<longrightarrow> (?I j, ?H j)\<in> ctran \<longrightarrow> (?P j)" in allE)
-apply(rule_tac x=ia in exI)
-apply(simp add:same_state_def)
-apply(erule_tac x=ia and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in all_dupE,simp)
-apply(erule_tac x=ia and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in allE,simp)
-apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in all_dupE)
-apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in all_dupE,simp)
-apply(erule_tac x="Suc i" and P="\<lambda>j. ?H j \<longrightarrow> (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]:"[] \<notin> par_cptn"
-apply(force elim:par_cptn.cases)
-done
-
-lemma five: 
-  "\<lbrakk>xs\<noteq>[]; \<forall>i<length xs. rely \<union> (\<Union>j\<in>{j. j < length xs \<and> j \<noteq> i}. Guar (xs ! j))
-   \<subseteq> Rely (xs ! i);
-   pre \<subseteq> (\<Inter>i\<in>{i. i < length xs}. Pre (xs ! i)); 
-   (\<Inter>i\<in>{i. i < length xs}. Post (xs ! i)) \<subseteq> post;
-   \<forall>i < length xs.
-    \<Turnstile> Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)];
-    x \<in> par_cp (ParallelCom xs) s; x \<in> par_assum (pre, rely); 
-   All_None (fst (last x)) \<rbrakk> \<Longrightarrow> snd (last x) \<in> post"
-apply(simp add: ParallelCom_def del: Un_subset_iff)
-apply(subgoal_tac "(map (Some \<circ> fst) xs)\<noteq>[]")
- prefer 2
- apply simp
-apply(frule rev_subsetD)
- apply(erule one [THEN equalityD1])
-apply(erule subsetD)
-apply(simp del: Un_subset_iff)
-apply clarify
-apply(subgoal_tac "\<forall>i<length clist. clist!i\<in>assum(Pre(xs!i), Rely(xs!i))")
- apply(erule_tac x=i and P="\<lambda>i. ?H i \<longrightarrow> \<Turnstile> (?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="\<lambda>j. ?H j \<longrightarrow> (?I j) \<in> 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="\<lambda>j. ?H j \<longrightarrow> fst(?I j)=(?J j)" in allE)
- apply (simp add:All_None_def same_length_def)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> length(?J j)=(?K j)" in allE)
- apply(subgoal_tac "length x - 1 < length x",simp)
-  apply(case_tac "x\<noteq>[]")
-   apply(simp add: last_conv_nth)
-   apply(erule_tac x="clist!i" in ballE)
-    apply(simp add:same_state_def)
-    apply(subgoal_tac "clist!i\<noteq>[]")
-     apply(simp add: last_conv_nth)
-    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="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in allE,simp)
- apply(erule_tac x=0 and P="\<lambda>j. ?H j \<longrightarrow> (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(rule subsetD, erule mp, assumption)
-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]: 
-  "\<forall>i s. x \<in> par_cp (ParallelCom []) s \<longrightarrow> 
-  Suc i < length x \<longrightarrow> (x ! i, x ! Suc i) \<notin> 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_ctranE,simp)
-apply(simp add:par_cp_def ParallelCom_def)
-apply clarify
-apply(erule par_cptn.cases,simp)
- apply simp
-apply(erule par_ctranE)
-back
-apply simp
-done
-
-theorem par_rgsound: 
-  "\<turnstile> c SAT [pre, rely, guar, post] \<Longrightarrow> 
-   \<Turnstile> (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\<noteq>[]")
- 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
--- a/src/HOL/HoareParallel/RG_Syntax.thy	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,95 +0,0 @@
-header {* \section{Concrete Syntax} *}
-
-theory RG_Syntax
-imports RG_Hoare Quote_Antiquote
-begin
-
-syntax
-  "_Assign"    :: "idt \<Rightarrow> 'b \<Rightarrow> 'a com"                     ("(\<acute>_ :=/ _)" [70, 65] 61)
-  "_skip"      :: "'a com"                                  ("SKIP")
-  "_Seq"       :: "'a com \<Rightarrow> 'a com \<Rightarrow> 'a com"              ("(_;;/ _)" [60,61] 60)
-  "_Cond"      :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com \<Rightarrow> 'a com"   ("(0IF _/ THEN _/ ELSE _/FI)" [0, 0, 0] 61)
-  "_Cond2"     :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com"             ("(0IF _ THEN _ FI)" [0,0] 56)
-  "_While"     :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com"             ("(0WHILE _ /DO _ /OD)"  [0, 0] 61)
-  "_Await"     :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com"             ("(0AWAIT _ /THEN /_ /END)"  [0,0] 61)
-  "_Atom"      :: "'a com \<Rightarrow> 'a com"                        ("(\<langle>_\<rangle>)" 61)
-  "_Wait"      :: "'a bexp \<Rightarrow> 'a com"                       ("(0WAIT _ END)" 61)
-
-translations
-  "\<acute>\<spacespace>x := a" \<rightharpoonup> "Basic \<guillemotleft>\<acute>\<spacespace>(_update_name x (\<lambda>_. a))\<guillemotright>"
-  "SKIP" \<rightleftharpoons> "Basic id"
-  "c1;; c2" \<rightleftharpoons> "Seq c1 c2"
-  "IF b THEN c1 ELSE c2 FI" \<rightharpoonup> "Cond .{b}. c1 c2"
-  "IF b THEN c FI" \<rightleftharpoons> "IF b THEN c ELSE SKIP FI"
-  "WHILE b DO c OD" \<rightharpoonup> "While .{b}. c"
-  "AWAIT b THEN c END" \<rightleftharpoons> "Await .{b}. c"
-  "\<langle>c\<rangle>" \<rightleftharpoons> "AWAIT True THEN c END"
-  "WAIT b END" \<rightleftharpoons> "AWAIT b THEN SKIP END"
-
-nonterminals
-  prgs
-
-syntax
-  "_PAR"        :: "prgs \<Rightarrow> 'a"              ("COBEGIN//_//COEND" 60)
-  "_prg"        :: "'a \<Rightarrow> prgs"              ("_" 57)
-  "_prgs"       :: "['a, prgs] \<Rightarrow> prgs"      ("_//\<parallel>//_" [60,57] 57)
-
-translations
-  "_prg a" \<rightharpoonup> "[a]"
-  "_prgs a ps" \<rightharpoonup> "a # ps"
-  "_PAR ps" \<rightharpoonup> "ps"
-
-syntax
-  "_prg_scheme" :: "['a, 'a, 'a, 'a] \<Rightarrow> prgs"  ("SCHEME [_ \<le> _ < _] _" [0,0,0,60] 57)
-
-translations
-  "_prg_scheme j i k c" \<rightleftharpoons> "(map (\<lambda>i. c) [j..<k])"
-
-text {* Translations for variables before and after a transition: *}
-
-syntax 
-  "_before" :: "id \<Rightarrow> 'a" ("\<ordmasculine>_")
-  "_after"  :: "id \<Rightarrow> 'a" ("\<ordfeminine>_")
- 
-translations
-  "\<ordmasculine>x" \<rightleftharpoons> "x \<acute>fst"
-  "\<ordfeminine>x" \<rightleftharpoons> "x \<acute>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 Record.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 K_tr' (Abs (_,_,t)) = if null (loose_bnos t) then t else raise Match
-      | K_tr' (Abs (_,_,Abs (_,_,t)$Bound 0)) = if null (loose_bnos t) then t else raise Match
-      | K_tr' _ = raise Match;
-
-    fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) =
-          quote_tr' (Syntax.const "_Assign" $ update_name_tr' f)
-            (Abs (x, dummyT, K_tr' k) :: 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
--- a/src/HOL/HoareParallel/RG_Tran.thy	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1075 +0,0 @@
-header {* \section{Operational Semantics} *}
-
-theory RG_Tran
-imports RG_Com
-begin
-
-subsection {* Semantics of Component Programs *}
-
-subsubsection {* Environment transitions *}
-
-types 'a conf = "(('a com) option) \<times> 'a"
-
-inductive_set
-  etran :: "('a conf \<times> 'a conf) set" 
-  and etran' :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool"  ("_ -e\<rightarrow> _" [81,81] 80)
-where
-  "P -e\<rightarrow> Q \<equiv> (P,Q) \<in> etran"
-| Env: "(P, s) -e\<rightarrow> (P, t)"
-
-lemma etranE: "c -e\<rightarrow> c' \<Longrightarrow> (\<And>P s t. c = (P, s) \<Longrightarrow> c' = (P, t) \<Longrightarrow> Q) \<Longrightarrow> Q"
-  by (induct c, induct c', erule etran.cases, blast)
-
-subsubsection {* Component transitions *}
-
-inductive_set
-  ctran :: "('a conf \<times> 'a conf) set"
-  and ctran' :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool"   ("_ -c\<rightarrow> _" [81,81] 80)
-  and ctrans :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool"   ("_ -c*\<rightarrow> _" [81,81] 80)
-where
-  "P -c\<rightarrow> Q \<equiv> (P,Q) \<in> ctran"
-| "P -c*\<rightarrow> Q \<equiv> (P,Q) \<in> ctran^*"
-
-| Basic:  "(Some(Basic f), s) -c\<rightarrow> (None, f s)"
-
-| Seq1:   "(Some P0, s) -c\<rightarrow> (None, t) \<Longrightarrow> (Some(Seq P0 P1), s) -c\<rightarrow> (Some P1, t)"
-
-| Seq2:   "(Some P0, s) -c\<rightarrow> (Some P2, t) \<Longrightarrow> (Some(Seq P0 P1), s) -c\<rightarrow> (Some(Seq P2 P1), t)"
-
-| CondT: "s\<in>b  \<Longrightarrow> (Some(Cond b P1 P2), s) -c\<rightarrow> (Some P1, s)"
-| CondF: "s\<notin>b \<Longrightarrow> (Some(Cond b P1 P2), s) -c\<rightarrow> (Some P2, s)"
-
-| WhileF: "s\<notin>b \<Longrightarrow> (Some(While b P), s) -c\<rightarrow> (None, s)"
-| WhileT: "s\<in>b  \<Longrightarrow> (Some(While b P), s) -c\<rightarrow> (Some(Seq P (While b P)), s)"
-
-| Await:  "\<lbrakk>s\<in>b; (Some P, s) -c*\<rightarrow> (None, t)\<rbrakk> \<Longrightarrow> (Some(Await b P), s) -c\<rightarrow> (None, t)" 
-
-monos "rtrancl_mono"
-
-subsection {* Semantics of Parallel Programs *}
-
-types 'a par_conf = "('a par_com) \<times> 'a"
-
-inductive_set
-  par_etran :: "('a par_conf \<times> 'a par_conf) set"
-  and par_etran' :: "['a par_conf,'a par_conf] \<Rightarrow> bool" ("_ -pe\<rightarrow> _" [81,81] 80)
-where
-  "P -pe\<rightarrow> Q \<equiv> (P,Q) \<in> par_etran"
-| ParEnv:  "(Ps, s) -pe\<rightarrow> (Ps, t)"
-
-inductive_set
-  par_ctran :: "('a par_conf \<times> 'a par_conf) set"
-  and par_ctran' :: "['a par_conf,'a par_conf] \<Rightarrow> bool" ("_ -pc\<rightarrow> _" [81,81] 80)
-where
-  "P -pc\<rightarrow> Q \<equiv> (P,Q) \<in> par_ctran"
-| ParComp: "\<lbrakk>i<length Ps; (Ps!i, s) -c\<rightarrow> (r, t)\<rbrakk> \<Longrightarrow> (Ps, s) -pc\<rightarrow> (Ps[i:=r], t)"
-
-lemma par_ctranE: "c -pc\<rightarrow> c' \<Longrightarrow>
-  (\<And>i Ps s r t. c = (Ps, s) \<Longrightarrow> c' = (Ps[i := r], t) \<Longrightarrow> i < length Ps \<Longrightarrow>
-     (Ps ! i, s) -c\<rightarrow> (r, t) \<Longrightarrow> P) \<Longrightarrow> P"
-  by (induct c, induct c', erule par_ctran.cases, blast)
-
-subsection {* Computations *}
-
-subsubsection {* Sequential computations *}
-
-types 'a confs = "('a conf) list"
-
-inductive_set cptn :: "('a confs) set"
-where
-  CptnOne: "[(P,s)] \<in> cptn"
-| CptnEnv: "(P, t)#xs \<in> cptn \<Longrightarrow> (P,s)#(P,t)#xs \<in> cptn"
-| CptnComp: "\<lbrakk>(P,s) -c\<rightarrow> (Q,t); (Q, t)#xs \<in> cptn \<rbrakk> \<Longrightarrow> (P,s)#(Q,t)#xs \<in> cptn"
-
-constdefs
-  cp :: "('a com) option \<Rightarrow> 'a \<Rightarrow> ('a confs) set"
-  "cp P s \<equiv> {l. l!0=(P,s) \<and> l \<in> cptn}"  
-
-subsubsection {* Parallel computations *}
-
-types  'a par_confs = "('a par_conf) list"
-
-inductive_set par_cptn :: "('a par_confs) set"
-where
-  ParCptnOne: "[(P,s)] \<in> par_cptn"
-| ParCptnEnv: "(P,t)#xs \<in> par_cptn \<Longrightarrow> (P,s)#(P,t)#xs \<in> par_cptn"
-| ParCptnComp: "\<lbrakk> (P,s) -pc\<rightarrow> (Q,t); (Q,t)#xs \<in> par_cptn \<rbrakk> \<Longrightarrow> (P,s)#(Q,t)#xs \<in> par_cptn"
-
-constdefs
-  par_cp :: "'a par_com \<Rightarrow> 'a \<Rightarrow> ('a par_confs) set"
-  "par_cp P s \<equiv> {l. l!0=(P,s) \<and> l \<in> par_cptn}"  
-
-subsection{* Modular Definition of Computation *}
-
-constdefs 
-  lift :: "'a com \<Rightarrow> 'a conf \<Rightarrow> 'a conf"
-  "lift Q \<equiv> \<lambda>(P, s). (if P=None then (Some Q,s) else (Some(Seq (the P) Q), s))"
-
-inductive_set cptn_mod :: "('a confs) set"
-where
-  CptnModOne: "[(P, s)] \<in> cptn_mod"
-| CptnModEnv: "(P, t)#xs \<in> cptn_mod \<Longrightarrow> (P, s)#(P, t)#xs \<in> cptn_mod"
-| CptnModNone: "\<lbrakk>(Some P, s) -c\<rightarrow> (None, t); (None, t)#xs \<in> cptn_mod \<rbrakk> \<Longrightarrow> (Some P,s)#(None, t)#xs \<in>cptn_mod"
-| CptnModCondT: "\<lbrakk>(Some P0, s)#ys \<in> cptn_mod; s \<in> b \<rbrakk> \<Longrightarrow> (Some(Cond b P0 P1), s)#(Some P0, s)#ys \<in> cptn_mod"
-| CptnModCondF: "\<lbrakk>(Some P1, s)#ys \<in> cptn_mod; s \<notin> b \<rbrakk> \<Longrightarrow> (Some(Cond b P0 P1), s)#(Some P1, s)#ys \<in> cptn_mod"
-| CptnModSeq1: "\<lbrakk>(Some P0, s)#xs \<in> cptn_mod; zs=map (lift P1) xs \<rbrakk>
-                 \<Longrightarrow> (Some(Seq P0 P1), s)#zs \<in> cptn_mod"
-| CptnModSeq2: 
-  "\<lbrakk>(Some P0, s)#xs \<in> cptn_mod; fst(last ((Some P0, s)#xs)) = None; 
-  (Some P1, snd(last ((Some P0, s)#xs)))#ys \<in> cptn_mod; 
-  zs=(map (lift P1) xs)@ys \<rbrakk> \<Longrightarrow> (Some(Seq P0 P1), s)#zs \<in> cptn_mod"
-
-| CptnModWhile1: 
-  "\<lbrakk> (Some P, s)#xs \<in> cptn_mod; s \<in> b; zs=map (lift (While b P)) xs \<rbrakk> 
-  \<Longrightarrow> (Some(While b P), s)#(Some(Seq P (While b P)), s)#zs \<in> cptn_mod"
-| CptnModWhile2: 
-  "\<lbrakk> (Some P, s)#xs \<in> cptn_mod; fst(last ((Some P, s)#xs))=None; s \<in> b; 
-  zs=(map (lift (While b P)) xs)@ys; 
-  (Some(While b P), snd(last ((Some P, s)#xs)))#ys \<in> cptn_mod\<rbrakk> 
-  \<Longrightarrow> (Some(While b P), s)#(Some(Seq P (While b P)), s)#zs \<in> 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 xs)
-apply simp_all
-done
-
-lemma div_seq [rule_format]: "list \<in> cptn_mod \<Longrightarrow>
- (\<forall>s P Q zs. list=(Some (Seq P Q), s)#zs \<longrightarrow>
-  (\<exists>xs. (Some P, s)#xs \<in> cptn_mod  \<and> (zs=(map (lift Q) xs) \<or>
-  ( fst(((Some P, s)#xs)!length xs)=None \<and> 
-  (\<exists>ys. (Some Q, snd(((Some P, s)#xs)!length xs))#ys \<in> cptn_mod  
-  \<and> 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.cases,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]:
-  "\<forall>s Q t xs.((Some a, s), Q, t) \<in> ctran \<longrightarrow> (Q, t) # xs \<in> cptn_mod 
-  \<longrightarrow> (Some a, s) # (Q, t) # xs \<in> cptn_mod"
-apply(induct a)
-apply simp_all
---{* basic *}
-apply clarify
-apply(erule ctran.cases,simp_all)
-apply(rule CptnModNone,rule Basic,simp)
-apply clarify
-apply(erule ctran.cases,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.cases,simp_all)
-apply(force elim: CptnModCondT)
-apply(force elim: CptnModCondF)
---{* While *}
-apply  clarify
-apply(erule ctran.cases,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.cases,simp_all)
-apply(rule CptnModNone,erule Await,simp+)
-done
-
-lemma cptn_onlyif_cptn_mod [rule_format]: "c \<in> cptn \<Longrightarrow> c \<in> cptn_mod"
-apply(erule cptn.induct)
-  apply(rule CptnModOne)
- apply(erule CptnModEnv)
-apply(case_tac P)
- apply simp
- apply(erule ctran.cases,simp_all)
-apply(force elim:cptn_onlyif_cptn_mod_aux)
-done
-
-lemma lift_is_cptn: "c\<in>cptn \<Longrightarrow> map (lift P) c \<in> 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.cases)
-done
-
-lemma cptn_append_is_cptn [rule_format]: 
- "\<forall>b a. b#c1\<in>cptn \<longrightarrow>  a#c2\<in>cptn \<longrightarrow> (b#c1)!length c1=a \<longrightarrow> b#c1@c2\<in>cptn"
-apply(induct c1)
- apply simp
-apply clarify
-apply(erule cptn.cases,simp_all)
- apply(force intro:CptnEnv)
-apply(force elim:CptnComp)
-done
-
-lemma last_lift: "\<lbrakk>xs\<noteq>[]; fst(xs!(length xs - (Suc 0)))=None\<rbrakk> 
- \<Longrightarrow> 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) \<longrightarrow> \<not>P a \<longrightarrow> P (x!(length x - (Suc 0)))" 
-apply(induct x,simp+)
-done
-
-lemma last_fst_esp: 
- "fst(((Some a,s)#xs)!(length xs))=None \<Longrightarrow> fst(xs!(length xs - (Suc 0)))=None" 
-apply(erule last_fst)
-apply simp
-done
-
-lemma last_snd: "xs\<noteq>[] \<Longrightarrow> 
-  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<length xs \<Longrightarrow> map (lift Q) xs ! i = lift Q  (xs! i)"
-by (simp add:lift_def)
-
-lemma snd_lift: "i< length xs \<Longrightarrow> 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 \<in> cptn_mod \<Longrightarrow> c \<in> 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.cases,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\<noteq>[]")
- 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\<noteq>[]")
- 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 \<in> cptn) = (c \<in> 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 \<times> 'a set \<times> ('a \<times> 'a) set \<times> ('a \<times> 'a) set \<times> 'a set"
-
-constdefs
-  assum :: "('a set \<times> ('a \<times> 'a) set) \<Rightarrow> ('a confs) set"
-  "assum \<equiv> \<lambda>(pre, rely). {c. snd(c!0) \<in> pre \<and> (\<forall>i. Suc i<length c \<longrightarrow> 
-               c!i -e\<rightarrow> c!(Suc i) \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> rely)}"
-
-  comm :: "(('a \<times> 'a) set \<times> 'a set) \<Rightarrow> ('a confs) set"
-  "comm \<equiv> \<lambda>(guar, post). {c. (\<forall>i. Suc i<length c \<longrightarrow> 
-               c!i -c\<rightarrow> c!(Suc i) \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> guar) \<and> 
-               (fst (last c) = None \<longrightarrow> snd (last c) \<in> post)}"
-
-  com_validity :: "'a com \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set \<Rightarrow> bool" 
-                 ("\<Turnstile> _ sat [_, _, _, _]" [60,0,0,0,0] 45)
-  "\<Turnstile> P sat [pre, rely, guar, post] \<equiv> 
-   \<forall>s. cp (Some P) s \<inter> assum(pre, rely) \<subseteq> comm(guar, post)"
-
-subsection {* Validity for Parallel Programs. *}
-
-constdefs
-  All_None :: "(('a com) option) list \<Rightarrow> bool"
-  "All_None xs \<equiv> \<forall>c\<in>set xs. c=None"
-
-  par_assum :: "('a set \<times> ('a \<times> 'a) set) \<Rightarrow> ('a par_confs) set"
-  "par_assum \<equiv> \<lambda>(pre, rely). {c. snd(c!0) \<in> pre \<and> (\<forall>i. Suc i<length c \<longrightarrow> 
-             c!i -pe\<rightarrow> c!Suc i \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> rely)}"
-
-  par_comm :: "(('a \<times> 'a) set \<times> 'a set) \<Rightarrow> ('a par_confs) set"
-  "par_comm \<equiv> \<lambda>(guar, post). {c. (\<forall>i. Suc i<length c \<longrightarrow>   
-        c!i -pc\<rightarrow> c!Suc i \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> guar) \<and> 
-         (All_None (fst (last c)) \<longrightarrow> snd( last c) \<in> post)}"
-
-  par_com_validity :: "'a  par_com \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set 
-\<Rightarrow> 'a set \<Rightarrow> bool"  ("\<Turnstile> _ SAT [_, _, _, _]" [60,0,0,0,0] 45)
-  "\<Turnstile> Ps SAT [pre, rely, guar, post] \<equiv> 
-   \<forall>s. par_cp Ps s \<inter> par_assum(pre, rely) \<subseteq> par_comm(guar, post)"
-
-subsection {* Compositionality of the Semantics *}
-
-subsubsection {* Definition of the conjoin operator *}
-
-constdefs
-  same_length :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"
-  "same_length c clist \<equiv> (\<forall>i<length clist. length(clist!i)=length c)"
- 
-  same_state :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"
-  "same_state c clist \<equiv> (\<forall>i <length clist. \<forall>j<length c. snd(c!j) = snd((clist!i)!j))"
-
-  same_program :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"
-  "same_program c clist \<equiv> (\<forall>j<length c. fst(c!j) = map (\<lambda>x. fst(nth x j)) clist)"
-
-  compat_label :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"
-  "compat_label c clist \<equiv> (\<forall>j. Suc j<length c \<longrightarrow> 
-         (c!j -pc\<rightarrow> c!Suc j \<and> (\<exists>i<length clist. (clist!i)!j -c\<rightarrow> (clist!i)! Suc j \<and> 
-                       (\<forall>l<length clist. l\<noteq>i \<longrightarrow> (clist!l)!j -e\<rightarrow> (clist!l)! Suc j))) \<or> 
-         (c!j -pe\<rightarrow> c!Suc j \<and> (\<forall>i<length clist. (clist!i)!j -e\<rightarrow> (clist!i)! Suc j)))"
-
-  conjoin :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"  ("_ \<propto> _" [65,65] 64)
-  "c \<propto> clist \<equiv> (same_length c clist) \<and> (same_state c clist) \<and> (same_program c clist) \<and> (compat_label c clist)"
-
-subsubsection {* Some previous lemmas *}
-
-lemma list_eq_if [rule_format]: 
-  "\<forall>ys. xs=ys \<longrightarrow> (length xs = length ys) \<longrightarrow> (\<forall>i<length xs. xs!i=ys!i)"
-apply (induct xs)
- apply simp
-apply clarify
-done
-
-lemma list_eq: "(length xs = length ys \<and> (\<forall>i<length xs. xs!i=ys!i)) = (xs=ys)"
-apply(rule iffI)
- apply clarify
- apply(erule nth_equalityI)
- apply simp+
-done
-
-lemma nth_tl: "\<lbrakk> ys!0=a; ys\<noteq>[] \<rbrakk> \<Longrightarrow> ys=(a#(tl ys))"
-apply(case_tac ys)
- apply simp+
-done
-
-lemma nth_tl_if [rule_format]: "ys\<noteq>[] \<longrightarrow> ys!0=a \<longrightarrow> P ys \<longrightarrow> P (a#(tl ys))"
-apply(induct ys)
- apply simp+
-done
-
-lemma nth_tl_onlyif [rule_format]: "ys\<noteq>[] \<longrightarrow> ys!0=a \<longrightarrow> P (a#(tl ys)) \<longrightarrow> P ys"
-apply(induct ys)
- apply simp+
-done
-
-lemma seq_not_eq1: "Seq c1 c2\<noteq>c1"
-apply(rule com.induct)
-apply simp_all
-apply clarify
-done
-
-lemma seq_not_eq2: "Seq c1 c2\<noteq>c2"
-apply(rule com.induct)
-apply simp_all
-apply clarify
-done
-
-lemma if_not_eq1: "Cond b c1 c2 \<noteq>c1"
-apply(rule com.induct)
-apply simp_all
-apply clarify
-done
-
-lemma if_not_eq2: "Cond b c1 c2\<noteq>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:
-  assumes c: "(P,s) -c\<rightarrow> (Q,t)"
-  shows "P\<noteq>Q" using c
-  by (induct x1 \<equiv> "(P,s)" x2 \<equiv> "(Q,t)" arbitrary: P s Q t) auto
-
-lemma prog_not_eq_in_ctran [simp]: "\<not> (P,s) -c\<rightarrow> (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\<rightarrow> (Q,t) \<Longrightarrow> (P\<noteq>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]: "\<not> (P,s) -pc\<rightarrow> (P,t)"
-apply clarify
-apply(drule prog_not_eq_in_par_ctran_aux)
-apply simp
-done
-
-lemma tl_in_cptn: "\<lbrakk> a#xs \<in>cptn; xs\<noteq>[] \<rbrakk> \<Longrightarrow> xs\<in>cptn"
-apply(force elim:cptn.cases)
-done
-
-lemma tl_zero[rule_format]: 
-  "P (ys!Suc j) \<longrightarrow> Suc j<length ys \<longrightarrow> ys\<noteq>[] \<longrightarrow> P (tl(ys)!j)"
-apply(induct ys)
- apply simp_all
-done
-
-subsection {* The Semantics is Compositional *}
-
-lemma aux_if [rule_format]: 
-  "\<forall>xs s clist. (length clist = length xs \<and> (\<forall>i<length xs. (xs!i,s)#clist!i \<in> cptn) 
-  \<and> ((xs, s)#ys \<propto> map (\<lambda>i. (fst i,s)#snd i) (zip xs clist)) 
-   \<longrightarrow> (xs, s)#ys \<in> 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="\<lambda>j. ?H j \<longrightarrow> (?P j \<or> ?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="\<lambda>j. (?H j) \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
-  prefer 2
-  apply(simp add:same_program_def)
-  apply(erule_tac x=1 and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
-  apply(drule_tac t=i in not_sym,simp)
-  apply(erule etranE,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="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
-   apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
-   apply(drule_tac t=i  in not_sym,simp)
-   apply(erule etranE,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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (?I j \<or> ?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="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE,assumption)
-   apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE,assumption)
-   apply simp
-   apply(case_tac j,simp)
-    apply(rule tl_zero)
-      apply(erule_tac x=l in allE, erule impE, assumption, 
-            erule_tac x=1 and P="\<lambda>j.  (?H j) \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
-      apply(force elim:etranE 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="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
-    apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
-    apply(drule_tac t=i  in not_sym,simp)
-    apply(erule etranE,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="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
-    apply simp
-   apply(erule_tac P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE,assumption,erule impE,assumption)
-   apply(erule tl_zero,force)
-   apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
-   apply(rule nth_tl_if)
-     apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
-    apply(erule_tac x=l  in allE, erule impE, assumption,
-          erule_tac x=1 and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
-    apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE, assumption,simp)
-    apply(erule etranE,simp)
-   apply(rule tl_zero)
-    apply force
-   apply force
-  apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> ?I j\<in>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="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
-    apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE, assumption,simp)
-    apply(force elim:etranE intro:Env)
-   apply force
-  apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
-  apply force
- apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
---{* first step is an environmental step *}
-apply clarify
-apply(erule par_etran.cases)
-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="\<lambda>j. ?H j \<longrightarrow> (?I ?s j) \<in> cptn" in allE,simp)
- apply(erule cptn.cases)
-   apply(simp add:same_length_def)
-   apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> ?J j \<in>etran" in allE,simp)
- apply(erule etranE,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="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
- apply(rule tl_zero)
-   apply(simp)
-  apply force
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (?I j \<or> ?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="\<lambda>i. ?H i \<longrightarrow> ?J i \<in>etran" in allE, erule impE, assumption)
-  apply(erule etranE,simp)
-  apply(erule_tac x=i  in allE, erule impE, assumption,
-        erule_tac x=1 and P="\<lambda>j.  (?H j) \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
-  apply(rule nth_tl_if)
-   apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
-  apply simp
- apply(erule tl_zero,force) 
-  apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
- apply clarify
- apply(erule_tac x=l and P="\<lambda>i. ?H i \<longrightarrow> ?J i \<in>etran" in allE, erule impE, assumption)
- apply(erule etranE,simp)
- apply(erule_tac x=l  in allE, erule impE, assumption,
-       erule_tac x=1 and P="\<lambda>j.  (?H j) \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
- apply(rule nth_tl_if)
-   apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
-  apply simp
-  apply(rule tl_zero,force)
-  apply force
- apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>i. ?H i \<longrightarrow> ?J i \<in>etran" in allE, erule impE, assumption)
-   apply(erule_tac x=i and P="\<lambda>i. ?H i \<longrightarrow> ?J i \<in>etran" in allE, erule impE, assumption)
-   apply(force elim:etranE intro:Env)
-  apply force
- apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
- apply force
-apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (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]: "\<forall>xs s. (xs, s)#ys \<in> par_cptn \<longrightarrow> 
-  (\<exists>clist. (length clist = length xs) \<and> 
-  (xs, s)#ys \<propto> map (\<lambda>i. (fst i,s)#(snd i)) (zip xs clist) \<and> 
-  (\<forall>i<length xs. (xs!i,s)#(clist!i) \<in> cptn))"
-apply(induct ys)
- apply(clarify)
- apply(rule_tac x="map (\<lambda>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.cases,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 (\<lambda>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.cases,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 (\<lambda>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="\<lambda>j. ?H j \<longrightarrow> (?I j \<in> 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 (no_asm_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="\<lambda>j. ?H j \<longrightarrow> (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(rule Env)
-apply simp
-apply(erule_tac x=nat and P="\<lambda>j. ?H j \<longrightarrow> (?P j \<or> ?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 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="\<lambda>j. ?H j \<longrightarrow> (?P j)\<in>etran" in allE, erule impE, assumption)
-apply(case_tac "i=ia",simp,simp)
-done
-
-lemma one_iff_aux: "xs\<noteq>[] \<Longrightarrow> (\<forall>ys. ((xs, s)#ys \<in> par_cptn) = 
- (\<exists>clist. length clist= length xs \<and> 
- ((xs, s)#ys \<propto> map (\<lambda>i. (fst i,s)#(snd i)) (zip xs clist)) \<and> 
- (\<forall>i<length xs. (xs!i,s)#(clist!i) \<in> cptn))) = 
- (par_cp (xs) s = {c. \<exists>clist. (length clist)=(length xs) \<and>
- (\<forall>i<length clist. (clist!i) \<in> cp(xs!i) s) \<and> c \<propto> 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.cases)
-  apply simp
-  apply(erule_tac x="list" in allE)
-  apply clarify
-  apply simp
-  apply(rule_tac x="map (\<lambda>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.cases,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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (\<forall>i. ?T i \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in allE,simp)
-  apply(erule_tac x=0 and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
-  apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
-  apply(case_tac "clist!i",simp,simp)
- apply(thin_tac "?H = (\<exists>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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
-    apply clarify
-    apply(simp add:cp_def)
-    apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
-    apply(case_tac "clist!i",simp,simp)
-   apply clarify
-   apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
-   apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (?P j)\<in>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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
-    apply(case_tac j,simp)
-    apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
-      apply(case_tac "clist!i",simp,simp)
-     apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
-     apply(case_tac "clist!i",simp,simp)
-    apply clarify
-    apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
-    apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
-    apply force
-   apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (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 (\<lambda>i. (fst i, s) # snd i) (zip xs clist)" in allE)
-apply simp
-apply clarify
-apply(simp add:cp_def)
-done
-
-theorem one: "xs\<noteq>[] \<Longrightarrow> 
- par_cp xs s = {c. \<exists>clist. (length clist)=(length xs) \<and> 
-               (\<forall>i<length clist. (clist!i) \<in> cp(xs!i) s) \<and> c \<propto> 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
--- a/src/HOL/HoareParallel/ROOT.ML	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,3 +0,0 @@
-(* $Id$ *)
-
-use_thys ["OG_Examples", "Gar_Coll", "Mul_Gar_Coll", "RG_Examples"];
--- a/src/HOL/HoareParallel/document/root.bib	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,28 +0,0 @@
-@inproceedings{NipkowP-FASE99,author={Tobias Nipkow and Prensa Nieto, Leonor},
-title={{Owicki/Gries} in {Isabelle/HOL}},
-booktitle={Fundamental Approaches to Software Engineering (FASE'99)},
-editor={J.-P. Finance},publisher="Springer",series="LNCS",volume=1577,
-pages={188--203},year=1999}
-
-@InProceedings{PrenEsp00,
-  author =    {Prensa Nieto, Leonor and Javier Esparza}, 
-  title =     {Verifying Single and Multi-mutator Garbage Collectors
-                with {Owicki/Gries} in {Isabelle/HOL}},
-  booktitle = {Mathematical Foundations of Computer Science (MFCS 2000)},
-  editor =    {M. Nielsen and B. Rovan},
-  publisher = {Springer-Verlag},
-  series =    {LNCS},
-  volume =    1893, 
-  pages =     {619--628},
-  year =      2000
-}
-
-@PhdThesis{Prensa-PhD,author={Leonor Prensa Nieto},
-title={Verification of Parallel Programs with the Owicki-Gries and
-Rely-Guarantee Methods in Isabelle/HOL},
-school={Technische Universit{\"a}t M{\"u}nchen},year=2002}
-
-@inproceedings{Prensa-ESOP03,author={Prensa Nieto, Leonor},
-title={The {Rely-Guarantee} Method in {Isabelle/HOL}},
-booktitle={European Symposium on Programming (ESOP'03)},editor={P. Degano},
-publisher=Springer,series=LNCS,volume=2618,pages={348--362},year=2003}
--- a/src/HOL/HoareParallel/document/root.tex	Mon Sep 21 11:15:21 2009 +0200
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,62 +0,0 @@
-
-% $Id$
-
-\documentclass[11pt,a4paper]{report}
-\usepackage{graphicx}
-\usepackage[english]{babel}
-\usepackage{isabelle,isabellesym}
-\usepackage{pdfsetup}
-
-\urlstyle{rm}
-\isabellestyle{it}
-
-\renewcommand{\isamarkupheader}[1]{#1}
-
-\begin{document}
-
-\title{Hoare Logic for Parallel Programs}
-\author{Leonor Prensa Nieto}
-\maketitle
-
-\begin{abstract}\noindent
-  In the following theories a formalization of the Owicki-Gries and
-  the rely-guarantee methods is presented. These methods are widely
-  used for correctness proofs of parallel imperative programs with
-  shared variables.  We define syntax, semantics and proof rules in
-  Isabelle/HOL.  The proof rules also provide for programs
-  parameterized in the number of parallel components. Their
-  correctness w.r.t.\ the semantics is proven.  Completeness proofs
-  for both methods are extended to the new case of parameterized
-  programs. (These proofs have not been formalized in Isabelle. They
-  can be found in~\cite{Prensa-PhD}.)  Using this formalizations we
-  verify several non-trivial examples for parameterized and
-  non-parameterized programs.  For the automatic generation of
-  verification conditions with the Owicki-Gries method we define a
-  tactic based on the proof rules.  The most involved examples are the
-  verification of two garbage-collection algorithms, the second one
-  parameterized in the number of mutators.
-
-For excellent descriptions of this work see
-\cite{NipkowP-FASE99,PrenEsp00,Prensa-PhD,Prensa-ESOP03}.
-
-\end{abstract}
-
-\pagestyle{plain}
-\thispagestyle{empty}
-\tableofcontents
-
-\clearpage
-
-\begin{center}
-  \includegraphics[scale=0.7]{session_graph}  
-\end{center}
-
-\newpage
-
-\parindent 0pt\parskip 0.5ex
-\input{session}
-
-\bibliographystyle{plain}
-\bibliography{root}
-
-\end{document}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/Gar_Coll.thy	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,846 @@
+
+header {* \section{The Single Mutator Case} *}
+
+theory Gar_Coll imports Graph OG_Syntax begin
+
+declare psubsetE [rule del]
+
+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 \<Rightarrow> bool"
+  "Mut_init \<equiv> \<guillemotleft> T \<in> Reach \<acute>E \<and> R < length \<acute>E \<and> T < length \<acute>M \<guillemotright>"
+
+text {* \noindent For the mutator we
+consider two modules, one for each action.  An auxiliary variable
+@{text "\<acute>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 \<equiv> .{\<acute>Mut_init \<and> \<acute>z}. \<langle>\<acute>E:=\<acute>E[R:=(fst(\<acute>E!R), T)],, \<acute>z:= (\<not>\<acute>z)\<rangle>"
+
+  Color_Target :: "gar_coll_state ann_com"
+  "Color_Target \<equiv> .{\<acute>Mut_init \<and> \<not>\<acute>z}. \<langle>\<acute>M:=\<acute>M[T:=Black],, \<acute>z:= (\<not>\<acute>z)\<rangle>"
+
+  Mutator :: "gar_coll_state ann_com"
+  "Mutator \<equiv>
+  .{\<acute>Mut_init \<and> \<acute>z}. 
+  WHILE True INV .{\<acute>Mut_init \<and> \<acute>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: 
+  "\<turnstile> Redirect_Edge pre(Color_Target)"
+apply (unfold mutator_defs)
+apply annhoare
+apply(simp_all)
+apply(force elim:Graph2)
+done
+
+lemma Color_Target:
+  "\<turnstile> Color_Target .{\<acute>Mut_init \<and> \<acute>z}."
+apply (unfold mutator_defs)
+apply annhoare
+apply(simp_all)
+done
+
+lemma Mutator: 
+ "\<turnstile> 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 "\<acute>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 \<Rightarrow> bool"
+  "Proper_M_init \<equiv>  \<guillemotleft> Blacks M_init=Roots \<and> length M_init=length \<acute>M \<guillemotright>"
+ 
+  Proper :: "gar_coll_state \<Rightarrow> bool"
+  "Proper \<equiv> \<guillemotleft> Proper_Roots \<acute>M \<and> Proper_Edges(\<acute>M, \<acute>E) \<and> \<acute>Proper_M_init \<guillemotright>"
+
+  Safe :: "gar_coll_state \<Rightarrow> bool"
+  "Safe \<equiv> \<guillemotleft> Reach \<acute>E \<subseteq> Blacks \<acute>M \<guillemotright>"
+
+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 \<equiv> 
+  .{\<acute>Proper}.
+  \<acute>ind:=0;;
+  .{\<acute>Proper \<and> \<acute>ind=0}.
+  WHILE \<acute>ind<length \<acute>M 
+   INV .{\<acute>Proper \<and> (\<forall>i<\<acute>ind. i \<in> Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind\<le>length \<acute>M}.
+  DO .{\<acute>Proper \<and> (\<forall>i<\<acute>ind. i \<in> Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M}.
+   IF \<acute>ind\<in>Roots THEN 
+   .{\<acute>Proper \<and> (\<forall>i<\<acute>ind. i \<in> Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M \<and> \<acute>ind\<in>Roots}. 
+    \<acute>M:=\<acute>M[\<acute>ind:=Black] FI;;
+   .{\<acute>Proper \<and> (\<forall>i<\<acute>ind+1. i \<in> Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M}.
+    \<acute>ind:=\<acute>ind+1 
+  OD"
+
+lemma Blacken_Roots: 
+ "\<turnstile> Blacken_Roots .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>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 force
+apply force
+done
+
+subsubsection {* Propagating black *}
+
+constdefs
+  PBInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
+  "PBInv \<equiv> \<guillemotleft> \<lambda>ind. \<acute>obc < Blacks \<acute>M \<or> (\<forall>i <ind. \<not>BtoW (\<acute>E!i, \<acute>M) \<or>
+   (\<not>\<acute>z \<and> i=R \<and> (snd(\<acute>E!R)) = T \<and> (\<exists>r. ind \<le> r \<and> r < length \<acute>E \<and> BtoW(\<acute>E!r,\<acute>M))))\<guillemotright>"
+
+constdefs  
+  Propagate_Black_aux :: "gar_coll_state ann_com"
+  "Propagate_Black_aux \<equiv>
+  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M}.
+  \<acute>ind:=0;;
+  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M \<and> \<acute>ind=0}. 
+  WHILE \<acute>ind<length \<acute>E 
+   INV .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+         \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind\<le>length \<acute>E}.
+  DO .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+       \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E}. 
+   IF \<acute>M!(fst (\<acute>E!\<acute>ind)) = Black THEN 
+    .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+       \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E \<and> \<acute>M!fst(\<acute>E!\<acute>ind)=Black}.
+     \<acute>M:=\<acute>M[snd(\<acute>E!\<acute>ind):=Black];;
+    .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+       \<and> \<acute>PBInv (\<acute>ind + 1) \<and> \<acute>ind<length \<acute>E}.
+     \<acute>ind:=\<acute>ind+1
+   FI
+  OD"
+
+lemma Propagate_Black_aux: 
+  "\<turnstile>  Propagate_Black_aux
+  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+    \<and> ( \<acute>obc < Blacks \<acute>M \<or> \<acute>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_tac y = r in le_imp_less_or_eq)
+  apply (erule disjE)
+   apply (subgoal_tac "Suc (ind x)\<le>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 (asm_lr))
+ 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 \<Rightarrow> bool"
+  "Auxk \<equiv> \<guillemotleft>\<acute>k<length \<acute>M \<and> (\<acute>M!\<acute>k\<noteq>Black \<or> \<not>BtoW(\<acute>E!\<acute>ind, \<acute>M) \<or> 
+          \<acute>obc<Blacks \<acute>M \<or> (\<not>\<acute>z \<and> \<acute>ind=R \<and> snd(\<acute>E!R)=T  
+          \<and> (\<exists>r. \<acute>ind<r \<and> r<length \<acute>E \<and> BtoW(\<acute>E!r, \<acute>M))))\<guillemotright>"
+
+constdefs  
+  Propagate_Black :: " gar_coll_state ann_com"
+  "Propagate_Black \<equiv>
+  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M}.
+  \<acute>ind:=0;;
+  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M \<and> \<acute>ind=0}.
+  WHILE \<acute>ind<length \<acute>E 
+   INV .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+         \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind\<le>length \<acute>E}.
+  DO .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+       \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E}. 
+   IF (\<acute>M!(fst (\<acute>E!\<acute>ind)))=Black THEN 
+    .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+      \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E \<and> (\<acute>M!fst(\<acute>E!\<acute>ind))=Black}.
+     \<acute>k:=(snd(\<acute>E!\<acute>ind));;
+    .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+      \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E \<and> (\<acute>M!fst(\<acute>E!\<acute>ind))=Black 
+      \<and> \<acute>Auxk}.
+     \<langle>\<acute>M:=\<acute>M[\<acute>k:=Black],, \<acute>ind:=\<acute>ind+1\<rangle>
+   ELSE .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+          \<and> \<acute>PBInv \<acute>ind \<and> \<acute>ind<length \<acute>E}. 
+         \<langle>IF (\<acute>M!(fst (\<acute>E!\<acute>ind)))\<noteq>Black THEN \<acute>ind:=\<acute>ind+1 FI\<rangle> 
+   FI
+  OD"
+
+lemma Propagate_Black: 
+  "\<turnstile>  Propagate_Black
+  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+    \<and> ( \<acute>obc < Blacks \<acute>M \<or> \<acute>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_tac y = r in le_imp_less_or_eq)
+   apply (erule disjE)
+    apply (subgoal_tac "Suc (ind x)\<le>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\<le>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_tac y = r in le_imp_less_or_eq)
+ apply (erule disjE)
+  apply (subgoal_tac "Suc (ind x)\<le>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 (asm_lr))
+ 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 \<Rightarrow> nat \<Rightarrow> bool"
+  "CountInv \<equiv> \<guillemotleft> \<lambda>ind. {i. i<ind \<and> \<acute>Ma!i=Black}\<subseteq>\<acute>bc \<guillemotright>"
+
+constdefs
+  Count :: " gar_coll_state ann_com"
+  "Count \<equiv>
+  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
+    \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+    \<and> length \<acute>Ma=length \<acute>M \<and> (\<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>bc={}}.
+  \<acute>ind:=0;;
+  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
+    \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+   \<and> length \<acute>Ma=length \<acute>M \<and> (\<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>bc={} 
+   \<and> \<acute>ind=0}.
+   WHILE \<acute>ind<length \<acute>M 
+     INV .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
+           \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+           \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>CountInv \<acute>ind
+           \<and> ( \<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>ind\<le>length \<acute>M}.
+   DO .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
+         \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+         \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>CountInv \<acute>ind 
+         \<and> ( \<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>ind<length \<acute>M}. 
+       IF \<acute>M!\<acute>ind=Black 
+          THEN .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
+                 \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+                 \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>CountInv \<acute>ind
+                 \<and> ( \<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>ind<length \<acute>M \<and> \<acute>M!\<acute>ind=Black}.
+          \<acute>bc:=insert \<acute>ind \<acute>bc
+       FI;;
+      .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
+        \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+        \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>CountInv (\<acute>ind+1)
+        \<and> ( \<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe) \<and> \<acute>ind<length \<acute>M}.
+      \<acute>ind:=\<acute>ind+1
+   OD"
+
+lemma Count: 
+  "\<turnstile> Count 
+  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
+   \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>\<acute>bc \<and> \<acute>bc\<subseteq>Blacks \<acute>M \<and> length \<acute>Ma=length \<acute>M
+   \<and> (\<acute>obc < Blacks \<acute>Ma \<or> \<acute>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 \<times> edges \<Rightarrow> edges"
+
+axioms
+  Append_to_free0: "length (Append_to_free (i, e)) = length e"
+  Append_to_free1: "Proper_Edges (m, e) 
+                   \<Longrightarrow> Proper_Edges (m, Append_to_free(i, e))"
+  Append_to_free2: "i \<notin> Reach e 
+     \<Longrightarrow> n \<in> Reach (Append_to_free(i, e)) = ( n = i \<or> n \<in> Reach e)"
+
+constdefs
+  AppendInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
+  "AppendInv \<equiv> \<guillemotleft>\<lambda>ind. \<forall>i<length \<acute>M. ind\<le>i \<longrightarrow> i\<in>Reach \<acute>E \<longrightarrow> \<acute>M!i=Black\<guillemotright>"
+
+constdefs
+  Append :: " gar_coll_state ann_com"
+   "Append \<equiv>
+  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe}.
+  \<acute>ind:=0;;
+  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe \<and> \<acute>ind=0}.
+    WHILE \<acute>ind<length \<acute>M 
+      INV .{\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind\<le>length \<acute>M}.
+    DO .{\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M}.
+       IF \<acute>M!\<acute>ind=Black THEN 
+          .{\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>M!\<acute>ind=Black}. 
+          \<acute>M:=\<acute>M[\<acute>ind:=White] 
+       ELSE .{\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>ind\<notin>Reach \<acute>E}.
+              \<acute>E:=Append_to_free(\<acute>ind,\<acute>E)
+       FI;;
+     .{\<acute>Proper \<and> \<acute>AppendInv (\<acute>ind+1) \<and> \<acute>ind<length \<acute>M}. 
+       \<acute>ind:=\<acute>ind+1
+    OD"
+
+lemma Append: 
+  "\<turnstile> Append .{\<acute>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 \<equiv>
+.{\<acute>Proper}.  
+ WHILE True INV .{\<acute>Proper}. 
+ DO  
+  Blacken_Roots;; 
+  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M}.  
+   \<acute>obc:={};; 
+  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={}}. 
+   \<acute>bc:=Roots;; 
+  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots}. 
+   \<acute>Ma:=M_init;;  
+  .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots \<and> \<acute>Ma=M_init}. 
+   WHILE \<acute>obc\<noteq>\<acute>bc  
+     INV .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M 
+           \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>\<acute>bc \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+           \<and> length \<acute>Ma=length \<acute>M \<and> (\<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe)}. 
+   DO .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M}.
+       \<acute>obc:=\<acute>bc;;
+       Propagate_Black;; 
+      .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+        \<and> (\<acute>obc < Blacks \<acute>M \<or> \<acute>Safe)}. 
+       \<acute>Ma:=\<acute>M;;
+      .{\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma 
+        \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M \<and> length \<acute>Ma=length \<acute>M 
+        \<and> ( \<acute>obc < Blacks \<acute>Ma \<or> \<acute>Safe)}.
+       \<acute>bc:={};;
+       Count 
+   OD;; 
+  Append  
+ OD"
+
+lemma Collector: 
+  "\<turnstile> 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)
+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+
+--{* 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+
+--{* 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 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(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+
+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 subgoal 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 @{claset}) *})
+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 @{clasimpset}, assume_tac]) *})
+apply(tactic {* TRYALL(EVERY'[rtac disjI2,etac subset_trans,rtac @{thm Graph9},force_tac @{clasimpset}]) *})
+apply(tactic {* TRYALL(EVERY'[rtac disjI1,etac @{thm psubset_subset_trans},rtac @{thm Graph9},force_tac @{clasimpset}]) *})
+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 @{claset}) *})
+--{* 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: 
+  "\<parallel>- .{\<acute>Proper \<and> \<acute>Mut_init \<and> \<acute>z}.  
+  COBEGIN  
+   Collector
+  .{False}.
+ \<parallel>  
+   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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/Graph.thy	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,410 @@
+header {* \chapter{Case Study: Single and Multi-Mutator Garbage Collection Algorithms}
+
+\section {Formalization of the Memory} *}
+
+theory Graph imports Main begin
+
+datatype node = Black | White
+
+types 
+  nodes = "node list"
+  edge  = "nat \<times> nat"
+  edges = "edge list"
+
+consts Roots :: "nat set"
+
+constdefs
+  Proper_Roots :: "nodes \<Rightarrow> bool"
+  "Proper_Roots M \<equiv> Roots\<noteq>{} \<and> Roots \<subseteq> {i. i<length M}"
+
+  Proper_Edges :: "(nodes \<times> edges) \<Rightarrow> bool"
+  "Proper_Edges \<equiv> (\<lambda>(M,E). \<forall>i<length E. fst(E!i)<length M \<and> snd(E!i)<length M)"
+
+  BtoW :: "(edge \<times> nodes) \<Rightarrow> bool"
+  "BtoW \<equiv> (\<lambda>(e,M). (M!fst e)=Black \<and> (M!snd e)\<noteq>Black)"
+
+  Blacks :: "nodes \<Rightarrow> nat set"
+  "Blacks M \<equiv> {i. i<length M \<and> M!i=Black}"
+
+  Reach :: "edges \<Rightarrow> nat set"
+  "Reach E \<equiv> {x. (\<exists>path. 1<length path \<and> path!(length path - 1)\<in>Roots \<and> x=path!0
+              \<and> (\<forall>i<length path - 1. (\<exists>j<length E. E!j=(path!(i+1), path!i))))
+	      \<or> x\<in>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]: 
+  "\<lbrakk> Roots\<subseteq>Blacks M; \<forall>i<length E. \<not>BtoW(E!i,M)\<rbrakk>
+  \<Longrightarrow> 1< length path \<longrightarrow> (path!(length path - 1))\<in>Roots \<longrightarrow>  
+  (\<forall>i<length path - 1. (\<exists>j. j < length E \<and> E!j=(path!(Suc i), path!i))) 
+  \<longrightarrow> 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 -2)
+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: 
+  "\<lbrakk>Roots\<subseteq>Blacks M; Proper_Edges(M, E); \<forall>i<length E. \<not>BtoW(E!i,M) \<rbrakk> 
+  \<Longrightarrow> Reach E\<subseteq>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) \<longrightarrow> (\<exists>m. P m \<and> (\<forall>i. i<m \<longrightarrow> \<not> P i))";
+apply(rule nat_less_induct)
+apply clarify
+apply(case_tac "\<forall>m. m<n \<longrightarrow> \<not> P m")
+apply auto
+done
+
+lemma Compl_lemma: "(n::nat)\<le>l \<Longrightarrow> (\<exists>m. m\<le>l \<and> n=l - m)"
+apply(rule_tac x = "l - n" in exI)
+apply arith
+done
+
+lemma Ex_last_occurrence: 
+  "\<lbrakk>P (n::nat); n\<le>l\<rbrakk> \<Longrightarrow> (\<exists>m. P (l - m) \<and> (\<forall>i. i<m \<longrightarrow> \<not>P (l - i)))"
+apply(drule Compl_lemma)
+apply clarify
+apply(erule Ex_first_occurrence)
+done
+
+lemma Graph2: 
+  "\<lbrakk>T \<in> Reach E; R<length E\<rbrakk> \<Longrightarrow> T \<in> Reach (E[R:=(fst(E!R), T)])"
+apply (unfold Reach_def)
+apply clarify
+apply simp
+apply(case_tac "\<forall>z<length path. fst(E!R)\<noteq>path!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 (force simp add:nth_list_update)
+apply simp
+apply(erule exE)
+apply(subgoal_tac "z \<le> length path - Suc 0")
+ prefer 2 apply arith
+apply(drule_tac P = "\<lambda>m. m<length path \<and> 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 \<le> 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 \<le> length path")
+ prefer 2 apply arith
+apply(subgoal_tac "(length path - Suc m) + (Suc nata) \<le> 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(simp del: diff_is_0_eq)
+ apply(subgoal_tac "Suc nata\<le>nat")
+ prefer 2 apply arith
+ apply(drule_tac n = "Suc nata" in Compl_lemma)
+ apply clarify
+ using [[linarith_split_limit = 0]]
+ apply force
+ using [[linarith_split_limit = 9]]
+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 simp
+done
+
+
+subsubsection{* Graph 3 *}
+
+lemma Graph3: 
+  "\<lbrakk> T\<in>Reach E; R<length E \<rbrakk> \<Longrightarrow> Reach(E[R:=(fst(E!R),T)]) \<subseteq> Reach E"
+apply (unfold Reach_def)
+apply clarify
+apply simp
+apply(case_tac "\<exists>i<length path - 1. (fst(E!R),T)=(path!(Suc i),path!i)")
+--{* the changed edge is part of the path *}
+ apply(erule exE)
+ apply(drule_tac P = "\<lambda>i. i<length path - 1 \<and> (fst(E!R),T)=(path!Suc i,path!i)" in Ex_first_occurrence)
+ apply clarify
+ apply(erule disjE)
+--{* T is NOT a root *}
+  apply clarify
+  apply(rule_tac x = "(take m path)@patha" in exI)
+  apply(subgoal_tac "\<not>(length path\<le>m)")
+   prefer 2 apply arith
+  apply(simp)
+  apply(rule conjI)
+   apply(subgoal_tac "\<not>(m + length patha - 1 < m)")
+    prefer 2 apply arith
+   apply(simp add: nth_append)
+  apply(rule conjI)
+   apply(case_tac "m")
+    apply force
+   apply(case_tac "path")
+    apply force
+   apply force
+  apply clarify
+  apply(case_tac "Suc i\<le>m")
+   apply(erule_tac x = "i" in allE)
+   apply simp
+   apply clarify
+   apply(rule_tac x = "j" in exI)
+   apply(case_tac "Suc i<m")
+    apply(simp add: nth_append)
+    apply(case_tac "R=j")
+     apply(simp add: nth_list_update)
+     apply(case_tac "i=m")
+      apply force
+     apply(erule_tac x = "i" in allE)
+     apply force
+    apply(force simp add: nth_list_update)
+   apply(simp add: nth_append)
+   apply(subgoal_tac "i=m - 1")
+    prefer 2 apply arith
+   apply(case_tac "R=j")
+    apply(erule_tac x = "m - 1" in allE)
+    apply(simp add: nth_list_update)
+   apply(force simp add: nth_list_update)
+  apply(simp add: nth_append)
+  apply(rotate_tac -4)
+  apply(erule_tac x = "i - m" in allE)
+  apply(subgoal_tac "Suc (i - m)=(Suc i - m)" )
+    prefer 2 apply arith
+   apply simp
+--{* T is a root *}
+ apply(case_tac "m=0")
+  apply force
+ apply(rule_tac x = "take (Suc m) path" in exI)
+ apply(subgoal_tac "\<not>(length path\<le>Suc m)" )
+  prefer 2 apply arith
+ apply clarsimp
+ 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: 
+  "\<lbrakk>T \<in> Reach E; Roots\<subseteq>Blacks M; I\<le>length E; T<length M; R<length E; 
+  \<forall>i<I. \<not>BtoW(E!i,M); R<I; M!fst(E!R)=Black; M!T\<noteq>Black\<rbrakk> \<Longrightarrow> 
+  (\<exists>r. I\<le>r \<and> r<length E \<and> 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 "\<exists>m<length path. M!(path!m)=Black")
+ apply(erule exE)
+ apply(drule_tac P = "\<lambda>m. m<length path \<and> M!(path!m)=Black" in Ex_first_occurrence)
+ apply clarify
+ apply(case_tac "ma")
+  apply force
+ apply simp
+ apply(case_tac "length path")
+  apply force
+ apply simp
+ apply(erule_tac P = "\<lambda>i. i < nata \<longrightarrow> ?P i" and x = "nat" in allE)
+ apply simp
+ apply clarify
+ apply(erule_tac P = "\<lambda>i. i < Suc nat \<longrightarrow> ?P i" and x = "nat" in allE)
+ apply simp
+ apply(case_tac "j<I")
+  apply(erule_tac x = "j" in allE)
+  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 5 *}
+
+lemma Graph5: 
+  "\<lbrakk> T \<in> Reach E ; Roots \<subseteq> Blacks M; \<forall>i<R. \<not>BtoW(E!i,M); T<length M; 
+    R<length E; M!fst(E!R)=Black; M!snd(E!R)=Black; M!T \<noteq> Black\<rbrakk> 
+   \<Longrightarrow> (\<exists>r. R<r \<and> r<length E \<and> 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 "\<exists>m<length path. M!(path!m)=Black")
+ apply(erule exE)
+ apply(drule_tac P = "\<lambda>m. m<length path \<and> M!(path!m)=Black" in Ex_first_occurrence)
+ apply clarify
+ apply(case_tac "ma")
+  apply force
+ apply simp
+ apply(case_tac "length path")
+  apply force
+ apply simp
+ apply(erule_tac P = "\<lambda>i. i < nata \<longrightarrow> ?P i" and x = "nat" in allE)
+ apply simp
+ apply clarify
+ apply(erule_tac P = "\<lambda>i. i < Suc nat \<longrightarrow> ?P i" and x = "nat" in allE)
+ apply simp
+ apply(case_tac "j\<le>R")
+  apply(drule le_imp_less_or_eq [of _ R])
+  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 {* Other lemmas about graphs *}
+
+lemma Graph6: 
+ "\<lbrakk>Proper_Edges(M,E); R<length E ; T<length M\<rbrakk> \<Longrightarrow> Proper_Edges(M,E[R:=(fst(E!R),T)])"
+apply (unfold Proper_Edges_def)
+ apply(force  simp add: nth_list_update)
+done
+
+lemma Graph7: 
+ "\<lbrakk>Proper_Edges(M,E)\<rbrakk> \<Longrightarrow> Proper_Edges(M[T:=a],E)"
+apply (unfold Proper_Edges_def)
+apply force
+done
+
+lemma Graph8: 
+ "\<lbrakk>Proper_Roots(M)\<rbrakk> \<Longrightarrow> 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<length M \<Longrightarrow> Blacks M\<subseteq>Blacks (M[j := Black])"
+apply (unfold Blacks_def)
+ apply(force simp add: nth_list_update)
+done
+
+lemma Graph10 [rule_format (no_asm)]: "\<forall>i. M!i=a \<longrightarrow>M[i:=a]=M"
+apply(induct_tac "M")
+apply auto
+apply(case_tac "i")
+apply auto
+done
+
+lemma Graph11 [rule_format (no_asm)]: 
+  "\<lbrakk> M!j\<noteq>Black;j<length M\<rbrakk> \<Longrightarrow> Blacks M \<subset> 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: "\<lbrakk>a\<subseteq>Blacks M;j<length M\<rbrakk> \<Longrightarrow> a\<subseteq>Blacks (M[j := Black])"
+apply (unfold Blacks_def)
+apply(force simp add: nth_list_update)
+done
+
+lemma Graph13: "\<lbrakk>a\<subset> Blacks M;j<length M\<rbrakk> \<Longrightarrow> a \<subset> 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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/Hoare_Parallel.thy	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,5 @@
+theory Hoare_Parallel
+imports OG_Examples Gar_Coll Mul_Gar_Coll RG_Examples
+begin
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,1283 @@
+
+header {* \section{The Multi-Mutator Case} *}
+
+theory Mul_Gar_Coll imports Graph OG_Syntax begin
+
+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 \<Rightarrow> nat \<Rightarrow> bool"
+  "Mul_mut_init \<equiv> \<guillemotleft> \<lambda>n. n=length \<acute>Muts \<and> (\<forall>i<n. R (\<acute>Muts!i)<length \<acute>E 
+                          \<and> T (\<acute>Muts!i)<length \<acute>M) \<guillemotright>"
+
+  Mul_Redirect_Edge  :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com"
+  "Mul_Redirect_Edge j n \<equiv>
+  .{\<acute>Mul_mut_init n \<and> Z (\<acute>Muts!j)}.
+  \<langle>IF T(\<acute>Muts!j) \<in> Reach \<acute>E THEN  
+  \<acute>E:= \<acute>E[R (\<acute>Muts!j):= (fst (\<acute>E!R(\<acute>Muts!j)), T (\<acute>Muts!j))] FI,, 
+  \<acute>Muts:= \<acute>Muts[j:= (\<acute>Muts!j) \<lparr>Z:=False\<rparr>]\<rangle>"
+
+  Mul_Color_Target :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com"
+  "Mul_Color_Target j n \<equiv>
+  .{\<acute>Mul_mut_init n \<and> \<not> Z (\<acute>Muts!j)}. 
+  \<langle>\<acute>M:=\<acute>M[T (\<acute>Muts!j):=Black],, \<acute>Muts:=\<acute>Muts[j:= (\<acute>Muts!j) \<lparr>Z:=True\<rparr>]\<rangle>"
+
+  Mul_Mutator :: "nat \<Rightarrow> nat \<Rightarrow>  mul_gar_coll_state ann_com"
+  "Mul_Mutator j n \<equiv>
+  .{\<acute>Mul_mut_init n \<and> Z (\<acute>Muts!j)}.  
+  WHILE True  
+    INV .{\<acute>Mul_mut_init n \<and> Z (\<acute>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\<le>j \<and> j<n \<Longrightarrow> 
+  \<turnstile> 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\<le>j \<and> j<n \<Longrightarrow> 
+  \<turnstile>  Mul_Color_Target j n  
+    .{\<acute>Mul_mut_init n \<and> Z (\<acute>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\<le>j \<and> j<n \<Longrightarrow>  
+ \<turnstile> 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: 
+  "\<lbrakk>0\<le>i; i<n; 0\<le>j; j<n; i\<noteq>j\<rbrakk> \<Longrightarrow>  
+  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)
+done
+
+lemma Mul_interfree_Redirect_Edge_Color_Target: 
+  "\<lbrakk>0\<le>i; i<n; 0\<le>j; j<n; i\<noteq>j\<rbrakk> \<Longrightarrow>  
+  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: 
+  "\<lbrakk>0\<le>i; i<n; 0\<le>j; j<n; i\<noteq>j\<rbrakk> \<Longrightarrow> 
+  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)
+done
+
+lemma Mul_interfree_Color_Target_Color_Target: 
+  " \<lbrakk>0\<le>i; i<n; 0\<le>j; j<n; i\<noteq>j\<rbrakk> \<Longrightarrow> 
+  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)
+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: "\<lbrakk>i < n; j < n; i \<noteq> j\<rbrakk> \<Longrightarrow> 
+  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 @{claset}) *})
+apply (simp_all add:nth_list_update)
+done
+
+subsubsection {* Modular Parameterized Mutators *}
+
+lemma Mul_Parameterized_Mutators: "0<n \<Longrightarrow>
+ \<parallel>- .{\<acute>Mul_mut_init n \<and> (\<forall>i<n. Z (\<acute>Muts!i))}.
+ COBEGIN
+ SCHEME  [0\<le> 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 \<Rightarrow> nat"
+ "Queue \<equiv> \<guillemotleft> length (filter (\<lambda>i. \<not> Z i \<and> \<acute>M!(T i) \<noteq> Black) \<acute>Muts) \<guillemotright>"
+
+consts  M_init :: nodes
+
+constdefs
+  Proper_M_init :: "mul_gar_coll_state \<Rightarrow> bool"
+  "Proper_M_init \<equiv> \<guillemotleft> Blacks M_init=Roots \<and> length M_init=length \<acute>M \<guillemotright>"
+
+  Mul_Proper :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
+  "Mul_Proper \<equiv> \<guillemotleft> \<lambda>n. Proper_Roots \<acute>M \<and> Proper_Edges (\<acute>M, \<acute>E) \<and> \<acute>Proper_M_init \<and> n=length \<acute>Muts \<guillemotright>"
+
+  Safe :: "mul_gar_coll_state \<Rightarrow> bool"
+  "Safe \<equiv> \<guillemotleft> Reach \<acute>E \<subseteq> Blacks \<acute>M \<guillemotright>"
+
+lemmas mul_collector_defs = Proper_M_init_def Mul_Proper_def Safe_def
+
+subsubsection {* Blackening Roots *}
+
+constdefs
+  Mul_Blacken_Roots :: "nat \<Rightarrow>  mul_gar_coll_state ann_com"
+  "Mul_Blacken_Roots n \<equiv>
+  .{\<acute>Mul_Proper n}.
+  \<acute>ind:=0;;
+  .{\<acute>Mul_Proper n \<and> \<acute>ind=0}.
+  WHILE \<acute>ind<length \<acute>M 
+    INV .{\<acute>Mul_Proper n \<and> (\<forall>i<\<acute>ind. i\<in>Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind\<le>length \<acute>M}.
+  DO .{\<acute>Mul_Proper n \<and> (\<forall>i<\<acute>ind. i\<in>Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M}.
+       IF \<acute>ind\<in>Roots THEN 
+     .{\<acute>Mul_Proper n \<and> (\<forall>i<\<acute>ind. i\<in>Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M \<and> \<acute>ind\<in>Roots}. 
+       \<acute>M:=\<acute>M[\<acute>ind:=Black] FI;;
+     .{\<acute>Mul_Proper n \<and> (\<forall>i<\<acute>ind+1. i\<in>Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M}.
+       \<acute>ind:=\<acute>ind+1 
+  OD"
+
+lemma Mul_Blacken_Roots: 
+  "\<turnstile> Mul_Blacken_Roots n  
+  .{\<acute>Mul_Proper n \<and> Roots \<subseteq> Blacks \<acute>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 force
+apply force
+done
+
+subsubsection {* Propagating Black *} 
+
+constdefs
+  Mul_PBInv :: "mul_gar_coll_state \<Rightarrow> bool"
+  "Mul_PBInv \<equiv>  \<guillemotleft>\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>M \<or> \<acute>l<\<acute>Queue 
+                 \<or> (\<forall>i<\<acute>ind. \<not>BtoW(\<acute>E!i,\<acute>M)) \<and> \<acute>l\<le>\<acute>Queue\<guillemotright>"
+
+  Mul_Auxk :: "mul_gar_coll_state \<Rightarrow> bool"
+  "Mul_Auxk \<equiv> \<guillemotleft>\<acute>l<\<acute>Queue \<or> \<acute>M!\<acute>k\<noteq>Black \<or> \<not>BtoW(\<acute>E!\<acute>ind, \<acute>M) \<or> \<acute>obc\<subset>Blacks \<acute>M\<guillemotright>"
+
+constdefs
+  Mul_Propagate_Black :: "nat \<Rightarrow>  mul_gar_coll_state ann_com"
+  "Mul_Propagate_Black n \<equiv>
+ .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+  \<and> (\<acute>Safe \<or> \<acute>l\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M)}. 
+ \<acute>ind:=0;;
+ .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+   \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> Blacks \<acute>M\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+   \<and> (\<acute>Safe \<or> \<acute>l\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M) \<and> \<acute>ind=0}. 
+ WHILE \<acute>ind<length \<acute>E 
+  INV .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+        \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+        \<and> \<acute>Mul_PBInv \<and> \<acute>ind\<le>length \<acute>E}.
+ DO .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+     \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+     \<and> \<acute>Mul_PBInv \<and> \<acute>ind<length \<acute>E}.
+   IF \<acute>M!(fst (\<acute>E!\<acute>ind))=Black THEN 
+   .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+     \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+     \<and> \<acute>Mul_PBInv \<and> (\<acute>M!fst(\<acute>E!\<acute>ind))=Black \<and> \<acute>ind<length \<acute>E}.
+    \<acute>k:=snd(\<acute>E!\<acute>ind);;
+   .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+     \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+     \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>M \<or> \<acute>l<\<acute>Queue \<or> (\<forall>i<\<acute>ind. \<not>BtoW(\<acute>E!i,\<acute>M)) 
+        \<and> \<acute>l\<le>\<acute>Queue \<and> \<acute>Mul_Auxk ) \<and> \<acute>k<length \<acute>M \<and> \<acute>M!fst(\<acute>E!\<acute>ind)=Black 
+     \<and> \<acute>ind<length \<acute>E}.
+   \<langle>\<acute>M:=\<acute>M[\<acute>k:=Black],,\<acute>ind:=\<acute>ind+1\<rangle>
+   ELSE .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+         \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+         \<and> \<acute>Mul_PBInv \<and> \<acute>ind<length \<acute>E}.
+	 \<langle>IF \<acute>M!(fst (\<acute>E!\<acute>ind))\<noteq>Black THEN \<acute>ind:=\<acute>ind+1 FI\<rangle> FI
+ OD"
+
+lemma Mul_Propagate_Black: 
+  "\<turnstile> Mul_Propagate_Black n  
+   .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+     \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>M \<or> \<acute>l<\<acute>Queue \<and> (\<acute>l\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>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 subgoal 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 \<Rightarrow> nat \<Rightarrow> bool"
+ "Mul_CountInv \<equiv> \<guillemotleft> \<lambda>ind. {i. i<ind \<and> \<acute>Ma!i=Black}\<subseteq>\<acute>bc \<guillemotright>"
+
+  Mul_Count :: "nat \<Rightarrow>  mul_gar_coll_state ann_com"
+  "Mul_Count n \<equiv> 
+  .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+    \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+    \<and> length \<acute>Ma=length \<acute>M 
+    \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M) ) 
+    \<and> \<acute>q<n+1 \<and> \<acute>bc={}}.
+  \<acute>ind:=0;;
+  .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+    \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+    \<and> length \<acute>Ma=length \<acute>M 
+    \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M) ) 
+    \<and> \<acute>q<n+1 \<and> \<acute>bc={} \<and> \<acute>ind=0}.
+  WHILE \<acute>ind<length \<acute>M 
+     INV .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+          \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M  
+          \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>Mul_CountInv \<acute>ind 
+          \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M))
+	  \<and> \<acute>q<n+1 \<and> \<acute>ind\<le>length \<acute>M}.
+  DO .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+       \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+       \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>Mul_CountInv \<acute>ind 
+       \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M))
+       \<and> \<acute>q<n+1 \<and> \<acute>ind<length \<acute>M}. 
+     IF \<acute>M!\<acute>ind=Black 
+     THEN .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+            \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M  
+            \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>Mul_CountInv \<acute>ind 
+            \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M))
+            \<and> \<acute>q<n+1 \<and> \<acute>ind<length \<acute>M \<and> \<acute>M!\<acute>ind=Black}.
+          \<acute>bc:=insert \<acute>ind \<acute>bc
+     FI;;
+  .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+    \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+    \<and> length \<acute>Ma=length \<acute>M \<and> \<acute>Mul_CountInv (\<acute>ind+1) 
+    \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M))
+    \<and> \<acute>q<n+1 \<and> \<acute>ind<length \<acute>M}.
+  \<acute>ind:=\<acute>ind+1
+  OD"
+ 
+lemma Mul_Count: 
+  "\<turnstile> Mul_Count n  
+  .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+    \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+    \<and> length \<acute>Ma=length \<acute>M \<and> Blacks \<acute>Ma\<subseteq>\<acute>bc 
+    \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M)) 
+    \<and> \<acute>q<n+1}."
+apply (unfold Mul_Count_def)
+apply annhoare
+apply(simp_all add:Mul_CountInv_def mul_collector_defs Mul_Auxk_def Graph6 Graph7 Graph8 Graph12 mul_collector_defs Queue_def)
+--{* 7 subgoals left *}
+apply force
+apply force
+apply force
+--{* 4 subgoals left *}
+apply clarify
+apply(conjI_tac)
+apply(disjE_tac)
+ apply simp_all
+apply(simp add:Blacks_def)
+apply clarify
+apply(erule less_SucE)
+ back
+ apply force
+apply force
+--{* 3 subgoals left *}
+apply clarify
+apply(conjI_tac)
+apply(disjE_tac)
+ apply simp_all
+apply clarify
+apply(erule less_SucE)
+ back
+ apply force
+apply simp
+apply(rotate_tac -1)
+apply (force simp add:Blacks_def)
+--{* 2 subgoals left *}
+apply force
+--{* 1 subgoal left *}
+apply clarify
+apply(drule_tac x = "ind x" in le_imp_less_or_eq)
+apply (simp_all add:Blacks_def)
+done
+
+subsubsection {* Appending garbage nodes to the free list *}
+
+consts  Append_to_free :: "nat \<times> edges \<Rightarrow> edges"
+
+axioms
+  Append_to_free0: "length (Append_to_free (i, e)) = length e"
+  Append_to_free1: "Proper_Edges (m, e) 
+                    \<Longrightarrow> Proper_Edges (m, Append_to_free(i, e))"
+  Append_to_free2: "i \<notin> Reach e 
+           \<Longrightarrow> n \<in> Reach (Append_to_free(i, e)) = ( n = i \<or> n \<in> Reach e)"
+
+constdefs
+  Mul_AppendInv :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool"
+  "Mul_AppendInv \<equiv> \<guillemotleft> \<lambda>ind. (\<forall>i. ind\<le>i \<longrightarrow> i<length \<acute>M \<longrightarrow> i\<in>Reach \<acute>E \<longrightarrow> \<acute>M!i=Black)\<guillemotright>"
+
+  Mul_Append :: "nat \<Rightarrow>  mul_gar_coll_state ann_com"
+  "Mul_Append n \<equiv> 
+  .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe}.
+  \<acute>ind:=0;;
+  .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe \<and> \<acute>ind=0}.
+  WHILE \<acute>ind<length \<acute>M 
+    INV .{\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv \<acute>ind \<and> \<acute>ind\<le>length \<acute>M}.
+  DO .{\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M}.
+      IF \<acute>M!\<acute>ind=Black THEN 
+     .{\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>M!\<acute>ind=Black}. 
+      \<acute>M:=\<acute>M[\<acute>ind:=White] 
+      ELSE 
+     .{\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>ind\<notin>Reach \<acute>E}. 
+      \<acute>E:=Append_to_free(\<acute>ind,\<acute>E)
+      FI;;
+  .{\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv (\<acute>ind+1) \<and> \<acute>ind<length \<acute>M}. 
+   \<acute>ind:=\<acute>ind+1
+  OD"
+
+lemma Mul_Append: 
+  "\<turnstile> Mul_Append n  
+     .{\<acute>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 \<Rightarrow>  mul_gar_coll_state ann_com"
+  "Mul_Collector n \<equiv>
+.{\<acute>Mul_Proper n}.  
+WHILE True INV .{\<acute>Mul_Proper n}. 
+DO  
+Mul_Blacken_Roots n ;; 
+.{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M}.  
+ \<acute>obc:={};; 
+.{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={}}.  
+ \<acute>bc:=Roots;; 
+.{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots}. 
+ \<acute>l:=0;; 
+.{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots \<and> \<acute>l=0}. 
+ WHILE \<acute>l<n+1  
+   INV .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M \<and>  
+         (\<acute>Safe \<or> (\<acute>l\<le>\<acute>Queue \<or> \<acute>bc\<subset>Blacks \<acute>M) \<and> \<acute>l<n+1)}. 
+ DO .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+      \<and> (\<acute>Safe \<or> \<acute>l\<le>\<acute>Queue \<or> \<acute>bc\<subset>Blacks \<acute>M)}.
+    \<acute>obc:=\<acute>bc;;
+    Mul_Propagate_Black n;; 
+    .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+      \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+      \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>M \<or> \<acute>l<\<acute>Queue 
+      \<and> (\<acute>l\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M))}. 
+    \<acute>bc:={};;
+    .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+      \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+      \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>M \<or> \<acute>l<\<acute>Queue 
+      \<and> (\<acute>l\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M)) \<and> \<acute>bc={}}. 
+       \<langle> \<acute>Ma:=\<acute>M,, \<acute>q:=\<acute>Queue \<rangle>;;
+    Mul_Count n;; 
+    .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+      \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+      \<and> length \<acute>Ma=length \<acute>M \<and> Blacks \<acute>Ma\<subseteq>\<acute>bc 
+      \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M)) 
+      \<and> \<acute>q<n+1}. 
+    IF \<acute>obc=\<acute>bc THEN
+    .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+      \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+      \<and> length \<acute>Ma=length \<acute>M \<and> Blacks \<acute>Ma\<subseteq>\<acute>bc 
+      \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M)) 
+      \<and> \<acute>q<n+1 \<and> \<acute>obc=\<acute>bc}.  
+    \<acute>l:=\<acute>l+1  
+    ELSE .{\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M 
+          \<and> \<acute>obc\<subseteq>Blacks \<acute>Ma \<and> Blacks \<acute>Ma\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M 
+          \<and> length \<acute>Ma=length \<acute>M \<and> Blacks \<acute>Ma\<subseteq>\<acute>bc 
+          \<and> (\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>Ma \<or> \<acute>l<\<acute>q \<and> (\<acute>q\<le>\<acute>Queue \<or> \<acute>obc\<subset>Blacks \<acute>M)) 
+          \<and> \<acute>q<n+1 \<and> \<acute>obc\<noteq>\<acute>bc}.  
+        \<acute>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:
+  "\<turnstile> 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)
+apply force
+apply (force dest:subset_antisym)
+apply force
+apply force
+apply force
+done
+
+subsection {* Interference Freedom *}
+
+lemma le_length_filter_update[rule_format]: 
+ "\<forall>i. (\<not>P (list!i) \<or> P j) \<and> i<length list 
+ \<longrightarrow> length(filter P list) \<le> 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]: 
+ "\<forall>i. P j \<and> \<not>(P (list!i)) \<and> i<length list 
+ \<longrightarrow> 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: "\<lbrakk>0\<le>j; j<n\<rbrakk> \<Longrightarrow>  
+  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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow> 
+  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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
+  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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
+  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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
+  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 subgoal 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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
+  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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
+  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 subgoal 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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
+  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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
+  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 subgoal left *}
+apply(simp add:mul_mutator_defs nth_list_update)
+done
+
+lemma Mul_interfree_Redirect_Edge_Count: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
+  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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
+  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 subgoal 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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
+  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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
+  interfree_aux (Some(Mul_Append n),{}, Some(Mul_Redirect_Edge j n))"
+apply (unfold mul_modules)
+apply interfree_aux
+apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
+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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
+  interfree_aux (Some(Mul_Redirect_Edge j n),{},Some(Mul_Append n))"
+apply (unfold mul_modules)
+apply interfree_aux
+apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
+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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
+  interfree_aux (Some(Mul_Append n),{}, Some(Mul_Color_Target j n))"
+apply (unfold mul_modules)
+apply interfree_aux
+apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
+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: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>  
+  interfree_aux (Some(Mul_Color_Target j n),{}, Some(Mul_Append n))"
+apply (unfold mul_modules)
+apply interfree_aux
+apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
+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<n  \<Longrightarrow> 
+  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 @{claset}) *})
+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 @{clasimpset}, 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(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 (thm "subset_psubset_trans"),etac (thm "Graph11"),force_tac @{clasimpset}]) *})
+--{* 41 subgoals left *}
+apply(tactic {* TRYALL(EVERY'[rtac disjI2, rtac disjI1, etac @{thm le_trans}, force_tac (@{claset},@{simpset} addsimps [@{thm Queue_def}, @{thm less_Suc_eq_le}, @{thm le_length_filter_update}])]) *})
+--{* 35 subgoals left *}
+apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac disjI1,etac (thm "psubset_subset_trans"),rtac (thm "Graph9"),force_tac @{clasimpset}]) *})
+--{* 31 subgoals left *}
+apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac disjI1,etac (thm "subset_psubset_trans"),etac (thm "Graph11"),force_tac @{clasimpset}]) *})
+--{* 29 subgoals left *}
+apply(tactic {* TRYALL(EVERY'[REPEAT o (rtac disjI2),etac (thm "subset_psubset_trans"),etac (thm "subset_psubset_trans"),etac (thm "Graph11"),force_tac @{clasimpset}]) *})
+--{* 25 subgoals left *}
+apply(tactic {* TRYALL(EVERY'[rtac disjI2, rtac disjI2, rtac disjI1, etac @{thm le_trans}, force_tac (@{claset},@{simpset} addsimps [@{thm Queue_def}, @{thm 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 \<Longrightarrow> 
+  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: 
+ "\<parallel>- .{\<acute>Mul_Proper n \<and> \<acute>Mul_mut_init n \<and> (\<forall>i<n. Z (\<acute>Muts!i))}.  
+ COBEGIN  
+  Mul_Collector n
+ .{False}.
+ \<parallel>  
+ SCHEME  [0\<le> 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 i)
+ 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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/OG_Com.thy	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,55 @@
+
+header {* \chapter{The Owicki-Gries Method} 
+
+\section{Abstract Syntax} *} 
+
+theory OG_Com imports Main begin
+
+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 \<Rightarrow> '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 \<times> 'a assn) list"
+   | Basic "('a \<Rightarrow> '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 \<Rightarrow> '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 \<Rightarrow> bool"
+primrec  
+  "atom_com (Parallel Ts) = False"
+  "atom_com (Basic f) = True"
+  "atom_com (Seq c1 c2) = (atom_com c1 \<and> atom_com c2)"
+  "atom_com (Cond b c1 c2) = (atom_com c1 \<and> atom_com c2)"
+  "atom_com (While b i c) = atom_com c"
+  
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/OG_Examples.thy	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,549 @@
+
+header {* \section{Examples} *}
+
+theory OG_Examples imports OG_Syntax begin
+
+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: 
+  "\<parallel>- .{\<acute>pr1=0 \<and> \<not>\<acute>in1 \<and> \<acute>pr2=0 \<and> \<not>\<acute>in2 }.  
+  COBEGIN .{\<acute>pr1=0 \<and> \<not>\<acute>in1}.  
+  WHILE True INV .{\<acute>pr1=0 \<and> \<not>\<acute>in1}.  
+  DO  
+  .{\<acute>pr1=0 \<and> \<not>\<acute>in1}. \<langle> \<acute>in1:=True,,\<acute>pr1:=1 \<rangle>;;  
+  .{\<acute>pr1=1 \<and> \<acute>in1}.  \<langle> \<acute>hold:=1,,\<acute>pr1:=2 \<rangle>;;  
+  .{\<acute>pr1=2 \<and> \<acute>in1 \<and> (\<acute>hold=1 \<or> \<acute>hold=2 \<and> \<acute>pr2=2)}.  
+  AWAIT (\<not>\<acute>in2 \<or> \<not>(\<acute>hold=1)) THEN \<acute>pr1:=3 END;;    
+  .{\<acute>pr1=3 \<and> \<acute>in1 \<and> (\<acute>hold=1 \<or> \<acute>hold=2 \<and> \<acute>pr2=2)}. 
+   \<langle>\<acute>in1:=False,,\<acute>pr1:=0\<rangle> 
+  OD .{\<acute>pr1=0 \<and> \<not>\<acute>in1}.  
+  \<parallel>  
+  .{\<acute>pr2=0 \<and> \<not>\<acute>in2}.  
+  WHILE True INV .{\<acute>pr2=0 \<and> \<not>\<acute>in2}.  
+  DO  
+  .{\<acute>pr2=0 \<and> \<not>\<acute>in2}. \<langle> \<acute>in2:=True,,\<acute>pr2:=1 \<rangle>;;  
+  .{\<acute>pr2=1 \<and> \<acute>in2}. \<langle>  \<acute>hold:=2,,\<acute>pr2:=2 \<rangle>;;  
+  .{\<acute>pr2=2 \<and> \<acute>in2 \<and> (\<acute>hold=2 \<or> (\<acute>hold=1 \<and> \<acute>pr1=2))}.  
+  AWAIT (\<not>\<acute>in1 \<or> \<not>(\<acute>hold=2)) THEN \<acute>pr2:=3  END;;    
+  .{\<acute>pr2=3 \<and> \<acute>in2 \<and> (\<acute>hold=2 \<or> (\<acute>hold=1 \<and> \<acute>pr1=2))}. 
+    \<langle>\<acute>in2:=False,,\<acute>pr2:=0\<rangle> 
+  OD .{\<acute>pr2=0 \<and> \<not>\<acute>in2}.  
+  COEND  
+  .{\<acute>pr1=0 \<and> \<not>\<acute>in1 \<and> \<acute>pr2=0 \<and> \<not>\<acute>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: 
+ "\<parallel>-  .{True}.  
+  \<acute>flag1:=False,, \<acute>flag2:=False,,  
+  COBEGIN .{\<not>\<acute>flag1}.  
+        WHILE True  
+        INV .{\<not>\<acute>flag1}.  
+        DO .{\<not>\<acute>flag1}. \<langle> \<acute>flag1:=True,,\<acute>after1:=False \<rangle>;;  
+           .{\<acute>flag1 \<and> \<not>\<acute>after1}. \<langle> \<acute>turn:=1,,\<acute>after1:=True \<rangle>;;  
+           .{\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}.  
+            WHILE \<not>(\<acute>flag2 \<longrightarrow> \<acute>turn=2)  
+            INV .{\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}.  
+            DO .{\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}. SKIP OD;; 
+           .{\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>flag2 \<and> \<acute>after2 \<longrightarrow> \<acute>turn=2)}.
+            \<acute>flag1:=False  
+        OD  
+       .{False}.  
+  \<parallel>  
+     .{\<not>\<acute>flag2}.  
+        WHILE True  
+        INV .{\<not>\<acute>flag2}.  
+        DO .{\<not>\<acute>flag2}. \<langle> \<acute>flag2:=True,,\<acute>after2:=False \<rangle>;;  
+           .{\<acute>flag2 \<and> \<not>\<acute>after2}. \<langle> \<acute>turn:=2,,\<acute>after2:=True \<rangle>;;  
+           .{\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}.  
+            WHILE \<not>(\<acute>flag1 \<longrightarrow> \<acute>turn=1)  
+            INV .{\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}.  
+            DO .{\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)}. SKIP OD;;  
+           .{\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>flag1 \<and> \<acute>after1 \<longrightarrow> \<acute>turn=1)}. 
+            \<acute>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: 
+ "\<parallel>- .{i\<noteq>j}.  
+  \<acute>out:=True ,,  
+  COBEGIN .{i\<noteq>j}.  
+       WHILE True INV .{i\<noteq>j}.  
+       DO .{i\<noteq>j}. AWAIT \<acute>out THEN  \<acute>out:=False,, \<acute>who:=i END;;  
+          .{\<not>\<acute>out \<and> \<acute>who=i \<and> i\<noteq>j}. \<acute>out:=True OD  
+       .{False}.  
+  \<parallel>  
+       .{i\<noteq>j}.  
+       WHILE True INV .{i\<noteq>j}.  
+       DO .{i\<noteq>j}. AWAIT \<acute>out THEN  \<acute>out:=False,,\<acute>who:=j END;;  
+          .{\<not>\<acute>out \<and> \<acute>who=j \<and> i\<noteq>j}. \<acute>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<n \<Longrightarrow> \<parallel>- .{True}.  
+  \<acute>out:=True ,,  
+ COBEGIN
+  SCHEME [0\<le> i< n]
+    .{True}.  
+     WHILE True INV .{True}.  
+      DO .{True}. AWAIT \<acute>out THEN  \<acute>out:=False,, \<acute>who:=i END;;  
+         .{\<not>\<acute>out \<and> \<acute>who=i}. \<acute>out:=True OD
+    .{False}. 
+ COEND
+  .{False}." 
+apply oghoare
+--{* 20 vc *}
+apply auto
+done
+
+subsubsection{* The Ticket Algorithm *}
+
+record Ticket_mutex =
+ num :: nat
+ nextv :: nat
+ turn :: "nat list"
+ index :: nat 
+
+lemma Ticket_mutex: 
+ "\<lbrakk> 0<n; I=\<guillemotleft>n=length \<acute>turn \<and> 0<\<acute>nextv \<and> (\<forall>k l. k<n \<and> l<n \<and> k\<noteq>l 
+    \<longrightarrow> \<acute>turn!k < \<acute>num \<and> (\<acute>turn!k =0 \<or> \<acute>turn!k\<noteq>\<acute>turn!l))\<guillemotright> \<rbrakk>
+   \<Longrightarrow> \<parallel>- .{n=length \<acute>turn}.  
+   \<acute>index:= 0,,
+   WHILE \<acute>index < n INV .{n=length \<acute>turn \<and> (\<forall>i<\<acute>index. \<acute>turn!i=0)}. 
+    DO \<acute>turn:= \<acute>turn[\<acute>index:=0],, \<acute>index:=\<acute>index +1 OD,,
+  \<acute>num:=1 ,, \<acute>nextv:=1 ,, 
+ COBEGIN
+  SCHEME [0\<le> i< n]
+    .{\<acute>I}.  
+     WHILE True INV .{\<acute>I}.  
+      DO .{\<acute>I}. \<langle> \<acute>turn :=\<acute>turn[i:=\<acute>num],, \<acute>num:=\<acute>num+1 \<rangle>;;  
+         .{\<acute>I}. WAIT \<acute>turn!i=\<acute>nextv END;;
+         .{\<acute>I \<and> \<acute>turn!i=\<acute>nextv}. \<acute>nextv:=\<acute>nextv+1
+      OD
+    .{False}. 
+ COEND
+  .{False}." 
+apply oghoare
+--{* 35 vc *}
+apply simp_all
+--{* 21 vc *}
+apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
+--{* 11 vc *}
+apply simp_all
+apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
+--{* 10 subgoals left *}
+apply(erule less_SucE)
+ apply simp
+apply simp
+--{* 9 subgoals left *}
+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 fastsimp
+--{* 5 subgoals left *}
+prefer 5
+apply(case_tac [!] "j=k")
+--{* 10 subgoals left *}
+apply simp_all
+apply(erule_tac x=k in allE)
+apply force
+--{* 9 subgoals left *}
+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
+--{* 5 subgoals left *}
+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
+--{* 3 subgoals left *}
+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
+--{* 1 subgoals left *}
+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: 
+  "\<lbrakk>I1= \<guillemotleft> a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> (a<\<acute>x \<and> f(\<acute>x)=0) \<or> (\<acute>y\<le>a \<and> f(\<acute>y)=0)) 
+      \<and> (\<not>\<acute>found \<and> a<\<acute> x \<longrightarrow> f(\<acute>x)\<noteq>0) \<guillemotright> ;  
+    I2= \<guillemotleft>\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> (a<\<acute>x \<and> f(\<acute>x)=0) \<or> (\<acute>y\<le>a \<and> f(\<acute>y)=0)) 
+      \<and> (\<not>\<acute>found \<and> \<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0) \<guillemotright> \<rbrakk> \<Longrightarrow>  
+  \<parallel>- .{\<exists> u. f(u)=0}.  
+  \<acute>turn:=1,, \<acute>found:= False,,  
+  \<acute>x:=a,, \<acute>y:=a+1 ,,  
+  COBEGIN .{\<acute>I1}.  
+       WHILE \<not>\<acute>found  
+       INV .{\<acute>I1}.  
+       DO .{a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> \<acute>y\<le>a \<and> f(\<acute>y)=0) \<and> (a<\<acute>x \<longrightarrow> f(\<acute>x)\<noteq>0)}.  
+          WAIT \<acute>turn=1 END;;  
+          .{a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> \<acute>y\<le>a \<and> f(\<acute>y)=0) \<and> (a<\<acute>x \<longrightarrow> f(\<acute>x)\<noteq>0)}.  
+          \<acute>turn:=2;;  
+          .{a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> \<acute>y\<le>a \<and> f(\<acute>y)=0) \<and> (a<\<acute>x \<longrightarrow> f(\<acute>x)\<noteq>0)}.    
+          \<langle> \<acute>x:=\<acute>x+1,,  
+            IF f(\<acute>x)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>  
+       OD;;  
+       .{\<acute>I1  \<and> \<acute>found}.  
+       \<acute>turn:=2  
+       .{\<acute>I1 \<and> \<acute>found}.  
+  \<parallel>  
+      .{\<acute>I2}.  
+       WHILE \<not>\<acute>found  
+       INV .{\<acute>I2}.  
+       DO .{\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> a<\<acute>x \<and> f(\<acute>x)=0) \<and> (\<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0)}.  
+          WAIT \<acute>turn=2 END;;  
+          .{\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> a<\<acute>x \<and> f(\<acute>x)=0) \<and> (\<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0)}.  
+          \<acute>turn:=1;;  
+          .{\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> a<\<acute>x \<and> f(\<acute>x)=0) \<and> (\<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0)}.  
+          \<langle> \<acute>y:=(\<acute>y - 1),,  
+            IF f(\<acute>y)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>  
+       OD;;  
+       .{\<acute>I2 \<and> \<acute>found}.  
+       \<acute>turn:=1  
+       .{\<acute>I2 \<and> \<acute>found}.  
+  COEND  
+  .{f(\<acute>x)=0 \<or> f(\<acute>y)=0}."
+apply oghoare
+--{* 98 verification conditions *}
+apply auto 
+--{* auto takes about 3 minutes !! *}
+done
+
+text {* Easier Version: without AWAIT.  Apt and Olderog. page 256: *}
+
+lemma Zero_Search_2: 
+"\<lbrakk>I1=\<guillemotleft> a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> (a<\<acute>x \<and> f(\<acute>x)=0) \<or> (\<acute>y\<le>a \<and> f(\<acute>y)=0)) 
+    \<and> (\<not>\<acute>found \<and> a<\<acute>x \<longrightarrow> f(\<acute>x)\<noteq>0)\<guillemotright>;  
+ I2= \<guillemotleft>\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> (a<\<acute>x \<and> f(\<acute>x)=0) \<or> (\<acute>y\<le>a \<and> f(\<acute>y)=0)) 
+    \<and> (\<not>\<acute>found \<and> \<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0)\<guillemotright>\<rbrakk> \<Longrightarrow>  
+  \<parallel>- .{\<exists>u. f(u)=0}.  
+  \<acute>found:= False,,  
+  \<acute>x:=a,, \<acute>y:=a+1,,  
+  COBEGIN .{\<acute>I1}.  
+       WHILE \<not>\<acute>found  
+       INV .{\<acute>I1}.  
+       DO .{a\<le>\<acute>x \<and> (\<acute>found \<longrightarrow> \<acute>y\<le>a \<and> f(\<acute>y)=0) \<and> (a<\<acute>x \<longrightarrow> f(\<acute>x)\<noteq>0)}.  
+          \<langle> \<acute>x:=\<acute>x+1,,IF f(\<acute>x)=0 THEN  \<acute>found:=True ELSE  SKIP FI\<rangle>  
+       OD  
+       .{\<acute>I1 \<and> \<acute>found}.  
+  \<parallel>  
+      .{\<acute>I2}.  
+       WHILE \<not>\<acute>found  
+       INV .{\<acute>I2}.  
+       DO .{\<acute>y\<le>a+1 \<and> (\<acute>found \<longrightarrow> a<\<acute>x \<and> f(\<acute>x)=0) \<and> (\<acute>y\<le>a \<longrightarrow> f(\<acute>y)\<noteq>0)}.  
+          \<langle> \<acute>y:=(\<acute>y - 1),,IF f(\<acute>y)=0 THEN  \<acute>found:=True ELSE  SKIP FI\<rangle>  
+       OD  
+       .{\<acute>I2 \<and> \<acute>found}.  
+  COEND  
+  .{f(\<acute>x)=0 \<or> f(\<acute>y)=0}."
+apply oghoare
+--{* 20 vc *}
+apply auto
+--{* auto takes aprox. 2 minutes. *}
+done
+
+subsection {* Producer/Consumer *}
+
+subsubsection {* Previous lemmas *}
+
+lemma nat_lemma2: "\<lbrakk> b = m*(n::nat) + t; a = s*n + u; t=u; b-a < n \<rbrakk> \<Longrightarrow> m \<le> s"
+proof -
+  assume "b = m*(n::nat) + t" "a = s*n + u" "t=u"
+  hence "(m - s) * n = b - a" by (simp add: diff_mult_distrib)
+  also assume "\<dots> < n"
+  finally have "m - s < 1" by simp
+  thus ?thesis by arith
+qed
+
+lemma mod_lemma: "\<lbrakk> (c::nat) \<le> a; a < b; b - c < n \<rbrakk> \<Longrightarrow> b mod n \<noteq> 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(subgoal_tac "b - a \<le> b - c")
+ prefer 2 apply arith
+apply(drule le_less_trans)
+back
+ apply assumption
+apply(frule less_not_refl2)
+apply(drule less_imp_le)
+apply (drule_tac m = "a" and k = n in div_le_mono)
+apply(safe)
+apply(frule_tac b = "b" and a = "a" and n = "n" in nat_lemma2, assumption, assumption)
+apply assumption
+apply(drule order_antisym, assumption)
+apply(rotate_tac -3)
+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: 
+  "\<lbrakk>INIT= \<guillemotleft>0<length a \<and> 0<length \<acute>buffer \<and> length \<acute>b=length a\<guillemotright> ;  
+    I= \<guillemotleft>(\<forall>k<\<acute>ins. \<acute>outs\<le>k \<longrightarrow> (a ! k) = \<acute>buffer ! (k mod (length \<acute>buffer))) \<and>  
+            \<acute>outs\<le>\<acute>ins \<and> \<acute>ins-\<acute>outs\<le>length \<acute>buffer\<guillemotright> ;  
+    I1= \<guillemotleft>\<acute>I \<and> \<acute>li\<le>length a\<guillemotright> ;  
+    p1= \<guillemotleft>\<acute>I1 \<and> \<acute>li=\<acute>ins\<guillemotright> ;  
+    I2 = \<guillemotleft>\<acute>I \<and> (\<forall>k<\<acute>lj. (a ! k)=(\<acute>b ! k)) \<and> \<acute>lj\<le>length a\<guillemotright> ;
+    p2 = \<guillemotleft>\<acute>I2 \<and> \<acute>lj=\<acute>outs\<guillemotright> \<rbrakk> \<Longrightarrow>   
+  \<parallel>- .{\<acute>INIT}.  
+ \<acute>ins:=0,, \<acute>outs:=0,, \<acute>li:=0,, \<acute>lj:=0,,
+ COBEGIN .{\<acute>p1 \<and> \<acute>INIT}. 
+   WHILE \<acute>li <length a 
+     INV .{\<acute>p1 \<and> \<acute>INIT}.   
+   DO .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a}.  
+       \<acute>vx:= (a ! \<acute>li);;  
+      .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a \<and> \<acute>vx=(a ! \<acute>li)}. 
+        WAIT \<acute>ins-\<acute>outs < length \<acute>buffer END;; 
+      .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a \<and> \<acute>vx=(a ! \<acute>li) 
+         \<and> \<acute>ins-\<acute>outs < length \<acute>buffer}. 
+       \<acute>buffer:=(list_update \<acute>buffer (\<acute>ins mod (length \<acute>buffer)) \<acute>vx);; 
+      .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a 
+         \<and> (a ! \<acute>li)=(\<acute>buffer ! (\<acute>ins mod (length \<acute>buffer))) 
+         \<and> \<acute>ins-\<acute>outs <length \<acute>buffer}.  
+       \<acute>ins:=\<acute>ins+1;; 
+      .{\<acute>I1 \<and> \<acute>INIT \<and> (\<acute>li+1)=\<acute>ins \<and> \<acute>li<length a}.  
+       \<acute>li:=\<acute>li+1  
+   OD  
+  .{\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li=length a}.  
+  \<parallel>  
+  .{\<acute>p2 \<and> \<acute>INIT}.  
+   WHILE \<acute>lj < length a  
+     INV .{\<acute>p2 \<and> \<acute>INIT}.  
+   DO .{\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>INIT}.  
+        WAIT \<acute>outs<\<acute>ins END;; 
+      .{\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>outs<\<acute>ins \<and> \<acute>INIT}.  
+       \<acute>vy:=(\<acute>buffer ! (\<acute>outs mod (length \<acute>buffer)));; 
+      .{\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>outs<\<acute>ins \<and> \<acute>vy=(a ! \<acute>lj) \<and> \<acute>INIT}.  
+       \<acute>outs:=\<acute>outs+1;;  
+      .{\<acute>I2 \<and> (\<acute>lj+1)=\<acute>outs \<and> \<acute>lj<length a \<and> \<acute>vy=(a ! \<acute>lj) \<and> \<acute>INIT}.  
+       \<acute>b:=(list_update \<acute>b \<acute>lj \<acute>vy);; 
+      .{\<acute>I2 \<and> (\<acute>lj+1)=\<acute>outs \<and> \<acute>lj<length a \<and> (a ! \<acute>lj)=(\<acute>b ! \<acute>lj) \<and> \<acute>INIT}.  
+       \<acute>lj:=\<acute>lj+1  
+   OD  
+  .{\<acute>p2 \<and> \<acute>lj=length a \<and> \<acute>INIT}.  
+ COEND  
+ .{ \<forall>k<length a. (a ! k)=(\<acute>b ! k)}."
+apply oghoare
+--{* 138 vc  *}
+apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
+--{* 112 subgoals left *}
+apply(simp_all (no_asm))
+apply(tactic {*ALLGOALS (conjI_Tac (K all_tac)) *})
+--{* 930 subgoals left *}
+apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
+apply(simp_all (asm_lr) only:length_0_conv [THEN sym])
+--{* 44 subgoals left *}
+apply (simp_all (asm_lr) del:length_0_conv add: neq0_conv nth_list_update mod_less_divisor mod_lemma)
+--{* 32 subgoals left *}
+apply(tactic {* ALLGOALS (clarify_tac @{claset}) *})
+
+apply(tactic {* TRYALL (Lin_Arith.tac @{context}) *})
+--{* 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 Example1 =
+  a :: "nat \<Rightarrow> nat"
+
+lemma Example1: 
+ "\<parallel>- .{True}.
+   COBEGIN SCHEME [0\<le>i<n] .{True}. \<acute>a:=\<acute>a (i:=0) .{\<acute>a i=0}. COEND 
+  .{\<forall>i < n. \<acute>a i = 0}."
+apply oghoare
+apply simp_all
+done
+
+text {* Same example with lists as auxiliary variables. *}
+record Example1_list =
+  A :: "nat list"
+lemma Example1_list: 
+ "\<parallel>- .{n < length \<acute>A}. 
+   COBEGIN 
+     SCHEME [0\<le>i<n] .{n < length \<acute>A}. \<acute>A:=\<acute>A[i:=0] .{\<acute>A!i=0}. 
+   COEND 
+    .{\<forall>i < n. \<acute>A!i = 0}."
+apply oghoare
+apply force+
+done
+
+subsubsection {* Increment a Variable in Parallel *}
+
+text {* First some lemmas about summation properties. *}
+(*
+lemma Example2_lemma1: "!!b. j<n \<Longrightarrow> (\<Sum>i::nat<n. b i) = (0::nat) \<Longrightarrow> b j = 0 "
+apply(induct n)
+ apply simp_all
+apply(force simp add: less_Suc_eq)
+done
+*)
+lemma Example2_lemma2_aux: "!!b. j<n \<Longrightarrow> 
+ (\<Sum>i=0..<n. (b i::nat)) =
+ (\<Sum>i=0..<j. b i) + b j + (\<Sum>i=0..<n-(Suc j) . b (Suc j + i))"
+apply(induct n)
+ apply simp_all
+apply(simp add:less_Suc_eq)
+ apply(auto)
+apply(subgoal_tac "n - j = Suc(n- Suc j)")
+  apply simp
+apply arith
+done
+
+lemma Example2_lemma2_aux2: 
+  "!!b. j\<le> s \<Longrightarrow> (\<Sum>i::nat=0..<j. (b (s:=t)) i) = (\<Sum>i=0..<j. b i)"
+apply(induct j) 
+ apply simp_all
+done
+
+lemma Example2_lemma2: 
+ "!!b. \<lbrakk>j<n; b j=0\<rbrakk> \<Longrightarrow> Suc (\<Sum>i::nat=0..<n. b i)=(\<Sum>i=0..<n. (b (j := Suc 0)) i)"
+apply(frule_tac b="(b (j:=(Suc 0)))" in Example2_lemma2_aux)
+apply(erule_tac  t="setsum (b(j := (Suc 0))) {0..<n}" in ssubst)
+apply(frule_tac b=b in Example2_lemma2_aux)
+apply(erule_tac  t="setsum b {0..<n}" in ssubst)
+apply(subgoal_tac "Suc (setsum b {0..<j} + b j + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))=(setsum b {0..<j} + Suc (b j) + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))")
+apply(rotate_tac -1)
+apply(erule ssubst)
+apply(subgoal_tac "j\<le>j")
+ apply(drule_tac b="b" and t="(Suc 0)" in Example2_lemma2_aux2)
+apply(rotate_tac -1)
+apply(erule ssubst)
+apply simp_all
+done
+
+
+record Example2 = 
+ c :: "nat \<Rightarrow> nat" 
+ x :: nat
+
+lemma Example_2: "0<n \<Longrightarrow> 
+ \<parallel>- .{\<acute>x=0 \<and> (\<Sum>i=0..<n. \<acute>c i)=0}.  
+ COBEGIN 
+   SCHEME [0\<le>i<n] 
+  .{\<acute>x=(\<Sum>i=0..<n. \<acute>c i) \<and> \<acute>c i=0}. 
+   \<langle> \<acute>x:=\<acute>x+(Suc 0),, \<acute>c:=\<acute>c (i:=(Suc 0)) \<rangle>
+  .{\<acute>x=(\<Sum>i=0..<n. \<acute>c i) \<and> \<acute>c i=(Suc 0)}.
+ COEND 
+ .{\<acute>x=n}."
+apply oghoare
+apply (simp_all cong del: strong_setsum_cong)
+apply (tactic {* ALLGOALS (clarify_tac @{claset}) *})
+apply (simp_all cong del: strong_setsum_cong)
+   apply(erule (1) Example2_lemma2)
+  apply(erule (1) Example2_lemma2)
+ apply(erule (1) Example2_lemma2)
+apply(simp)
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/OG_Hoare.thy	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,469 @@
+
+header {* \section{The Proof System} *}
+
+theory OG_Hoare imports OG_Tran begin
+
+consts assertions :: "'a ann_com \<Rightarrow> ('a assn) set"
+primrec
+  "assertions (AnnBasic r f) = {r}"
+  "assertions (AnnSeq c1 c2) = assertions c1 \<union> assertions c2"
+  "assertions (AnnCond1 r b c1 c2) = {r} \<union> assertions c1 \<union> assertions c2"
+  "assertions (AnnCond2 r b c) = {r} \<union> assertions c"
+  "assertions (AnnWhile r b i c) = {r, i} \<union> assertions c"
+  "assertions (AnnAwait r b c) = {r}" 
+
+consts atomics :: "'a ann_com \<Rightarrow> ('a assn \<times> 'a com) set"       
+primrec
+  "atomics (AnnBasic r f) = {(r, Basic f)}"
+  "atomics (AnnSeq c1 c2) = atomics c1 \<union> atomics c2"
+  "atomics (AnnCond1 r b c1 c2) = atomics c1 \<union> atomics c2"
+  "atomics (AnnCond2 r b c) = atomics c"
+  "atomics (AnnWhile r b i c) = atomics c" 
+  "atomics (AnnAwait r b c) = {(r \<inter> b, c)}"
+
+consts com :: "'a ann_triple_op \<Rightarrow> 'a ann_com_op"
+primrec "com (c, q) = c"
+
+consts post :: "'a ann_triple_op \<Rightarrow> 'a assn"
+primrec "post (c, q) = q"
+
+constdefs  interfree_aux :: "('a ann_com_op \<times> 'a assn \<times> 'a ann_com_op) \<Rightarrow> bool"
+  "interfree_aux \<equiv> \<lambda>(co, q, co'). co'= None \<or>  
+                    (\<forall>(r,a) \<in> atomics (the co'). \<parallel>= (q \<inter> r) a q \<and>
+                    (co = None \<or> (\<forall>p \<in> assertions (the co). \<parallel>= (p \<inter> r) a p)))"
+
+constdefs interfree :: "(('a ann_triple_op) list) \<Rightarrow> bool" 
+  "interfree Ts \<equiv> \<forall>i j. i < length Ts \<and> j < length Ts \<and> i \<noteq> j \<longrightarrow> 
+                         interfree_aux (com (Ts!i), post (Ts!i), com (Ts!j)) "
+
+inductive
+  oghoare :: "'a assn \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool"  ("(3\<parallel>- _//_//_)" [90,55,90] 50)
+  and ann_hoare :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool"  ("(2\<turnstile> _// _)" [60,90] 45)
+where
+  AnnBasic: "r \<subseteq> {s. f s \<in> q} \<Longrightarrow> \<turnstile> (AnnBasic r f) q"
+
+| AnnSeq:   "\<lbrakk> \<turnstile> c0 pre c1; \<turnstile> c1 q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnSeq c0 c1) q"
+  
+| AnnCond1: "\<lbrakk> r \<inter> b \<subseteq> pre c1; \<turnstile> c1 q; r \<inter> -b \<subseteq> pre c2; \<turnstile> c2 q\<rbrakk> 
+              \<Longrightarrow> \<turnstile> (AnnCond1 r b c1 c2) q"
+| AnnCond2: "\<lbrakk> r \<inter> b \<subseteq> pre c; \<turnstile> c q; r \<inter> -b \<subseteq> q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnCond2 r b c) q"
+  
+| AnnWhile: "\<lbrakk> r \<subseteq> i; i \<inter> b \<subseteq> pre c; \<turnstile> c i; i \<inter> -b \<subseteq> q \<rbrakk> 
+              \<Longrightarrow> \<turnstile> (AnnWhile r b i c) q"
+  
+| AnnAwait:  "\<lbrakk> atom_com c; \<parallel>- (r \<inter> b) c q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnAwait r b c) q"
+  
+| AnnConseq: "\<lbrakk>\<turnstile> c q; q \<subseteq> q' \<rbrakk> \<Longrightarrow> \<turnstile> c q'"
+
+
+| Parallel: "\<lbrakk> \<forall>i<length Ts. \<exists>c q. Ts!i = (Some c, q) \<and> \<turnstile> c q; interfree Ts \<rbrakk>
+	   \<Longrightarrow> \<parallel>- (\<Inter>i\<in>{i. i<length Ts}. pre(the(com(Ts!i)))) 
+                     Parallel Ts 
+                  (\<Inter>i\<in>{i. i<length Ts}. post(Ts!i))"
+
+| Basic:   "\<parallel>- {s. f s \<in>q} (Basic f) q"
+  
+| Seq:    "\<lbrakk> \<parallel>- p c1 r; \<parallel>- r c2 q \<rbrakk> \<Longrightarrow> \<parallel>- p (Seq c1 c2) q "
+
+| Cond:   "\<lbrakk> \<parallel>- (p \<inter> b) c1 q; \<parallel>- (p \<inter> -b) c2 q \<rbrakk> \<Longrightarrow> \<parallel>- p (Cond b c1 c2) q"
+
+| While:  "\<lbrakk> \<parallel>- (p \<inter> b) c p \<rbrakk> \<Longrightarrow> \<parallel>- p (While b i c) (p \<inter> -b)"
+
+| Conseq: "\<lbrakk> p' \<subseteq> p; \<parallel>- p c q ; q \<subseteq> q' \<rbrakk> \<Longrightarrow> \<parallel>- 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\<rightarrow> (Parallel Ts, t) \<longrightarrow> All_None Ts \<longrightarrow> 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 \<inter> b) \<subseteq> p \<Longrightarrow> SEM (fwhile b S k) p \<subseteq> (p \<inter> -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 conjI)
+ apply (blast dest: L3_5i) 
+apply(simp add: SEM_def sem_def id_def)
+apply (blast dest: Basic_ntran rtrancl_imp_UN_rel_pow) 
+done
+
+lemma atom_hoare_sound [rule_format]: 
+ " \<parallel>- p c q \<longrightarrow> atom_com(c) \<longrightarrow> \<parallel>= 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(simp add: L3_5iv)
+--{* While *}
+ apply (force simp add: L3_5v dest: SEM_fwhile) 
+--{* Conseq *}
+apply(force simp add: SEM_def sem_def)
+done
+    
+subsection {* Soundness of the System for Component Programs *}
+
+inductive_cases ann_transition_cases:
+    "(None,s) -1\<rightarrow> (c', s')"
+    "(Some (AnnBasic r f),s) -1\<rightarrow> (c', s')"
+    "(Some (AnnSeq c1 c2), s) -1\<rightarrow> (c', s')"
+    "(Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> (c', s')"
+    "(Some (AnnCond2 r b c), s) -1\<rightarrow> (c', s')"
+    "(Some (AnnWhile r b I c), s) -1\<rightarrow> (c', s')"
+    "(Some (AnnAwait r b c),s) -1\<rightarrow> (c', s')"
+
+text {* Strong Soundness for Component Programs:*}
+
+lemma ann_hoare_case_analysis [rule_format]: 
+  defines I: "I \<equiv> \<lambda>C q'.
+  ((\<forall>r f. C = AnnBasic r f \<longrightarrow> (\<exists>q. r \<subseteq> {s. f s \<in> q} \<and> q \<subseteq> q')) \<and>  
+  (\<forall>c0 c1. C = AnnSeq c0 c1 \<longrightarrow> (\<exists>q. q \<subseteq> q' \<and> \<turnstile> c0 pre c1 \<and> \<turnstile> c1 q)) \<and>  
+  (\<forall>r b c1 c2. C = AnnCond1 r b c1 c2 \<longrightarrow> (\<exists>q. q \<subseteq> q' \<and>  
+  r \<inter> b \<subseteq> pre c1 \<and> \<turnstile> c1 q \<and> r \<inter> -b \<subseteq> pre c2 \<and> \<turnstile> c2 q)) \<and>  
+  (\<forall>r b c. C = AnnCond2 r b c \<longrightarrow> 
+  (\<exists>q. q \<subseteq> q' \<and> r \<inter> b \<subseteq> pre c  \<and> \<turnstile> c q \<and> r \<inter> -b \<subseteq> q)) \<and>  
+  (\<forall>r i b c. C = AnnWhile r b i c \<longrightarrow>  
+  (\<exists>q. q \<subseteq> q' \<and> r \<subseteq> i \<and> i \<inter> b \<subseteq> pre c \<and> \<turnstile> c i \<and> i \<inter> -b \<subseteq> q)) \<and>  
+  (\<forall>r b c. C = AnnAwait r b c \<longrightarrow> (\<exists>q. q \<subseteq> q' \<and> \<parallel>- (r \<inter> b) c q)))"
+  shows "\<turnstile> C q' \<longrightarrow> I C q'"
+apply(rule ann_hoare_induct)
+apply (simp_all add: I)
+ 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 \<inter> {(x,y). True}) = (transition)"
+apply force
+done
+
+lemma Strong_Soundness_aux_aux [rule_format]: 
+ "(co, s) -1\<rightarrow> (co', t) \<longrightarrow> (\<forall>c. co = Some c \<longrightarrow> s\<in> pre c \<longrightarrow> 
+ (\<forall>q. \<turnstile> c q \<longrightarrow> (if co' = None then t\<in>q else t \<in> pre(the co') \<and> \<turnstile> (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
+--{* 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: "\<lbrakk> (Some c, s) -*\<rightarrow> (co, t); s \<in> pre c; \<turnstile> c q \<rbrakk>  
+  \<Longrightarrow> if co = None then t \<in> q else t \<in> pre (the co) \<and> \<turnstile> (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: "\<lbrakk> (Some c, s)-*\<rightarrow>(co, t); s \<in> pre c; \<turnstile> c q \<rbrakk>  
+  \<Longrightarrow> if co = None then t\<in>q else t \<in> pre (the co)"
+apply(force dest:Strong_Soundness_aux)
+done
+
+lemma ann_hoare_sound: "\<turnstile> c q  \<Longrightarrow> \<Turnstile> 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\<rightarrow> (R', t) \<Longrightarrow>  
+  (\<exists>Rs. R' = (Parallel Rs) \<and> (length Rs) = (length Ts) \<and>
+  (\<forall>i. i<length Ts \<longrightarrow> 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*\<rightarrow> (R',t) \<Longrightarrow>   
+  (\<exists>Rs. R' = (Parallel Rs) \<and> (length Rs) = (length Ts) \<and>  
+  (\<forall>i. i<length Ts \<longrightarrow> 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 \<in> assertions c"
+apply(rule ann_com_com.induct [THEN conjunct1])
+apply auto
+done
+
+lemma interfree_aux1 [rule_format]: 
+  "(c,s) -1\<rightarrow> (r,t)  \<longrightarrow> (interfree_aux(c1, q1, c) \<longrightarrow> 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\<rightarrow> (r,t) \<longrightarrow> (interfree_aux(c, q, a)  \<longrightarrow> interfree_aux(r, q, a) )"
+apply (rule ann_transition_transition.induct [THEN conjunct1])
+apply(force simp add:interfree_aux_def)+
+done
+
+lemma interfree_lemma: "\<lbrakk> (Some c, s) -1\<rightarrow> (r, t);interfree Ts ; i<length Ts;  
+           Ts!i = (Some c, q) \<rbrakk> \<Longrightarrow> 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: 
+  "\<lbrakk>interfree Ts; i<length Ts; com(Ts ! i) = Some(AnnSeq c0 c1) \<rbrakk> 
+  \<Longrightarrow>  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)]: 
+ "\<lbrakk> \<forall>i<length Ts. (if com(Ts!i) = None then b \<in> post(Ts!i) 
+  else b \<in> pre(the(com(Ts!i))) \<and> \<turnstile> the(com(Ts!i)) post(Ts!i));  
+  com(Ts ! i) = Some(AnnSeq c0 c1); i<length Ts; interfree Ts \<rbrakk> \<Longrightarrow> 
+ (\<forall>ia<length Ts. (if com(Ts[i:=(Some c0, pre c1)]! ia) = None  
+  then b \<in> post(Ts[i:=(Some c0, pre c1)]! ia) 
+ else b \<in> pre(the(com(Ts[i:=(Some c0, pre c1)]! ia))) \<and>  
+ \<turnstile> the(com(Ts[i:=(Some c0, pre c1)]! ia)) post(Ts[i:=(Some c0, pre c1)]! ia))) 
+  \<and> 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\<rightarrow> (co, t) \<longrightarrow>  
+  (\<forall>Ts. i<length Ts \<longrightarrow> com(Ts ! i) = Some c \<longrightarrow>  
+  (\<forall>i<length Ts. (if com(Ts ! i) = None then b\<in>post(Ts!i)  
+  else b\<in>pre(the(com(Ts!i))) \<and> \<turnstile> the(com(Ts!i)) post(Ts!i))) \<longrightarrow>  
+ interfree Ts \<longrightarrow>  
+  (\<forall>j. j<length Ts \<and> i\<noteq>j \<longrightarrow> (if com(Ts!j) = None then t\<in>post(Ts!j)  
+  else t\<in>pre(the(com(Ts!j))) \<and> \<turnstile> 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(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(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]: 
+ "\<lbrakk>(Ts',s) -P*\<rightarrow> (Rs',t);  Ts' = (Parallel Ts); interfree Ts;
+ \<forall>i. i<length Ts \<longrightarrow> (\<exists>c q. (Ts ! i) = (Some c, q) \<and> s\<in>(pre c) \<and> \<turnstile> c q ) \<rbrakk> \<Longrightarrow>  
+  \<forall>Rs. Rs' = (Parallel Rs) \<longrightarrow> (\<forall>j. j<length Rs \<longrightarrow> 
+  (if com(Rs ! j) = None then t\<in>post(Ts ! j) 
+  else t\<in>pre(the(com(Rs ! j))) \<and> \<turnstile> the(com(Rs ! j)) post(Ts ! j))) \<and> 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\<rightarrow> (Parallel Rs, t)" for Ts s 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: 
+ "\<lbrakk>(Parallel Ts, s) -P*\<rightarrow> (Parallel Rs, t); interfree Ts; j<length Rs; 
+  \<forall>i. i<length Ts \<longrightarrow> (\<exists>c q. Ts ! i = (Some c, q) \<and> s\<in>pre c \<and> \<turnstile> c q) \<rbrakk> \<Longrightarrow>  
+  if com(Rs ! j) = None then t\<in>post(Ts ! j) else t\<in>pre (the(com(Rs ! j)))"
+apply(drule  Parallel_Strong_Soundness_aux)
+apply simp+
+done
+
+lemma oghoare_sound [rule_format]: "\<parallel>- p c q \<longrightarrow> \<parallel>= 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=xb in Parallel_Strong_Soundness)
+         apply clarify
+        apply simp
+       apply force
+      apply simp
+      apply(erule_tac V = "\<forall>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)
+--{* While *}
+ apply(simp add: L3_5v)
+ apply (blast dest: SEM_fwhile) 
+--{* Conseq *}
+apply(auto simp add: SEM_def sem_def)
+done
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/OG_Syntax.thy	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,140 @@
+theory OG_Syntax
+imports OG_Tactics Quote_Antiquote
+begin
+
+text{* Syntax for commands and for assertions and boolean expressions in 
+ commands @{text com} and annotated commands @{text ann_com}. *}
+
+syntax
+  "_Assign"      :: "idt \<Rightarrow> 'b \<Rightarrow> 'a com"    ("(\<acute>_ :=/ _)" [70, 65] 61)
+  "_AnnAssign"   :: "'a assn \<Rightarrow> idt \<Rightarrow> 'b \<Rightarrow> 'a com"    ("(_ \<acute>_ :=/ _)" [90,70,65] 61)
+
+translations
+  "\<acute>\<spacespace>x := a" \<rightharpoonup> "Basic \<guillemotleft>\<acute>\<spacespace>(_update_name x (\<lambda>_. a))\<guillemotright>"
+  "r \<acute>\<spacespace>x := a" \<rightharpoonup> "AnnBasic r \<guillemotleft>\<acute>\<spacespace>(_update_name x (\<lambda>_. a))\<guillemotright>"
+
+syntax
+  "_AnnSkip"     :: "'a assn \<Rightarrow> 'a ann_com"              ("_//SKIP" [90] 63)
+  "_AnnSeq"      :: "'a ann_com \<Rightarrow> 'a ann_com \<Rightarrow> 'a ann_com"  ("_;;/ _" [60,61] 60)
+  
+  "_AnnCond1"    :: "'a assn \<Rightarrow> 'a bexp  \<Rightarrow> 'a ann_com  \<Rightarrow> 'a ann_com \<Rightarrow> 'a ann_com"
+                    ("_ //IF _ /THEN _ /ELSE _ /FI"  [90,0,0,0] 61)
+  "_AnnCond2"    :: "'a assn \<Rightarrow> 'a bexp  \<Rightarrow> 'a ann_com \<Rightarrow> 'a ann_com"
+                    ("_ //IF _ /THEN _ /FI"  [90,0,0] 61)
+  "_AnnWhile"    :: "'a assn \<Rightarrow> 'a bexp  \<Rightarrow> 'a assn \<Rightarrow> 'a ann_com \<Rightarrow> 'a ann_com" 
+                    ("_ //WHILE _ /INV _ //DO _//OD"  [90,0,0,0] 61)
+  "_AnnAwait"    :: "'a assn \<Rightarrow> 'a bexp  \<Rightarrow> 'a com \<Rightarrow> 'a ann_com"
+                    ("_ //AWAIT _ /THEN /_ /END"  [90,0,0] 61)
+  "_AnnAtom"     :: "'a assn  \<Rightarrow> 'a com \<Rightarrow> 'a ann_com"   ("_//\<langle>_\<rangle>" [90,0] 61)
+  "_AnnWait"     :: "'a assn \<Rightarrow> 'a bexp \<Rightarrow> 'a ann_com"   ("_//WAIT _ END" [90,0] 61)
+
+  "_Skip"        :: "'a com"                 ("SKIP" 63)
+  "_Seq"         :: "'a com \<Rightarrow> 'a com \<Rightarrow> 'a com" ("_,,/ _" [55, 56] 55)
+  "_Cond"        :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com \<Rightarrow> 'a com" 
+                                  ("(0IF _/ THEN _/ ELSE _/ FI)" [0, 0, 0] 61)
+  "_Cond2"       :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com"   ("IF _ THEN _ FI" [0,0] 56)
+  "_While_inv"   :: "'a bexp \<Rightarrow> 'a assn \<Rightarrow> 'a com \<Rightarrow> 'a com"
+                    ("(0WHILE _/ INV _ //DO _ /OD)"  [0, 0, 0] 61)
+  "_While"       :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com"
+                    ("(0WHILE _ //DO _ /OD)"  [0, 0] 61)
+
+translations
+  "SKIP" \<rightleftharpoons> "Basic id"
+  "c_1,, c_2" \<rightleftharpoons> "Seq c_1 c_2"
+
+  "IF b THEN c1 ELSE c2 FI" \<rightharpoonup> "Cond .{b}. c1 c2"
+  "IF b THEN c FI" \<rightleftharpoons> "IF b THEN c ELSE SKIP FI"
+  "WHILE b INV i DO c OD" \<rightharpoonup> "While .{b}. i c"
+  "WHILE b DO c OD" \<rightleftharpoons> "WHILE b INV CONST undefined DO c OD"
+
+  "r SKIP" \<rightleftharpoons> "AnnBasic r id"
+  "c_1;; c_2" \<rightleftharpoons> "AnnSeq c_1 c_2" 
+  "r IF b THEN c1 ELSE c2 FI" \<rightharpoonup> "AnnCond1 r .{b}. c1 c2"
+  "r IF b THEN c FI" \<rightharpoonup> "AnnCond2 r .{b}. c"
+  "r WHILE b INV i DO c OD" \<rightharpoonup> "AnnWhile r .{b}. i c"
+  "r AWAIT b THEN c END" \<rightharpoonup> "AnnAwait r .{b}. c"
+  "r \<langle>c\<rangle>" \<rightleftharpoons> "r AWAIT True THEN c END"
+  "r WAIT b END" \<rightleftharpoons> "r AWAIT b THEN SKIP END"
+ 
+nonterminals
+  prgs
+
+syntax
+  "_PAR" :: "prgs \<Rightarrow> 'a"              ("COBEGIN//_//COEND" [57] 56)
+  "_prg" :: "['a, 'a] \<Rightarrow> prgs"        ("_//_" [60, 90] 57)
+  "_prgs" :: "['a, 'a, prgs] \<Rightarrow> prgs"  ("_//_//\<parallel>//_" [60,90,57] 57)
+
+  "_prg_scheme" :: "['a, 'a, 'a, 'a, 'a] \<Rightarrow> prgs"  
+                  ("SCHEME [_ \<le> _ < _] _// _" [0,0,0,60, 90] 57)
+
+translations
+  "_prg c q" \<rightleftharpoons> "[(Some c, q)]"
+  "_prgs c q ps" \<rightleftharpoons> "(Some c, q) # ps"
+  "_PAR ps" \<rightleftharpoons> "Parallel ps"
+
+  "_prg_scheme j i k c q" \<rightleftharpoons> "map (\<lambda>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 Record.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 K_tr' (Abs (_,_,t)) = if null (loose_bnos t) then t else raise Match
+      | K_tr' (Abs (_,_,Abs (_,_,t)$Bound 0)) = if null (loose_bnos t) then t else raise Match
+      | K_tr' _ = raise Match;
+
+    fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) =
+          quote_tr' (Syntax.const "_Assign" $ update_name_tr' f)
+            (Abs (x, dummyT, K_tr' k) :: ts)
+      | assign_tr' _ = raise Match;
+
+    fun annassign_tr' (r :: Abs (x, _, f $ k $ Bound 0) :: ts) =
+          quote_tr' (Syntax.const "_AnnAssign" $ r $ update_name_tr' f)
+            (Abs (x, dummyT, K_tr' k) :: 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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/OG_Tactics.thy	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,496 @@
+header {* \section{Generation of Verification Conditions} *}
+
+theory OG_Tactics
+imports OG_Hoare
+begin
+
+lemmas ann_hoare_intros=AnnBasic AnnSeq AnnCond1 AnnCond2 AnnWhile AnnAwait AnnConseq
+lemmas oghoare_intros=Parallel Basic Seq Cond While Conseq
+
+lemma ParallelConseqRule: 
+ "\<lbrakk> p \<subseteq> (\<Inter>i\<in>{i. i<length Ts}. pre(the(com(Ts ! i))));  
+  \<parallel>- (\<Inter>i\<in>{i. i<length Ts}. pre(the(com(Ts ! i)))) 
+      (Parallel Ts) 
+     (\<Inter>i\<in>{i. i<length Ts}. post(Ts ! i));  
+  (\<Inter>i\<in>{i. i<length Ts}. post(Ts ! i)) \<subseteq> q \<rbrakk>  
+  \<Longrightarrow> \<parallel>- p (Parallel Ts) q"
+apply (rule Conseq)
+prefer 2 
+ apply fast
+apply assumption+
+done
+
+lemma SkipRule: "p \<subseteq> q \<Longrightarrow> \<parallel>- 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 \<subseteq> {s. (f s)\<in>q} \<Longrightarrow> \<parallel>- 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: "\<lbrakk> \<parallel>- p c1 r; \<parallel>- r c2 q \<rbrakk> \<Longrightarrow> \<parallel>- p (Seq c1 c2) q"
+apply(rule Seq)
+apply fast+
+done
+
+lemma CondRule: 
+ "\<lbrakk> p \<subseteq> {s. (s\<in>b \<longrightarrow> s\<in>w) \<and> (s\<notin>b \<longrightarrow> s\<in>w')}; \<parallel>- w c1 q; \<parallel>- w' c2 q \<rbrakk> 
+  \<Longrightarrow> \<parallel>- 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: "\<lbrakk> p \<subseteq> i; \<parallel>- (i \<inter> b) c i ; (i \<inter> (-b)) \<subseteq> q \<rbrakk>  
+        \<Longrightarrow> \<parallel>- 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:
+  "\<lbrakk> atom_com(c); \<parallel>- r c q \<rbrakk>  \<Longrightarrow> \<turnstile> (AnnAwait r {s. True} c) q"
+apply(rule AnnAwait)
+apply simp_all
+done
+
+lemma AnnskipRule:
+  "r \<subseteq> q \<Longrightarrow> \<turnstile> (AnnBasic r id) q"
+apply(rule AnnBasic)
+apply simp
+done
+
+lemma AnnwaitRule:
+  "\<lbrakk> (r \<inter> b) \<subseteq> q \<rbrakk> \<Longrightarrow> \<turnstile> (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: 
+  "\<forall>(R,r)\<in>(atomics a). \<parallel>- (q \<inter> R) r q \<Longrightarrow> interfree_aux(None, q, Some a)"
+apply(simp add:interfree_aux_def)
+apply(force elim:oghoare_sound)
+done
+
+lemma interfree_aux_rule3: 
+  "(\<forall>(R, r)\<in>(atomics a). \<parallel>- (q \<inter> R) r q \<and> (\<forall>p\<in>(assertions c). \<parallel>- (p \<inter> R) r p))
+  \<Longrightarrow> interfree_aux(Some c, q, Some a)"
+apply(simp add:interfree_aux_def)
+apply(force elim:oghoare_sound)
+done
+
+lemma AnnBasic_assertions: 
+  "\<lbrakk>interfree_aux(None, r, Some a); interfree_aux(None, q, Some a)\<rbrakk> \<Longrightarrow> 
+    interfree_aux(Some (AnnBasic r f), q, Some a)"
+apply(simp add: interfree_aux_def)
+by force
+
+lemma AnnSeq_assertions: 
+  "\<lbrakk> interfree_aux(Some c1, q, Some a); interfree_aux(Some c2, q, Some a)\<rbrakk>\<Longrightarrow> 
+   interfree_aux(Some (AnnSeq c1 c2), q, Some a)"
+apply(simp add: interfree_aux_def)
+by force
+
+lemma AnnCond1_assertions: 
+  "\<lbrakk> interfree_aux(None, r, Some a); interfree_aux(Some c1, q, Some a); 
+  interfree_aux(Some c2, q, Some a)\<rbrakk>\<Longrightarrow> 
+  interfree_aux(Some(AnnCond1 r b c1 c2), q, Some a)"
+apply(simp add: interfree_aux_def)
+by force
+
+lemma AnnCond2_assertions: 
+  "\<lbrakk> interfree_aux(None, r, Some a); interfree_aux(Some c, q, Some a)\<rbrakk>\<Longrightarrow> 
+  interfree_aux(Some (AnnCond2 r b c), q, Some a)"
+apply(simp add: interfree_aux_def)
+by force
+
+lemma AnnWhile_assertions: 
+  "\<lbrakk> interfree_aux(None, r, Some a); interfree_aux(None, i, Some a); 
+  interfree_aux(Some c, q, Some a)\<rbrakk>\<Longrightarrow> 
+  interfree_aux(Some (AnnWhile r b i c), q, Some a)"
+apply(simp add: interfree_aux_def)
+by force
+ 
+lemma AnnAwait_assertions: 
+  "\<lbrakk> interfree_aux(None, r, Some a); interfree_aux(None, q, Some a)\<rbrakk>\<Longrightarrow> 
+  interfree_aux(Some (AnnAwait r b c), q, Some a)"
+apply(simp add: interfree_aux_def)
+by force
+ 
+lemma AnnBasic_atomics: 
+  "\<parallel>- (q \<inter> r) (Basic f) q \<Longrightarrow> interfree_aux(None, q, Some (AnnBasic r f))"
+by(simp add: interfree_aux_def oghoare_sound)
+
+lemma AnnSeq_atomics: 
+  "\<lbrakk> interfree_aux(Any, q, Some a1); interfree_aux(Any, q, Some a2)\<rbrakk>\<Longrightarrow> 
+  interfree_aux(Any, q, Some (AnnSeq a1 a2))"
+apply(simp add: interfree_aux_def)
+by force
+
+lemma AnnCond1_atomics:
+  "\<lbrakk> interfree_aux(Any, q, Some a1); interfree_aux(Any, q, Some a2)\<rbrakk>\<Longrightarrow> 
+   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)\<Longrightarrow> interfree_aux(Any, q, Some (AnnCond2 r b a))"
+by(simp add: interfree_aux_def)
+
+lemma AnnWhile_atomics: "interfree_aux (Any, q, Some a) 
+     \<Longrightarrow> interfree_aux(Any, q, Some (AnnWhile r b i a))"
+by(simp add: interfree_aux_def)
+
+lemma Annatom_atomics: 
+  "\<parallel>- (q \<inter> r) a q \<Longrightarrow> interfree_aux (None, q, Some (AnnAwait r {x. True} a))"
+by(simp add: interfree_aux_def oghoare_sound) 
+
+lemma AnnAwait_atomics: 
+  "\<parallel>- (q \<inter> (r \<inter> b)) a q \<Longrightarrow> 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) \<Rightarrow> bool"
+  "interfree_swap == \<lambda>(x, xs). \<forall>y\<in>set xs. interfree_aux (com x, post x, com y)
+  \<and> 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:  
+  "\<lbrakk> interfree_aux (com x, post x, com y); 
+  interfree_aux (com y, post y ,com x); interfree_swap (x, xs) \<rbrakk> 
+  \<Longrightarrow> interfree_swap (x, y#xs)"
+by(simp add:interfree_swap_def)
+
+lemma interfree_swap_Map: "\<forall>k. i\<le>k \<and> k<j \<longrightarrow> interfree_aux (com x, post x, c k) 
+ \<and> interfree_aux (c k, Q k, com x)   
+ \<Longrightarrow> interfree_swap (x, map (\<lambda>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: 
+  "\<lbrakk> interfree_swap(x, xs); interfree xs \<rbrakk> \<Longrightarrow> 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: 
+  "(\<forall>i j. a\<le>i \<and> i<b \<and> a\<le>j \<and> j<b  \<and> i\<noteq>j \<longrightarrow> interfree_aux (c i, Q i, c j))  
+  \<Longrightarrow> interfree (map (\<lambda>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) \<Rightarrow> bool " ("[\<turnstile>] _" [0] 45)
+  "[\<turnstile>] Ts == (\<forall>i<length Ts. \<exists>c q. Ts!i=(Some c, q) \<and> \<turnstile> c q)"
+
+lemma MapAnnEmpty: "[\<turnstile>] []"
+by(simp add:map_ann_hoare_def)
+
+lemma MapAnnList: "\<lbrakk> \<turnstile> c q ; [\<turnstile>] xs \<rbrakk> \<Longrightarrow> [\<turnstile>] (Some c,q)#xs"
+apply(simp add:map_ann_hoare_def)
+apply clarify
+apply(case_tac i,simp+)
+done
+
+lemma MapAnnMap: 
+  "\<forall>k. i\<le>k \<and> k<j \<longrightarrow> \<turnstile> (c k) (Q k) \<Longrightarrow> [\<turnstile>] map (\<lambda>k. (Some (c k), Q k)) [i..<j]"
+apply(simp add: map_ann_hoare_def less_diff_conv)
+done
+
+lemma ParallelRule:"\<lbrakk> [\<turnstile>] Ts ; interfree Ts \<rbrakk>
+  \<Longrightarrow> \<parallel>- (\<Inter>i\<in>{i. i<length Ts}. pre(the(com(Ts!i)))) 
+          Parallel Ts 
+        (\<Inter>i\<in>{i. i<length Ts}. post(Ts!i))"
+apply(rule Parallel)
+ apply(simp add:map_ann_hoare_def)
+apply simp
+done
+(*
+lemma ParamParallelRule:
+ "\<lbrakk> \<forall>k<n. \<turnstile> (c k) (Q k); 
+   \<forall>k l. k<n \<and> l<n  \<and> k\<noteq>l \<longrightarrow> interfree_aux (Some(c k), Q k, Some(c l)) \<rbrakk>
+  \<Longrightarrow> \<parallel>- (\<Inter>i\<in>{i. i<n} . pre(c i)) COBEGIN SCHEME [0\<le>i<n] (c i) (Q i) COEND  (\<Inter>i\<in>{i. i<n} . Q i )"
+apply(rule ParallelConseqRule)
+  apply simp
+  apply clarify
+  apply force
+ apply(rule ParallelRule)
+  apply(rule MapAnnMap)
+  apply simp
+ apply(rule interfree_Map)
+ apply simp
+apply simp
+apply clarify
+apply force
+done
+*)
+
+text {* The following are some useful lemmas and simplification
+tactics to control which theorems are used to simplify at each moment,
+so that the original input does not suffer any unexpected
+transformation. *}
+
+lemma Compl_Collect: "-(Collect b) = {x. \<not>(b x)}"
+by fast
+lemma list_length: "length []=0 \<and> length (x#xs) = Suc(length xs)"
+by simp
+lemma list_lemmas: "length []=0 \<and> length (x#xs) = Suc(length xs) 
+\<and> (x#xs) ! 0=x \<and> (x#xs) ! Suc n = xs ! n"
+by simp
+lemma le_Suc_eq_insert: "{i. i <Suc n} = insert n {i. i< n}"
+by auto
+lemmas primrecdef_list = "pre.simps" "assertions.simps" "atomics.simps" "atom_com.simps"
+lemmas my_simp_list = list_lemmas fst_conv snd_conv
+not_less0 refl le_Suc_eq_insert Suc_not_Zero Zero_not_Suc nat.inject
+Collect_mem_eq ball_simps option.simps primrecdef_list
+lemmas ParallelConseq_list = INTER_def Collect_conj_eq length_map length_upt length_append list_length
+
+ML {*
+val before_interfree_simp_tac = (simp_tac (HOL_basic_ss addsimps [thm "com.simps", thm "post.simps"]))
+
+val  interfree_simp_tac = (asm_simp_tac (HOL_ss addsimps [thm "split", thm "ball_Un", thm "ball_empty"]@(thms "my_simp_list")))
+
+val ParallelConseq = (simp_tac (HOL_basic_ss addsimps (thms "ParallelConseq_list")@(thms "my_simp_list")))
+*}
+
+text {* The following tactic applies @{text tac} to each conjunct in a
+subgoal of the form @{text "A \<Longrightarrow> a1 \<and> a2 \<and> .. \<and> 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 "\<langle> \<rangle>"} 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 "\<parallel>- ?p c q"},
+ in this case we have precond=False and the generated  verification     
+ condition would have the form @{text "?p \<subseteq> \<dots>"} 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 {* 
+val oghoare_tac = SUBGOAL (fn (_, i) =>
+   (HoareRuleTac true i))
+*}
+
+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 {* val annhoare_tac = SUBGOAL (fn (_, i) =>
+  (AnnHoareRuleTac i))
+
+val interfree_aux_tac = SUBGOAL (fn (_, i) =>
+   (interfree_aux_Tac i))
+*}
+
+text {* The so defined ML tactics are then ``exported'' to be used in
+Isabelle proofs. *}
+
+method_setup oghoare = {*
+  Scan.succeed (K (SIMPLE_METHOD' oghoare_tac)) *}
+  "verification condition generator for the oghoare logic"
+
+method_setup annhoare = {*
+  Scan.succeed (K (SIMPLE_METHOD' annhoare_tac)) *}
+  "verification condition generator for the ann_hoare logic"
+
+method_setup interfree_aux = {*
+  Scan.succeed (K (SIMPLE_METHOD' interfree_aux_tac)) *}
+  "verification condition generator for interference freedom tests"
+
+text {* Tactics useful for dealing with the generated verification conditions: *}
+
+method_setup conjI_tac = {*
+  Scan.succeed (K (SIMPLE_METHOD' (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 = {*
+  Scan.succeed (K (SIMPLE_METHOD' (disjE_Tac (K all_tac)))) *}
+  "verification condition generator for interference freedom tests"
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/OG_Tran.thy	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,309 @@
+
+header {* \section{Operational Semantics} *}
+
+theory OG_Tran imports OG_Com begin
+
+types
+  'a ann_com_op = "('a ann_com) option"
+  'a ann_triple_op = "('a ann_com_op \<times> 'a assn)"
+  
+consts com :: "'a ann_triple_op \<Rightarrow> 'a ann_com_op"
+primrec "com (c, q) = c"
+
+consts post :: "'a ann_triple_op \<Rightarrow> 'a assn"
+primrec "post (c, q) = q"
+
+constdefs
+  All_None :: "'a ann_triple_op list \<Rightarrow> bool"
+  "All_None Ts \<equiv> \<forall>(c, q) \<in> set Ts. c = None"
+
+subsection {* The Transition Relation *}
+
+inductive_set
+  ann_transition :: "(('a ann_com_op \<times> 'a) \<times> ('a ann_com_op \<times> 'a)) set"        
+  and transition :: "(('a com \<times> 'a) \<times> ('a com \<times> 'a)) set"
+  and ann_transition' :: "('a ann_com_op \<times> 'a) \<Rightarrow> ('a ann_com_op \<times> 'a) \<Rightarrow> bool"
+    ("_ -1\<rightarrow> _"[81,81] 100)
+  and transition' :: "('a com \<times> 'a) \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"
+    ("_ -P1\<rightarrow> _"[81,81] 100)
+  and transitions :: "('a com \<times> 'a) \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"
+    ("_ -P*\<rightarrow> _"[81,81] 100)
+where
+  "con_0 -1\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> ann_transition"
+| "con_0 -P1\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> transition"
+| "con_0 -P*\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> transition\<^sup>*"
+
+| AnnBasic:  "(Some (AnnBasic r f), s) -1\<rightarrow> (None, f s)"
+
+| AnnSeq1: "(Some c0, s) -1\<rightarrow> (None, t) \<Longrightarrow> 
+               (Some (AnnSeq c0 c1), s) -1\<rightarrow> (Some c1, t)"
+| AnnSeq2: "(Some c0, s) -1\<rightarrow> (Some c2, t) \<Longrightarrow> 
+               (Some (AnnSeq c0 c1), s) -1\<rightarrow> (Some (AnnSeq c2 c1), t)"
+
+| AnnCond1T: "s \<in> b  \<Longrightarrow> (Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> (Some c1, s)"
+| AnnCond1F: "s \<notin> b \<Longrightarrow> (Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> (Some c2, s)"
+
+| AnnCond2T: "s \<in> b  \<Longrightarrow> (Some (AnnCond2 r b c), s) -1\<rightarrow> (Some c, s)"
+| AnnCond2F: "s \<notin> b \<Longrightarrow> (Some (AnnCond2 r b c), s) -1\<rightarrow> (None, s)"
+
+| AnnWhileF: "s \<notin> b \<Longrightarrow> (Some (AnnWhile r b i c), s) -1\<rightarrow> (None, s)"
+| AnnWhileT: "s \<in> b  \<Longrightarrow> (Some (AnnWhile r b i c), s) -1\<rightarrow> 
+                         (Some (AnnSeq c (AnnWhile i b i c)), s)"
+
+| AnnAwait: "\<lbrakk> s \<in> b; atom_com c; (c, s) -P*\<rightarrow> (Parallel [], t) \<rbrakk> \<Longrightarrow>
+	           (Some (AnnAwait r b c), s) -1\<rightarrow> (None, t)" 
+
+| Parallel: "\<lbrakk> i<length Ts; Ts!i = (Some c, q); (Some c, s) -1\<rightarrow> (r, t) \<rbrakk>
+              \<Longrightarrow> (Parallel Ts, s) -P1\<rightarrow> (Parallel (Ts [i:=(r, q)]), t)"
+
+| Basic:  "(Basic f, s) -P1\<rightarrow> (Parallel [], f s)"
+
+| Seq1:   "All_None Ts \<Longrightarrow> (Seq (Parallel Ts) c, s) -P1\<rightarrow> (c, s)"
+| Seq2:   "(c0, s) -P1\<rightarrow> (c2, t) \<Longrightarrow> (Seq c0 c1, s) -P1\<rightarrow> (Seq c2 c1, t)"
+
+| CondT: "s \<in> b \<Longrightarrow> (Cond b c1 c2, s) -P1\<rightarrow> (c1, s)"
+| CondF: "s \<notin> b \<Longrightarrow> (Cond b c1 c2, s) -P1\<rightarrow> (c2, s)"
+
+| WhileF: "s \<notin> b \<Longrightarrow> (While b i c, s) -P1\<rightarrow> (Parallel [], s)"
+| WhileT: "s \<in> b \<Longrightarrow> (While b i c, s) -P1\<rightarrow> (Seq c (While b i c), s)"
+
+monos "rtrancl_mono"
+
+text {* The corresponding syntax translations are: *}
+
+abbreviation
+  ann_transition_n :: "('a ann_com_op \<times> 'a) \<Rightarrow> nat \<Rightarrow> ('a ann_com_op \<times> 'a) 
+                           \<Rightarrow> bool"  ("_ -_\<rightarrow> _"[81,81] 100)  where
+  "con_0 -n\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> ann_transition ^^ n"
+
+abbreviation
+  ann_transitions :: "('a ann_com_op \<times> 'a) \<Rightarrow> ('a ann_com_op \<times> 'a) \<Rightarrow> bool"
+                           ("_ -*\<rightarrow> _"[81,81] 100)  where
+  "con_0 -*\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> ann_transition\<^sup>*"
+
+abbreviation
+  transition_n :: "('a com \<times> 'a) \<Rightarrow> nat \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"  
+                          ("_ -P_\<rightarrow> _"[81,81,81] 100)  where
+  "con_0 -Pn\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> transition ^^ n"
+
+subsection {* Definition of Semantics *}
+
+constdefs
+  ann_sem :: "'a ann_com \<Rightarrow> 'a \<Rightarrow> 'a set"
+  "ann_sem c \<equiv> \<lambda>s. {t. (Some c, s) -*\<rightarrow> (None, t)}"
+
+  ann_SEM :: "'a ann_com \<Rightarrow> 'a set \<Rightarrow> 'a set"
+  "ann_SEM c S \<equiv> \<Union>ann_sem c ` S"  
+
+  sem :: "'a com \<Rightarrow> 'a \<Rightarrow> 'a set"
+  "sem c \<equiv> \<lambda>s. {t. \<exists>Ts. (c, s) -P*\<rightarrow> (Parallel Ts, t) \<and> All_None Ts}"
+
+  SEM :: "'a com \<Rightarrow> 'a set \<Rightarrow> 'a set"
+  "SEM c S \<equiv> \<Union>sem c ` S "
+
+syntax "_Omega" :: "'a com"    ("\<Omega>" 63)
+translations  "\<Omega>" \<rightleftharpoons> "While CONST UNIV CONST UNIV (Basic id)"
+
+consts fwhile :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> nat \<Rightarrow> 'a com"
+primrec 
+   "fwhile b c 0 = \<Omega>"
+   "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\<rightarrow> t"  
+    "(Basic f, s) -P1\<rightarrow> t"
+    "(Seq c1 c2, s) -P1\<rightarrow> t" 
+    "(Cond b c1 c2, s) -P1\<rightarrow> t"
+    "(While b i c, s) -P1\<rightarrow> t"
+
+lemma Parallel_empty_lemma [rule_format (no_asm)]: 
+  "(Parallel [],s) -Pn\<rightarrow> (Parallel Ts,t) \<longrightarrow> Ts=[] \<and> n=0 \<and> 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 \<longrightarrow> (Parallel Ss,s) -Pn\<rightarrow> (Parallel Ts,t) \<longrightarrow> Ts=Ss \<and> n=0 \<and> 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 \<Longrightarrow> (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=[] \<Longrightarrow> (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\<subseteq>Y \<Longrightarrow> SEM c X \<subseteq> SEM c Y" 
+apply (unfold SEM_def)
+apply force
+done
+
+lemma L3_5ii_lemma1: 
+ "\<lbrakk> (c1, s1) -P*\<rightarrow> (Parallel Ts, s2); All_None Ts;  
+  (c2, s2) -P*\<rightarrow> (Parallel Ss, s3); All_None Ss \<rbrakk> 
+ \<Longrightarrow> (Seq c1 c2, s1) -P*\<rightarrow> (Parallel Ss, s3)"
+apply(erule converse_rtrancl_induct2)
+apply(force intro:converse_rtrancl_into_rtrancl)+
+done
+
+lemma L3_5ii_lemma2 [rule_format (no_asm)]: 
+ "\<forall>c1 c2 s t. (Seq c1 c2, s) -Pn\<rightarrow> (Parallel Ts, t) \<longrightarrow>  
+  (All_None Ts) \<longrightarrow> (\<exists>y m Rs. (c1,s) -P*\<rightarrow> (Parallel Rs, y) \<and> 
+  (All_None Rs) \<and> (c2, y) -Pm\<rightarrow> (Parallel Ts, t) \<and>  m \<le> 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: 
+ "\<lbrakk>(Seq c1 c2,s) -P*\<rightarrow> (Parallel Ts,t); All_None Ts\<rbrakk> \<Longrightarrow> 
+    (\<exists>y Rs. (c1,s) -P*\<rightarrow> (Parallel Rs,y) \<and> All_None Rs 
+   \<and> (c2,y) -P*\<rightarrow> (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 \<inter> b)) Un (SEM c2 (X \<inter> (-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\<rightarrow> (T,t) \<longrightarrow> S=\<Omega> \<longrightarrow> (\<not>(\<exists>Rs. T=(Parallel Rs) \<and> 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: "\<lbrakk>(\<Omega>, s) -P*\<rightarrow> (Parallel Ts, t); All_None Ts \<rbrakk> \<Longrightarrow> False"
+apply(fast dest: rtrancl_imp_UN_rel_pow L3_5v_lemma1)
+done
+
+lemma L3_5v_lemma3: "SEM (\<Omega>) S = {}"
+apply (unfold SEM_def sem_def)
+apply(fast dest: L3_5v_lemma2)
+done
+
+lemma L3_5v_lemma4 [rule_format]: 
+ "\<forall>s. (While b i c, s) -Pn\<rightarrow> (Parallel Ts, t) \<longrightarrow> All_None Ts \<longrightarrow>  
+  (\<exists>k. (fwhile b c k, s) -P*\<rightarrow> (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]: 
+ "\<forall>s. (fwhile b c k, s) -P*\<rightarrow> (Parallel Ts, t) \<longrightarrow> All_None Ts \<longrightarrow>  
+  (While b i c, s) -P*\<rightarrow> (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) = (\<lambda>x. (\<Union>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 \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool"  ("(3\<parallel>= _// _//_)" [90,55,90] 50)
+  "\<parallel>= p c q \<equiv> SEM c p \<subseteq> q"
+
+  ann_com_validity :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool"   ("\<Turnstile> _ _" [60,90] 45)
+  "\<Turnstile> c q \<equiv> ann_SEM c (pre c) \<subseteq> q"
+
+end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/Quote_Antiquote.thy	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,24 @@
+
+header {* \section{Concrete Syntax} *}
+
+theory Quote_Antiquote imports Main begin
+
+syntax
+  "_quote"     :: "'b \<Rightarrow> ('a \<Rightarrow> 'b)"                ("(\<guillemotleft>_\<guillemotright>)" [0] 1000)
+  "_antiquote" :: "('a \<Rightarrow> 'b) \<Rightarrow> 'b"                ("\<acute>_" [1000] 1000)
+  "_Assert"    :: "'a \<Rightarrow> 'a set"                    ("(.{_}.)" [0] 1000)
+
+syntax (xsymbols)
+  "_Assert"    :: "'a \<Rightarrow> 'a set"            ("(\<lbrace>_\<rbrace>)" [0] 1000)
+
+translations
+  ".{b}." \<rightharpoonup> "Collect \<guillemotleft>b\<guillemotright>"
+
+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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/RG_Com.thy	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,25 @@
+
+header {* \chapter{The Rely-Guarantee Method} 
+
+\section {Abstract Syntax}
+*}
+
+theory RG_Com imports Main begin
+
+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 \<Rightarrow>'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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/RG_Examples.thy	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,359 @@
+header {* \section{Examples} *}
+
+theory RG_Examples
+imports RG_Syntax
+begin
+
+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: "\<lbrakk>(j::nat)<k; i\<le> j\<rbrakk> \<Longrightarrow> i<k"
+by simp
+
+lemma add_le_less_mono: "\<lbrakk> (a::nat) < c; b\<le>d \<rbrakk> \<Longrightarrow> a + b < c + d"
+by simp
+
+record Example1 =
+  A :: "nat list"
+
+lemma Example1: 
+ "\<turnstile> COBEGIN
+      SCHEME [0 \<le> i < n]
+     (\<acute>A := \<acute>A [i := 0], 
+     \<lbrace> n < length \<acute>A \<rbrace>, 
+     \<lbrace> length \<ordmasculine>A = length \<ordfeminine>A \<and> \<ordmasculine>A ! i = \<ordfeminine>A ! i \<rbrace>, 
+     \<lbrace> length \<ordmasculine>A = length \<ordfeminine>A \<and> (\<forall>j<n. i \<noteq> j \<longrightarrow> \<ordmasculine>A ! j = \<ordfeminine>A ! j) \<rbrace>, 
+     \<lbrace> \<acute>A ! i = 0 \<rbrace>) 
+    COEND
+ SAT [\<lbrace> n < length \<acute>A \<rbrace>, \<lbrace> \<ordmasculine>A = \<ordfeminine>A \<rbrace>, \<lbrace> True \<rbrace>, \<lbrace> \<forall>i < n. \<acute>A ! i = 0 \<rbrace>]"
+apply(rule Parallel)
+apply (auto intro!: Basic) 
+done
+
+lemma Example1_parameterized: 
+"k < t \<Longrightarrow>
+  \<turnstile> COBEGIN 
+    SCHEME [k*n\<le>i<(Suc k)*n] (\<acute>A:=\<acute>A[i:=0], 
+   \<lbrace>t*n < length \<acute>A\<rbrace>, 
+   \<lbrace>t*n < length \<ordmasculine>A \<and> length \<ordmasculine>A=length \<ordfeminine>A \<and> \<ordmasculine>A!i = \<ordfeminine>A!i\<rbrace>, 
+   \<lbrace>t*n < length \<ordmasculine>A \<and> length \<ordmasculine>A=length \<ordfeminine>A \<and> (\<forall>j<length \<ordmasculine>A . i\<noteq>j \<longrightarrow> \<ordmasculine>A!j = \<ordfeminine>A!j)\<rbrace>, 
+   \<lbrace>\<acute>A!i=0\<rbrace>) 
+   COEND  
+ SAT [\<lbrace>t*n < length \<acute>A\<rbrace>, 
+      \<lbrace>t*n < length \<ordmasculine>A \<and> length \<ordmasculine>A=length \<ordfeminine>A \<and> (\<forall>i<n. \<ordmasculine>A!(k*n+i)=\<ordfeminine>A!(k*n+i))\<rbrace>, 
+      \<lbrace>t*n < length \<ordmasculine>A \<and> length \<ordmasculine>A=length \<ordfeminine>A \<and> 
+      (\<forall>i<length \<ordmasculine>A . (i<k*n \<longrightarrow> \<ordmasculine>A!i = \<ordfeminine>A!i) \<and> ((Suc k)*n \<le> i\<longrightarrow> \<ordmasculine>A!i = \<ordfeminine>A!i))\<rbrace>, 
+      \<lbrace>\<forall>i<n. \<acute>A!(k*n+i) = 0\<rbrace>]"
+apply(rule Parallel)
+    apply auto
+  apply(erule_tac x="k*n +i" in allE)
+  apply(subgoal_tac "k*n+i <length (A b)")
+   apply force
+  apply(erule le_less_trans2) 
+  apply(case_tac t,simp+)
+  apply (simp add:add_commute)
+  apply(simp add: add_le_mono)
+apply(rule Basic)
+   apply simp
+   apply clarify
+   apply (subgoal_tac "k*n+i< length (A x)")
+    apply simp
+   apply(erule le_less_trans2)
+   apply(case_tac t,simp+)
+   apply (simp add:add_commute)
+   apply(rule add_le_mono, auto)
+done
+
+
+subsection {* Increment a Variable in Parallel *}
+
+subsubsection {* Two components *}
+
+record Example2 =
+  x  :: nat
+  c_0 :: nat
+  c_1 :: nat
+
+lemma Example2: 
+ "\<turnstile>  COBEGIN
+    (\<langle> \<acute>x:=\<acute>x+1;; \<acute>c_0:=\<acute>c_0 + 1 \<rangle>, 
+     \<lbrace>\<acute>x=\<acute>c_0 + \<acute>c_1  \<and> \<acute>c_0=0\<rbrace>, 
+     \<lbrace>\<ordmasculine>c_0 = \<ordfeminine>c_0 \<and> 
+        (\<ordmasculine>x=\<ordmasculine>c_0 + \<ordmasculine>c_1 
+        \<longrightarrow> \<ordfeminine>x = \<ordfeminine>c_0 + \<ordfeminine>c_1)\<rbrace>,  
+     \<lbrace>\<ordmasculine>c_1 = \<ordfeminine>c_1 \<and> 
+         (\<ordmasculine>x=\<ordmasculine>c_0 + \<ordmasculine>c_1 
+         \<longrightarrow> \<ordfeminine>x =\<ordfeminine>c_0 + \<ordfeminine>c_1)\<rbrace>,
+     \<lbrace>\<acute>x=\<acute>c_0 + \<acute>c_1 \<and> \<acute>c_0=1 \<rbrace>)
+  \<parallel>
+      (\<langle> \<acute>x:=\<acute>x+1;; \<acute>c_1:=\<acute>c_1+1 \<rangle>, 
+     \<lbrace>\<acute>x=\<acute>c_0 + \<acute>c_1 \<and> \<acute>c_1=0 \<rbrace>, 
+     \<lbrace>\<ordmasculine>c_1 = \<ordfeminine>c_1 \<and> 
+        (\<ordmasculine>x=\<ordmasculine>c_0 + \<ordmasculine>c_1 
+        \<longrightarrow> \<ordfeminine>x = \<ordfeminine>c_0 + \<ordfeminine>c_1)\<rbrace>,  
+     \<lbrace>\<ordmasculine>c_0 = \<ordfeminine>c_0 \<and> 
+         (\<ordmasculine>x=\<ordmasculine>c_0 + \<ordmasculine>c_1 
+        \<longrightarrow> \<ordfeminine>x =\<ordfeminine>c_0 + \<ordfeminine>c_1)\<rbrace>,
+     \<lbrace>\<acute>x=\<acute>c_0 + \<acute>c_1 \<and> \<acute>c_1=1\<rbrace>)
+ COEND
+ SAT [\<lbrace>\<acute>x=0 \<and> \<acute>c_0=0 \<and> \<acute>c_1=0\<rbrace>, 
+      \<lbrace>\<ordmasculine>x=\<ordfeminine>x \<and>  \<ordmasculine>c_0= \<ordfeminine>c_0 \<and> \<ordmasculine>c_1=\<ordfeminine>c_1\<rbrace>,
+      \<lbrace>True\<rbrace>,
+      \<lbrace>\<acute>x=2\<rbrace>]"
+apply(rule Parallel)
+   apply simp_all
+   apply clarify
+   apply(case_tac i)
+    apply simp
+    apply(rule conjI)
+     apply clarify
+     apply simp
+    apply clarify
+    apply simp
+    apply(case_tac j,simp)
+    apply simp
+   apply simp
+   apply(rule conjI)
+    apply clarify
+    apply simp
+   apply clarify
+   apply simp
+   apply(subgoal_tac "j=0")
+    apply (rotate_tac -1)
+    apply (simp (asm_lr))
+   apply arith
+  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(auto intro!: Basic)
+done
+
+subsubsection {* Parameterized *}
+
+lemma Example2_lemma2_aux: "j<n \<Longrightarrow> 
+ (\<Sum>i=0..<n. (b i::nat)) =
+ (\<Sum>i=0..<j. b i) + b j + (\<Sum>i=0..<n-(Suc j) . b (Suc j + i))"
+apply(induct n)
+ apply simp_all
+apply(simp add:less_Suc_eq)
+ apply(auto)
+apply(subgoal_tac "n - j = Suc(n- Suc j)")
+  apply simp
+apply arith
+done
+
+lemma Example2_lemma2_aux2: 
+  "j\<le> s \<Longrightarrow> (\<Sum>i::nat=0..<j. (b (s:=t)) i) = (\<Sum>i=0..<j. b i)"
+apply(induct j)
+ apply (simp_all cong:setsum_cong)
+done
+
+lemma Example2_lemma2: 
+ "\<lbrakk>j<n; b j=0\<rbrakk> \<Longrightarrow> Suc (\<Sum>i::nat=0..<n. b i)=(\<Sum>i=0..<n. (b (j := Suc 0)) i)"
+apply(frule_tac b="(b (j:=(Suc 0)))" in Example2_lemma2_aux)
+apply(erule_tac  t="setsum (b(j := (Suc 0))) {0..<n}" in ssubst)
+apply(frule_tac b=b in Example2_lemma2_aux)
+apply(erule_tac  t="setsum b {0..<n}" in ssubst)
+apply(subgoal_tac "Suc (setsum b {0..<j} + b j + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))=(setsum b {0..<j} + Suc (b j) + (\<Sum>i=0..<n - Suc j. b (Suc j + i)))")
+apply(rotate_tac -1)
+apply(erule ssubst)
+apply(subgoal_tac "j\<le>j")
+ apply(drule_tac b="b" and t="(Suc 0)" in Example2_lemma2_aux2)
+apply(rotate_tac -1)
+apply(erule ssubst)
+apply simp_all
+done
+
+lemma Example2_lemma2_Suc0: "\<lbrakk>j<n; b j=0\<rbrakk> \<Longrightarrow>
+ Suc (\<Sum>i::nat=0..< n. b i)=(\<Sum>i=0..< n. (b (j:=Suc 0)) i)"
+by(simp add:Example2_lemma2)
+
+record Example2_parameterized =   
+  C :: "nat \<Rightarrow> nat"
+  y  :: nat
+
+lemma Example2_parameterized: "0<n \<Longrightarrow> 
+  \<turnstile> COBEGIN SCHEME  [0\<le>i<n]
+     (\<langle> \<acute>y:=\<acute>y+1;; \<acute>C:=\<acute>C (i:=1) \<rangle>, 
+     \<lbrace>\<acute>y=(\<Sum>i=0..<n. \<acute>C i) \<and> \<acute>C i=0\<rbrace>, 
+     \<lbrace>\<ordmasculine>C i = \<ordfeminine>C i \<and> 
+      (\<ordmasculine>y=(\<Sum>i=0..<n. \<ordmasculine>C i) \<longrightarrow> \<ordfeminine>y =(\<Sum>i=0..<n. \<ordfeminine>C i))\<rbrace>,  
+     \<lbrace>(\<forall>j<n. i\<noteq>j \<longrightarrow> \<ordmasculine>C j = \<ordfeminine>C j) \<and> 
+       (\<ordmasculine>y=(\<Sum>i=0..<n. \<ordmasculine>C i) \<longrightarrow> \<ordfeminine>y =(\<Sum>i=0..<n. \<ordfeminine>C i))\<rbrace>,
+     \<lbrace>\<acute>y=(\<Sum>i=0..<n. \<acute>C i) \<and> \<acute>C i=1\<rbrace>) 
+    COEND
+ SAT [\<lbrace>\<acute>y=0 \<and> (\<Sum>i=0..<n. \<acute>C i)=0 \<rbrace>, \<lbrace>\<ordmasculine>C=\<ordfeminine>C \<and> \<ordmasculine>y=\<ordfeminine>y\<rbrace>, \<lbrace>True\<rbrace>, \<lbrace>\<acute>y=n\<rbrace>]"
+apply(rule Parallel)
+apply force
+apply force
+apply(force)
+apply clarify
+apply simp
+apply(simp cong:setsum_ivl_cong)
+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(simp add:Example2_lemma2_Suc0 cong:if_cong)
+apply simp+
+done
+
+subsection {* Find Least Element *}
+
+text {* A previous lemma: *}
+
+lemma mod_aux :"\<lbrakk>i < (n::nat); a mod n = i;  j < a + n; j mod n = i; a < j\<rbrakk> \<Longrightarrow> False"
+apply(subgoal_tac "a=a div n*n + a mod n" )
+ prefer 2 apply (simp (no_asm_use))
+apply(subgoal_tac "j=j div n*n + j mod n")
+ prefer 2 apply (simp (no_asm_use))
+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 \<Rightarrow> nat"
+  Y :: "nat \<Rightarrow> nat"
+
+lemma Example3: "m mod n=0 \<Longrightarrow> 
+ \<turnstile> COBEGIN 
+ SCHEME [0\<le>i<n]
+ (WHILE (\<forall>j<n. \<acute>X i < \<acute>Y j)  DO 
+   IF P(B!(\<acute>X i)) THEN \<acute>Y:=\<acute>Y (i:=\<acute>X i) 
+   ELSE \<acute>X:= \<acute>X (i:=(\<acute>X i)+ n) FI 
+  OD,
+ \<lbrace>(\<acute>X i) mod n=i \<and> (\<forall>j<\<acute>X i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and> (\<acute>Y i<m \<longrightarrow> P(B!(\<acute>Y i)) \<and> \<acute>Y i\<le> m+i)\<rbrace>,
+ \<lbrace>(\<forall>j<n. i\<noteq>j \<longrightarrow> \<ordfeminine>Y j \<le> \<ordmasculine>Y j) \<and> \<ordmasculine>X i = \<ordfeminine>X i \<and> 
+   \<ordmasculine>Y i = \<ordfeminine>Y i\<rbrace>,
+ \<lbrace>(\<forall>j<n. i\<noteq>j \<longrightarrow> \<ordmasculine>X j = \<ordfeminine>X j \<and> \<ordmasculine>Y j = \<ordfeminine>Y j) \<and>   
+   \<ordfeminine>Y i \<le> \<ordmasculine>Y i\<rbrace>,
+ \<lbrace>(\<acute>X i) mod n=i \<and> (\<forall>j<\<acute>X i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and> (\<acute>Y i<m \<longrightarrow> P(B!(\<acute>Y i)) \<and> \<acute>Y i\<le> m+i) \<and> (\<exists>j<n. \<acute>Y j \<le> \<acute>X i) \<rbrace>) 
+ COEND
+ SAT [\<lbrace> \<forall>i<n. \<acute>X i=i \<and> \<acute>Y i=m+i \<rbrace>,\<lbrace>\<ordmasculine>X=\<ordfeminine>X \<and> \<ordmasculine>Y=\<ordfeminine>Y\<rbrace>,\<lbrace>True\<rbrace>,
+  \<lbrace>\<forall>i<n. (\<acute>X i) mod n=i \<and> (\<forall>j<\<acute>X i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and> 
+    (\<acute>Y i<m \<longrightarrow> P(B!(\<acute>Y i)) \<and> \<acute>Y i\<le> m+i) \<and> (\<exists>j<n. \<acute>Y j \<le> \<acute>X i)\<rbrace>]"
+apply(rule Parallel)
+--{*5 subgoals left *}
+apply force+
+apply clarify
+apply simp
+apply(rule While)
+    apply force
+   apply force
+  apply force
+ apply(rule_tac pre'="\<lbrace> \<acute>X i mod n = i \<and> (\<forall>j. j<\<acute>X i \<longrightarrow> j mod n = i \<longrightarrow> \<not>P(B!j)) \<and> (\<acute>Y i < n * q \<longrightarrow> P (B!(\<acute>Y i))) \<and> \<acute>X i<\<acute>Y i\<rbrace>" in Conseq)
+     apply force
+    apply(rule subset_refl)+
+ apply(rule Cond)
+    apply force
+   apply(rule Basic)
+      apply force
+     apply fastsimp
+    apply force
+   apply force
+  apply(rule Basic)
+     apply simp
+     apply clarify
+     apply simp
+     apply (case_tac "X x (j mod n) \<le> 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 auto
+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 \<Longrightarrow> \<turnstile> (COBEGIN SCHEME [0\<le>i<n]
+ (WHILE (\<forall>j<n. \<acute>X!i < \<acute>Y!j)  DO 
+     IF P(B!(\<acute>X!i)) THEN \<acute>Y:=\<acute>Y[i:=\<acute>X!i] ELSE \<acute>X:= \<acute>X[i:=(\<acute>X!i)+ n] FI 
+  OD,
+ \<lbrace>n<length \<acute>X \<and> n<length \<acute>Y \<and> (\<acute>X!i) mod n=i \<and> (\<forall>j<\<acute>X!i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and> (\<acute>Y!i<m \<longrightarrow> P(B!(\<acute>Y!i)) \<and> \<acute>Y!i\<le> m+i)\<rbrace>,
+ \<lbrace>(\<forall>j<n. i\<noteq>j \<longrightarrow> \<ordfeminine>Y!j \<le> \<ordmasculine>Y!j) \<and> \<ordmasculine>X!i = \<ordfeminine>X!i \<and> 
+   \<ordmasculine>Y!i = \<ordfeminine>Y!i \<and> length \<ordmasculine>X = length \<ordfeminine>X \<and> length \<ordmasculine>Y = length \<ordfeminine>Y\<rbrace>,
+ \<lbrace>(\<forall>j<n. i\<noteq>j \<longrightarrow> \<ordmasculine>X!j = \<ordfeminine>X!j \<and> \<ordmasculine>Y!j = \<ordfeminine>Y!j) \<and>   
+   \<ordfeminine>Y!i \<le> \<ordmasculine>Y!i \<and> length \<ordmasculine>X = length \<ordfeminine>X \<and> length \<ordmasculine>Y = length \<ordfeminine>Y\<rbrace>,
+ \<lbrace>(\<acute>X!i) mod n=i \<and> (\<forall>j<\<acute>X!i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and> (\<acute>Y!i<m \<longrightarrow> P(B!(\<acute>Y!i)) \<and> \<acute>Y!i\<le> m+i) \<and> (\<exists>j<n. \<acute>Y!j \<le> \<acute>X!i) \<rbrace>) COEND)
+ SAT [\<lbrace>n<length \<acute>X \<and> n<length \<acute>Y \<and> (\<forall>i<n. \<acute>X!i=i \<and> \<acute>Y!i=m+i) \<rbrace>,
+      \<lbrace>\<ordmasculine>X=\<ordfeminine>X \<and> \<ordmasculine>Y=\<ordfeminine>Y\<rbrace>,
+      \<lbrace>True\<rbrace>,
+      \<lbrace>\<forall>i<n. (\<acute>X!i) mod n=i \<and> (\<forall>j<\<acute>X!i. j mod n=i \<longrightarrow> \<not>P(B!j)) \<and> 
+        (\<acute>Y!i<m \<longrightarrow> P(B!(\<acute>Y!i)) \<and> \<acute>Y!i\<le> m+i) \<and> (\<exists>j<n. \<acute>Y!j \<le> \<acute>X!i)\<rbrace>]"
+apply(rule Parallel)
+--{* 5 subgoals left *}
+apply force+
+apply clarify
+apply simp
+apply(rule While)
+    apply force
+   apply force
+  apply force
+ apply(rule_tac pre'="\<lbrace>n<length \<acute>X \<and> n<length \<acute>Y \<and> \<acute>X ! i mod n = i \<and> (\<forall>j. j < \<acute>X ! i \<longrightarrow> j mod n = i \<longrightarrow> \<not> P (B ! j)) \<and> (\<acute>Y ! i < n * q \<longrightarrow> P (B ! (\<acute>Y ! i))) \<and> \<acute>X!i<\<acute>Y!i\<rbrace>" 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\<le> 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 auto
+done
+
+end
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/RG_Hoare.thy	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,1375 @@
+header {* \section{The Proof System} *}
+
+theory RG_Hoare imports RG_Tran begin
+
+subsection {* Proof System for Component Programs *}
+
+declare Un_subset_iff [iff del]
+declare Cons_eq_map_conv[iff]
+
+constdefs
+  stable :: "'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> bool"  
+  "stable \<equiv> \<lambda>f g. (\<forall>x y. x \<in> f \<longrightarrow> (x, y) \<in> g \<longrightarrow> y \<in> f)" 
+
+inductive
+  rghoare :: "['a com, 'a set, ('a \<times> 'a) set, ('a \<times> 'a) set, 'a set] \<Rightarrow> bool"  
+    ("\<turnstile> _ sat [_, _, _, _]" [60,0,0,0,0] 45)
+where
+  Basic: "\<lbrakk> pre \<subseteq> {s. f s \<in> post}; {(s,t). s \<in> pre \<and> (t=f s \<or> t=s)} \<subseteq> guar; 
+            stable pre rely; stable post rely \<rbrakk> 
+           \<Longrightarrow> \<turnstile> Basic f sat [pre, rely, guar, post]"
+
+| Seq: "\<lbrakk> \<turnstile> P sat [pre, rely, guar, mid]; \<turnstile> Q sat [mid, rely, guar, post] \<rbrakk> 
+           \<Longrightarrow> \<turnstile> Seq P Q sat [pre, rely, guar, post]"
+
+| Cond: "\<lbrakk> stable pre rely; \<turnstile> P1 sat [pre \<inter> b, rely, guar, post];
+           \<turnstile> P2 sat [pre \<inter> -b, rely, guar, post]; \<forall>s. (s,s)\<in>guar \<rbrakk>
+          \<Longrightarrow> \<turnstile> Cond b P1 P2 sat [pre, rely, guar, post]"
+
+| While: "\<lbrakk> stable pre rely; (pre \<inter> -b) \<subseteq> post; stable post rely;
+            \<turnstile> P sat [pre \<inter> b, rely, guar, pre]; \<forall>s. (s,s)\<in>guar \<rbrakk>
+          \<Longrightarrow> \<turnstile> While b P sat [pre, rely, guar, post]"
+
+| Await: "\<lbrakk> stable pre rely; stable post rely; 
+            \<forall>V. \<turnstile> P sat [pre \<inter> b \<inter> {V}, {(s, t). s = t}, 
+                UNIV, {s. (V, s) \<in> guar} \<inter> post] \<rbrakk>
+           \<Longrightarrow> \<turnstile> Await b P sat [pre, rely, guar, post]"
+  
+| Conseq: "\<lbrakk> pre \<subseteq> pre'; rely \<subseteq> rely'; guar' \<subseteq> guar; post' \<subseteq> post;
+             \<turnstile> P sat [pre', rely', guar', post'] \<rbrakk>
+            \<Longrightarrow> \<turnstile> P sat [pre, rely, guar, post]"
+
+constdefs 
+  Pre :: "'a rgformula \<Rightarrow> 'a set"
+  "Pre x \<equiv> fst(snd x)"
+  Post :: "'a rgformula \<Rightarrow> 'a set"
+  "Post x \<equiv> snd(snd(snd(snd x)))"
+  Rely :: "'a rgformula \<Rightarrow> ('a \<times> 'a) set"
+  "Rely x \<equiv> fst(snd(snd x))"
+  Guar :: "'a rgformula \<Rightarrow> ('a \<times> 'a) set"
+  "Guar x \<equiv> fst(snd(snd(snd x)))"
+  Com :: "'a rgformula \<Rightarrow> 'a com"
+  "Com x \<equiv> fst x"
+
+subsection {* Proof System for Parallel Programs *}
+
+types 'a par_rgformula = "('a rgformula) list \<times> 'a set \<times> ('a \<times> 'a) set \<times> ('a \<times> 'a) set \<times> 'a set"
+
+inductive
+  par_rghoare :: "('a rgformula) list \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set \<Rightarrow> bool"
+    ("\<turnstile> _ SAT [_, _, _, _]" [60,0,0,0,0] 45)
+where
+  Parallel: 
+  "\<lbrakk> \<forall>i<length xs. rely \<union> (\<Union>j\<in>{j. j<length xs \<and> j\<noteq>i}. Guar(xs!j)) \<subseteq> Rely(xs!i);
+    (\<Union>j\<in>{j. j<length xs}. Guar(xs!j)) \<subseteq> guar;
+     pre \<subseteq> (\<Inter>i\<in>{i. i<length xs}. Pre(xs!i)); 
+    (\<Inter>i\<in>{i. i<length xs}. Post(xs!i)) \<subseteq> post;
+    \<forall>i<length xs. \<turnstile> Com(xs!i) sat [Pre(xs!i),Rely(xs!i),Guar(xs!i),Post(xs!i)] \<rbrakk>
+   \<Longrightarrow>  \<turnstile> xs SAT [pre, rely, guar, post]"
+
+section {* Soundness*}
+
+subsubsection {* Some previous lemmas *}
+
+lemma tl_of_assum_in_assum: 
+  "(P, s) # (P, t) # xs \<in> assum (pre, rely) \<Longrightarrow> stable pre rely 
+  \<Longrightarrow> (P, t) # xs \<in> 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 \<in> comm(guar, post) \<Longrightarrow> (P, s) # (P, t) # xs \<in> comm(guar, post)"
+apply(simp add:comm_def)
+apply clarify
+apply(case_tac i,simp+)
+done
+
+lemma ctran_in_comm: 
+  "\<lbrakk>(s, s) \<in> guar; (Q, s) # xs \<in> comm(guar, post)\<rbrakk> 
+  \<Longrightarrow> (P, s) # (Q, s) # xs \<in> comm(guar, post)"
+apply(simp add:comm_def)
+apply clarify
+apply(case_tac i,simp+)
+done
+
+lemma takecptn_is_cptn [rule_format, elim!]: 
+  "\<forall>j. c \<in> cptn \<longrightarrow> take (Suc j) c \<in> cptn"
+apply(induct "c")
+ apply(force elim: cptn.cases)
+apply clarify
+apply(case_tac j) 
+ apply simp
+ apply(rule CptnOne)
+apply simp
+apply(force intro:cptn.intros elim:cptn.cases)
+done
+
+lemma dropcptn_is_cptn [rule_format,elim!]: 
+  "\<forall>j<length c. c \<in> cptn \<longrightarrow> drop j c \<in> cptn"
+apply(induct "c")
+ apply(force elim: cptn.cases)
+apply clarify
+apply(case_tac j,simp+) 
+apply(erule cptn.cases)
+  apply simp
+ apply force
+apply force
+done
+
+lemma takepar_cptn_is_par_cptn [rule_format,elim]: 
+  "\<forall>j. c \<in> par_cptn \<longrightarrow> take (Suc j) c \<in> par_cptn"
+apply(induct "c")
+ apply(force elim: cptn.cases)
+apply clarify
+apply(case_tac j,simp) 
+ apply(rule ParCptnOne)
+apply(force intro:par_cptn.intros elim:par_cptn.cases)
+done
+
+lemma droppar_cptn_is_par_cptn [rule_format]:
+  "\<forall>j<length c. c \<in> par_cptn \<longrightarrow> drop j c \<in> par_cptn"
+apply(induct "c")
+ apply(force elim: par_cptn.cases)
+apply clarify
+apply(case_tac j,simp+) 
+apply(erule par_cptn.cases)
+  apply simp
+ apply force
+apply force
+done
+
+lemma tl_of_cptn_is_cptn: "\<lbrakk>x # xs \<in> cptn; xs \<noteq> []\<rbrakk> \<Longrightarrow> xs  \<in> cptn"
+apply(subgoal_tac "1 < length (x # xs)") 
+ apply(drule dropcptn_is_cptn,simp+)
+done
+
+lemma not_ctran_None [rule_format]: 
+  "\<forall>s. (None, s)#xs \<in> cptn \<longrightarrow> (\<forall>i<length xs. ((None, s)#xs)!i -e\<rightarrow> xs!i)"
+apply(induct xs,simp+)
+apply clarify
+apply(erule cptn.cases,simp)
+ apply simp
+ apply(case_tac i,simp)
+  apply(rule Env)
+ apply simp
+apply(force elim:ctran.cases)
+done
+
+lemma cptn_not_empty [simp]:"[] \<notin> cptn"
+apply(force elim:cptn.cases)
+done
+
+lemma etran_or_ctran [rule_format]: 
+  "\<forall>m i. x\<in>cptn \<longrightarrow> m \<le> length x 
+   \<longrightarrow> (\<forall>i. Suc i < m \<longrightarrow> \<not> x!i -c\<rightarrow> x!Suc i) \<longrightarrow> Suc i < m 
+   \<longrightarrow> x!i -e\<rightarrow> x!Suc i"
+apply(induct x,simp)
+apply clarify
+apply(erule cptn.cases,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 "(\<forall>i. Suc i < nata \<longrightarrow> (((P, t) # xs) ! i, xs ! i) \<notin> ctran)")
+  apply force
+ apply clarify
+ apply(erule_tac x="Suc ia" in allE,simp)
+apply(erule_tac x="0" and P="\<lambda>j. ?H j \<longrightarrow> (?J j) \<notin> ctran" in allE,simp)
+done
+
+lemma etran_or_ctran2 [rule_format]: 
+  "\<forall>i. Suc i<length x \<longrightarrow> x\<in>cptn \<longrightarrow> (x!i -c\<rightarrow> x!Suc i \<longrightarrow> \<not> x!i -e\<rightarrow> x!Suc i)
+  \<or> (x!i -e\<rightarrow> x!Suc i \<longrightarrow> \<not> x!i -c\<rightarrow> x!Suc i)"
+apply(induct x)
+ apply simp
+apply clarify
+apply(erule cptn.cases,simp)
+ apply(case_tac i,simp+)
+apply(case_tac i,simp)
+ apply(force elim:etran.cases)
+apply simp
+done
+
+lemma etran_or_ctran2_disjI1: 
+  "\<lbrakk> x\<in>cptn; Suc i<length x; x!i -c\<rightarrow> x!Suc i\<rbrakk> \<Longrightarrow> \<not> x!i -e\<rightarrow> x!Suc i"
+by(drule etran_or_ctran2,simp_all)
+
+lemma etran_or_ctran2_disjI2: 
+  "\<lbrakk> x\<in>cptn; Suc i<length x; x!i -e\<rightarrow> x!Suc i\<rbrakk> \<Longrightarrow> \<not> x!i -c\<rightarrow> x!Suc i"
+by(drule etran_or_ctran2,simp_all)
+
+lemma not_ctran_None2 [rule_format]: 
+  "\<lbrakk> (None, s) # xs \<in>cptn; i<length xs\<rbrakk> \<Longrightarrow> \<not> ((None, s) # xs) ! i -c\<rightarrow> xs ! i"
+apply(frule not_ctran_None,simp)
+apply(case_tac i,simp)
+ apply(force elim:etranE)
+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) \<longrightarrow> (\<exists>m. P m \<and> (\<forall>i<m. \<not> P i))";
+apply(rule nat_less_induct)
+apply clarify
+apply(case_tac "\<forall>m. m<n \<longrightarrow> \<not> P m")
+apply auto
+done
+ 
+lemma stability [rule_format]: 
+  "\<forall>j k. x \<in> cptn \<longrightarrow> stable p rely \<longrightarrow> j\<le>k \<longrightarrow> k<length x \<longrightarrow> snd(x!j)\<in>p  \<longrightarrow>
+  (\<forall>i. (Suc i)<length x \<longrightarrow> 
+          (x!i -e\<rightarrow> x!(Suc i)) \<longrightarrow> (snd(x!i), snd(x!(Suc i))) \<in> rely) \<longrightarrow> 
+  (\<forall>i. j\<le>i \<and> i<k \<longrightarrow> x!i -e\<rightarrow> x!Suc i) \<longrightarrow> snd(x!k)\<in>p \<and> fst(x!j)=fst(x!k)"
+apply(induct x)
+ apply clarify
+ apply(force elim:cptn.cases)
+apply clarify
+apply(erule cptn.cases,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="\<lambda>j. (0\<le>j) \<longrightarrow> (?J j)" in allE,simp)
+  apply(subgoal_tac "t\<in>p")
+   apply(subgoal_tac "(\<forall>i. i < length xs \<longrightarrow> ((P, t) # xs) ! i -e\<rightarrow> xs ! i \<longrightarrow> (snd (((P, t) # xs) ! i), snd (xs ! i)) \<in> rely)")
+    apply clarify
+    apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j)\<in>etran" in allE,simp)
+   apply clarify
+   apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j) \<longrightarrow> (?T j)\<in>rely" in allE,simp)
+  apply(erule_tac x=0 and P="\<lambda>j. (?H j) \<longrightarrow> (?J j)\<in>etran \<longrightarrow> ?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="\<lambda>j. (?s\<le>j) \<longrightarrow> (?J j)" in allE,simp)
+ apply(subgoal_tac "(\<forall>i. i < length xs \<longrightarrow> ((P, t) # xs) ! i -e\<rightarrow> xs ! i \<longrightarrow> (snd (((P, t) # xs) ! i), snd (xs ! i)) \<in> rely)")
+  apply clarify
+  apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j)\<in>etran" in allE,simp)
+ apply clarify
+ apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j) \<longrightarrow> (?T j)\<in>rely" in allE,simp)
+apply(case_tac k,simp,simp)
+apply(case_tac j)
+ apply(erule_tac x=0 and P="\<lambda>j. (?H j) \<longrightarrow> (?J j)\<in>etran" in allE,simp)
+ apply(erule etran.cases,simp)
+apply(erule_tac x="nata" in allE)
+apply(erule_tac x="nat" and P="\<lambda>j. (?s\<le>j) \<longrightarrow> (?J j)" in allE,simp)
+apply(subgoal_tac "(\<forall>i. i < length xs \<longrightarrow> ((Q, t) # xs) ! i -e\<rightarrow> xs ! i \<longrightarrow> (snd (((Q, t) # xs) ! i), snd (xs ! i)) \<in> rely)")
+ apply clarify
+ apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j)\<in>etran" in allE,simp)
+apply clarify
+apply(erule_tac x="Suc i" and P="\<lambda>j. (?H j) \<longrightarrow> (?J j) \<longrightarrow> (?T j)\<in>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]: 
+  "\<forall>s i. x \<in> cptn \<longrightarrow> x ! 0 = (Some (Basic f), s) \<longrightarrow> 
+  Suc i<length x \<longrightarrow> x!i -c\<rightarrow> x!Suc i \<longrightarrow> 
+  (\<forall>j. Suc j<length x \<longrightarrow> i\<noteq>j \<longrightarrow> x!j -e\<rightarrow> x!Suc j)"
+apply(induct x,simp)
+apply simp
+apply clarify
+apply(erule cptn.cases,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.cases,simp_all)
+ apply(force elim: not_ctran_None)
+apply(ind_cases "((Some (Basic f), sa), Q, t) \<in> ctran" for sa Q t)
+apply simp
+apply(drule_tac i=nat in not_ctran_None,simp)
+apply(erule etranE,simp)
+done
+
+lemma exists_ctran_Basic_None [rule_format]: 
+  "\<forall>s i. x \<in> cptn \<longrightarrow> x ! 0 = (Some (Basic f), s) 
+  \<longrightarrow> i<length x \<longrightarrow> fst(x!i)=None \<longrightarrow> (\<exists>j<i. x!j -c\<rightarrow> x!Suc j)"
+apply(induct x,simp)
+apply simp
+apply clarify
+apply(erule cptn.cases,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: 
+  " \<lbrakk>pre \<subseteq> {s. f s \<in> post}; {(s, t). s \<in> pre \<and> t = f s} \<subseteq> guar; 
+  stable pre rely; stable post rely\<rbrakk>
+  \<Longrightarrow> \<Turnstile> 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 "\<forall>j. ?H j")
+ apply(force elim:ctran.cases)
+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(erule_tac x=i in allE,simp)
+ apply(erule_tac i=j and f=f in unique_ctran_Basic,simp_all)
+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.cases,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]: 
+  "\<forall>s i. x \<in> cptn \<longrightarrow> x ! 0 = (Some (Await b c), s) \<longrightarrow> 
+  Suc i<length x \<longrightarrow> x!i -c\<rightarrow> x!Suc i \<longrightarrow> 
+  (\<forall>j. Suc j<length x \<longrightarrow> i\<noteq>j \<longrightarrow> x!j -e\<rightarrow> x!Suc j)"
+apply(induct x,simp+)
+apply clarify
+apply(erule cptn.cases,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.cases,simp_all)
+ apply(force elim: not_ctran_None)
+apply(ind_cases "((Some (Await b c), sa), Q, t) \<in> ctran" for sa Q t,simp)
+apply(drule_tac i=nat in not_ctran_None,simp)
+apply(erule etranE,simp)
+done
+
+lemma exists_ctran_Await_None [rule_format]: 
+  "\<forall>s i.  x \<in> cptn \<longrightarrow> x ! 0 = (Some (Await b c), s) 
+  \<longrightarrow> i<length x \<longrightarrow> fst(x!i)=None \<longrightarrow> (\<exists>j<i. x!j -c\<rightarrow> x!Suc j)"
+apply(induct x,simp+)
+apply clarify
+apply(erule cptn.cases,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*\<rightarrow> (R, t) \<Longrightarrow> \<exists>l \<in> cp P s. (last l)=(R, t)
+  \<and> (\<forall>i. Suc i<length l \<longrightarrow> l!i -c\<rightarrow> 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.cases)
+apply simp
+apply(erule CptnComp)
+apply clarify
+done
+ 
+lemma Await_sound: 
+  "\<lbrakk>stable pre rely; stable post rely;
+  \<forall>V. \<turnstile> P sat [pre \<inter> b \<inter> {s. s = V}, {(s, t). s = t}, 
+                 UNIV, {s. (V, s) \<in> guar} \<inter> post] \<and>
+  \<Turnstile> P sat [pre \<inter> b \<inter> {s. s = V}, {(s, t). s = t}, 
+                 UNIV, {s. (V, s) \<in> guar} \<inter> post] \<rbrakk>
+  \<Longrightarrow> \<Turnstile> 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\<in> 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.cases,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="\<lambda>i. ?H i \<longrightarrow> (?J i,?I i)\<in>ctran" in allE,simp)
+  apply(erule etranE,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(erule_tac x=i in allE,simp)
+ apply(erule_tac i=j in unique_ctran_Await,force,simp_all)
+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.cases,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="\<lambda>i. ?H i \<longrightarrow> (?J i,?I i)\<in>ctran" in allE,simp)
+ apply(erule etranE,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 Cond_sound: 
+  "\<lbrakk> stable pre rely; \<Turnstile> P1 sat [pre \<inter> b, rely, guar, post]; 
+  \<Turnstile> P2 sat [pre \<inter> - b, rely, guar, post]; \<forall>s. (s,s)\<in>guar\<rbrakk>
+  \<Longrightarrow> \<Turnstile> (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 "\<exists>i. Suc i<length x \<and> x!i -c\<rightarrow> 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 add:last_length)+)
+apply(erule exE)
+apply(drule_tac n=i and P="\<lambda>i. ?H i \<and> (?J i,?I i)\<in> 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.cases,simp_all)
+ apply(erule_tac x="sa" in allE)
+ apply(drule_tac c="drop (Suc m) x" in subsetD)
+  apply simp
+  apply clarify
+ apply simp
+ apply clarify
+ apply(case_tac "i\<le>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="\<lambda>j. ?H j \<longrightarrow> ?J j \<longrightarrow> (?I j)\<in>guar" in allE)
+ apply(subgoal_tac "(Suc m)+(i - Suc m) \<le> length x")
+  apply(subgoal_tac "(Suc m)+Suc (i - Suc m) \<le> length x")
+   apply(rotate_tac -2)
+   apply simp
+  apply arith
+ apply arith
+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 clarify
+apply simp
+apply clarify
+apply(case_tac "i\<le>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="\<lambda>j. ?H j \<longrightarrow> ?J j \<longrightarrow> (?I j)\<in>guar" in allE)
+apply(subgoal_tac "(Suc m)+(i - Suc m) \<le> length x")
+ apply(subgoal_tac "(Suc m)+Suc (i - Suc m) \<le> length x")
+  apply(rotate_tac -2)
+  apply simp
+ apply arith
+apply arith
+done
+
+subsubsection{* Soundness of the Sequential rule *}
+
+inductive_cases Seq_cases [elim!]: "(Some (Seq P Q), s) -c\<rightarrow> t"
+
+lemma last_lift_not_None: "fst ((lift Q) ((x#xs)!(length xs))) \<noteq> None"
+apply(subgoal_tac "length xs<length (x # xs)")
+ apply(drule_tac Q=Q in  lift_nth)
+ apply(erule ssubst)
+ apply (simp add:lift_def)
+ apply(case_tac "(x # xs) ! length xs",simp)
+apply simp
+done
+
+declare map_eq_Cons_conv [simp del] Cons_eq_map_conv [simp del]
+lemma Seq_sound1 [rule_format]: 
+  "x\<in> cptn_mod \<Longrightarrow> \<forall>s P. x !0=(Some (Seq P Q), s) \<longrightarrow> 
+  (\<forall>i<length x. fst(x!i)\<noteq>Some Q) \<longrightarrow> 
+  (\<exists>xs\<in> 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 "(\<forall>i < Suc (length xs). fst (((Some (Seq Pa Q), t) # xs) ! i) \<noteq> Some Q)")
+    apply clarify
+    apply(rule_tac x="(Some Pa, sa) #(Some Pa, t) # zs" in exI,simp)
+    apply(rule conjI,erule CptnEnv)
+    apply(simp (no_asm_use) add:lift_def)
+   apply clarify
+   apply(erule_tac x="Suc i" in allE, simp)
+  apply(ind_cases "((Some (Seq Pa Q), sa), None, t) \<in> ctran" for Pa sa t)
+ 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
+declare map_eq_Cons_conv [simp del] Cons_eq_map_conv [simp del]
+
+lemma Seq_sound2 [rule_format]: 
+  "x \<in> cptn \<Longrightarrow> \<forall>s P i. x!0=(Some (Seq P Q), s) \<longrightarrow> i<length x 
+  \<longrightarrow> fst(x!i)=Some Q \<longrightarrow>
+  (\<forall>j<i. fst(x!j)\<noteq>(Some Q)) \<longrightarrow>
+  (\<exists>xs ys. xs \<in> cp (Some P) s \<and> length xs=Suc i 
+   \<and> ys \<in> cp (Some Q) (snd(xs !i)) \<and> 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 "(\<forall>j < nat. fst (((Some (Seq Pa Q), t) # xs) ! j) \<noteq> 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 (no_asm_use) add:lift_def)
+ apply(rule_tac x=ys in exI,simp)
+apply(ind_cases "((Some (Seq Pa Q), sa), t) \<in> ctran" for Pa sa t)
+ 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 "(\<forall>j<Suc nata. fst (((Some (Seq P2 Q), ta) # xs) ! j) \<noteq> 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))) \<noteq> None"
+apply(simp only:last_length [THEN sym])
+apply(subgoal_tac "length xs<length (x # xs)")
+ apply(drule_tac Q=Q in  lift_nth)
+ apply(erule ssubst)
+ apply (simp add:lift_def)
+ apply(case_tac "(x # xs) ! length xs",simp)
+apply simp
+done
+*)
+
+lemma last_lift_not_None2: "fst ((lift Q) (last (x#xs))) \<noteq> None"
+apply(simp only:last_length [THEN sym])
+apply(subgoal_tac "length xs<length (x # xs)")
+ apply(drule_tac Q=Q in  lift_nth)
+ apply(erule ssubst)
+ apply (simp add:lift_def)
+ apply(case_tac "(x # xs) ! length xs",simp)
+apply simp
+done
+
+lemma Seq_sound: 
+  "\<lbrakk>\<Turnstile> P sat [pre, rely, guar, mid]; \<Turnstile> Q sat [mid, rely, guar, post]\<rbrakk>
+  \<Longrightarrow> \<Turnstile> Seq P Q sat [pre, rely, guar, post]"
+apply(unfold com_validity_def)
+apply clarify
+apply(case_tac "\<exists>i<length x. fst(x!i)=Some Q")
+ prefer 2
+ apply (simp add:cp_def cptn_iff_cptn_mod)
+ apply clarify
+ apply(frule_tac Seq_sound1,force)
+  apply force
+ apply clarify
+ apply(erule_tac x=s in allE,simp)
+ apply(drule_tac c=xs in subsetD,simp add:cp_def cptn_iff_cptn_mod)
+  apply(simp add:assum_def)
+  apply clarify
+  apply(erule_tac P="\<lambda>j. ?H j \<longrightarrow> ?J j \<longrightarrow> ?I j" in allE,erule impE, assumption)
+  apply(simp add:snd_lift)
+  apply(erule mp)
+  apply(force elim:etranE intro:Env simp add:lift_def)
+ apply(simp add:comm_def)
+ apply(rule conjI)
+  apply clarify
+  apply(erule_tac P="\<lambda>j. ?H j \<longrightarrow> ?J j \<longrightarrow> ?I j" in allE,erule impE, assumption)
+  apply(simp add:snd_lift)
+  apply(erule mp)
+  apply(case_tac "(xs!i)")
+  apply(case_tac "(xs! Suc i)")
+  apply(case_tac "fst(xs!i)")
+   apply(erule_tac x=i in allE, simp add:lift_def)
+  apply(case_tac "fst(xs!Suc i)")
+   apply(force simp add:lift_def)
+  apply(force simp add:lift_def)
+ apply clarify
+ apply(case_tac xs,simp add:cp_def)
+ apply clarify
+ apply (simp del:map.simps)
+ apply(subgoal_tac "(map (lift Q) ((a, b) # list))\<noteq>[]")
+  apply(drule last_conv_nth)
+  apply (simp del:map.simps)
+  apply(simp only:last_lift_not_None)
+ apply simp
+--{* @{text "\<exists>i<length x. fst (x ! i) = Some Q"} *}
+apply(erule exE)
+apply(drule_tac n=i and P="\<lambda>i. i < length x \<and> fst (x ! i) = Some Q" in Ex_first_occurrence)
+apply clarify
+apply (simp add:cp_def)
+ apply clarify
+ apply(frule_tac i=m in Seq_sound2,force)
+  apply simp+
+apply clarify
+apply(simp add:comm_def)
+apply(erule_tac x=s in allE)
+apply(drule_tac c=xs in subsetD,simp)
+ apply(case_tac "xs=[]",simp)
+ apply(simp add:cp_def assum_def nth_append)
+ apply clarify
+ apply(erule_tac x=i in allE)
+  back 
+ apply(simp add:snd_lift)
+ apply(erule mp)
+ apply(force elim:etranE intro:Env simp add:lift_def)
+apply simp
+apply clarify
+apply(erule_tac x="snd(xs!m)" in allE)
+apply(drule_tac c=ys in subsetD,simp add:cp_def assum_def)
+ apply(case_tac "xs\<noteq>[]")
+ apply(drule last_conv_nth,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.cases intro:Env simp add:lift_def)
+ apply simp 
+apply simp
+apply clarify
+apply(rule conjI,clarify)
+ apply(case_tac "i<m",simp add:nth_append)
+  apply(simp add:snd_lift)
+  apply(erule allE, erule impE, assumption, erule mp)
+  apply(case_tac "(xs ! i)")
+  apply(case_tac "(xs ! Suc i)")   
+  apply(case_tac "fst(xs ! i)",force simp add:lift_def)   
+  apply(case_tac "fst(xs ! Suc i)")
+   apply (force simp add:lift_def)
+  apply (force simp add:lift_def)
+ apply(erule_tac x="i-m" in allE) 
+ back
+ back
+ apply(subgoal_tac "Suc (i - m) < length ys",simp)
+  prefer 2
+  apply arith
+ apply(simp add:nth_append snd_lift)
+ apply(rule conjI,clarify)
+  apply(subgoal_tac "i=m")
+   prefer 2
+   apply arith
+  apply clarify
+  apply(simp add:cp_def)
+  apply(rule tl_zero)
+    apply(erule mp)
+    apply(case_tac "lift Q (xs!m)",simp add:snd_lift)
+    apply(case_tac "xs!m",case_tac "fst(xs!m)",simp add:lift_def snd_lift)
+     apply(case_tac ys,simp+)
+    apply(simp add:lift_def)
+   apply simp 
+  apply force
+ apply clarify
+ apply(rule tl_zero)
+   apply(rule tl_zero)
+     apply (subgoal_tac "i-m=Suc(i-Suc m)")
+      apply simp
+      apply(erule mp)
+      apply(case_tac ys,simp+)
+   apply force
+  apply arith
+ apply force
+apply clarify
+apply(case_tac "(map (lift Q) xs @ tl ys)\<noteq>[]")
+ apply(drule last_conv_nth)
+ apply(simp add: snd_lift nth_append)
+ apply(rule conjI,clarify)
+  apply(case_tac ys,simp+)
+ apply clarify
+ apply(case_tac ys,simp+)
+done
+
+subsubsection{* Soundness of the While rule *}
+
+lemma last_append[rule_format]:
+  "\<forall>xs. ys\<noteq>[] \<longrightarrow> ((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: 
+  "\<lbrakk> \<Turnstile> P sat [pre \<inter> b, rely, guar, pre]; 
+  (Some P, s) # xs \<in> cptn_mod; fst (last ((Some P, s) # xs)) = None; s \<in> b;
+  (Some (While b P), s) # (Some (Seq P (While b P)), s) # 
+   map (lift (While b P)) xs @ ys \<in> assum (pre, rely)\<rbrakk>
+  \<Longrightarrow> (Some (While b P), snd (last ((Some P, s) # xs))) # ys \<in> 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 etranE,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 While_sound_aux [rule_format]: 
+  "\<lbrakk> pre \<inter> - b \<subseteq> post; \<Turnstile> P sat [pre \<inter> b, rely, guar, pre]; \<forall>s. (s, s) \<in> guar;
+   stable pre rely;  stable post rely; x \<in> cptn_mod \<rbrakk> 
+  \<Longrightarrow>  \<forall>s xs. x=(Some(While b P),s)#xs \<longrightarrow> x\<in>assum(pre, rely) \<longrightarrow> x \<in> 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) \<in> ctran" for s t)
+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 "\<forall>i. Suc i < length ((None, t) # xs) \<longrightarrow> (((None, t) # xs) ! i, ((None, t) # xs) ! Suc i)\<in> 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 \<longrightarrow> ?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\<rightarrow> t" for ba 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 etranE,simp add:lift_def)
+   apply(rule Env)
+  apply(erule etranE,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))\<noteq>[]")
+ apply(drule last_conv_nth)
+ apply (simp del:map.simps)
+ apply(simp only:last_lift_not_None)
+apply simp
+--{* WhileMore *}
+apply(thin_tac "P = While b P \<longrightarrow> ?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 \<in> 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<length xs")
+  apply(simp add:nth_append del:map.simps last.simps)
+  apply(case_tac "fst(((Some P, sa) # xs) ! i)")
+   apply(case_tac "((Some P, sa) # xs) ! i")
+   apply (simp add:lift_def del:last.simps)
+   apply(ind_cases "(Some (While b P), ba) -c\<rightarrow> t" for ba t)
+    apply simp
+   apply simp
+  apply(simp add:snd_lift del:map.simps last.simps)
+  apply(thin_tac " \<forall>i. i < length ys \<longrightarrow> ?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 etranE,simp add:lift_def)
+    apply(rule Env)
+   apply(erule etranE,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 \<ge> length xs"} *}
+apply(subgoal_tac "i-length xs <length ys") 
+ prefer 2
+ apply arith
+apply(erule_tac x="i-length xs" in allE,clarify)
+apply(case_tac "i=length xs")
+ apply (simp add:nth_append snd_lift del:map.simps last.simps)
+ apply(simp add:last_length del:last.simps)
+ apply(erule mp)
+ apply(case_tac "last((Some P, sa) # xs)")
+ apply(simp add:lift_def del:last.simps)
+--{* @{text "i>length xs"} *} 
+apply(case_tac "i-length xs")
+ apply arith
+apply(simp add:nth_append del:map.simps last.simps)
+apply(rotate_tac -3)
+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))\<noteq>[]")
+  apply(drule last_conv_nth)
+  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)\<noteq>[]")
+ apply(drule last_conv_nth)
+ 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)\<noteq>[]")
+  apply(drule last_conv_nth)
+  apply (simp del:map.simps last.simps)
+ apply simp
+apply simp
+done
+
+lemma While_sound: 
+  "\<lbrakk>stable pre rely; pre \<inter> - b \<subseteq> post; stable post rely;
+    \<Turnstile> P sat [pre \<inter> b, rely, guar, pre]; \<forall>s. (s,s)\<in>guar\<rbrakk>
+  \<Longrightarrow> \<Turnstile> 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: 
+  "\<lbrakk>pre \<subseteq> pre'; rely \<subseteq> rely'; guar' \<subseteq> guar; post' \<subseteq> post; 
+  \<Turnstile> P sat [pre', rely', guar', post']\<rbrakk>
+  \<Longrightarrow> \<Turnstile> 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: 
+  "\<turnstile> P sat [pre, rely, guar, post] \<Longrightarrow> \<Turnstile> 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 \<Rightarrow> 'a par_com"
+  "ParallelCom Ps \<equiv> map (Some \<circ> fst) Ps" 
+
+lemma two: 
+  "\<lbrakk> \<forall>i<length xs. rely \<union> (\<Union>j\<in>{j. j < length xs \<and> j \<noteq> i}. Guar (xs ! j)) 
+     \<subseteq> Rely (xs ! i);
+   pre \<subseteq> (\<Inter>i\<in>{i. i < length xs}. Pre (xs ! i));
+   \<forall>i<length xs. 
+   \<Turnstile> Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)];
+   length xs=length clist; x \<in> par_cp (ParallelCom xs) s; x\<in>par_assum(pre, rely);
+  \<forall>i<length clist. clist!i\<in>cp (Some(Com(xs!i))) s; x \<propto> clist \<rbrakk>
+  \<Longrightarrow> \<forall>j i. i<length clist \<and> Suc j<length x \<longrightarrow> (clist!i!j) -c\<rightarrow> (clist!i!Suc j) 
+  \<longrightarrow> (snd(clist!i!j), snd(clist!i!Suc j)) \<in> Guar(xs!i)"
+apply(unfold par_cp_def)
+apply (rule ccontr) 
+--{* By contradiction: *}
+apply (simp del: Un_subset_iff)
+apply(erule exE)
+--{* the first c-tran that does not satisfy the guarantee-condition is from @{text "\<sigma>_i"} at step @{text "m"}. *}
+apply(drule_tac n=j and P="\<lambda>j. \<exists>i. ?H i j" in Ex_first_occurrence)
+apply(erule exE)
+apply clarify
+--{* @{text "\<sigma>_i \<in> A(pre, rely_1)"} *}
+apply(subgoal_tac "take (Suc (Suc m)) (clist!i) \<in> assum(Pre(xs!i), Rely(xs!i))")
+--{* but this contradicts @{text "\<Turnstile> \<sigma>_i sat [pre_i,rely_i,guar_i,post_i]"} *}
+ apply(erule_tac x=i and P="\<lambda>i. ?H i \<longrightarrow> \<Turnstile> (?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 (blast intro: takecptn_is_cptn) 
+ apply simp
+ apply clarify
+ apply(erule_tac x=m and P="\<lambda>j. ?I j \<and> ?J j \<longrightarrow> ?H j" in allE)
+ apply (simp add:conjoin_def same_length_def)
+apply(simp add:assum_def del: Un_subset_iff)
+apply(rule conjI)
+ apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow>  ?I j \<in>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="\<lambda>j. ?H j \<longrightarrow> ?M \<union> UNION (?S j) (?T j) \<subseteq>  (?L j)" in allE)
+apply(simp del: Un_subset_iff)
+apply(erule subsetD)
+apply simp
+apply(simp add:conjoin_def compat_label_def)
+apply clarify
+apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (?P j) \<or> ?Q j" in allE,simp)
+--{* each etran in @{text "\<sigma>_1[0\<dots>m]"} corresponds to  *}
+apply(erule disjE)
+--{* a c-tran in some @{text "\<sigma>_{ib}"}  *}
+ apply clarify
+ apply(case_tac "i=ib",simp)
+  apply(erule etranE,simp)
+ apply(erule_tac x="ib" and P="\<lambda>i. ?H i \<longrightarrow> (?I i) \<or> (?J i)" in allE)
+ apply (erule etranE)
+ apply(case_tac "ia=m",simp)
+ apply simp
+ apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (\<forall> i. ?P i j)" in allE)
+ apply(subgoal_tac "ia<m",simp)
+  prefer 2
+  apply arith
+ apply(erule_tac x=ib and P="\<lambda>j. (?I j, ?H j)\<in> ctran \<longrightarrow> (?P i j)" in allE,simp)
+ apply(simp add:same_state_def)
+ apply(erule_tac x=i and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in all_dupE)
+ apply(erule_tac x=ib and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in allE,simp)
+--{* or an e-tran in @{text "\<sigma>"}, 
+therefore it satisfies @{text "rely \<or> guar_{ib}"} *}
+apply (force simp add:par_assum_def same_state_def)
+done
+
+
+lemma three [rule_format]: 
+  "\<lbrakk> xs\<noteq>[]; \<forall>i<length xs. rely \<union> (\<Union>j\<in>{j. j < length xs \<and> j \<noteq> i}. Guar (xs ! j)) 
+   \<subseteq> Rely (xs ! i);
+   pre \<subseteq> (\<Inter>i\<in>{i. i < length xs}. Pre (xs ! i));
+   \<forall>i<length xs. 
+    \<Turnstile> Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)];
+   length xs=length clist; x \<in> par_cp (ParallelCom xs) s; x \<in> par_assum(pre, rely);
+  \<forall>i<length clist. clist!i\<in>cp (Some(Com(xs!i))) s; x \<propto> clist \<rbrakk>
+  \<Longrightarrow> \<forall>j i. i<length clist \<and> Suc j<length x \<longrightarrow> (clist!i!j) -e\<rightarrow> (clist!i!Suc j) 
+  \<longrightarrow> (snd(clist!i!j), snd(clist!i!Suc j)) \<in> rely \<union> (\<Union>j\<in>{j. j < length xs \<and> j \<noteq> 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="\<lambda>j. ?H j \<longrightarrow> (?J j \<and> (\<exists>i. ?P i j)) \<or> ?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 etranE,simp)
+apply(erule_tac x="ia" and P="\<lambda>i. ?H i \<longrightarrow> (?I i) \<or> (?J i)" in allE,simp)
+apply(erule_tac x=j and P="\<lambda>j. \<forall>i. ?S j i \<longrightarrow> (?I j i, ?H j i)\<in> ctran \<longrightarrow> (?P i j)" in allE)
+apply(erule_tac x=ia and P="\<lambda>j. ?S j \<longrightarrow> (?I j, ?H j)\<in> ctran \<longrightarrow> (?P j)" in allE)
+apply(simp add:same_state_def)
+apply(erule_tac x=i and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in all_dupE)
+apply(erule_tac x=ia and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in allE,simp)
+done
+
+lemma four: 
+  "\<lbrakk>xs\<noteq>[]; \<forall>i < length xs. rely \<union> (\<Union>j\<in>{j. j < length xs \<and> j \<noteq> i}. Guar (xs ! j)) 
+    \<subseteq> Rely (xs ! i);
+   (\<Union>j\<in>{j. j < length xs}. Guar (xs ! j)) \<subseteq> guar; 
+   pre \<subseteq> (\<Inter>i\<in>{i. i < length xs}. Pre (xs ! i));
+   \<forall>i < length xs.
+    \<Turnstile> Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)];
+   x \<in> par_cp (ParallelCom xs) s; x \<in> par_assum (pre, rely); Suc i < length x; 
+   x ! i -pc\<rightarrow> x ! Suc i\<rbrakk>
+  \<Longrightarrow> (snd (x ! i), snd (x ! Suc i)) \<in> guar"
+apply(simp add: ParallelCom_def del: Un_subset_iff)
+apply(subgoal_tac "(map (Some \<circ> fst) xs)\<noteq>[]")
+ prefer 2
+ apply simp
+apply(frule rev_subsetD)
+ apply(erule one [THEN equalityD1])
+apply(erule subsetD)
+apply (simp del: Un_subset_iff)
+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="\<lambda>j. ?H j \<longrightarrow> fst(?I j)=(?J j)" in all_dupE)
+apply(erule_tac x="Suc i" and P="\<lambda>j. ?H j \<longrightarrow> fst(?I j)=(?J j)" in allE)
+apply(erule par_ctranE,simp)
+apply(erule_tac x=i and P="\<lambda>j. \<forall>i. ?S j i \<longrightarrow> (?I j i, ?H j i)\<in> ctran \<longrightarrow> (?P i j)" in allE)
+apply(erule_tac x=ia and P="\<lambda>j. ?S j \<longrightarrow> (?I j, ?H j)\<in> ctran \<longrightarrow> (?P j)" in allE)
+apply(rule_tac x=ia in exI)
+apply(simp add:same_state_def)
+apply(erule_tac x=ia and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in all_dupE,simp)
+apply(erule_tac x=ia and P="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in allE,simp)
+apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in all_dupE)
+apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in all_dupE,simp)
+apply(erule_tac x="Suc i" and P="\<lambda>j. ?H j \<longrightarrow> (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]:"[] \<notin> par_cptn"
+apply(force elim:par_cptn.cases)
+done
+
+lemma five: 
+  "\<lbrakk>xs\<noteq>[]; \<forall>i<length xs. rely \<union> (\<Union>j\<in>{j. j < length xs \<and> j \<noteq> i}. Guar (xs ! j))
+   \<subseteq> Rely (xs ! i);
+   pre \<subseteq> (\<Inter>i\<in>{i. i < length xs}. Pre (xs ! i)); 
+   (\<Inter>i\<in>{i. i < length xs}. Post (xs ! i)) \<subseteq> post;
+   \<forall>i < length xs.
+    \<Turnstile> Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)];
+    x \<in> par_cp (ParallelCom xs) s; x \<in> par_assum (pre, rely); 
+   All_None (fst (last x)) \<rbrakk> \<Longrightarrow> snd (last x) \<in> post"
+apply(simp add: ParallelCom_def del: Un_subset_iff)
+apply(subgoal_tac "(map (Some \<circ> fst) xs)\<noteq>[]")
+ prefer 2
+ apply simp
+apply(frule rev_subsetD)
+ apply(erule one [THEN equalityD1])
+apply(erule subsetD)
+apply(simp del: Un_subset_iff)
+apply clarify
+apply(subgoal_tac "\<forall>i<length clist. clist!i\<in>assum(Pre(xs!i), Rely(xs!i))")
+ apply(erule_tac x=i and P="\<lambda>i. ?H i \<longrightarrow> \<Turnstile> (?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="\<lambda>j. ?H j \<longrightarrow> (?I j) \<in> 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="\<lambda>j. ?H j \<longrightarrow> fst(?I j)=(?J j)" in allE)
+ apply (simp add:All_None_def same_length_def)
+ apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> length(?J j)=(?K j)" in allE)
+ apply(subgoal_tac "length x - 1 < length x",simp)
+  apply(case_tac "x\<noteq>[]")
+   apply(simp add: last_conv_nth)
+   apply(erule_tac x="clist!i" in ballE)
+    apply(simp add:same_state_def)
+    apply(subgoal_tac "clist!i\<noteq>[]")
+     apply(simp add: last_conv_nth)
+    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="\<lambda>j. (?T j) \<longrightarrow> (\<forall>i. (?H j i) \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in allE,simp)
+ apply(erule_tac x=0 and P="\<lambda>j. ?H j \<longrightarrow> (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(rule subsetD, erule mp, assumption)
+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]: 
+  "\<forall>i s. x \<in> par_cp (ParallelCom []) s \<longrightarrow> 
+  Suc i < length x \<longrightarrow> (x ! i, x ! Suc i) \<notin> 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_ctranE,simp)
+apply(simp add:par_cp_def ParallelCom_def)
+apply clarify
+apply(erule par_cptn.cases,simp)
+ apply simp
+apply(erule par_ctranE)
+back
+apply simp
+done
+
+theorem par_rgsound: 
+  "\<turnstile> c SAT [pre, rely, guar, post] \<Longrightarrow> 
+   \<Turnstile> (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\<noteq>[]")
+ 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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/RG_Syntax.thy	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,95 @@
+header {* \section{Concrete Syntax} *}
+
+theory RG_Syntax
+imports RG_Hoare Quote_Antiquote
+begin
+
+syntax
+  "_Assign"    :: "idt \<Rightarrow> 'b \<Rightarrow> 'a com"                     ("(\<acute>_ :=/ _)" [70, 65] 61)
+  "_skip"      :: "'a com"                                  ("SKIP")
+  "_Seq"       :: "'a com \<Rightarrow> 'a com \<Rightarrow> 'a com"              ("(_;;/ _)" [60,61] 60)
+  "_Cond"      :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com \<Rightarrow> 'a com"   ("(0IF _/ THEN _/ ELSE _/FI)" [0, 0, 0] 61)
+  "_Cond2"     :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com"             ("(0IF _ THEN _ FI)" [0,0] 56)
+  "_While"     :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com"             ("(0WHILE _ /DO _ /OD)"  [0, 0] 61)
+  "_Await"     :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com"             ("(0AWAIT _ /THEN /_ /END)"  [0,0] 61)
+  "_Atom"      :: "'a com \<Rightarrow> 'a com"                        ("(\<langle>_\<rangle>)" 61)
+  "_Wait"      :: "'a bexp \<Rightarrow> 'a com"                       ("(0WAIT _ END)" 61)
+
+translations
+  "\<acute>\<spacespace>x := a" \<rightharpoonup> "Basic \<guillemotleft>\<acute>\<spacespace>(_update_name x (\<lambda>_. a))\<guillemotright>"
+  "SKIP" \<rightleftharpoons> "Basic id"
+  "c1;; c2" \<rightleftharpoons> "Seq c1 c2"
+  "IF b THEN c1 ELSE c2 FI" \<rightharpoonup> "Cond .{b}. c1 c2"
+  "IF b THEN c FI" \<rightleftharpoons> "IF b THEN c ELSE SKIP FI"
+  "WHILE b DO c OD" \<rightharpoonup> "While .{b}. c"
+  "AWAIT b THEN c END" \<rightleftharpoons> "Await .{b}. c"
+  "\<langle>c\<rangle>" \<rightleftharpoons> "AWAIT True THEN c END"
+  "WAIT b END" \<rightleftharpoons> "AWAIT b THEN SKIP END"
+
+nonterminals
+  prgs
+
+syntax
+  "_PAR"        :: "prgs \<Rightarrow> 'a"              ("COBEGIN//_//COEND" 60)
+  "_prg"        :: "'a \<Rightarrow> prgs"              ("_" 57)
+  "_prgs"       :: "['a, prgs] \<Rightarrow> prgs"      ("_//\<parallel>//_" [60,57] 57)
+
+translations
+  "_prg a" \<rightharpoonup> "[a]"
+  "_prgs a ps" \<rightharpoonup> "a # ps"
+  "_PAR ps" \<rightharpoonup> "ps"
+
+syntax
+  "_prg_scheme" :: "['a, 'a, 'a, 'a] \<Rightarrow> prgs"  ("SCHEME [_ \<le> _ < _] _" [0,0,0,60] 57)
+
+translations
+  "_prg_scheme j i k c" \<rightleftharpoons> "(map (\<lambda>i. c) [j..<k])"
+
+text {* Translations for variables before and after a transition: *}
+
+syntax 
+  "_before" :: "id \<Rightarrow> 'a" ("\<ordmasculine>_")
+  "_after"  :: "id \<Rightarrow> 'a" ("\<ordfeminine>_")
+ 
+translations
+  "\<ordmasculine>x" \<rightleftharpoons> "x \<acute>fst"
+  "\<ordfeminine>x" \<rightleftharpoons> "x \<acute>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 Record.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 K_tr' (Abs (_,_,t)) = if null (loose_bnos t) then t else raise Match
+      | K_tr' (Abs (_,_,Abs (_,_,t)$Bound 0)) = if null (loose_bnos t) then t else raise Match
+      | K_tr' _ = raise Match;
+
+    fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) =
+          quote_tr' (Syntax.const "_Assign" $ update_name_tr' f)
+            (Abs (x, dummyT, K_tr' k) :: 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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/RG_Tran.thy	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,1075 @@
+header {* \section{Operational Semantics} *}
+
+theory RG_Tran
+imports RG_Com
+begin
+
+subsection {* Semantics of Component Programs *}
+
+subsubsection {* Environment transitions *}
+
+types 'a conf = "(('a com) option) \<times> 'a"
+
+inductive_set
+  etran :: "('a conf \<times> 'a conf) set" 
+  and etran' :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool"  ("_ -e\<rightarrow> _" [81,81] 80)
+where
+  "P -e\<rightarrow> Q \<equiv> (P,Q) \<in> etran"
+| Env: "(P, s) -e\<rightarrow> (P, t)"
+
+lemma etranE: "c -e\<rightarrow> c' \<Longrightarrow> (\<And>P s t. c = (P, s) \<Longrightarrow> c' = (P, t) \<Longrightarrow> Q) \<Longrightarrow> Q"
+  by (induct c, induct c', erule etran.cases, blast)
+
+subsubsection {* Component transitions *}
+
+inductive_set
+  ctran :: "('a conf \<times> 'a conf) set"
+  and ctran' :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool"   ("_ -c\<rightarrow> _" [81,81] 80)
+  and ctrans :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool"   ("_ -c*\<rightarrow> _" [81,81] 80)
+where
+  "P -c\<rightarrow> Q \<equiv> (P,Q) \<in> ctran"
+| "P -c*\<rightarrow> Q \<equiv> (P,Q) \<in> ctran^*"
+
+| Basic:  "(Some(Basic f), s) -c\<rightarrow> (None, f s)"
+
+| Seq1:   "(Some P0, s) -c\<rightarrow> (None, t) \<Longrightarrow> (Some(Seq P0 P1), s) -c\<rightarrow> (Some P1, t)"
+
+| Seq2:   "(Some P0, s) -c\<rightarrow> (Some P2, t) \<Longrightarrow> (Some(Seq P0 P1), s) -c\<rightarrow> (Some(Seq P2 P1), t)"
+
+| CondT: "s\<in>b  \<Longrightarrow> (Some(Cond b P1 P2), s) -c\<rightarrow> (Some P1, s)"
+| CondF: "s\<notin>b \<Longrightarrow> (Some(Cond b P1 P2), s) -c\<rightarrow> (Some P2, s)"
+
+| WhileF: "s\<notin>b \<Longrightarrow> (Some(While b P), s) -c\<rightarrow> (None, s)"
+| WhileT: "s\<in>b  \<Longrightarrow> (Some(While b P), s) -c\<rightarrow> (Some(Seq P (While b P)), s)"
+
+| Await:  "\<lbrakk>s\<in>b; (Some P, s) -c*\<rightarrow> (None, t)\<rbrakk> \<Longrightarrow> (Some(Await b P), s) -c\<rightarrow> (None, t)" 
+
+monos "rtrancl_mono"
+
+subsection {* Semantics of Parallel Programs *}
+
+types 'a par_conf = "('a par_com) \<times> 'a"
+
+inductive_set
+  par_etran :: "('a par_conf \<times> 'a par_conf) set"
+  and par_etran' :: "['a par_conf,'a par_conf] \<Rightarrow> bool" ("_ -pe\<rightarrow> _" [81,81] 80)
+where
+  "P -pe\<rightarrow> Q \<equiv> (P,Q) \<in> par_etran"
+| ParEnv:  "(Ps, s) -pe\<rightarrow> (Ps, t)"
+
+inductive_set
+  par_ctran :: "('a par_conf \<times> 'a par_conf) set"
+  and par_ctran' :: "['a par_conf,'a par_conf] \<Rightarrow> bool" ("_ -pc\<rightarrow> _" [81,81] 80)
+where
+  "P -pc\<rightarrow> Q \<equiv> (P,Q) \<in> par_ctran"
+| ParComp: "\<lbrakk>i<length Ps; (Ps!i, s) -c\<rightarrow> (r, t)\<rbrakk> \<Longrightarrow> (Ps, s) -pc\<rightarrow> (Ps[i:=r], t)"
+
+lemma par_ctranE: "c -pc\<rightarrow> c' \<Longrightarrow>
+  (\<And>i Ps s r t. c = (Ps, s) \<Longrightarrow> c' = (Ps[i := r], t) \<Longrightarrow> i < length Ps \<Longrightarrow>
+     (Ps ! i, s) -c\<rightarrow> (r, t) \<Longrightarrow> P) \<Longrightarrow> P"
+  by (induct c, induct c', erule par_ctran.cases, blast)
+
+subsection {* Computations *}
+
+subsubsection {* Sequential computations *}
+
+types 'a confs = "('a conf) list"
+
+inductive_set cptn :: "('a confs) set"
+where
+  CptnOne: "[(P,s)] \<in> cptn"
+| CptnEnv: "(P, t)#xs \<in> cptn \<Longrightarrow> (P,s)#(P,t)#xs \<in> cptn"
+| CptnComp: "\<lbrakk>(P,s) -c\<rightarrow> (Q,t); (Q, t)#xs \<in> cptn \<rbrakk> \<Longrightarrow> (P,s)#(Q,t)#xs \<in> cptn"
+
+constdefs
+  cp :: "('a com) option \<Rightarrow> 'a \<Rightarrow> ('a confs) set"
+  "cp P s \<equiv> {l. l!0=(P,s) \<and> l \<in> cptn}"  
+
+subsubsection {* Parallel computations *}
+
+types  'a par_confs = "('a par_conf) list"
+
+inductive_set par_cptn :: "('a par_confs) set"
+where
+  ParCptnOne: "[(P,s)] \<in> par_cptn"
+| ParCptnEnv: "(P,t)#xs \<in> par_cptn \<Longrightarrow> (P,s)#(P,t)#xs \<in> par_cptn"
+| ParCptnComp: "\<lbrakk> (P,s) -pc\<rightarrow> (Q,t); (Q,t)#xs \<in> par_cptn \<rbrakk> \<Longrightarrow> (P,s)#(Q,t)#xs \<in> par_cptn"
+
+constdefs
+  par_cp :: "'a par_com \<Rightarrow> 'a \<Rightarrow> ('a par_confs) set"
+  "par_cp P s \<equiv> {l. l!0=(P,s) \<and> l \<in> par_cptn}"  
+
+subsection{* Modular Definition of Computation *}
+
+constdefs 
+  lift :: "'a com \<Rightarrow> 'a conf \<Rightarrow> 'a conf"
+  "lift Q \<equiv> \<lambda>(P, s). (if P=None then (Some Q,s) else (Some(Seq (the P) Q), s))"
+
+inductive_set cptn_mod :: "('a confs) set"
+where
+  CptnModOne: "[(P, s)] \<in> cptn_mod"
+| CptnModEnv: "(P, t)#xs \<in> cptn_mod \<Longrightarrow> (P, s)#(P, t)#xs \<in> cptn_mod"
+| CptnModNone: "\<lbrakk>(Some P, s) -c\<rightarrow> (None, t); (None, t)#xs \<in> cptn_mod \<rbrakk> \<Longrightarrow> (Some P,s)#(None, t)#xs \<in>cptn_mod"
+| CptnModCondT: "\<lbrakk>(Some P0, s)#ys \<in> cptn_mod; s \<in> b \<rbrakk> \<Longrightarrow> (Some(Cond b P0 P1), s)#(Some P0, s)#ys \<in> cptn_mod"
+| CptnModCondF: "\<lbrakk>(Some P1, s)#ys \<in> cptn_mod; s \<notin> b \<rbrakk> \<Longrightarrow> (Some(Cond b P0 P1), s)#(Some P1, s)#ys \<in> cptn_mod"
+| CptnModSeq1: "\<lbrakk>(Some P0, s)#xs \<in> cptn_mod; zs=map (lift P1) xs \<rbrakk>
+                 \<Longrightarrow> (Some(Seq P0 P1), s)#zs \<in> cptn_mod"
+| CptnModSeq2: 
+  "\<lbrakk>(Some P0, s)#xs \<in> cptn_mod; fst(last ((Some P0, s)#xs)) = None; 
+  (Some P1, snd(last ((Some P0, s)#xs)))#ys \<in> cptn_mod; 
+  zs=(map (lift P1) xs)@ys \<rbrakk> \<Longrightarrow> (Some(Seq P0 P1), s)#zs \<in> cptn_mod"
+
+| CptnModWhile1: 
+  "\<lbrakk> (Some P, s)#xs \<in> cptn_mod; s \<in> b; zs=map (lift (While b P)) xs \<rbrakk> 
+  \<Longrightarrow> (Some(While b P), s)#(Some(Seq P (While b P)), s)#zs \<in> cptn_mod"
+| CptnModWhile2: 
+  "\<lbrakk> (Some P, s)#xs \<in> cptn_mod; fst(last ((Some P, s)#xs))=None; s \<in> b; 
+  zs=(map (lift (While b P)) xs)@ys; 
+  (Some(While b P), snd(last ((Some P, s)#xs)))#ys \<in> cptn_mod\<rbrakk> 
+  \<Longrightarrow> (Some(While b P), s)#(Some(Seq P (While b P)), s)#zs \<in> 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 xs)
+apply simp_all
+done
+
+lemma div_seq [rule_format]: "list \<in> cptn_mod \<Longrightarrow>
+ (\<forall>s P Q zs. list=(Some (Seq P Q), s)#zs \<longrightarrow>
+  (\<exists>xs. (Some P, s)#xs \<in> cptn_mod  \<and> (zs=(map (lift Q) xs) \<or>
+  ( fst(((Some P, s)#xs)!length xs)=None \<and> 
+  (\<exists>ys. (Some Q, snd(((Some P, s)#xs)!length xs))#ys \<in> cptn_mod  
+  \<and> 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.cases,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]:
+  "\<forall>s Q t xs.((Some a, s), Q, t) \<in> ctran \<longrightarrow> (Q, t) # xs \<in> cptn_mod 
+  \<longrightarrow> (Some a, s) # (Q, t) # xs \<in> cptn_mod"
+apply(induct a)
+apply simp_all
+--{* basic *}
+apply clarify
+apply(erule ctran.cases,simp_all)
+apply(rule CptnModNone,rule Basic,simp)
+apply clarify
+apply(erule ctran.cases,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.cases,simp_all)
+apply(force elim: CptnModCondT)
+apply(force elim: CptnModCondF)
+--{* While *}
+apply  clarify
+apply(erule ctran.cases,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.cases,simp_all)
+apply(rule CptnModNone,erule Await,simp+)
+done
+
+lemma cptn_onlyif_cptn_mod [rule_format]: "c \<in> cptn \<Longrightarrow> c \<in> cptn_mod"
+apply(erule cptn.induct)
+  apply(rule CptnModOne)
+ apply(erule CptnModEnv)
+apply(case_tac P)
+ apply simp
+ apply(erule ctran.cases,simp_all)
+apply(force elim:cptn_onlyif_cptn_mod_aux)
+done
+
+lemma lift_is_cptn: "c\<in>cptn \<Longrightarrow> map (lift P) c \<in> 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.cases)
+done
+
+lemma cptn_append_is_cptn [rule_format]: 
+ "\<forall>b a. b#c1\<in>cptn \<longrightarrow>  a#c2\<in>cptn \<longrightarrow> (b#c1)!length c1=a \<longrightarrow> b#c1@c2\<in>cptn"
+apply(induct c1)
+ apply simp
+apply clarify
+apply(erule cptn.cases,simp_all)
+ apply(force intro:CptnEnv)
+apply(force elim:CptnComp)
+done
+
+lemma last_lift: "\<lbrakk>xs\<noteq>[]; fst(xs!(length xs - (Suc 0)))=None\<rbrakk> 
+ \<Longrightarrow> 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) \<longrightarrow> \<not>P a \<longrightarrow> P (x!(length x - (Suc 0)))" 
+apply(induct x,simp+)
+done
+
+lemma last_fst_esp: 
+ "fst(((Some a,s)#xs)!(length xs))=None \<Longrightarrow> fst(xs!(length xs - (Suc 0)))=None" 
+apply(erule last_fst)
+apply simp
+done
+
+lemma last_snd: "xs\<noteq>[] \<Longrightarrow> 
+  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<length xs \<Longrightarrow> map (lift Q) xs ! i = lift Q  (xs! i)"
+by (simp add:lift_def)
+
+lemma snd_lift: "i< length xs \<Longrightarrow> 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 \<in> cptn_mod \<Longrightarrow> c \<in> 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.cases,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\<noteq>[]")
+ 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\<noteq>[]")
+ 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 \<in> cptn) = (c \<in> 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 \<times> 'a set \<times> ('a \<times> 'a) set \<times> ('a \<times> 'a) set \<times> 'a set"
+
+constdefs
+  assum :: "('a set \<times> ('a \<times> 'a) set) \<Rightarrow> ('a confs) set"
+  "assum \<equiv> \<lambda>(pre, rely). {c. snd(c!0) \<in> pre \<and> (\<forall>i. Suc i<length c \<longrightarrow> 
+               c!i -e\<rightarrow> c!(Suc i) \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> rely)}"
+
+  comm :: "(('a \<times> 'a) set \<times> 'a set) \<Rightarrow> ('a confs) set"
+  "comm \<equiv> \<lambda>(guar, post). {c. (\<forall>i. Suc i<length c \<longrightarrow> 
+               c!i -c\<rightarrow> c!(Suc i) \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> guar) \<and> 
+               (fst (last c) = None \<longrightarrow> snd (last c) \<in> post)}"
+
+  com_validity :: "'a com \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set \<Rightarrow> bool" 
+                 ("\<Turnstile> _ sat [_, _, _, _]" [60,0,0,0,0] 45)
+  "\<Turnstile> P sat [pre, rely, guar, post] \<equiv> 
+   \<forall>s. cp (Some P) s \<inter> assum(pre, rely) \<subseteq> comm(guar, post)"
+
+subsection {* Validity for Parallel Programs. *}
+
+constdefs
+  All_None :: "(('a com) option) list \<Rightarrow> bool"
+  "All_None xs \<equiv> \<forall>c\<in>set xs. c=None"
+
+  par_assum :: "('a set \<times> ('a \<times> 'a) set) \<Rightarrow> ('a par_confs) set"
+  "par_assum \<equiv> \<lambda>(pre, rely). {c. snd(c!0) \<in> pre \<and> (\<forall>i. Suc i<length c \<longrightarrow> 
+             c!i -pe\<rightarrow> c!Suc i \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> rely)}"
+
+  par_comm :: "(('a \<times> 'a) set \<times> 'a set) \<Rightarrow> ('a par_confs) set"
+  "par_comm \<equiv> \<lambda>(guar, post). {c. (\<forall>i. Suc i<length c \<longrightarrow>   
+        c!i -pc\<rightarrow> c!Suc i \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> guar) \<and> 
+         (All_None (fst (last c)) \<longrightarrow> snd( last c) \<in> post)}"
+
+  par_com_validity :: "'a  par_com \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set 
+\<Rightarrow> 'a set \<Rightarrow> bool"  ("\<Turnstile> _ SAT [_, _, _, _]" [60,0,0,0,0] 45)
+  "\<Turnstile> Ps SAT [pre, rely, guar, post] \<equiv> 
+   \<forall>s. par_cp Ps s \<inter> par_assum(pre, rely) \<subseteq> par_comm(guar, post)"
+
+subsection {* Compositionality of the Semantics *}
+
+subsubsection {* Definition of the conjoin operator *}
+
+constdefs
+  same_length :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"
+  "same_length c clist \<equiv> (\<forall>i<length clist. length(clist!i)=length c)"
+ 
+  same_state :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"
+  "same_state c clist \<equiv> (\<forall>i <length clist. \<forall>j<length c. snd(c!j) = snd((clist!i)!j))"
+
+  same_program :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"
+  "same_program c clist \<equiv> (\<forall>j<length c. fst(c!j) = map (\<lambda>x. fst(nth x j)) clist)"
+
+  compat_label :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"
+  "compat_label c clist \<equiv> (\<forall>j. Suc j<length c \<longrightarrow> 
+         (c!j -pc\<rightarrow> c!Suc j \<and> (\<exists>i<length clist. (clist!i)!j -c\<rightarrow> (clist!i)! Suc j \<and> 
+                       (\<forall>l<length clist. l\<noteq>i \<longrightarrow> (clist!l)!j -e\<rightarrow> (clist!l)! Suc j))) \<or> 
+         (c!j -pe\<rightarrow> c!Suc j \<and> (\<forall>i<length clist. (clist!i)!j -e\<rightarrow> (clist!i)! Suc j)))"
+
+  conjoin :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"  ("_ \<propto> _" [65,65] 64)
+  "c \<propto> clist \<equiv> (same_length c clist) \<and> (same_state c clist) \<and> (same_program c clist) \<and> (compat_label c clist)"
+
+subsubsection {* Some previous lemmas *}
+
+lemma list_eq_if [rule_format]: 
+  "\<forall>ys. xs=ys \<longrightarrow> (length xs = length ys) \<longrightarrow> (\<forall>i<length xs. xs!i=ys!i)"
+apply (induct xs)
+ apply simp
+apply clarify
+done
+
+lemma list_eq: "(length xs = length ys \<and> (\<forall>i<length xs. xs!i=ys!i)) = (xs=ys)"
+apply(rule iffI)
+ apply clarify
+ apply(erule nth_equalityI)
+ apply simp+
+done
+
+lemma nth_tl: "\<lbrakk> ys!0=a; ys\<noteq>[] \<rbrakk> \<Longrightarrow> ys=(a#(tl ys))"
+apply(case_tac ys)
+ apply simp+
+done
+
+lemma nth_tl_if [rule_format]: "ys\<noteq>[] \<longrightarrow> ys!0=a \<longrightarrow> P ys \<longrightarrow> P (a#(tl ys))"
+apply(induct ys)
+ apply simp+
+done
+
+lemma nth_tl_onlyif [rule_format]: "ys\<noteq>[] \<longrightarrow> ys!0=a \<longrightarrow> P (a#(tl ys)) \<longrightarrow> P ys"
+apply(induct ys)
+ apply simp+
+done
+
+lemma seq_not_eq1: "Seq c1 c2\<noteq>c1"
+apply(rule com.induct)
+apply simp_all
+apply clarify
+done
+
+lemma seq_not_eq2: "Seq c1 c2\<noteq>c2"
+apply(rule com.induct)
+apply simp_all
+apply clarify
+done
+
+lemma if_not_eq1: "Cond b c1 c2 \<noteq>c1"
+apply(rule com.induct)
+apply simp_all
+apply clarify
+done
+
+lemma if_not_eq2: "Cond b c1 c2\<noteq>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:
+  assumes c: "(P,s) -c\<rightarrow> (Q,t)"
+  shows "P\<noteq>Q" using c
+  by (induct x1 \<equiv> "(P,s)" x2 \<equiv> "(Q,t)" arbitrary: P s Q t) auto
+
+lemma prog_not_eq_in_ctran [simp]: "\<not> (P,s) -c\<rightarrow> (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\<rightarrow> (Q,t) \<Longrightarrow> (P\<noteq>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]: "\<not> (P,s) -pc\<rightarrow> (P,t)"
+apply clarify
+apply(drule prog_not_eq_in_par_ctran_aux)
+apply simp
+done
+
+lemma tl_in_cptn: "\<lbrakk> a#xs \<in>cptn; xs\<noteq>[] \<rbrakk> \<Longrightarrow> xs\<in>cptn"
+apply(force elim:cptn.cases)
+done
+
+lemma tl_zero[rule_format]: 
+  "P (ys!Suc j) \<longrightarrow> Suc j<length ys \<longrightarrow> ys\<noteq>[] \<longrightarrow> P (tl(ys)!j)"
+apply(induct ys)
+ apply simp_all
+done
+
+subsection {* The Semantics is Compositional *}
+
+lemma aux_if [rule_format]: 
+  "\<forall>xs s clist. (length clist = length xs \<and> (\<forall>i<length xs. (xs!i,s)#clist!i \<in> cptn) 
+  \<and> ((xs, s)#ys \<propto> map (\<lambda>i. (fst i,s)#snd i) (zip xs clist)) 
+   \<longrightarrow> (xs, s)#ys \<in> 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="\<lambda>j. ?H j \<longrightarrow> (?P j \<or> ?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="\<lambda>j. (?H j) \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
+  prefer 2
+  apply(simp add:same_program_def)
+  apply(erule_tac x=1 and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
+  apply(drule_tac t=i in not_sym,simp)
+  apply(erule etranE,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="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
+   apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
+   apply(drule_tac t=i  in not_sym,simp)
+   apply(erule etranE,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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (?I j \<or> ?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="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE,assumption)
+   apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE,assumption)
+   apply simp
+   apply(case_tac j,simp)
+    apply(rule tl_zero)
+      apply(erule_tac x=l in allE, erule impE, assumption, 
+            erule_tac x=1 and P="\<lambda>j.  (?H j) \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
+      apply(force elim:etranE 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="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
+    apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
+    apply(drule_tac t=i  in not_sym,simp)
+    apply(erule etranE,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="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+    apply simp
+   apply(erule_tac P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE,assumption,erule impE,assumption)
+   apply(erule tl_zero,force)
+   apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+   apply(rule nth_tl_if)
+     apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+    apply(erule_tac x=l  in allE, erule impE, assumption,
+          erule_tac x=1 and P="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
+    apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE, assumption,simp)
+    apply(erule etranE,simp)
+   apply(rule tl_zero)
+    apply force
+   apply force
+  apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> ?I j\<in>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="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
+    apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE,erule impE, assumption,simp)
+    apply(force elim:etranE intro:Env)
+   apply force
+  apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+  apply force
+ apply(erule_tac x=ia and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+--{* first step is an environmental step *}
+apply clarify
+apply(erule par_etran.cases)
+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="\<lambda>j. ?H j \<longrightarrow> (?I ?s j) \<in> cptn" in allE,simp)
+ apply(erule cptn.cases)
+   apply(simp add:same_length_def)
+   apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
+ apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> ?J j \<in>etran" in allE,simp)
+ apply(erule etranE,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="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
+ apply(rule tl_zero)
+   apply(simp)
+  apply force
+ apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (?I j \<or> ?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="\<lambda>i. ?H i \<longrightarrow> ?J i \<in>etran" in allE, erule impE, assumption)
+  apply(erule etranE,simp)
+  apply(erule_tac x=i  in allE, erule impE, assumption,
+        erule_tac x=1 and P="\<lambda>j.  (?H j) \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
+  apply(rule nth_tl_if)
+   apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+  apply simp
+ apply(erule tl_zero,force) 
+  apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply clarify
+ apply(erule_tac x=l and P="\<lambda>i. ?H i \<longrightarrow> ?J i \<in>etran" in allE, erule impE, assumption)
+ apply(erule etranE,simp)
+ apply(erule_tac x=l  in allE, erule impE, assumption,
+       erule_tac x=1 and P="\<lambda>j.  (?H j) \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE,simp)
+ apply(rule nth_tl_if)
+   apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+  apply simp
+  apply(rule tl_zero,force)
+  apply force
+ apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>i. ?H i \<longrightarrow> ?J i \<in>etran" in allE, erule impE, assumption)
+   apply(erule_tac x=i and P="\<lambda>i. ?H i \<longrightarrow> ?J i \<in>etran" in allE, erule impE, assumption)
+   apply(force elim:etranE intro:Env)
+  apply force
+ apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+ apply force
+apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (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]: "\<forall>xs s. (xs, s)#ys \<in> par_cptn \<longrightarrow> 
+  (\<exists>clist. (length clist = length xs) \<and> 
+  (xs, s)#ys \<propto> map (\<lambda>i. (fst i,s)#(snd i)) (zip xs clist) \<and> 
+  (\<forall>i<length xs. (xs!i,s)#(clist!i) \<in> cptn))"
+apply(induct ys)
+ apply(clarify)
+ apply(rule_tac x="map (\<lambda>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.cases,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 (\<lambda>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.cases,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 (\<lambda>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="\<lambda>j. ?H j \<longrightarrow> (?I j \<in> 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 (no_asm_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="\<lambda>j. ?H j \<longrightarrow> (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(rule Env)
+apply simp
+apply(erule_tac x=nat and P="\<lambda>j. ?H j \<longrightarrow> (?P j \<or> ?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 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="\<lambda>j. ?H j \<longrightarrow> (?P j)\<in>etran" in allE, erule impE, assumption)
+apply(case_tac "i=ia",simp,simp)
+done
+
+lemma one_iff_aux: "xs\<noteq>[] \<Longrightarrow> (\<forall>ys. ((xs, s)#ys \<in> par_cptn) = 
+ (\<exists>clist. length clist= length xs \<and> 
+ ((xs, s)#ys \<propto> map (\<lambda>i. (fst i,s)#(snd i)) (zip xs clist)) \<and> 
+ (\<forall>i<length xs. (xs!i,s)#(clist!i) \<in> cptn))) = 
+ (par_cp (xs) s = {c. \<exists>clist. (length clist)=(length xs) \<and>
+ (\<forall>i<length clist. (clist!i) \<in> cp(xs!i) s) \<and> c \<propto> 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.cases)
+  apply simp
+  apply(erule_tac x="list" in allE)
+  apply clarify
+  apply simp
+  apply(rule_tac x="map (\<lambda>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.cases,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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (\<forall>i. ?T i \<longrightarrow> (snd (?d j i))=(snd (?e j i)))" in allE,simp)
+  apply(erule_tac x=0 and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
+  apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+  apply(case_tac "clist!i",simp,simp)
+ apply(thin_tac "?H = (\<exists>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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
+    apply clarify
+    apply(simp add:cp_def)
+    apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+    apply(case_tac "clist!i",simp,simp)
+   apply clarify
+   apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
+   apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (?P j)\<in>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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (snd (?d j))=(snd (?e j))" in allE)
+    apply(case_tac j,simp)
+    apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+      apply(case_tac "clist!i",simp,simp)
+     apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE)
+     apply(case_tac "clist!i",simp,simp)
+    apply clarify
+    apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> ?I j \<longrightarrow> ?J j" in allE)
+    apply(erule_tac x=l and P="\<lambda>j. ?H j \<longrightarrow> (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="\<lambda>j. ?H j \<longrightarrow> (length (?s j) = ?t)" in allE,force)
+    apply force
+   apply(erule_tac x=i and P="\<lambda>j. ?H j \<longrightarrow> (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 (\<lambda>i. (fst i, s) # snd i) (zip xs clist)" in allE)
+apply simp
+apply clarify
+apply(simp add:cp_def)
+done
+
+theorem one: "xs\<noteq>[] \<Longrightarrow> 
+ par_cp xs s = {c. \<exists>clist. (length clist)=(length xs) \<and> 
+               (\<forall>i<length clist. (clist!i) \<in> cp(xs!i) s) \<and> c \<propto> 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
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/ROOT.ML	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,2 @@
+
+use_thy "Hoare_Parallel";
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/document/root.bib	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,28 @@
+@inproceedings{NipkowP-FASE99,author={Tobias Nipkow and Prensa Nieto, Leonor},
+title={{Owicki/Gries} in {Isabelle/HOL}},
+booktitle={Fundamental Approaches to Software Engineering (FASE'99)},
+editor={J.-P. Finance},publisher="Springer",series="LNCS",volume=1577,
+pages={188--203},year=1999}
+
+@InProceedings{PrenEsp00,
+  author =    {Prensa Nieto, Leonor and Javier Esparza}, 
+  title =     {Verifying Single and Multi-mutator Garbage Collectors
+                with {Owicki/Gries} in {Isabelle/HOL}},
+  booktitle = {Mathematical Foundations of Computer Science (MFCS 2000)},
+  editor =    {M. Nielsen and B. Rovan},
+  publisher = {Springer-Verlag},
+  series =    {LNCS},
+  volume =    1893, 
+  pages =     {619--628},
+  year =      2000
+}
+
+@PhdThesis{Prensa-PhD,author={Leonor Prensa Nieto},
+title={Verification of Parallel Programs with the Owicki-Gries and
+Rely-Guarantee Methods in Isabelle/HOL},
+school={Technische Universit{\"a}t M{\"u}nchen},year=2002}
+
+@inproceedings{Prensa-ESOP03,author={Prensa Nieto, Leonor},
+title={The {Rely-Guarantee} Method in {Isabelle/HOL}},
+booktitle={European Symposium on Programming (ESOP'03)},editor={P. Degano},
+publisher=Springer,series=LNCS,volume=2618,pages={348--362},year=2003}
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Hoare_Parallel/document/root.tex	Mon Sep 21 11:15:55 2009 +0200
@@ -0,0 +1,62 @@
+
+% $Id$
+
+\documentclass[11pt,a4paper]{report}
+\usepackage{graphicx}
+\usepackage[english]{babel}
+\usepackage{isabelle,isabellesym}
+\usepackage{pdfsetup}
+
+\urlstyle{rm}
+\isabellestyle{it}
+
+\renewcommand{\isamarkupheader}[1]{#1}
+
+\begin{document}
+
+\title{Hoare Logic for Parallel Programs}
+\author{Leonor Prensa Nieto}
+\maketitle
+
+\begin{abstract}\noindent
+  In the following theories a formalization of the Owicki-Gries and
+  the rely-guarantee methods is presented. These methods are widely
+  used for correctness proofs of parallel imperative programs with
+  shared variables.  We define syntax, semantics and proof rules in
+  Isabelle/HOL.  The proof rules also provide for programs
+  parameterized in the number of parallel components. Their
+  correctness w.r.t.\ the semantics is proven.  Completeness proofs
+  for both methods are extended to the new case of parameterized
+  programs. (These proofs have not been formalized in Isabelle. They
+  can be found in~\cite{Prensa-PhD}.)  Using this formalizations we
+  verify several non-trivial examples for parameterized and
+  non-parameterized programs.  For the automatic generation of
+  verification conditions with the Owicki-Gries method we define a
+  tactic based on the proof rules.  The most involved examples are the
+  verification of two garbage-collection algorithms, the second one
+  parameterized in the number of mutators.
+
+For excellent descriptions of this work see
+\cite{NipkowP-FASE99,PrenEsp00,Prensa-PhD,Prensa-ESOP03}.
+
+\end{abstract}
+
+\pagestyle{plain}
+\thispagestyle{empty}
+\tableofcontents
+
+\clearpage
+
+\begin{center}
+  \includegraphics[scale=0.7]{session_graph}  
+\end{center}
+
+\newpage
+
+\parindent 0pt\parskip 0.5ex
+\input{session}
+
+\bibliographystyle{plain}
+\bibliography{root}
+
+\end{document}
--- a/src/HOL/IsaMakefile	Mon Sep 21 11:15:21 2009 +0200
+++ b/src/HOL/IsaMakefile	Mon Sep 21 11:15:55 2009 +0200
@@ -18,7 +18,7 @@
   HOL-Extraction \
   HOL-Hahn_Banach \
   HOL-Hoare \
-  HOL-HoareParallel \
+  HOL-Hoare_Parallel \
   HOL-Import \
   HOL-IMP \
   HOL-IMPP \
@@ -537,21 +537,22 @@
 	@$(ISABELLE_TOOL) usedir $(OUT)/HOL Hoare
 
 
-## HOL-HoareParallel
+## HOL-Hoare_Parallel
 
-HOL-HoareParallel: HOL $(LOG)/HOL-HoareParallel.gz
+HOL-Hoare_Parallel: HOL $(LOG)/HOL-Hoare_Parallel.gz
 
-$(LOG)/HOL-HoareParallel.gz: $(OUT)/HOL HoareParallel/Gar_Coll.thy	\
-  HoareParallel/Graph.thy HoareParallel/Mul_Gar_Coll.thy		\
-  HoareParallel/OG_Com.thy HoareParallel/OG_Examples.thy		\
-  HoareParallel/OG_Hoare.thy HoareParallel/OG_Syntax.thy		\
-  HoareParallel/OG_Tactics.thy HoareParallel/OG_Tran.thy		\
-  HoareParallel/Quote_Antiquote.thy HoareParallel/RG_Com.thy		\
-  HoareParallel/RG_Examples.thy HoareParallel/RG_Hoare.thy		\
-  HoareParallel/RG_Syntax.thy HoareParallel/RG_Tran.thy			\
-  HoareParallel/ROOT.ML HoareParallel/document/root.tex			\
-  HoareParallel/document/root.bib
-	@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL HoareParallel
+$(LOG)/HOL-Hoare_Parallel.gz: $(OUT)/HOL Hoare_Parallel/Gar_Coll.thy	\
+  Hoare_Parallel/Graph.thy Hoare_Parallel/Hoare_Parallel.thy	\
+  Hoare_Parallel/Mul_Gar_Coll.thy		\
+  Hoare_Parallel/OG_Com.thy Hoare_Parallel/OG_Examples.thy		\
+  Hoare_Parallel/OG_Hoare.thy Hoare_Parallel/OG_Syntax.thy		\
+  Hoare_Parallel/OG_Tactics.thy Hoare_Parallel/OG_Tran.thy		\
+  Hoare_Parallel/Quote_Antiquote.thy Hoare_Parallel/RG_Com.thy		\
+  Hoare_Parallel/RG_Examples.thy Hoare_Parallel/RG_Hoare.thy		\
+  Hoare_Parallel/RG_Syntax.thy Hoare_Parallel/RG_Tran.thy			\
+  Hoare_Parallel/ROOT.ML Hoare_Parallel/document/root.tex			\
+  Hoare_Parallel/document/root.bib
+	@$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL Hoare_Parallel
 
 
 ## HOL-MetisExamples
@@ -1157,7 +1158,7 @@
 		$(LOG)/HOL-Induct.gz $(LOG)/HOL-ex.gz			\
 		$(LOG)/HOL-Subst.gz $(LOG)/HOL-IMP.gz			\
 		$(LOG)/HOL-IMPP.gz $(LOG)/HOL-Hoare.gz			\
-		$(LOG)/HOL-HoareParallel.gz $(LOG)/HOL-Lex.gz		\
+		$(LOG)/HOL-Hoare_Parallel.gz $(LOG)/HOL-Lex.gz		\
 		$(LOG)/HOL-Algebra.gz $(LOG)/HOL-Auth.gz		\
 		$(LOG)/HOL-UNITY.gz $(LOG)/HOL-Modelcheck.gz		\
 		$(LOG)/HOL-Lambda.gz $(LOG)/HOL-Bali.gz			\