--- a/src/HOL/Hoare_Parallel/Gar_Coll.thy Sat Dec 27 19:51:55 2014 +0100
+++ b/src/HOL/Hoare_Parallel/Gar_Coll.thy Sat Dec 27 20:32:06 2014 +0100
@@ -1,10 +1,10 @@
-section {* The Single Mutator Case *}
+section \<open>The Single Mutator Case\<close>
theory Gar_Coll imports Graph OG_Syntax begin
declare psubsetE [rule del]
-text {* Declaration of variables: *}
+text \<open>Declaration of variables:\<close>
record gar_coll_state =
M :: nodes
@@ -12,32 +12,32 @@
bc :: "nat set"
obc :: "nat set"
Ma :: nodes
- ind :: nat
+ ind :: nat
k :: nat
z :: bool
-subsection {* The Mutator *}
+subsection \<open>The Mutator\<close>
-text {* The mutator first redirects an arbitrary edge @{text "R"} from
+text \<open>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.
+@{text "T"}. It then colors the new target @{text "T"} black.
-We declare the arbitrarily selected node and edge as constants:*}
+We declare the arbitrarily selected node and edge as constants:\<close>
consts R :: nat T :: nat
-text {* \noindent The following predicate states, given a list of
+text \<open>\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: *}
+valid:\<close>
definition Mut_init :: "gar_coll_state \<Rightarrow> bool" where
"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
+text \<open>\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. *}
+edge but has not yet colored the new target.\<close>
definition Redirect_Edge :: "gar_coll_state ann_com" where
"Redirect_Edge \<equiv> \<lbrace>\<acute>Mut_init \<and> \<acute>z\<rbrace> \<langle>\<acute>E:=\<acute>E[R:=(fst(\<acute>E!R), T)],, \<acute>z:= (\<not>\<acute>z)\<rangle>"
@@ -47,15 +47,15 @@
definition Mutator :: "gar_coll_state ann_com" where
"Mutator \<equiv>
- \<lbrace>\<acute>Mut_init \<and> \<acute>z\<rbrace>
- WHILE True INV \<lbrace>\<acute>Mut_init \<and> \<acute>z\<rbrace>
+ \<lbrace>\<acute>Mut_init \<and> \<acute>z\<rbrace>
+ WHILE True INV \<lbrace>\<acute>Mut_init \<and> \<acute>z\<rbrace>
DO Redirect_Edge ;; Color_Target OD"
-subsubsection {* Correctness of the mutator *}
+subsubsection \<open>Correctness of the mutator\<close>
lemmas mutator_defs = Mut_init_def Redirect_Edge_def Color_Target_def
-lemma Redirect_Edge:
+lemma Redirect_Edge:
"\<turnstile> Redirect_Edge pre(Color_Target)"
apply (unfold mutator_defs)
apply annhoare
@@ -70,7 +70,7 @@
apply(simp_all)
done
-lemma Mutator:
+lemma Mutator:
"\<turnstile> Mutator \<lbrace>False\<rbrace>"
apply(unfold Mutator_def)
apply annhoare
@@ -78,17 +78,17 @@
apply(simp add:mutator_defs)
done
-subsection {* The Collector *}
+subsection \<open>The Collector\<close>
-text {* \noindent A constant @{text "M_init"} is used to give @{text "\<acute>Ma"} a
+text \<open>\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. *}
+"Roots"} are black.\<close>
consts M_init :: nodes
definition Proper_M_init :: "gar_coll_state \<Rightarrow> bool" where
"Proper_M_init \<equiv> \<guillemotleft> Blacks M_init=Roots \<and> length M_init=length \<acute>M \<guillemotright>"
-
+
definition Proper :: "gar_coll_state \<Rightarrow> bool" where
"Proper \<equiv> \<guillemotleft> Proper_Roots \<acute>M \<and> Proper_Edges(\<acute>M, \<acute>E) \<and> \<acute>Proper_M_init \<guillemotright>"
@@ -97,24 +97,24 @@
lemmas collector_defs = Proper_M_init_def Proper_def Safe_def
-subsubsection {* Blackening the roots *}
+subsubsection \<open>Blackening the roots\<close>
definition Blacken_Roots :: " gar_coll_state ann_com" where
- "Blacken_Roots \<equiv>
+ "Blacken_Roots \<equiv>
\<lbrace>\<acute>Proper\<rbrace>
\<acute>ind:=0;;
\<lbrace>\<acute>Proper \<and> \<acute>ind=0\<rbrace>
- WHILE \<acute>ind<length \<acute>M
+ WHILE \<acute>ind<length \<acute>M
INV \<lbrace>\<acute>Proper \<and> (\<forall>i<\<acute>ind. i \<in> Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind\<le>length \<acute>M\<rbrace>
DO \<lbrace>\<acute>Proper \<and> (\<forall>i<\<acute>ind. i \<in> Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M\<rbrace>
- IF \<acute>ind\<in>Roots THEN
- \<lbrace>\<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\<rbrace>
+ IF \<acute>ind\<in>Roots THEN
+ \<lbrace>\<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\<rbrace>
\<acute>M:=\<acute>M[\<acute>ind:=Black] FI;;
\<lbrace>\<acute>Proper \<and> (\<forall>i<\<acute>ind+1. i \<in> Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M\<rbrace>
- \<acute>ind:=\<acute>ind+1
+ \<acute>ind:=\<acute>ind+1
OD"
-lemma Blacken_Roots:
+lemma Blacken_Roots:
"\<turnstile> Blacken_Roots \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M\<rbrace>"
apply (unfold Blacken_Roots_def)
apply annhoare
@@ -127,7 +127,7 @@
apply force
done
-subsubsection {* Propagating black *}
+subsubsection \<open>Propagating black\<close>
definition PBInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
"PBInv \<equiv> \<guillemotleft> \<lambda>ind. \<acute>obc < Blacks \<acute>M \<or> (\<forall>i <ind. \<not>BtoW (\<acute>E!i, \<acute>M) \<or>
@@ -137,25 +137,25 @@
"Propagate_Black_aux \<equiv>
\<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M\<rbrace>
\<acute>ind:=0;;
- \<lbrace>\<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\<rbrace>
- WHILE \<acute>ind<length \<acute>E
- INV \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M
+ \<lbrace>\<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\<rbrace>
+ WHILE \<acute>ind<length \<acute>E
+ INV \<lbrace>\<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\<rbrace>
- DO \<lbrace>\<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\<rbrace>
- IF \<acute>M!(fst (\<acute>E!\<acute>ind)) = Black THEN
- \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M
+ DO \<lbrace>\<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\<rbrace>
+ IF \<acute>M!(fst (\<acute>E!\<acute>ind)) = Black THEN
+ \<lbrace>\<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\<rbrace>
\<acute>M:=\<acute>M[snd(\<acute>E!\<acute>ind):=Black];;
- \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M
+ \<lbrace>\<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\<rbrace>
\<acute>ind:=\<acute>ind+1
FI
OD"
-lemma Propagate_Black_aux:
+lemma Propagate_Black_aux:
"\<turnstile> Propagate_Black_aux
- \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M
+ \<lbrace>\<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)\<rbrace>"
apply (unfold Propagate_Black_aux_def PBInv_def collector_defs)
apply annhoare
@@ -163,7 +163,7 @@
apply force
apply force
apply force
---{* 4 subgoals left *}
+--\<open>4 subgoals left\<close>
apply clarify
apply(simp add:Proper_Edges_def Proper_Roots_def Graph6 Graph7 Graph8 Graph12)
apply (erule disjE)
@@ -189,10 +189,10 @@
apply(erule subset_psubset_trans)
apply(erule Graph11)
apply fast
---{* 3 subgoals left *}
+--\<open>3 subgoals left\<close>
apply force
apply force
---{* last *}
+--\<open>last\<close>
apply clarify
apply simp
apply(subgoal_tac "ind x = length (E x)")
@@ -206,11 +206,11 @@
apply arith
done
-subsubsection {* Refining propagating black *}
+subsubsection \<open>Refining propagating black\<close>
definition Auxk :: "gar_coll_state \<Rightarrow> bool" where
- "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
+ "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>"
definition Propagate_Black :: " gar_coll_state ann_com" where
@@ -218,28 +218,28 @@
\<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M\<rbrace>
\<acute>ind:=0;;
\<lbrace>\<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\<rbrace>
- WHILE \<acute>ind<length \<acute>E
- INV \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M
+ WHILE \<acute>ind<length \<acute>E
+ INV \<lbrace>\<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\<rbrace>
- DO \<lbrace>\<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\<rbrace>
- IF (\<acute>M!(fst (\<acute>E!\<acute>ind)))=Black THEN
- \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M
+ DO \<lbrace>\<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\<rbrace>
+ IF (\<acute>M!(fst (\<acute>E!\<acute>ind)))=Black THEN
+ \<lbrace>\<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\<rbrace>
\<acute>k:=(snd(\<acute>E!\<acute>ind));;
- \<lbrace>\<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
+ \<lbrace>\<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\<rbrace>
\<langle>\<acute>M:=\<acute>M[\<acute>k:=Black],, \<acute>ind:=\<acute>ind+1\<rangle>
- ELSE \<lbrace>\<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\<rbrace>
- \<langle>IF (\<acute>M!(fst (\<acute>E!\<acute>ind)))\<noteq>Black THEN \<acute>ind:=\<acute>ind+1 FI\<rangle>
+ ELSE \<lbrace>\<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\<rbrace>
+ \<langle>IF (\<acute>M!(fst (\<acute>E!\<acute>ind)))\<noteq>Black THEN \<acute>ind:=\<acute>ind+1 FI\<rangle>
FI
OD"
-lemma Propagate_Black:
+lemma Propagate_Black:
"\<turnstile> Propagate_Black
- \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M
+ \<lbrace>\<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)\<rbrace>"
apply (unfold Propagate_Black_def PBInv_def Auxk_def collector_defs)
apply annhoare
@@ -247,10 +247,10 @@
apply force
apply force
apply force
---{* 5 subgoals left *}
+--\<open>5 subgoals left\<close>
apply clarify
apply(simp add:BtoW_def Proper_Edges_def)
---{* 4 subgoals left *}
+--\<open>4 subgoals left\<close>
apply clarify
apply(simp add:Proper_Edges_def Graph6 Graph7 Graph8 Graph12)
apply (erule disjE)
@@ -287,7 +287,7 @@
apply(erule subset_psubset_trans)
apply(erule Graph11)
apply fast
---{* 2 subgoals left *}
+--\<open>2 subgoals left\<close>
apply clarify
apply(simp add:Proper_Edges_def Graph6 Graph7 Graph8 Graph12)
apply (erule disjE)
@@ -304,61 +304,61 @@
apply arith
apply (simp add: BtoW_def)
apply (simp add: BtoW_def)
---{* last *}
+--\<open>last\<close>
apply clarify
apply simp
apply(subgoal_tac "ind x = length (E x)")
apply (simp)
apply(drule Graph1)
apply simp
- apply clarify
+ apply clarify
apply(erule allE, erule impE, assumption)
apply force
apply force
apply arith
done
-subsubsection {* Counting black nodes *}
+subsubsection \<open>Counting black nodes\<close>
definition CountInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
"CountInv \<equiv> \<guillemotleft> \<lambda>ind. {i. i<ind \<and> \<acute>Ma!i=Black}\<subseteq>\<acute>bc \<guillemotright>"
definition Count :: " gar_coll_state ann_com" where
"Count \<equiv>
- \<lbrace>\<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
+ \<lbrace>\<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={}\<rbrace>
\<acute>ind:=0;;
- \<lbrace>\<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={}
+ \<lbrace>\<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\<rbrace>
- WHILE \<acute>ind<length \<acute>M
- INV \<lbrace>\<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
+ WHILE \<acute>ind<length \<acute>M
+ INV \<lbrace>\<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\<rbrace>
- DO \<lbrace>\<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\<rbrace>
- IF \<acute>M!\<acute>ind=Black
- THEN \<lbrace>\<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
+ DO \<lbrace>\<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\<rbrace>
+ IF \<acute>M!\<acute>ind=Black
+ THEN \<lbrace>\<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\<rbrace>
\<acute>bc:=insert \<acute>ind \<acute>bc
FI;;
- \<lbrace>\<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
+ \<lbrace>\<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\<rbrace>
\<acute>ind:=\<acute>ind+1
OD"
-lemma Count:
- "\<turnstile> Count
- \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M
+lemma Count:
+ "\<turnstile> Count
+ \<lbrace>\<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)\<rbrace>"
apply(unfold Count_def)
@@ -377,14 +377,14 @@
apply force
done
-subsubsection {* Appending garbage nodes to the free list *}
+subsubsection \<open>Appending garbage nodes to the free list\<close>
axiomatization Append_to_free :: "nat \<times> edges \<Rightarrow> edges"
where
Append_to_free0: "length (Append_to_free (i, e)) = length e" and
- Append_to_free1: "Proper_Edges (m, e)
+ Append_to_free1: "Proper_Edges (m, e)
\<Longrightarrow> Proper_Edges (m, Append_to_free(i, e))" and
- Append_to_free2: "i \<notin> Reach e
+ Append_to_free2: "i \<notin> Reach e
\<Longrightarrow> n \<in> Reach (Append_to_free(i, e)) = ( n = i \<or> n \<in> Reach e)"
definition AppendInv :: "gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
@@ -395,20 +395,20 @@
\<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe\<rbrace>
\<acute>ind:=0;;
\<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe \<and> \<acute>ind=0\<rbrace>
- WHILE \<acute>ind<length \<acute>M
+ WHILE \<acute>ind<length \<acute>M
INV \<lbrace>\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind\<le>length \<acute>M\<rbrace>
DO \<lbrace>\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M\<rbrace>
- IF \<acute>M!\<acute>ind=Black THEN
- \<lbrace>\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>M!\<acute>ind=Black\<rbrace>
- \<acute>M:=\<acute>M[\<acute>ind:=White]
+ IF \<acute>M!\<acute>ind=Black THEN
+ \<lbrace>\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>M!\<acute>ind=Black\<rbrace>
+ \<acute>M:=\<acute>M[\<acute>ind:=White]
ELSE \<lbrace>\<acute>Proper \<and> \<acute>AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>ind\<notin>Reach \<acute>E\<rbrace>
\<acute>E:=Append_to_free(\<acute>ind,\<acute>E)
FI;;
- \<lbrace>\<acute>Proper \<and> \<acute>AppendInv (\<acute>ind+1) \<and> \<acute>ind<length \<acute>M\<rbrace>
+ \<lbrace>\<acute>Proper \<and> \<acute>AppendInv (\<acute>ind+1) \<and> \<acute>ind<length \<acute>M\<rbrace>
\<acute>ind:=\<acute>ind+1
OD"
-lemma Append:
+lemma Append:
"\<turnstile> Append \<lbrace>\<acute>Proper\<rbrace>"
apply(unfold Append_def AppendInv_def)
apply annhoare
@@ -429,41 +429,41 @@
apply force
done
-subsubsection {* Correctness of the Collector *}
+subsubsection \<open>Correctness of the Collector\<close>
definition Collector :: " gar_coll_state ann_com" where
"Collector \<equiv>
-\<lbrace>\<acute>Proper\<rbrace>
- WHILE True INV \<lbrace>\<acute>Proper\<rbrace>
- DO
- Blacken_Roots;;
- \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M\<rbrace>
- \<acute>obc:={};;
- \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={}\<rbrace>
- \<acute>bc:=Roots;;
- \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots\<rbrace>
- \<acute>Ma:=M_init;;
- \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots \<and> \<acute>Ma=M_init\<rbrace>
- WHILE \<acute>obc\<noteq>\<acute>bc
- INV \<lbrace>\<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)\<rbrace>
+\<lbrace>\<acute>Proper\<rbrace>
+ WHILE True INV \<lbrace>\<acute>Proper\<rbrace>
+ DO
+ Blacken_Roots;;
+ \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M\<rbrace>
+ \<acute>obc:={};;
+ \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={}\<rbrace>
+ \<acute>bc:=Roots;;
+ \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots\<rbrace>
+ \<acute>Ma:=M_init;;
+ \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots \<and> \<acute>Ma=M_init\<rbrace>
+ WHILE \<acute>obc\<noteq>\<acute>bc
+ INV \<lbrace>\<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)\<rbrace>
DO \<lbrace>\<acute>Proper \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M\<rbrace>
\<acute>obc:=\<acute>bc;;
- Propagate_Black;;
- \<lbrace>\<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)\<rbrace>
+ Propagate_Black;;
+ \<lbrace>\<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)\<rbrace>
\<acute>Ma:=\<acute>M;;
- \<lbrace>\<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
+ \<lbrace>\<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)\<rbrace>
\<acute>bc:={};;
- Count
- OD;;
- Append
+ Count
+ OD;;
+ Append
OD"
-lemma Collector:
+lemma Collector:
"\<turnstile> Collector \<lbrace>False\<rbrace>"
apply(unfold Collector_def)
apply annhoare
@@ -478,14 +478,14 @@
apply(force dest:subset_antisym)
done
-subsection {* Interference Freedom *}
+subsection \<open>Interference Freedom\<close>
-lemmas modules = Redirect_Edge_def Color_Target_def Blacken_Roots_def
+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:
+lemma interfree_Blacken_Roots_Redirect_Edge:
"interfree_aux (Some Blacken_Roots, {}, Some Redirect_Edge)"
apply (unfold modules)
apply interfree_aux
@@ -493,7 +493,7 @@
apply (simp_all add:Graph6 Graph12 abbrev)
done
-lemma interfree_Redirect_Edge_Blacken_Roots:
+lemma interfree_Redirect_Edge_Blacken_Roots:
"interfree_aux (Some Redirect_Edge, {}, Some Blacken_Roots)"
apply (unfold modules)
apply interfree_aux
@@ -501,7 +501,7 @@
apply(simp add:abbrev)+
done
-lemma interfree_Blacken_Roots_Color_Target:
+lemma interfree_Blacken_Roots_Color_Target:
"interfree_aux (Some Blacken_Roots, {}, Some Color_Target)"
apply (unfold modules)
apply interfree_aux
@@ -509,7 +509,7 @@
apply(simp_all add:Graph7 Graph8 nth_list_update abbrev)
done
-lemma interfree_Color_Target_Blacken_Roots:
+lemma interfree_Color_Target_Blacken_Roots:
"interfree_aux (Some Color_Target, {}, Some Blacken_Roots)"
apply (unfold modules )
apply interfree_aux
@@ -517,18 +517,18 @@
apply(simp add:abbrev)+
done
-lemma interfree_Propagate_Black_Redirect_Edge:
+lemma interfree_Propagate_Black_Redirect_Edge:
"interfree_aux (Some Propagate_Black, {}, Some Redirect_Edge)"
apply (unfold modules )
apply interfree_aux
---{* 11 subgoals left *}
+--\<open>11 subgoals left\<close>
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(erule Graph4)
apply(simp)+
apply (simp add:BtoW_def)
apply (simp add:BtoW_def)
@@ -536,11 +536,11 @@
apply (force simp add:BtoW_def)
apply(erule Graph4)
apply simp+
---{* 7 subgoals left *}
+--\<open>7 subgoals left\<close>
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(erule Graph4)
apply(simp)+
apply (simp add:BtoW_def)
apply (simp add:BtoW_def)
@@ -548,12 +548,12 @@
apply (force simp add:BtoW_def)
apply(erule Graph4)
apply simp+
---{* 6 subgoals left *}
+--\<open>6 subgoals left\<close>
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(erule Graph4)
apply(simp)+
apply (simp add:BtoW_def)
apply (simp add:BtoW_def)
@@ -561,15 +561,15 @@
apply (force simp add:BtoW_def)
apply(erule Graph4)
apply simp+
-apply(simp add:BtoW_def nth_list_update)
+apply(simp add:BtoW_def nth_list_update)
apply force
---{* 5 subgoals left *}
+--\<open>5 subgoals left\<close>
apply(clarify, simp add:abbrev Graph6 Graph12)
---{* 4 subgoals left *}
+--\<open>4 subgoals left\<close>
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(erule Graph4)
apply(simp)+
apply (simp add:BtoW_def)
apply (simp add:BtoW_def)
@@ -589,13 +589,13 @@
apply simp+
apply(force simp add:BtoW_def)
apply(force simp add:BtoW_def)
---{* 3 subgoals left *}
+--\<open>3 subgoals left\<close>
apply(clarify, simp add:abbrev Graph6 Graph12)
---{* 2 subgoals left *}
+--\<open>2 subgoals left\<close>
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(erule Graph4)
apply(simp)+
apply (simp add:BtoW_def)
apply (simp add:BtoW_def)
@@ -605,84 +605,84 @@
apply simp+
done
-lemma interfree_Redirect_Edge_Propagate_Black:
+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:
+lemma interfree_Propagate_Black_Color_Target:
"interfree_aux (Some Propagate_Black, {}, Some Color_Target)"
apply (unfold modules )
apply interfree_aux
---{* 11 subgoals left *}
+--\<open>11 subgoals left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)+
apply(erule conjE)+
-apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9,
+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 *}
+ erule allE, erule impE, assumption, erule impE, assumption,
+ simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force)
+--\<open>7 subgoals left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)
apply(erule conjE)+
-apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9,
+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 *}
+ erule allE, erule impE, assumption, erule impE, assumption,
+ simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force)
+--\<open>6 subgoals left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)
apply clarify
apply (rule conjI)
- apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9,
+ 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)
+ 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 *}
+--\<open>5 subgoals left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)
---{* 4 subgoals left *}
+--\<open>4 subgoals left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)
apply (rule conjI)
- apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9,
+ 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)
+ 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,
+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 *}
+--\<open>3 subgoals left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)
---{* 2 subgoals left *}
+--\<open>2 subgoals left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)
-apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9,
+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 *}
+ erule allE, erule impE, assumption, erule impE, assumption,
+ simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force)
+--\<open>3 subgoals left\<close>
apply(simp add:abbrev)
done
-lemma interfree_Color_Target_Propagate_Black:
+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:
+lemma interfree_Count_Redirect_Edge:
"interfree_aux (Some Count, {}, Some Redirect_Edge)"
apply (unfold modules)
apply interfree_aux
---{* 9 subgoals left *}
+--\<open>9 subgoals left\<close>
apply(simp_all add:abbrev Graph6 Graph12)
---{* 6 subgoals left *}
+--\<open>6 subgoals left\<close>
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:
+lemma interfree_Redirect_Edge_Count:
"interfree_aux (Some Redirect_Edge, {}, Some Count)"
apply (unfold modules )
apply interfree_aux
@@ -690,26 +690,26 @@
apply(simp add:abbrev)
done
-lemma interfree_Count_Color_Target:
+lemma interfree_Count_Color_Target:
"interfree_aux (Some Count, {}, Some Color_Target)"
apply (unfold modules )
apply interfree_aux
---{* 9 subgoals left *}
+--\<open>9 subgoals left\<close>
apply(simp_all add:abbrev Graph7 Graph8 Graph12)
---{* 6 subgoals left *}
+--\<open>6 subgoals left\<close>
apply(clarify,simp add:abbrev Graph7 Graph8 Graph12,
erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9)+
---{* 2 subgoals left *}
+--\<open>2 subgoals left\<close>
apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)
apply(rule conjI)
- apply(erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9)
+ apply(erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9)
apply(simp add:nth_list_update)
---{* 1 subgoal left *}
+--\<open>1 subgoal left\<close>
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:
+lemma interfree_Color_Target_Count:
"interfree_aux (Some Color_Target, {}, Some Count)"
apply (unfold modules )
apply interfree_aux
@@ -717,7 +717,7 @@
apply(simp add:abbrev)
done
-lemma interfree_Append_Redirect_Edge:
+lemma interfree_Append_Redirect_Edge:
"interfree_aux (Some Append, {}, Some Redirect_Edge)"
apply (unfold modules )
apply interfree_aux
@@ -725,7 +725,7 @@
apply(clarify, simp add:abbrev Graph6 Append_to_free0 Append_to_free1 Graph12, force dest:Graph3)+
done
-lemma interfree_Redirect_Edge_Append:
+lemma interfree_Redirect_Edge_Append:
"interfree_aux (Some Redirect_Edge, {}, Some Append)"
apply (unfold modules )
apply interfree_aux
@@ -734,7 +734,7 @@
apply(clarify, simp add:abbrev Append_to_free0)+
done
-lemma interfree_Append_Color_Target:
+lemma interfree_Append_Color_Target:
"interfree_aux (Some Append, {}, Some Color_Target)"
apply (unfold modules )
apply interfree_aux
@@ -742,7 +742,7 @@
apply(simp add:abbrev Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12 nth_list_update)
done
-lemma interfree_Color_Target_Append:
+lemma interfree_Color_Target_Append:
"interfree_aux (Some Color_Target, {}, Some Append)"
apply (unfold modules )
apply interfree_aux
@@ -751,17 +751,17 @@
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
+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 *}
+subsubsection \<open>Interference freedom Collector-Mutator\<close>
lemma interfree_Collector_Mutator:
"interfree_aux (Some Collector, {}, Some Mutator)"
@@ -769,20 +769,20 @@
apply interfree_aux
apply(simp_all add:collector_mutator_interfree)
apply(unfold modules collector_defs Mut_init_def)
-apply(tactic {* TRYALL (interfree_aux_tac @{context}) *})
---{* 32 subgoals left *}
+apply(tactic \<open>TRYALL (interfree_aux_tac @{context})\<close>)
+--\<open>32 subgoals left\<close>
apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)
---{* 20 subgoals left *}
-apply(tactic{* TRYALL (clarify_tac @{context}) *})
+--\<open>20 subgoals left\<close>
+apply(tactic\<open>TRYALL (clarify_tac @{context})\<close>)
apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)
-apply(tactic {* TRYALL (etac disjE) *})
+apply(tactic \<open>TRYALL (etac disjE)\<close>)
apply simp_all
-apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac subset_trans,etac @{thm Graph3},force_tac @{context}, assume_tac @{context}]) *})
-apply(tactic {* TRYALL(EVERY'[rtac disjI2,etac subset_trans,rtac @{thm Graph9},force_tac @{context}]) *})
-apply(tactic {* TRYALL(EVERY'[rtac disjI1,etac @{thm psubset_subset_trans},rtac @{thm Graph9},force_tac @{context}]) *})
+apply(tactic \<open>TRYALL(EVERY'[rtac disjI2,rtac subset_trans,etac @{thm Graph3},force_tac @{context}, assume_tac @{context}])\<close>)
+apply(tactic \<open>TRYALL(EVERY'[rtac disjI2,etac subset_trans,rtac @{thm Graph9},force_tac @{context}])\<close>)
+apply(tactic \<open>TRYALL(EVERY'[rtac disjI1,etac @{thm psubset_subset_trans},rtac @{thm Graph9},force_tac @{context}])\<close>)
done
-subsubsection {* Interference freedom Mutator-Collector *}
+subsubsection \<open>Interference freedom Mutator-Collector\<close>
lemma interfree_Mutator_Collector:
"interfree_aux (Some Mutator, {}, Some Collector)"
@@ -790,30 +790,30 @@
apply interfree_aux
apply(simp_all add:collector_mutator_interfree)
apply(unfold modules collector_defs Mut_init_def)
-apply(tactic {* TRYALL (interfree_aux_tac @{context}) *})
---{* 64 subgoals left *}
+apply(tactic \<open>TRYALL (interfree_aux_tac @{context})\<close>)
+--\<open>64 subgoals left\<close>
apply(simp_all add:nth_list_update Invariants Append_to_free0)+
-apply(tactic{* TRYALL (clarify_tac @{context}) *})
---{* 4 subgoals left *}
+apply(tactic\<open>TRYALL (clarify_tac @{context})\<close>)
+--\<open>4 subgoals left\<close>
apply force
apply(simp add:Append_to_free2)
apply force
apply(simp add:Append_to_free2)
done
-subsubsection {* The Garbage Collection algorithm *}
+subsubsection \<open>The Garbage Collection algorithm\<close>
-text {* In total there are 289 verification conditions. *}
+text \<open>In total there are 289 verification conditions.\<close>
-lemma Gar_Coll:
- "\<parallel>- \<lbrace>\<acute>Proper \<and> \<acute>Mut_init \<and> \<acute>z\<rbrace>
- COBEGIN
+lemma Gar_Coll:
+ "\<parallel>- \<lbrace>\<acute>Proper \<and> \<acute>Mut_init \<and> \<acute>z\<rbrace>
+ COBEGIN
Collector
\<lbrace>False\<rbrace>
- \<parallel>
+ \<parallel>
Mutator
- \<lbrace>False\<rbrace>
- COEND
+ \<lbrace>False\<rbrace>
+ COEND
\<lbrace>False\<rbrace>"
apply oghoare
apply(force simp add: Mutator_def Collector_def modules)
--- a/src/HOL/Hoare_Parallel/Graph.thy Sat Dec 27 19:51:55 2014 +0100
+++ b/src/HOL/Hoare_Parallel/Graph.thy Sat Dec 27 20:32:06 2014 +0100
@@ -1,6 +1,6 @@
-chapter {* Case Study: Single and Multi-Mutator Garbage Collection Algorithms *}
+chapter \<open>Case Study: Single and Multi-Mutator Garbage Collection Algorithms\<close>
-section {* Formalization of the Memory *}
+section \<open>Formalization of the Memory\<close>
theory Graph imports Main begin
@@ -29,22 +29,22 @@
\<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
+text\<open>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. *}
+consecutive nodes correspond to an edge in E.\<close>
-subsection {* Proofs about Graphs *}
+subsection \<open>Proofs about Graphs\<close>
lemmas Graph_defs= Blacks_def Proper_Roots_def Proper_Edges_def BtoW_def
declare Graph_defs [simp]
-subsubsection{* Graph 1 *}
+subsubsection\<open>Graph 1\<close>
-lemma Graph1_aux [rule_format]:
+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> 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
@@ -70,8 +70,8 @@
apply force
done
-lemma Graph1:
- "\<lbrakk>Roots\<subseteq>Blacks M; Proper_Edges(M, E); \<forall>i<length E. \<not>BtoW(E!i,M) \<rbrakk>
+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
@@ -87,9 +87,9 @@
apply auto
done
-subsubsection{* Graph 2 *}
+subsubsection\<open>Graph 2\<close>
-lemma Ex_first_occurrence [rule_format]:
+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
@@ -102,14 +102,14 @@
apply arith
done
-lemma Ex_last_occurrence:
+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:
+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
@@ -141,7 +141,7 @@
apply(subgoal_tac "(length path - Suc m) + nat \<le> length path")
prefer 2 apply arith
apply(subgoal_tac "length path - Suc m + nat = length path - Suc 0")
- prefer 2 apply arith
+ prefer 2 apply arith
apply clarify
apply(case_tac "i")
apply(force simp add: nth_list_update)
@@ -181,22 +181,22 @@
done
-subsubsection{* Graph 3 *}
+subsubsection\<open>Graph 3\<close>
declare min.absorb1 [simp] min.absorb2 [simp]
-lemma Graph3:
+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 *}
+--\<open>the changed edge is part of the path\<close>
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 *}
+--\<open>T is NOT a root\<close>
apply clarify
apply(rule_tac x = "(take m path)@patha" in exI)
apply(subgoal_tac "\<not>(length path\<le>m)")
@@ -240,7 +240,7 @@
apply(subgoal_tac "Suc (i - m)=(Suc i - m)" )
prefer 2 apply arith
apply simp
---{* T is a root *}
+--\<open>T is a root\<close>
apply(case_tac "m=0")
apply force
apply(rule_tac x = "take (Suc m) path" in exI)
@@ -253,7 +253,7 @@
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 *}
+--\<open>the changed edge is not part of the path\<close>
apply(rule_tac x = "path" in exI)
apply simp
apply clarify
@@ -265,18 +265,18 @@
apply(force simp add: nth_list_update)
done
-subsubsection{* Graph 4 *}
+subsubsection\<open>Graph 4\<close>
-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>
+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 *}
+--\<open>there exist a black node in the path to T\<close>
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)
@@ -307,18 +307,18 @@
declare min.absorb1 [simp del] min.absorb2 [simp del]
-subsubsection {* Graph 5 *}
+subsubsection \<open>Graph 5\<close>
-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>
+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*}
+--\<open>there exist a black node in the path to T\<close>
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)
@@ -350,27 +350,27 @@
apply force
done
-subsubsection {* Other lemmas about graphs *}
+subsubsection \<open>Other lemmas about graphs\<close>
-lemma Graph6:
+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:
+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:
+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. *}
+text\<open>Some specific lemmata for the verification of garbage collection algorithms.\<close>
lemma Graph9: "j<length M \<Longrightarrow> Blacks M\<subseteq>Blacks (M[j := Black])"
apply (unfold Blacks_def)
@@ -384,7 +384,7 @@
apply auto
done
-lemma Graph11 [rule_format (no_asm)]:
+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)
--- a/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy Sat Dec 27 19:51:55 2014 +0100
+++ b/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy Sat Dec 27 20:32:06 2014 +0100
@@ -1,15 +1,15 @@
-section {* The Multi-Mutator Case *}
+section \<open>The Multi-Mutator Case\<close>
theory Mul_Gar_Coll imports Graph OG_Syntax begin
-text {* The full theory takes aprox. 18 minutes. *}
+text \<open>The full theory takes aprox. 18 minutes.\<close>
record mut =
Z :: bool
R :: nat
T :: nat
-text {* Declaration of variables: *}
+text \<open>Declaration of variables:\<close>
record mul_gar_coll_state =
M :: nodes
@@ -17,45 +17,45 @@
bc :: "nat set"
obc :: "nat set"
Ma :: nodes
- ind :: nat
+ ind :: nat
k :: nat
q :: nat
l :: nat
Muts :: "mut list"
-subsection {* The Mutators *}
+subsection \<open>The Mutators\<close>
definition Mul_mut_init :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
- "Mul_mut_init \<equiv> \<guillemotleft> \<lambda>n. n=length \<acute>Muts \<and> (\<forall>i<n. R (\<acute>Muts!i)<length \<acute>E
+ "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>"
definition Mul_Redirect_Edge :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com" where
"Mul_Redirect_Edge j n \<equiv>
\<lbrace>\<acute>Mul_mut_init n \<and> Z (\<acute>Muts!j)\<rbrace>
- \<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,,
+ \<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>"
definition Mul_Color_Target :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com" where
"Mul_Color_Target j n \<equiv>
- \<lbrace>\<acute>Mul_mut_init n \<and> \<not> Z (\<acute>Muts!j)\<rbrace>
+ \<lbrace>\<acute>Mul_mut_init n \<and> \<not> Z (\<acute>Muts!j)\<rbrace>
\<langle>\<acute>M:=\<acute>M[T (\<acute>Muts!j):=Black],, \<acute>Muts:=\<acute>Muts[j:= (\<acute>Muts!j) \<lparr>Z:=True\<rparr>]\<rangle>"
definition Mul_Mutator :: "nat \<Rightarrow> nat \<Rightarrow> mul_gar_coll_state ann_com" where
"Mul_Mutator j n \<equiv>
- \<lbrace>\<acute>Mul_mut_init n \<and> Z (\<acute>Muts!j)\<rbrace>
- WHILE True
- INV \<lbrace>\<acute>Mul_mut_init n \<and> Z (\<acute>Muts!j)\<rbrace>
- DO Mul_Redirect_Edge j n ;;
- Mul_Color_Target j n
+ \<lbrace>\<acute>Mul_mut_init n \<and> Z (\<acute>Muts!j)\<rbrace>
+ WHILE True
+ INV \<lbrace>\<acute>Mul_mut_init n \<and> Z (\<acute>Muts!j)\<rbrace>
+ 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
+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 *}
+subsubsection \<open>Correctness of the proof outline of one mutator\<close>
-lemma Mul_Redirect_Edge: "0\<le>j \<and> j<n \<Longrightarrow>
- \<turnstile> Mul_Redirect_Edge j n
+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
@@ -64,8 +64,8 @@
apply(simp add:nth_list_update)
done
-lemma Mul_Color_Target: "0\<le>j \<and> j<n \<Longrightarrow>
- \<turnstile> Mul_Color_Target j n
+lemma Mul_Color_Target: "0\<le>j \<and> j<n \<Longrightarrow>
+ \<turnstile> Mul_Color_Target j n
\<lbrace>\<acute>Mul_mut_init n \<and> Z (\<acute>Muts!j)\<rbrace>"
apply (unfold mul_mutator_defs)
apply annhoare
@@ -74,7 +74,7 @@
apply(simp add:nth_list_update)
done
-lemma Mul_Mutator: "0\<le>j \<and> j<n \<Longrightarrow>
+lemma Mul_Mutator: "0\<le>j \<and> j<n \<Longrightarrow>
\<turnstile> Mul_Mutator j n \<lbrace>False\<rbrace>"
apply(unfold Mul_Mutator_def)
apply annhoare
@@ -82,10 +82,10 @@
apply(simp add:mul_mutator_defs Mul_Redirect_Edge_def)
done
-subsubsection {* Interference freedom between mutators *}
+subsubsection \<open>Interference freedom between mutators\<close>
-lemma Mul_interfree_Redirect_Edge_Redirect_Edge:
- "\<lbrakk>0\<le>i; i<n; 0\<le>j; j<n; i\<noteq>j\<rbrakk> \<Longrightarrow>
+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
@@ -93,8 +93,8 @@
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>
+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
@@ -102,8 +102,8 @@
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>
+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
@@ -111,8 +111,8 @@
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>
+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
@@ -120,22 +120,22 @@
apply(simp_all add: nth_list_update)
done
-lemmas mul_mutator_interfree =
+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>
+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 @{context}) *})
-apply(tactic {* ALLGOALS (clarify_tac @{context}) *})
+apply(tactic \<open>TRYALL (interfree_aux_tac @{context})\<close>)
+apply(tactic \<open>ALLGOALS (clarify_tac @{context})\<close>)
apply (simp_all add:nth_list_update)
done
-subsubsection {* Modular Parameterized Mutators *}
+subsubsection \<open>Modular Parameterized Mutators\<close>
lemma Mul_Parameterized_Mutators: "0<n \<Longrightarrow>
\<parallel>- \<lbrace>\<acute>Mul_mut_init n \<and> (\<forall>i<n. Z (\<acute>Muts!i))\<rbrace>
@@ -152,7 +152,7 @@
apply(force simp add:Mul_Mutator_def mul_mutator_defs nth_list_update)
done
-subsection {* The Collector *}
+subsection \<open>The Collector\<close>
definition Queue :: "mul_gar_coll_state \<Rightarrow> nat" where
"Queue \<equiv> \<guillemotleft> length (filter (\<lambda>i. \<not> Z i \<and> \<acute>M!(T i) \<noteq> Black) \<acute>Muts) \<guillemotright>"
@@ -170,25 +170,25 @@
lemmas mul_collector_defs = Proper_M_init_def Mul_Proper_def Safe_def
-subsubsection {* Blackening Roots *}
+subsubsection \<open>Blackening Roots\<close>
definition Mul_Blacken_Roots :: "nat \<Rightarrow> mul_gar_coll_state ann_com" where
"Mul_Blacken_Roots n \<equiv>
\<lbrace>\<acute>Mul_Proper n\<rbrace>
\<acute>ind:=0;;
\<lbrace>\<acute>Mul_Proper n \<and> \<acute>ind=0\<rbrace>
- WHILE \<acute>ind<length \<acute>M
+ WHILE \<acute>ind<length \<acute>M
INV \<lbrace>\<acute>Mul_Proper n \<and> (\<forall>i<\<acute>ind. i\<in>Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind\<le>length \<acute>M\<rbrace>
DO \<lbrace>\<acute>Mul_Proper n \<and> (\<forall>i<\<acute>ind. i\<in>Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M\<rbrace>
- IF \<acute>ind\<in>Roots THEN
- \<lbrace>\<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\<rbrace>
+ IF \<acute>ind\<in>Roots THEN
+ \<lbrace>\<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\<rbrace>
\<acute>M:=\<acute>M[\<acute>ind:=Black] FI;;
\<lbrace>\<acute>Mul_Proper n \<and> (\<forall>i<\<acute>ind+1. i\<in>Roots \<longrightarrow> \<acute>M!i=Black) \<and> \<acute>ind<length \<acute>M\<rbrace>
- \<acute>ind:=\<acute>ind+1
+ \<acute>ind:=\<acute>ind+1
OD"
-lemma Mul_Blacken_Roots:
- "\<turnstile> Mul_Blacken_Roots n
+lemma Mul_Blacken_Roots:
+ "\<turnstile> Mul_Blacken_Roots n
\<lbrace>\<acute>Mul_Proper n \<and> Roots \<subseteq> Blacks \<acute>M\<rbrace>"
apply (unfold Mul_Blacken_Roots_def)
apply annhoare
@@ -201,10 +201,10 @@
apply force
done
-subsubsection {* Propagating Black *}
+subsubsection \<open>Propagating Black\<close>
definition Mul_PBInv :: "mul_gar_coll_state \<Rightarrow> bool" where
- "Mul_PBInv \<equiv> \<guillemotleft>\<acute>Safe \<or> \<acute>obc\<subset>Blacks \<acute>M \<or> \<acute>l<\<acute>Queue
+ "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>"
definition Mul_Auxk :: "mul_gar_coll_state \<Rightarrow> bool" where
@@ -212,49 +212,49 @@
definition Mul_Propagate_Black :: "nat \<Rightarrow> mul_gar_coll_state ann_com" where
"Mul_Propagate_Black n \<equiv>
- \<lbrace>\<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)\<rbrace>
+ \<lbrace>\<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)\<rbrace>
\<acute>ind:=0;;
- \<lbrace>\<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\<rbrace>
- WHILE \<acute>ind<length \<acute>E
- INV \<lbrace>\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M
- \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M
+ \<lbrace>\<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\<rbrace>
+ WHILE \<acute>ind<length \<acute>E
+ INV \<lbrace>\<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\<rbrace>
- DO \<lbrace>\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M
- \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M
+ DO \<lbrace>\<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\<rbrace>
- IF \<acute>M!(fst (\<acute>E!\<acute>ind))=Black THEN
- \<lbrace>\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M
- \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M
+ IF \<acute>M!(fst (\<acute>E!\<acute>ind))=Black THEN
+ \<lbrace>\<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\<rbrace>
\<acute>k:=snd(\<acute>E!\<acute>ind);;
- \<lbrace>\<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
+ \<lbrace>\<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\<rbrace>
\<langle>\<acute>M:=\<acute>M[\<acute>k:=Black],,\<acute>ind:=\<acute>ind+1\<rangle>
- ELSE \<lbrace>\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M
- \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M
+ ELSE \<lbrace>\<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\<rbrace>
\<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
- \<lbrace>\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M
+lemma Mul_Propagate_Black:
+ "\<turnstile> Mul_Propagate_Black n
+ \<lbrace>\<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))\<rbrace>"
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 *}
+--\<open>8 subgoals left\<close>
apply force
apply force
apply force
apply(force simp add:BtoW_def Graph_defs)
---{* 4 subgoals left *}
+--\<open>4 subgoals left\<close>
apply clarify
apply(simp add: mul_collector_defs Graph12 Graph6 Graph7 Graph8)
apply(disjE_tac)
@@ -269,7 +269,7 @@
apply(force)
apply(force)
apply(rule disjI2, rule disjI1, erule subset_psubset_trans, erule Graph11, force)
---{* 2 subgoals left *}
+--\<open>2 subgoals left\<close>
apply clarify
apply(conjI_tac)
apply(disjE_tac)
@@ -278,7 +278,7 @@
apply(erule less_SucE)
apply force
apply (simp add:BtoW_def)
---{* 1 subgoal left *}
+--\<open>1 subgoal left\<close>
apply clarify
apply simp
apply(disjE_tac)
@@ -287,66 +287,66 @@
apply simp_all
done
-subsubsection {* Counting Black Nodes *}
+subsubsection \<open>Counting Black Nodes\<close>
definition Mul_CountInv :: "mul_gar_coll_state \<Rightarrow> nat \<Rightarrow> bool" where
"Mul_CountInv \<equiv> \<guillemotleft> \<lambda>ind. {i. i<ind \<and> \<acute>Ma!i=Black}\<subseteq>\<acute>bc \<guillemotright>"
definition Mul_Count :: "nat \<Rightarrow> mul_gar_coll_state ann_com" where
- "Mul_Count n \<equiv>
- \<lbrace>\<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) )
+ "Mul_Count n \<equiv>
+ \<lbrace>\<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={}\<rbrace>
\<acute>ind:=0;;
- \<lbrace>\<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) )
+ \<lbrace>\<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\<rbrace>
- WHILE \<acute>ind<length \<acute>M
- INV \<lbrace>\<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
+ WHILE \<acute>ind<length \<acute>M
+ INV \<lbrace>\<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\<rbrace>
- DO \<lbrace>\<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
+ DO \<lbrace>\<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\<rbrace>
- IF \<acute>M!\<acute>ind=Black
- THEN \<lbrace>\<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>q<n+1 \<and> \<acute>ind<length \<acute>M\<rbrace>
+ IF \<acute>M!\<acute>ind=Black
+ THEN \<lbrace>\<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\<rbrace>
\<acute>bc:=insert \<acute>ind \<acute>bc
FI;;
- \<lbrace>\<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)
+ \<lbrace>\<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\<rbrace>
\<acute>ind:=\<acute>ind+1
OD"
-
-lemma Mul_Count:
- "\<turnstile> Mul_Count n
- \<lbrace>\<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))
+
+lemma Mul_Count:
+ "\<turnstile> Mul_Count n
+ \<lbrace>\<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\<rbrace>"
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 *}
+--\<open>7 subgoals left\<close>
apply force
apply force
apply force
---{* 4 subgoals left *}
+--\<open>4 subgoals left\<close>
apply clarify
apply(conjI_tac)
apply(disjE_tac)
@@ -357,7 +357,7 @@
back
apply force
apply force
---{* 3 subgoals left *}
+--\<open>3 subgoals left\<close>
apply clarify
apply(conjI_tac)
apply(disjE_tac)
@@ -369,15 +369,15 @@
apply simp
apply(rotate_tac -1)
apply (force simp add:Blacks_def)
---{* 2 subgoals left *}
+--\<open>2 subgoals left\<close>
apply force
---{* 1 subgoal left *}
+--\<open>1 subgoal left\<close>
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 *}
+subsubsection \<open>Appending garbage nodes to the free list\<close>
axiomatization Append_to_free :: "nat \<times> edges \<Rightarrow> edges"
where
@@ -391,30 +391,30 @@
"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>"
definition Mul_Append :: "nat \<Rightarrow> mul_gar_coll_state ann_com" where
- "Mul_Append n \<equiv>
+ "Mul_Append n \<equiv>
\<lbrace>\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe\<rbrace>
\<acute>ind:=0;;
\<lbrace>\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>Safe \<and> \<acute>ind=0\<rbrace>
- WHILE \<acute>ind<length \<acute>M
+ WHILE \<acute>ind<length \<acute>M
INV \<lbrace>\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv \<acute>ind \<and> \<acute>ind\<le>length \<acute>M\<rbrace>
DO \<lbrace>\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M\<rbrace>
- IF \<acute>M!\<acute>ind=Black THEN
- \<lbrace>\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>M!\<acute>ind=Black\<rbrace>
- \<acute>M:=\<acute>M[\<acute>ind:=White]
- ELSE
- \<lbrace>\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>ind\<notin>Reach \<acute>E\<rbrace>
+ IF \<acute>M!\<acute>ind=Black THEN
+ \<lbrace>\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>M!\<acute>ind=Black\<rbrace>
+ \<acute>M:=\<acute>M[\<acute>ind:=White]
+ ELSE
+ \<lbrace>\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv \<acute>ind \<and> \<acute>ind<length \<acute>M \<and> \<acute>ind\<notin>Reach \<acute>E\<rbrace>
\<acute>E:=Append_to_free(\<acute>ind,\<acute>E)
FI;;
- \<lbrace>\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv (\<acute>ind+1) \<and> \<acute>ind<length \<acute>M\<rbrace>
+ \<lbrace>\<acute>Mul_Proper n \<and> \<acute>Mul_AppendInv (\<acute>ind+1) \<and> \<acute>ind<length \<acute>M\<rbrace>
\<acute>ind:=\<acute>ind+1
OD"
-lemma Mul_Append:
- "\<turnstile> Mul_Append n
+lemma Mul_Append:
+ "\<turnstile> Mul_Append n
\<lbrace>\<acute>Mul_Proper n\<rbrace>"
apply(unfold Mul_Append_def)
apply annhoare
-apply(simp_all add: mul_collector_defs Mul_AppendInv_def
+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)
@@ -426,71 +426,71 @@
apply force
done
-subsubsection {* Collector *}
+subsubsection \<open>Collector\<close>
definition Mul_Collector :: "nat \<Rightarrow> mul_gar_coll_state ann_com" where
"Mul_Collector n \<equiv>
-\<lbrace>\<acute>Mul_Proper n\<rbrace>
-WHILE True INV \<lbrace>\<acute>Mul_Proper n\<rbrace>
-DO
-Mul_Blacken_Roots n ;;
-\<lbrace>\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M\<rbrace>
- \<acute>obc:={};;
-\<lbrace>\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={}\<rbrace>
- \<acute>bc:=Roots;;
-\<lbrace>\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots\<rbrace>
- \<acute>l:=0;;
-\<lbrace>\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots \<and> \<acute>l=0\<rbrace>
- WHILE \<acute>l<n+1
- INV \<lbrace>\<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)\<rbrace>
- DO \<lbrace>\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>bc\<subseteq>Blacks \<acute>M
+\<lbrace>\<acute>Mul_Proper n\<rbrace>
+WHILE True INV \<lbrace>\<acute>Mul_Proper n\<rbrace>
+DO
+Mul_Blacken_Roots n ;;
+\<lbrace>\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M\<rbrace>
+ \<acute>obc:={};;
+\<lbrace>\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={}\<rbrace>
+ \<acute>bc:=Roots;;
+\<lbrace>\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots\<rbrace>
+ \<acute>l:=0;;
+\<lbrace>\<acute>Mul_Proper n \<and> Roots\<subseteq>Blacks \<acute>M \<and> \<acute>obc={} \<and> \<acute>bc=Roots \<and> \<acute>l=0\<rbrace>
+ WHILE \<acute>l<n+1
+ INV \<lbrace>\<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)\<rbrace>
+ DO \<lbrace>\<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)\<rbrace>
\<acute>obc:=\<acute>bc;;
- Mul_Propagate_Black n;;
- \<lbrace>\<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))\<rbrace>
+ Mul_Propagate_Black n;;
+ \<lbrace>\<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))\<rbrace>
\<acute>bc:={};;
- \<lbrace>\<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={}\<rbrace>
+ \<lbrace>\<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={}\<rbrace>
\<langle> \<acute>Ma:=\<acute>M,, \<acute>q:=\<acute>Queue \<rangle>;;
- Mul_Count n;;
- \<lbrace>\<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\<rbrace>
+ Mul_Count n;;
+ \<lbrace>\<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\<rbrace>
IF \<acute>obc=\<acute>bc THEN
- \<lbrace>\<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\<rbrace>
- \<acute>l:=\<acute>l+1
- ELSE \<lbrace>\<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\<rbrace>
- \<acute>l:=0 FI
- OD;;
- Mul_Append n
+ \<lbrace>\<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\<rbrace>
+ \<acute>l:=\<acute>l+1
+ ELSE \<lbrace>\<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\<rbrace>
+ \<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
+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
+ "\<turnstile> Mul_Collector n
\<lbrace>False\<rbrace>"
apply(unfold Mul_Collector_def)
apply annhoare
-apply(simp_all only:pre.simps Mul_Blacken_Roots
+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)
@@ -505,10 +505,10 @@
apply force
done
-subsection {* Interference Freedom *}
+subsection \<open>Interference Freedom\<close>
-lemma le_length_filter_update[rule_format]:
- "\<forall>i. (\<not>P (list!i) \<or> P j) \<and> i<length list
+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)
@@ -518,8 +518,8 @@
apply(simp)
done
-lemma less_length_filter_update [rule_format]:
- "\<forall>i. P j \<and> \<not>(P (list!i)) \<and> i<length list
+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)
@@ -529,7 +529,7 @@
apply(simp)
done
-lemma Mul_interfree_Blacken_Roots_Redirect_Edge: "\<lbrakk>0\<le>j; j<n\<rbrakk> \<Longrightarrow>
+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
@@ -537,7 +537,7 @@
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>
+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
@@ -545,7 +545,7 @@
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>
+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
@@ -553,7 +553,7 @@
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>
+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
@@ -561,12 +561,12 @@
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>
+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 *}
+--\<open>7 subgoals left\<close>
apply clarify
apply(disjE_tac)
apply(simp_all add:Graph6)
@@ -574,7 +574,7 @@
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 *}
+--\<open>6 subgoals left\<close>
apply clarify
apply(disjE_tac)
apply(simp_all add:Graph6)
@@ -582,7 +582,7 @@
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 *}
+--\<open>5 subgoals left\<close>
apply clarify
apply(disjE_tac)
apply(simp_all add:Graph6)
@@ -606,7 +606,7 @@
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 *}
+--\<open>4 subgoals left\<close>
apply clarify
apply(disjE_tac)
apply(simp_all add:Graph6)
@@ -630,7 +630,7 @@
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 *}
+--\<open>3 subgoals left\<close>
apply clarify
apply(disjE_tac)
apply(simp_all add:Graph6)
@@ -686,7 +686,7 @@
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 *}
+--\<open>2 subgoals left\<close>
apply clarify
apply(rule conjI)
apply(disjE_tac)
@@ -741,7 +741,7 @@
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(disjE_tac)
apply simp_all
apply(conjI_tac)
apply(rule impI)
@@ -756,7 +756,7 @@
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 *}
+--\<open>1 subgoal left\<close>
apply clarify
apply(disjE_tac)
apply(simp_all add:Graph6)
@@ -782,7 +782,7 @@
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>
+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
@@ -790,12 +790,12 @@
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>
+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 *}
+--\<open>7 subgoals left\<close>
apply clarify
apply (simp add:Graph7 Graph8 Graph12)
apply(disjE_tac)
@@ -803,9 +803,9 @@
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 subset_psubset_trans, erule Graph11, simp)
apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9)
---{* 6 subgoals left *}
+--\<open>6 subgoals left\<close>
apply clarify
apply (simp add:Graph7 Graph8 Graph12)
apply(disjE_tac)
@@ -813,13 +813,13 @@
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 subset_psubset_trans, erule Graph11, simp)
apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9)
---{* 5 subgoals left *}
+--\<open>5 subgoals left\<close>
apply clarify
apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12)
apply(disjE_tac)
- apply(simp add:Graph7 Graph8 Graph12)
+ 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)
@@ -832,8 +832,8 @@
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(rule disjI2,rule disjI1,erule subset_psubset_trans, erule Graph11, simp)
+--\<open>4 subgoals left\<close>
apply clarify
apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12)
apply(disjE_tac)
@@ -850,8 +850,8 @@
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(rule disjI2,rule disjI1,erule subset_psubset_trans, erule Graph11, simp)
+--\<open>3 subgoals left\<close>
apply clarify
apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12)
apply(case_tac "M x!(T (Muts x!j))=Black")
@@ -864,10 +864,10 @@
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(rule disjI2,rule disjI1, erule subset_psubset_trans,simp add:Graph11)
apply (force simp add:nth_list_update)
---{* 2 subgoals left *}
-apply clarify
+--\<open>2 subgoals left\<close>
+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)
@@ -881,13 +881,13 @@
apply((rule impI)+)
apply simp
apply(erule disjE)
- apply(rule disjI1, erule less_le_trans)
+ 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(rule disjI2,rule disjI1, erule subset_psubset_trans,simp add:Graph11)
apply (force simp add:nth_list_update)
---{* 1 subgoal left *}
+--\<open>1 subgoal left\<close>
apply clarify
apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12)
apply(case_tac "M x!(T (Muts x!j))=Black")
@@ -899,10 +899,10 @@
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)
+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>
+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
@@ -910,11 +910,11 @@
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>
+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 *}
+--\<open>9 subgoals left\<close>
apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def Graph6)
apply clarify
apply disjE_tac
@@ -928,9 +928,9 @@
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 *}
+--\<open>8 subgoals left\<close>
apply(simp add:mul_mutator_defs nth_list_update)
---{* 7 subgoals left *}
+--\<open>7 subgoals left\<close>
apply(simp add:mul_mutator_defs mul_collector_defs)
apply clarify
apply disjE_tac
@@ -944,7 +944,7 @@
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 *}
+--\<open>6 subgoals left\<close>
apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def)
apply clarify
apply disjE_tac
@@ -958,7 +958,7 @@
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 *}
+--\<open>5 subgoals left\<close>
apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def)
apply clarify
apply disjE_tac
@@ -972,7 +972,7 @@
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 *}
+--\<open>4 subgoals left\<close>
apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def)
apply clarify
apply disjE_tac
@@ -986,9 +986,9 @@
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 *}
+--\<open>3 subgoals left\<close>
apply(simp add:mul_mutator_defs nth_list_update)
---{* 2 subgoals left *}
+--\<open>2 subgoals left\<close>
apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def)
apply clarify
apply disjE_tac
@@ -1002,11 +1002,11 @@
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 *}
+--\<open>1 subgoal left\<close>
apply(simp add:mul_mutator_defs nth_list_update)
done
-lemma Mul_interfree_Redirect_Edge_Count: "\<lbrakk>0\<le>j; j<n\<rbrakk>\<Longrightarrow>
+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
@@ -1014,12 +1014,12 @@
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>
+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 *}
+--\<open>6 subgoals left\<close>
apply clarify
apply disjE_tac
apply (simp add: Graph7 Graph8 Graph12)
@@ -1033,7 +1033,7 @@
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 *}
+--\<open>5 subgoals left\<close>
apply clarify
apply disjE_tac
apply (simp add: Graph7 Graph8 Graph12)
@@ -1047,7 +1047,7 @@
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 *}
+--\<open>4 subgoals left\<close>
apply clarify
apply disjE_tac
apply (simp add: Graph7 Graph8 Graph12)
@@ -1061,7 +1061,7 @@
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 *}
+--\<open>3 subgoals left\<close>
apply clarify
apply disjE_tac
apply (simp add: Graph7 Graph8 Graph12)
@@ -1075,7 +1075,7 @@
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 *}
+--\<open>2 subgoals left\<close>
apply clarify
apply disjE_tac
apply (simp add: Graph7 Graph8 Graph12 nth_list_update)
@@ -1093,7 +1093,7 @@
apply(rule conjI)
apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9)
apply (simp add: nth_list_update)
---{* 1 subgoal left *}
+--\<open>1 subgoal left\<close>
apply clarify
apply disjE_tac
apply (simp add: Graph7 Graph8 Graph12)
@@ -1109,7 +1109,7 @@
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>
+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
@@ -1117,137 +1117,137 @@
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>
+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 @{context}) *})
+apply(tactic \<open>ALLGOALS (clarify_tac @{context})\<close>)
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>
+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 @{context}) *})
+apply(tactic \<open>ALLGOALS (clarify_tac @{context})\<close>)
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>
+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 @{context}) *})
-apply(simp_all add:mul_mutator_defs mul_collector_defs Mul_AppendInv_def Graph7 Graph8 Append_to_free0 Append_to_free1
+apply(tactic \<open>ALLGOALS (clarify_tac @{context})\<close>)
+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>
+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 @{context}) *})
+apply(tactic \<open>ALLGOALS (clarify_tac @{context})\<close>)
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 *}
+subsubsection \<open>Interference freedom Collector-Mutator\<close>
-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
+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>
+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 @{context}) *})
---{* 42 subgoals left *}
+apply(tactic \<open>TRYALL (interfree_aux_tac @{context})\<close>)
+--\<open>42 subgoals left\<close>
apply (clarify,simp add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)+
---{* 24 subgoals left *}
+--\<open>24 subgoals left\<close>
apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)
---{* 14 subgoals left *}
-apply(tactic {* TRYALL (clarify_tac @{context}) *})
+--\<open>14 subgoals left\<close>
+apply(tactic \<open>TRYALL (clarify_tac @{context})\<close>)
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(tactic \<open>TRYALL (rtac conjI)\<close>)
+apply(tactic \<open>TRYALL (rtac impI)\<close>)
+apply(tactic \<open>TRYALL (etac disjE)\<close>)
+apply(tactic \<open>TRYALL (etac conjE)\<close>)
+apply(tactic \<open>TRYALL (etac disjE)\<close>)
+apply(tactic \<open>TRYALL (etac disjE)\<close>)
+--\<open>72 subgoals left\<close>
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 @{context}, assume_tac @{context}]) *})
---{* 28 subgoals left *}
-apply(tactic {* TRYALL (etac conjE) *})
-apply(tactic {* TRYALL (etac disjE) *})
---{* 34 subgoals left *}
+--\<open>35 subgoals left\<close>
+apply(tactic \<open>TRYALL(EVERY'[rtac disjI1,rtac subset_trans,etac @{thm Graph3},force_tac @{context}, assume_tac @{context}])\<close>)
+--\<open>28 subgoals left\<close>
+apply(tactic \<open>TRYALL (etac conjE)\<close>)
+apply(tactic \<open>TRYALL (etac disjE)\<close>)
+--\<open>34 subgoals left\<close>
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 @{context}]) *})
---{* 41 subgoals left *}
-apply(tactic {* TRYALL(EVERY'[rtac disjI2, rtac disjI1, etac @{thm le_trans},
+--\<open>47 subgoals left\<close>
+apply(tactic \<open>TRYALL(EVERY'[REPEAT o (rtac disjI2),etac @{thm subset_psubset_trans}, etac @{thm Graph11},force_tac @{context}])\<close>)
+--\<open>41 subgoals left\<close>
+apply(tactic \<open>TRYALL(EVERY'[rtac disjI2, rtac disjI1, etac @{thm le_trans},
force_tac (@{context} 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 @{context}]) *})
---{* 31 subgoals left *}
-apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac disjI1,etac @{thm subset_psubset_trans},etac @{thm Graph11},force_tac @{context}]) *})
---{* 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 @{context}]) *})
---{* 25 subgoals left *}
-apply(tactic {* TRYALL(EVERY'[rtac disjI2, rtac disjI2, rtac disjI1, etac @{thm le_trans},
+ [@{thm Queue_def}, @{thm less_Suc_eq_le}, @{thm le_length_filter_update}])])\<close>)
+--\<open>35 subgoals left\<close>
+apply(tactic \<open>TRYALL(EVERY'[rtac disjI2,rtac disjI1,etac @{thm psubset_subset_trans},rtac @{thm Graph9},force_tac @{context}])\<close>)
+--\<open>31 subgoals left\<close>
+apply(tactic \<open>TRYALL(EVERY'[rtac disjI2,rtac disjI1,etac @{thm subset_psubset_trans},etac @{thm Graph11},force_tac @{context}])\<close>)
+--\<open>29 subgoals left\<close>
+apply(tactic \<open>TRYALL(EVERY'[REPEAT o (rtac disjI2),etac @{thm subset_psubset_trans},etac @{thm subset_psubset_trans},etac @{thm Graph11},force_tac @{context}])\<close>)
+--\<open>25 subgoals left\<close>
+apply(tactic \<open>TRYALL(EVERY'[rtac disjI2, rtac disjI2, rtac disjI1, etac @{thm le_trans},
force_tac (@{context} addsimps
- [@{thm Queue_def}, @{thm less_Suc_eq_le}, @{thm le_length_filter_update}])]) *})
---{* 10 subgoals left *}
+ [@{thm Queue_def}, @{thm less_Suc_eq_le}, @{thm le_length_filter_update}])])\<close>)
+--\<open>10 subgoals left\<close>
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 *}
+subsubsection \<open>Interference freedom Mutator-Collector\<close>
-lemma Mul_interfree_Mutator_Collector: " j < n \<Longrightarrow>
+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 @{context}) *})
---{* 76 subgoals left *}
+apply(tactic \<open>TRYALL (interfree_aux_tac @{context})\<close>)
+--\<open>76 subgoals left\<close>
apply (clarsimp simp add: nth_list_update)+
---{* 56 subgoals left *}
+--\<open>56 subgoals left\<close>
apply (clarsimp simp add: Mul_AppendInv_def Append_to_free0 nth_list_update)+
done
-subsubsection {* The Multi-Mutator Garbage Collection Algorithm *}
+subsubsection \<open>The Multi-Mutator Garbage Collection Algorithm\<close>
-text {* The total number of verification conditions is 328 *}
+text \<open>The total number of verification conditions is 328\<close>
-lemma Mul_Gar_Coll:
- "\<parallel>- \<lbrace>\<acute>Mul_Proper n \<and> \<acute>Mul_mut_init n \<and> (\<forall>i<n. Z (\<acute>Muts!i))\<rbrace>
- COBEGIN
+lemma Mul_Gar_Coll:
+ "\<parallel>- \<lbrace>\<acute>Mul_Proper n \<and> \<acute>Mul_mut_init n \<and> (\<forall>i<n. Z (\<acute>Muts!i))\<rbrace>
+ COBEGIN
Mul_Collector n
\<lbrace>False\<rbrace>
- \<parallel>
+ \<parallel>
SCHEME [0\<le> j< n]
Mul_Mutator j n
- \<lbrace>False\<rbrace>
- COEND
+ \<lbrace>False\<rbrace>
+ COEND
\<lbrace>False\<rbrace>"
apply oghoare
---{* Strengthening the precondition *}
+--\<open>Strengthening the precondition\<close>
apply(rule Int_greatest)
apply (case_tac n)
apply(force simp add: Mul_Collector_def mul_mutator_defs mul_collector_defs nth_append)
@@ -1257,15 +1257,15 @@
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 *}
+--\<open>Collector\<close>
apply(rule Mul_Collector)
---{* Mutator *}
+--\<open>Mutator\<close>
apply(erule Mul_Mutator)
---{* Interference freedom *}
+--\<open>Interference freedom\<close>
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 *}
+--\<open>Weakening of the postcondition\<close>
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)
--- a/src/HOL/Hoare_Parallel/OG_Com.thy Sat Dec 27 19:51:55 2014 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Com.thy Sat Dec 27 20:32:06 2014 +0100
@@ -1,34 +1,34 @@
-chapter {* The Owicki-Gries Method *}
+chapter \<open>The Owicki-Gries Method\<close>
-section {* Abstract Syntax *}
+section \<open>Abstract Syntax\<close>
theory OG_Com imports Main begin
-text {* Type abbreviations for boolean expressions and assertions: *}
+text \<open>Type abbreviations for boolean expressions and assertions:\<close>
type_synonym 'a bexp = "'a set"
type_synonym 'a assn = "'a set"
-text {* The syntax of commands is defined by two mutually recursive
+text \<open>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. *}
+com"} for non-annotated commands.\<close>
-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 =
+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)"
+ | 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: *}
+text \<open>The function @{text pre} extracts the precondition of an
+annotated command:\<close>
primrec pre ::"'a ann_com \<Rightarrow> 'a assn" where
"pre (AnnBasic r f) = r"
@@ -38,7 +38,7 @@
| "pre (AnnWhile r b i c) = r"
| "pre (AnnAwait r b c) = r"
-text {* Well-formedness predicate for atomic programs: *}
+text \<open>Well-formedness predicate for atomic programs:\<close>
primrec atom_com :: "'a com \<Rightarrow> bool" where
"atom_com (Parallel Ts) = False"
@@ -46,5 +46,5 @@
| "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/Hoare_Parallel/OG_Examples.thy Sat Dec 27 19:51:55 2014 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Examples.thy Sat Dec 27 20:32:06 2014 +0100
@@ -1,206 +1,206 @@
-section {* Examples *}
+section \<open>Examples\<close>
theory OG_Examples imports OG_Syntax begin
-subsection {* Mutual Exclusion *}
+subsection \<open>Mutual Exclusion\<close>
-subsubsection {* Peterson's Algorithm I*}
+subsubsection \<open>Peterson's Algorithm I\<close>
-text {* Eike Best. "Semantics of Sequential and Parallel Programs", page 217. *}
+text \<open>Eike Best. "Semantics of Sequential and Parallel Programs", page 217.\<close>
record Petersons_mutex_1 =
pr1 :: nat
pr2 :: nat
in1 :: bool
- in2 :: bool
+ in2 :: bool
hold :: nat
-lemma Petersons_mutex_1:
- "\<parallel>- \<lbrace>\<acute>pr1=0 \<and> \<not>\<acute>in1 \<and> \<acute>pr2=0 \<and> \<not>\<acute>in2 \<rbrace>
- COBEGIN \<lbrace>\<acute>pr1=0 \<and> \<not>\<acute>in1\<rbrace>
- WHILE True INV \<lbrace>\<acute>pr1=0 \<and> \<not>\<acute>in1\<rbrace>
- DO
- \<lbrace>\<acute>pr1=0 \<and> \<not>\<acute>in1\<rbrace> \<langle> \<acute>in1:=True,,\<acute>pr1:=1 \<rangle>;;
- \<lbrace>\<acute>pr1=1 \<and> \<acute>in1\<rbrace> \<langle> \<acute>hold:=1,,\<acute>pr1:=2 \<rangle>;;
- \<lbrace>\<acute>pr1=2 \<and> \<acute>in1 \<and> (\<acute>hold=1 \<or> \<acute>hold=2 \<and> \<acute>pr2=2)\<rbrace>
- AWAIT (\<not>\<acute>in2 \<or> \<not>(\<acute>hold=1)) THEN \<acute>pr1:=3 END;;
- \<lbrace>\<acute>pr1=3 \<and> \<acute>in1 \<and> (\<acute>hold=1 \<or> \<acute>hold=2 \<and> \<acute>pr2=2)\<rbrace>
- \<langle>\<acute>in1:=False,,\<acute>pr1:=0\<rangle>
- OD \<lbrace>\<acute>pr1=0 \<and> \<not>\<acute>in1\<rbrace>
- \<parallel>
- \<lbrace>\<acute>pr2=0 \<and> \<not>\<acute>in2\<rbrace>
- WHILE True INV \<lbrace>\<acute>pr2=0 \<and> \<not>\<acute>in2\<rbrace>
- DO
- \<lbrace>\<acute>pr2=0 \<and> \<not>\<acute>in2\<rbrace> \<langle> \<acute>in2:=True,,\<acute>pr2:=1 \<rangle>;;
- \<lbrace>\<acute>pr2=1 \<and> \<acute>in2\<rbrace> \<langle> \<acute>hold:=2,,\<acute>pr2:=2 \<rangle>;;
- \<lbrace>\<acute>pr2=2 \<and> \<acute>in2 \<and> (\<acute>hold=2 \<or> (\<acute>hold=1 \<and> \<acute>pr1=2))\<rbrace>
- AWAIT (\<not>\<acute>in1 \<or> \<not>(\<acute>hold=2)) THEN \<acute>pr2:=3 END;;
- \<lbrace>\<acute>pr2=3 \<and> \<acute>in2 \<and> (\<acute>hold=2 \<or> (\<acute>hold=1 \<and> \<acute>pr1=2))\<rbrace>
- \<langle>\<acute>in2:=False,,\<acute>pr2:=0\<rangle>
- OD \<lbrace>\<acute>pr2=0 \<and> \<not>\<acute>in2\<rbrace>
- COEND
+lemma Petersons_mutex_1:
+ "\<parallel>- \<lbrace>\<acute>pr1=0 \<and> \<not>\<acute>in1 \<and> \<acute>pr2=0 \<and> \<not>\<acute>in2 \<rbrace>
+ COBEGIN \<lbrace>\<acute>pr1=0 \<and> \<not>\<acute>in1\<rbrace>
+ WHILE True INV \<lbrace>\<acute>pr1=0 \<and> \<not>\<acute>in1\<rbrace>
+ DO
+ \<lbrace>\<acute>pr1=0 \<and> \<not>\<acute>in1\<rbrace> \<langle> \<acute>in1:=True,,\<acute>pr1:=1 \<rangle>;;
+ \<lbrace>\<acute>pr1=1 \<and> \<acute>in1\<rbrace> \<langle> \<acute>hold:=1,,\<acute>pr1:=2 \<rangle>;;
+ \<lbrace>\<acute>pr1=2 \<and> \<acute>in1 \<and> (\<acute>hold=1 \<or> \<acute>hold=2 \<and> \<acute>pr2=2)\<rbrace>
+ AWAIT (\<not>\<acute>in2 \<or> \<not>(\<acute>hold=1)) THEN \<acute>pr1:=3 END;;
+ \<lbrace>\<acute>pr1=3 \<and> \<acute>in1 \<and> (\<acute>hold=1 \<or> \<acute>hold=2 \<and> \<acute>pr2=2)\<rbrace>
+ \<langle>\<acute>in1:=False,,\<acute>pr1:=0\<rangle>
+ OD \<lbrace>\<acute>pr1=0 \<and> \<not>\<acute>in1\<rbrace>
+ \<parallel>
+ \<lbrace>\<acute>pr2=0 \<and> \<not>\<acute>in2\<rbrace>
+ WHILE True INV \<lbrace>\<acute>pr2=0 \<and> \<not>\<acute>in2\<rbrace>
+ DO
+ \<lbrace>\<acute>pr2=0 \<and> \<not>\<acute>in2\<rbrace> \<langle> \<acute>in2:=True,,\<acute>pr2:=1 \<rangle>;;
+ \<lbrace>\<acute>pr2=1 \<and> \<acute>in2\<rbrace> \<langle> \<acute>hold:=2,,\<acute>pr2:=2 \<rangle>;;
+ \<lbrace>\<acute>pr2=2 \<and> \<acute>in2 \<and> (\<acute>hold=2 \<or> (\<acute>hold=1 \<and> \<acute>pr1=2))\<rbrace>
+ AWAIT (\<not>\<acute>in1 \<or> \<not>(\<acute>hold=2)) THEN \<acute>pr2:=3 END;;
+ \<lbrace>\<acute>pr2=3 \<and> \<acute>in2 \<and> (\<acute>hold=2 \<or> (\<acute>hold=1 \<and> \<acute>pr1=2))\<rbrace>
+ \<langle>\<acute>in2:=False,,\<acute>pr2:=0\<rangle>
+ OD \<lbrace>\<acute>pr2=0 \<and> \<not>\<acute>in2\<rbrace>
+ COEND
\<lbrace>\<acute>pr1=0 \<and> \<not>\<acute>in1 \<and> \<acute>pr2=0 \<and> \<not>\<acute>in2\<rbrace>"
apply oghoare
---{* 104 verification conditions. *}
+--\<open>104 verification conditions.\<close>
apply auto
done
-subsubsection {*Peterson's Algorithm II: A Busy Wait Solution *}
-
-text {* Apt and Olderog. "Verification of sequential and concurrent Programs", page 282. *}
+subsubsection \<open>Peterson's Algorithm II: A Busy Wait Solution\<close>
+
+text \<open>Apt and Olderog. "Verification of sequential and concurrent Programs", page 282.\<close>
record Busy_wait_mutex =
flag1 :: bool
flag2 :: bool
turn :: nat
- after1 :: bool
+ after1 :: bool
after2 :: bool
-lemma Busy_wait_mutex:
- "\<parallel>- \<lbrace>True\<rbrace>
- \<acute>flag1:=False,, \<acute>flag2:=False,,
- COBEGIN \<lbrace>\<not>\<acute>flag1\<rbrace>
- WHILE True
- INV \<lbrace>\<not>\<acute>flag1\<rbrace>
- DO \<lbrace>\<not>\<acute>flag1\<rbrace> \<langle> \<acute>flag1:=True,,\<acute>after1:=False \<rangle>;;
- \<lbrace>\<acute>flag1 \<and> \<not>\<acute>after1\<rbrace> \<langle> \<acute>turn:=1,,\<acute>after1:=True \<rangle>;;
- \<lbrace>\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)\<rbrace>
- WHILE \<not>(\<acute>flag2 \<longrightarrow> \<acute>turn=2)
- INV \<lbrace>\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)\<rbrace>
- DO \<lbrace>\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)\<rbrace> SKIP OD;;
+lemma Busy_wait_mutex:
+ "\<parallel>- \<lbrace>True\<rbrace>
+ \<acute>flag1:=False,, \<acute>flag2:=False,,
+ COBEGIN \<lbrace>\<not>\<acute>flag1\<rbrace>
+ WHILE True
+ INV \<lbrace>\<not>\<acute>flag1\<rbrace>
+ DO \<lbrace>\<not>\<acute>flag1\<rbrace> \<langle> \<acute>flag1:=True,,\<acute>after1:=False \<rangle>;;
+ \<lbrace>\<acute>flag1 \<and> \<not>\<acute>after1\<rbrace> \<langle> \<acute>turn:=1,,\<acute>after1:=True \<rangle>;;
+ \<lbrace>\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)\<rbrace>
+ WHILE \<not>(\<acute>flag2 \<longrightarrow> \<acute>turn=2)
+ INV \<lbrace>\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)\<rbrace>
+ DO \<lbrace>\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)\<rbrace> SKIP OD;;
\<lbrace>\<acute>flag1 \<and> \<acute>after1 \<and> (\<acute>flag2 \<and> \<acute>after2 \<longrightarrow> \<acute>turn=2)\<rbrace>
- \<acute>flag1:=False
- OD
- \<lbrace>False\<rbrace>
- \<parallel>
- \<lbrace>\<not>\<acute>flag2\<rbrace>
- WHILE True
- INV \<lbrace>\<not>\<acute>flag2\<rbrace>
- DO \<lbrace>\<not>\<acute>flag2\<rbrace> \<langle> \<acute>flag2:=True,,\<acute>after2:=False \<rangle>;;
- \<lbrace>\<acute>flag2 \<and> \<not>\<acute>after2\<rbrace> \<langle> \<acute>turn:=2,,\<acute>after2:=True \<rangle>;;
- \<lbrace>\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)\<rbrace>
- WHILE \<not>(\<acute>flag1 \<longrightarrow> \<acute>turn=1)
- INV \<lbrace>\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)\<rbrace>
- DO \<lbrace>\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)\<rbrace> SKIP OD;;
- \<lbrace>\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>flag1 \<and> \<acute>after1 \<longrightarrow> \<acute>turn=1)\<rbrace>
- \<acute>flag2:=False
- OD
- \<lbrace>False\<rbrace>
- COEND
+ \<acute>flag1:=False
+ OD
+ \<lbrace>False\<rbrace>
+ \<parallel>
+ \<lbrace>\<not>\<acute>flag2\<rbrace>
+ WHILE True
+ INV \<lbrace>\<not>\<acute>flag2\<rbrace>
+ DO \<lbrace>\<not>\<acute>flag2\<rbrace> \<langle> \<acute>flag2:=True,,\<acute>after2:=False \<rangle>;;
+ \<lbrace>\<acute>flag2 \<and> \<not>\<acute>after2\<rbrace> \<langle> \<acute>turn:=2,,\<acute>after2:=True \<rangle>;;
+ \<lbrace>\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)\<rbrace>
+ WHILE \<not>(\<acute>flag1 \<longrightarrow> \<acute>turn=1)
+ INV \<lbrace>\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)\<rbrace>
+ DO \<lbrace>\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>turn=1 \<or> \<acute>turn=2)\<rbrace> SKIP OD;;
+ \<lbrace>\<acute>flag2 \<and> \<acute>after2 \<and> (\<acute>flag1 \<and> \<acute>after1 \<longrightarrow> \<acute>turn=1)\<rbrace>
+ \<acute>flag2:=False
+ OD
+ \<lbrace>False\<rbrace>
+ COEND
\<lbrace>False\<rbrace>"
apply oghoare
---{* 122 vc *}
+--\<open>122 vc\<close>
apply auto
done
-subsubsection {* Peterson's Algorithm III: A Solution using Semaphores *}
+subsubsection \<open>Peterson's Algorithm III: A Solution using Semaphores\<close>
record Semaphores_mutex =
out :: bool
who :: nat
-lemma Semaphores_mutex:
- "\<parallel>- \<lbrace>i\<noteq>j\<rbrace>
- \<acute>out:=True ,,
- COBEGIN \<lbrace>i\<noteq>j\<rbrace>
- WHILE True INV \<lbrace>i\<noteq>j\<rbrace>
- DO \<lbrace>i\<noteq>j\<rbrace> AWAIT \<acute>out THEN \<acute>out:=False,, \<acute>who:=i END;;
- \<lbrace>\<not>\<acute>out \<and> \<acute>who=i \<and> i\<noteq>j\<rbrace> \<acute>out:=True OD
- \<lbrace>False\<rbrace>
- \<parallel>
- \<lbrace>i\<noteq>j\<rbrace>
- WHILE True INV \<lbrace>i\<noteq>j\<rbrace>
- DO \<lbrace>i\<noteq>j\<rbrace> AWAIT \<acute>out THEN \<acute>out:=False,,\<acute>who:=j END;;
- \<lbrace>\<not>\<acute>out \<and> \<acute>who=j \<and> i\<noteq>j\<rbrace> \<acute>out:=True OD
- \<lbrace>False\<rbrace>
- COEND
+lemma Semaphores_mutex:
+ "\<parallel>- \<lbrace>i\<noteq>j\<rbrace>
+ \<acute>out:=True ,,
+ COBEGIN \<lbrace>i\<noteq>j\<rbrace>
+ WHILE True INV \<lbrace>i\<noteq>j\<rbrace>
+ DO \<lbrace>i\<noteq>j\<rbrace> AWAIT \<acute>out THEN \<acute>out:=False,, \<acute>who:=i END;;
+ \<lbrace>\<not>\<acute>out \<and> \<acute>who=i \<and> i\<noteq>j\<rbrace> \<acute>out:=True OD
+ \<lbrace>False\<rbrace>
+ \<parallel>
+ \<lbrace>i\<noteq>j\<rbrace>
+ WHILE True INV \<lbrace>i\<noteq>j\<rbrace>
+ DO \<lbrace>i\<noteq>j\<rbrace> AWAIT \<acute>out THEN \<acute>out:=False,,\<acute>who:=j END;;
+ \<lbrace>\<not>\<acute>out \<and> \<acute>who=j \<and> i\<noteq>j\<rbrace> \<acute>out:=True OD
+ \<lbrace>False\<rbrace>
+ COEND
\<lbrace>False\<rbrace>"
apply oghoare
---{* 38 vc *}
+--\<open>38 vc\<close>
apply auto
done
-subsubsection {* Peterson's Algorithm III: Parameterized version: *}
+subsubsection \<open>Peterson's Algorithm III: Parameterized version:\<close>
-lemma Semaphores_parameterized_mutex:
- "0<n \<Longrightarrow> \<parallel>- \<lbrace>True\<rbrace>
- \<acute>out:=True ,,
+lemma Semaphores_parameterized_mutex:
+ "0<n \<Longrightarrow> \<parallel>- \<lbrace>True\<rbrace>
+ \<acute>out:=True ,,
COBEGIN
SCHEME [0\<le> i< n]
- \<lbrace>True\<rbrace>
- WHILE True INV \<lbrace>True\<rbrace>
- DO \<lbrace>True\<rbrace> AWAIT \<acute>out THEN \<acute>out:=False,, \<acute>who:=i END;;
+ \<lbrace>True\<rbrace>
+ WHILE True INV \<lbrace>True\<rbrace>
+ DO \<lbrace>True\<rbrace> AWAIT \<acute>out THEN \<acute>out:=False,, \<acute>who:=i END;;
\<lbrace>\<not>\<acute>out \<and> \<acute>who=i\<rbrace> \<acute>out:=True OD
- \<lbrace>False\<rbrace>
+ \<lbrace>False\<rbrace>
COEND
- \<lbrace>False\<rbrace>"
+ \<lbrace>False\<rbrace>"
apply oghoare
---{* 20 vc *}
+--\<open>20 vc\<close>
apply auto
done
-subsubsection{* The Ticket Algorithm *}
+subsubsection\<open>The Ticket Algorithm\<close>
record Ticket_mutex =
num :: nat
nextv :: nat
turn :: "nat list"
- index :: nat
+ 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
+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>- \<lbrace>n=length \<acute>turn\<rbrace>
+ \<Longrightarrow> \<parallel>- \<lbrace>n=length \<acute>turn\<rbrace>
\<acute>index:= 0,,
- WHILE \<acute>index < n INV \<lbrace>n=length \<acute>turn \<and> (\<forall>i<\<acute>index. \<acute>turn!i=0)\<rbrace>
+ WHILE \<acute>index < n INV \<lbrace>n=length \<acute>turn \<and> (\<forall>i<\<acute>index. \<acute>turn!i=0)\<rbrace>
DO \<acute>turn:= \<acute>turn[\<acute>index:=0],, \<acute>index:=\<acute>index +1 OD,,
- \<acute>num:=1 ,, \<acute>nextv:=1 ,,
+ \<acute>num:=1 ,, \<acute>nextv:=1 ,,
COBEGIN
SCHEME [0\<le> i< n]
- \<lbrace>\<acute>I\<rbrace>
- WHILE True INV \<lbrace>\<acute>I\<rbrace>
- DO \<lbrace>\<acute>I\<rbrace> \<langle> \<acute>turn :=\<acute>turn[i:=\<acute>num],, \<acute>num:=\<acute>num+1 \<rangle>;;
+ \<lbrace>\<acute>I\<rbrace>
+ WHILE True INV \<lbrace>\<acute>I\<rbrace>
+ DO \<lbrace>\<acute>I\<rbrace> \<langle> \<acute>turn :=\<acute>turn[i:=\<acute>num],, \<acute>num:=\<acute>num+1 \<rangle>;;
\<lbrace>\<acute>I\<rbrace> WAIT \<acute>turn!i=\<acute>nextv END;;
\<lbrace>\<acute>I \<and> \<acute>turn!i=\<acute>nextv\<rbrace> \<acute>nextv:=\<acute>nextv+1
OD
- \<lbrace>False\<rbrace>
+ \<lbrace>False\<rbrace>
COEND
- \<lbrace>False\<rbrace>"
+ \<lbrace>False\<rbrace>"
apply oghoare
---{* 35 vc *}
+--\<open>35 vc\<close>
apply simp_all
---{* 21 vc *}
-apply(tactic {* ALLGOALS (clarify_tac @{context}) *})
---{* 11 vc *}
+--\<open>21 vc\<close>
+apply(tactic \<open>ALLGOALS (clarify_tac @{context})\<close>)
+--\<open>11 vc\<close>
apply simp_all
-apply(tactic {* ALLGOALS (clarify_tac @{context}) *})
---{* 10 subgoals left *}
+apply(tactic \<open>ALLGOALS (clarify_tac @{context})\<close>)
+--\<open>10 subgoals left\<close>
apply(erule less_SucE)
apply simp
apply simp
---{* 9 subgoals left *}
+--\<open>9 subgoals left\<close>
apply(case_tac "i=k")
apply force
apply simp
apply(case_tac "i=l")
apply force
apply force
---{* 8 subgoals left *}
+--\<open>8 subgoals left\<close>
prefer 8
apply force
apply force
---{* 6 subgoals left *}
+--\<open>6 subgoals left\<close>
prefer 6
apply(erule_tac x=i in allE)
apply fastforce
---{* 5 subgoals left *}
+--\<open>5 subgoals left\<close>
prefer 5
apply(case_tac [!] "j=k")
---{* 10 subgoals left *}
+--\<open>10 subgoals left\<close>
apply simp_all
apply(erule_tac x=k in allE)
apply force
---{* 9 subgoals left *}
+--\<open>9 subgoals left\<close>
apply(case_tac "j=l")
apply simp
apply(erule_tac x=k in allE)
@@ -211,7 +211,7 @@
apply(erule_tac x=k in allE)
apply(erule_tac x=l in allE)
apply force
---{* 8 subgoals left *}
+--\<open>8 subgoals left\<close>
apply force
apply(case_tac "j=l")
apply simp
@@ -220,21 +220,21 @@
apply force
apply force
apply force
---{* 5 subgoals left *}
+--\<open>5 subgoals left\<close>
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 *}
+--\<open>3 subgoals left\<close>
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 *}
+--\<open>1 subgoals left\<close>
apply(erule_tac x=k in allE)
apply(erule_tac x=l in allE)
apply(case_tac "j=l")
@@ -242,11 +242,11 @@
apply force
done
-subsection{* Parallel Zero Search *}
+subsection\<open>Parallel Zero Search\<close>
-text {* Synchronized Zero Search. Zero-6 *}
+text \<open>Synchronized Zero Search. Zero-6\<close>
-text {*Apt and Olderog. "Verification of sequential and concurrent Programs" page 294: *}
+text \<open>Apt and Olderog. "Verification of sequential and concurrent Programs" page 294:\<close>
record Zero_search =
turn :: nat
@@ -254,87 +254,87 @@
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>- \<lbrace>\<exists> u. f(u)=0\<rbrace>
- \<acute>turn:=1,, \<acute>found:= False,,
- \<acute>x:=a,, \<acute>y:=a+1 ,,
- COBEGIN \<lbrace>\<acute>I1\<rbrace>
- WHILE \<not>\<acute>found
- INV \<lbrace>\<acute>I1\<rbrace>
- DO \<lbrace>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)\<rbrace>
- WAIT \<acute>turn=1 END;;
- \<lbrace>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)\<rbrace>
- \<acute>turn:=2;;
- \<lbrace>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)\<rbrace>
- \<langle> \<acute>x:=\<acute>x+1,,
- IF f(\<acute>x)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>
- OD;;
- \<lbrace>\<acute>I1 \<and> \<acute>found\<rbrace>
- \<acute>turn:=2
- \<lbrace>\<acute>I1 \<and> \<acute>found\<rbrace>
- \<parallel>
- \<lbrace>\<acute>I2\<rbrace>
- WHILE \<not>\<acute>found
- INV \<lbrace>\<acute>I2\<rbrace>
- DO \<lbrace>\<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)\<rbrace>
- WAIT \<acute>turn=2 END;;
- \<lbrace>\<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)\<rbrace>
- \<acute>turn:=1;;
- \<lbrace>\<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)\<rbrace>
- \<langle> \<acute>y:=(\<acute>y - 1),,
- IF f(\<acute>y)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>
- OD;;
- \<lbrace>\<acute>I2 \<and> \<acute>found\<rbrace>
- \<acute>turn:=1
- \<lbrace>\<acute>I2 \<and> \<acute>found\<rbrace>
- COEND
+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>- \<lbrace>\<exists> u. f(u)=0\<rbrace>
+ \<acute>turn:=1,, \<acute>found:= False,,
+ \<acute>x:=a,, \<acute>y:=a+1 ,,
+ COBEGIN \<lbrace>\<acute>I1\<rbrace>
+ WHILE \<not>\<acute>found
+ INV \<lbrace>\<acute>I1\<rbrace>
+ DO \<lbrace>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)\<rbrace>
+ WAIT \<acute>turn=1 END;;
+ \<lbrace>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)\<rbrace>
+ \<acute>turn:=2;;
+ \<lbrace>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)\<rbrace>
+ \<langle> \<acute>x:=\<acute>x+1,,
+ IF f(\<acute>x)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>
+ OD;;
+ \<lbrace>\<acute>I1 \<and> \<acute>found\<rbrace>
+ \<acute>turn:=2
+ \<lbrace>\<acute>I1 \<and> \<acute>found\<rbrace>
+ \<parallel>
+ \<lbrace>\<acute>I2\<rbrace>
+ WHILE \<not>\<acute>found
+ INV \<lbrace>\<acute>I2\<rbrace>
+ DO \<lbrace>\<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)\<rbrace>
+ WAIT \<acute>turn=2 END;;
+ \<lbrace>\<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)\<rbrace>
+ \<acute>turn:=1;;
+ \<lbrace>\<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)\<rbrace>
+ \<langle> \<acute>y:=(\<acute>y - 1),,
+ IF f(\<acute>y)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>
+ OD;;
+ \<lbrace>\<acute>I2 \<and> \<acute>found\<rbrace>
+ \<acute>turn:=1
+ \<lbrace>\<acute>I2 \<and> \<acute>found\<rbrace>
+ COEND
\<lbrace>f(\<acute>x)=0 \<or> f(\<acute>y)=0\<rbrace>"
apply oghoare
---{* 98 verification conditions *}
-apply auto
---{* auto takes about 3 minutes !! *}
+--\<open>98 verification conditions\<close>
+apply auto
+--\<open>auto takes about 3 minutes !!\<close>
done
-text {* Easier Version: without AWAIT. Apt and Olderog. page 256: *}
+text \<open>Easier Version: without AWAIT. Apt and Olderog. page 256:\<close>
-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>- \<lbrace>\<exists>u. f(u)=0\<rbrace>
- \<acute>found:= False,,
- \<acute>x:=a,, \<acute>y:=a+1,,
- COBEGIN \<lbrace>\<acute>I1\<rbrace>
- WHILE \<not>\<acute>found
- INV \<lbrace>\<acute>I1\<rbrace>
- DO \<lbrace>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)\<rbrace>
- \<langle> \<acute>x:=\<acute>x+1,,IF f(\<acute>x)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>
- OD
- \<lbrace>\<acute>I1 \<and> \<acute>found\<rbrace>
- \<parallel>
- \<lbrace>\<acute>I2\<rbrace>
- WHILE \<not>\<acute>found
- INV \<lbrace>\<acute>I2\<rbrace>
- DO \<lbrace>\<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)\<rbrace>
- \<langle> \<acute>y:=(\<acute>y - 1),,IF f(\<acute>y)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>
- OD
- \<lbrace>\<acute>I2 \<and> \<acute>found\<rbrace>
- COEND
+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>- \<lbrace>\<exists>u. f(u)=0\<rbrace>
+ \<acute>found:= False,,
+ \<acute>x:=a,, \<acute>y:=a+1,,
+ COBEGIN \<lbrace>\<acute>I1\<rbrace>
+ WHILE \<not>\<acute>found
+ INV \<lbrace>\<acute>I1\<rbrace>
+ DO \<lbrace>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)\<rbrace>
+ \<langle> \<acute>x:=\<acute>x+1,,IF f(\<acute>x)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>
+ OD
+ \<lbrace>\<acute>I1 \<and> \<acute>found\<rbrace>
+ \<parallel>
+ \<lbrace>\<acute>I2\<rbrace>
+ WHILE \<not>\<acute>found
+ INV \<lbrace>\<acute>I2\<rbrace>
+ DO \<lbrace>\<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)\<rbrace>
+ \<langle> \<acute>y:=(\<acute>y - 1),,IF f(\<acute>y)=0 THEN \<acute>found:=True ELSE SKIP FI\<rangle>
+ OD
+ \<lbrace>\<acute>I2 \<and> \<acute>found\<rbrace>
+ COEND
\<lbrace>f(\<acute>x)=0 \<or> f(\<acute>y)=0\<rbrace>"
apply oghoare
---{* 20 vc *}
+--\<open>20 vc\<close>
apply auto
---{* auto takes aprox. 2 minutes. *}
+--\<open>auto takes aprox. 2 minutes.\<close>
done
-subsection {* Producer/Consumer *}
+subsection \<open>Producer/Consumer\<close>
-subsubsection {* Previous lemmas *}
+subsubsection \<open>Previous lemmas\<close>
lemma nat_lemma2: "\<lbrakk> b = m*(n::nat) + t; a = s*n + u; t=u; b-a < n \<rbrakk> \<Longrightarrow> m \<le> s"
proof -
@@ -368,7 +368,7 @@
done
-subsubsection {* Producer/Consumer Algorithm *}
+subsubsection \<open>Producer/Consumer Algorithm\<close>
record Producer_consumer =
ins :: nat
@@ -380,104 +380,104 @@
buffer :: "nat list"
b :: "nat list"
-text {* The whole proof takes aprox. 4 minutes. *}
+text \<open>The whole proof takes aprox. 4 minutes.\<close>
-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> ;
+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>- \<lbrace>\<acute>INIT\<rbrace>
+ p2 = \<guillemotleft>\<acute>I2 \<and> \<acute>lj=\<acute>outs\<guillemotright> \<rbrakk> \<Longrightarrow>
+ \<parallel>- \<lbrace>\<acute>INIT\<rbrace>
\<acute>ins:=0,, \<acute>outs:=0,, \<acute>li:=0,, \<acute>lj:=0,,
- COBEGIN \<lbrace>\<acute>p1 \<and> \<acute>INIT\<rbrace>
- WHILE \<acute>li <length a
- INV \<lbrace>\<acute>p1 \<and> \<acute>INIT\<rbrace>
- DO \<lbrace>\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a\<rbrace>
- \<acute>vx:= (a ! \<acute>li);;
- \<lbrace>\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a \<and> \<acute>vx=(a ! \<acute>li)\<rbrace>
- WAIT \<acute>ins-\<acute>outs < length \<acute>buffer END;;
- \<lbrace>\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a \<and> \<acute>vx=(a ! \<acute>li)
- \<and> \<acute>ins-\<acute>outs < length \<acute>buffer\<rbrace>
- \<acute>buffer:=(list_update \<acute>buffer (\<acute>ins mod (length \<acute>buffer)) \<acute>vx);;
- \<lbrace>\<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\<rbrace>
- \<acute>ins:=\<acute>ins+1;;
- \<lbrace>\<acute>I1 \<and> \<acute>INIT \<and> (\<acute>li+1)=\<acute>ins \<and> \<acute>li<length a\<rbrace>
- \<acute>li:=\<acute>li+1
- OD
- \<lbrace>\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li=length a\<rbrace>
- \<parallel>
- \<lbrace>\<acute>p2 \<and> \<acute>INIT\<rbrace>
- WHILE \<acute>lj < length a
- INV \<lbrace>\<acute>p2 \<and> \<acute>INIT\<rbrace>
- DO \<lbrace>\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>INIT\<rbrace>
- WAIT \<acute>outs<\<acute>ins END;;
- \<lbrace>\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>outs<\<acute>ins \<and> \<acute>INIT\<rbrace>
- \<acute>vy:=(\<acute>buffer ! (\<acute>outs mod (length \<acute>buffer)));;
- \<lbrace>\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>outs<\<acute>ins \<and> \<acute>vy=(a ! \<acute>lj) \<and> \<acute>INIT\<rbrace>
- \<acute>outs:=\<acute>outs+1;;
- \<lbrace>\<acute>I2 \<and> (\<acute>lj+1)=\<acute>outs \<and> \<acute>lj<length a \<and> \<acute>vy=(a ! \<acute>lj) \<and> \<acute>INIT\<rbrace>
- \<acute>b:=(list_update \<acute>b \<acute>lj \<acute>vy);;
- \<lbrace>\<acute>I2 \<and> (\<acute>lj+1)=\<acute>outs \<and> \<acute>lj<length a \<and> (a ! \<acute>lj)=(\<acute>b ! \<acute>lj) \<and> \<acute>INIT\<rbrace>
- \<acute>lj:=\<acute>lj+1
- OD
- \<lbrace>\<acute>p2 \<and> \<acute>lj=length a \<and> \<acute>INIT\<rbrace>
- COEND
+ COBEGIN \<lbrace>\<acute>p1 \<and> \<acute>INIT\<rbrace>
+ WHILE \<acute>li <length a
+ INV \<lbrace>\<acute>p1 \<and> \<acute>INIT\<rbrace>
+ DO \<lbrace>\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a\<rbrace>
+ \<acute>vx:= (a ! \<acute>li);;
+ \<lbrace>\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a \<and> \<acute>vx=(a ! \<acute>li)\<rbrace>
+ WAIT \<acute>ins-\<acute>outs < length \<acute>buffer END;;
+ \<lbrace>\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li<length a \<and> \<acute>vx=(a ! \<acute>li)
+ \<and> \<acute>ins-\<acute>outs < length \<acute>buffer\<rbrace>
+ \<acute>buffer:=(list_update \<acute>buffer (\<acute>ins mod (length \<acute>buffer)) \<acute>vx);;
+ \<lbrace>\<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\<rbrace>
+ \<acute>ins:=\<acute>ins+1;;
+ \<lbrace>\<acute>I1 \<and> \<acute>INIT \<and> (\<acute>li+1)=\<acute>ins \<and> \<acute>li<length a\<rbrace>
+ \<acute>li:=\<acute>li+1
+ OD
+ \<lbrace>\<acute>p1 \<and> \<acute>INIT \<and> \<acute>li=length a\<rbrace>
+ \<parallel>
+ \<lbrace>\<acute>p2 \<and> \<acute>INIT\<rbrace>
+ WHILE \<acute>lj < length a
+ INV \<lbrace>\<acute>p2 \<and> \<acute>INIT\<rbrace>
+ DO \<lbrace>\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>INIT\<rbrace>
+ WAIT \<acute>outs<\<acute>ins END;;
+ \<lbrace>\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>outs<\<acute>ins \<and> \<acute>INIT\<rbrace>
+ \<acute>vy:=(\<acute>buffer ! (\<acute>outs mod (length \<acute>buffer)));;
+ \<lbrace>\<acute>p2 \<and> \<acute>lj<length a \<and> \<acute>outs<\<acute>ins \<and> \<acute>vy=(a ! \<acute>lj) \<and> \<acute>INIT\<rbrace>
+ \<acute>outs:=\<acute>outs+1;;
+ \<lbrace>\<acute>I2 \<and> (\<acute>lj+1)=\<acute>outs \<and> \<acute>lj<length a \<and> \<acute>vy=(a ! \<acute>lj) \<and> \<acute>INIT\<rbrace>
+ \<acute>b:=(list_update \<acute>b \<acute>lj \<acute>vy);;
+ \<lbrace>\<acute>I2 \<and> (\<acute>lj+1)=\<acute>outs \<and> \<acute>lj<length a \<and> (a ! \<acute>lj)=(\<acute>b ! \<acute>lj) \<and> \<acute>INIT\<rbrace>
+ \<acute>lj:=\<acute>lj+1
+ OD
+ \<lbrace>\<acute>p2 \<and> \<acute>lj=length a \<and> \<acute>INIT\<rbrace>
+ COEND
\<lbrace> \<forall>k<length a. (a ! k)=(\<acute>b ! k)\<rbrace>"
apply oghoare
---{* 138 vc *}
-apply(tactic {* ALLGOALS (clarify_tac @{context}) *})
---{* 112 subgoals left *}
+--\<open>138 vc\<close>
+apply(tactic \<open>ALLGOALS (clarify_tac @{context})\<close>)
+--\<open>112 subgoals left\<close>
apply(simp_all (no_asm))
---{* 97 subgoals left *}
-apply(tactic {*ALLGOALS (conjI_Tac (K all_tac)) *})
---{* 930 subgoals left *}
-apply(tactic {* ALLGOALS (clarify_tac @{context}) *})
---{* 99 subgoals left *}
+--\<open>97 subgoals left\<close>
+apply(tactic \<open>ALLGOALS (conjI_Tac (K all_tac))\<close>)
+--\<open>930 subgoals left\<close>
+apply(tactic \<open>ALLGOALS (clarify_tac @{context})\<close>)
+--\<open>99 subgoals left\<close>
apply(simp_all (*asm_lr*) only:length_0_conv [THEN sym])
---{* 20 subgoals left *}
+--\<open>20 subgoals left\<close>
apply (simp_all (*asm_lr*) del:length_0_conv length_greater_0_conv add: nth_list_update mod_lemma)
---{* 9 subgoals left *}
+--\<open>9 subgoals left\<close>
apply (force simp add:less_Suc_eq)
apply(hypsubst_thin, drule sym)
apply (force simp add:less_Suc_eq)+
done
-subsection {* Parameterized Examples *}
+subsection \<open>Parameterized Examples\<close>
-subsubsection {* Set Elements of an Array to Zero *}
+subsubsection \<open>Set Elements of an Array to Zero\<close>
record Example1 =
a :: "nat \<Rightarrow> nat"
-lemma Example1:
+lemma Example1:
"\<parallel>- \<lbrace>True\<rbrace>
- COBEGIN SCHEME [0\<le>i<n] \<lbrace>True\<rbrace> \<acute>a:=\<acute>a (i:=0) \<lbrace>\<acute>a i=0\<rbrace> COEND
+ COBEGIN SCHEME [0\<le>i<n] \<lbrace>True\<rbrace> \<acute>a:=\<acute>a (i:=0) \<lbrace>\<acute>a i=0\<rbrace> COEND
\<lbrace>\<forall>i < n. \<acute>a i = 0\<rbrace>"
apply oghoare
apply simp_all
done
-text {* Same example with lists as auxiliary variables. *}
+text \<open>Same example with lists as auxiliary variables.\<close>
record Example1_list =
A :: "nat list"
-lemma Example1_list:
- "\<parallel>- \<lbrace>n < length \<acute>A\<rbrace>
- COBEGIN
- SCHEME [0\<le>i<n] \<lbrace>n < length \<acute>A\<rbrace> \<acute>A:=\<acute>A[i:=0] \<lbrace>\<acute>A!i=0\<rbrace>
- COEND
+lemma Example1_list:
+ "\<parallel>- \<lbrace>n < length \<acute>A\<rbrace>
+ COBEGIN
+ SCHEME [0\<le>i<n] \<lbrace>n < length \<acute>A\<rbrace> \<acute>A:=\<acute>A[i:=0] \<lbrace>\<acute>A!i=0\<rbrace>
+ COEND
\<lbrace>\<forall>i < n. \<acute>A!i = 0\<rbrace>"
apply oghoare
apply force+
done
-subsubsection {* Increment a Variable in Parallel *}
+subsubsection \<open>Increment a Variable in Parallel\<close>
-text {* First some lemmas about summation properties. *}
+text \<open>First some lemmas about summation properties.\<close>
(*
lemma Example2_lemma1: "!!b. j<n \<Longrightarrow> (\<Sum>i::nat<n. b i) = (0::nat) \<Longrightarrow> b j = 0 "
apply(induct n)
@@ -485,7 +485,7 @@
apply(force simp add: less_Suc_eq)
done
*)
-lemma Example2_lemma2_aux: "!!b. j<n \<Longrightarrow>
+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)
@@ -497,13 +497,13 @@
apply arith
done
-lemma Example2_lemma2_aux2:
+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(induct j)
apply simp_all
done
-lemma Example2_lemma2:
+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)
@@ -520,22 +520,22 @@
done
-record Example2 =
- c :: "nat \<Rightarrow> nat"
+record Example2 =
+ c :: "nat \<Rightarrow> nat"
x :: nat
-lemma Example_2: "0<n \<Longrightarrow>
- \<parallel>- \<lbrace>\<acute>x=0 \<and> (\<Sum>i=0..<n. \<acute>c i)=0\<rbrace>
- COBEGIN
- SCHEME [0\<le>i<n]
- \<lbrace>\<acute>x=(\<Sum>i=0..<n. \<acute>c i) \<and> \<acute>c i=0\<rbrace>
+lemma Example_2: "0<n \<Longrightarrow>
+ \<parallel>- \<lbrace>\<acute>x=0 \<and> (\<Sum>i=0..<n. \<acute>c i)=0\<rbrace>
+ COBEGIN
+ SCHEME [0\<le>i<n]
+ \<lbrace>\<acute>x=(\<Sum>i=0..<n. \<acute>c i) \<and> \<acute>c i=0\<rbrace>
\<langle> \<acute>x:=\<acute>x+(Suc 0),, \<acute>c:=\<acute>c (i:=(Suc 0)) \<rangle>
\<lbrace>\<acute>x=(\<Sum>i=0..<n. \<acute>c i) \<and> \<acute>c i=(Suc 0)\<rbrace>
- COEND
+ COEND
\<lbrace>\<acute>x=n\<rbrace>"
apply oghoare
apply (simp_all cong del: setsum.strong_cong)
-apply (tactic {* ALLGOALS (clarify_tac @{context}) *})
+apply (tactic \<open>ALLGOALS (clarify_tac @{context})\<close>)
apply (simp_all cong del: setsum.strong_cong)
apply(erule (1) Example2_lemma2)
apply(erule (1) Example2_lemma2)
--- a/src/HOL/Hoare_Parallel/OG_Hoare.thy Sat Dec 27 19:51:55 2014 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Hoare.thy Sat Dec 27 20:32:06 2014 +0100
@@ -1,4 +1,4 @@
-section {* The Proof System *}
+section \<open>The Proof System\<close>
theory OG_Hoare imports OG_Tran begin
@@ -8,14 +8,14 @@
| "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}"
+| "assertions (AnnAwait r b c) = {r}"
primrec atomics :: "'a ann_com \<Rightarrow> ('a assn \<times> 'a com) set" where
"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 (AnnWhile r b i c) = atomics c"
| "atomics (AnnAwait r b c) = {(r \<inter> b, c)}"
primrec com :: "'a ann_triple_op \<Rightarrow> 'a ann_com_op" where
@@ -25,12 +25,12 @@
"post (c, q) = q"
definition interfree_aux :: "('a ann_com_op \<times> 'a assn \<times> 'a ann_com_op) \<Rightarrow> bool" where
- "interfree_aux \<equiv> \<lambda>(co, q, co'). co'= None \<or>
+ "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)))"
-definition interfree :: "(('a ann_triple_op) list) \<Rightarrow> bool" where
- "interfree Ts \<equiv> \<forall>i j. i < length Ts \<and> j < length Ts \<and> i \<noteq> j \<longrightarrow>
+definition interfree :: "(('a ann_triple_op) list) \<Rightarrow> bool" where
+ "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
@@ -40,26 +40,26 @@
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>
+
+| 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>
+
+| 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
+ \<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"
@@ -68,7 +68,7 @@
| Conseq: "\<lbrakk> p' \<subseteq> p; \<parallel>- p c q ; q \<subseteq> q' \<rbrakk> \<Longrightarrow> \<parallel>- p' c q'"
-section {* Soundness *}
+section \<open>Soundness\<close>
(* 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
@@ -95,9 +95,9 @@
lemmas While = oghoare_ann_hoare.While
lemmas Conseq = oghoare_ann_hoare.Conseq
-subsection {* Soundness of the System for Atomic Programs *}
+subsection \<open>Soundness of the System for Atomic Programs\<close>
-lemma Basic_ntran [rule_format]:
+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))
@@ -109,33 +109,33 @@
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 (blast dest: L3_5i)
apply(simp add: SEM_def sem_def id_def)
-apply (blast dest: Basic_ntran rtrancl_imp_UN_relpow)
+apply (blast dest: Basic_ntran rtrancl_imp_UN_relpow)
done
-lemma atom_hoare_sound [rule_format]:
+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*}
+--\<open>Basic\<close>
apply(simp add: SEM_def sem_def)
apply(fast dest: rtrancl_imp_UN_relpow Basic_ntran)
---{* Seq *}
+--\<open>Seq\<close>
apply(rule impI)
apply(rule subset_trans)
prefer 2 apply simp
apply(simp add: L3_5ii L3_5i)
---{* Cond *}
+--\<open>Cond\<close>
apply(simp add: L3_5iv)
---{* While *}
- apply (force simp add: L3_5v dest: SEM_fwhile)
---{* Conseq *}
+--\<open>While\<close>
+ apply (force simp add: L3_5v dest: SEM_fwhile)
+--\<open>Conseq\<close>
apply(force simp add: SEM_def sem_def)
done
-
-subsection {* Soundness of the System for Component Programs *}
+
+subsection \<open>Soundness of the System for Component Programs\<close>
inductive_cases ann_transition_cases:
"(None,s) -1\<rightarrow> (c', s')"
@@ -146,17 +146,17 @@
"(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:*}
+text \<open>Strong Soundness for Component Programs:\<close>
lemma ann_hoare_case_analysis [rule_format]: "\<turnstile> C q' \<longrightarrow>
- ((\<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 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)))"
apply(rule ann_hoare_induct)
apply simp_all
@@ -169,16 +169,16 @@
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>
+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 simp_all
+--\<open>Basic\<close>
apply clarify
apply(frule ann_hoare_case_analysis)
apply force
---{* Seq *}
+--\<open>Seq\<close>
apply clarify
apply(frule ann_hoare_case_analysis,simp)
apply(fast intro: AnnConseq)
@@ -189,21 +189,21 @@
apply force
apply(rule AnnSeq,simp)
apply(fast intro: AnnConseq)
---{* Cond1 *}
+--\<open>Cond1\<close>
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 *}
+--\<open>Cond2\<close>
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 *}
+--\<open>While\<close>
apply clarify
apply(frule ann_hoare_case_analysis,simp)
apply force
@@ -214,17 +214,17 @@
apply simp
apply(rule AnnWhile)
apply simp_all
---{* Await *}
+--\<open>Await\<close>
apply(frule ann_hoare_case_analysis,simp)
apply clarify
apply(drule atom_hoare_sound)
- apply simp
+ 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>
+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
@@ -235,7 +235,7 @@
apply simp_all
done
-lemma Strong_Soundness: "\<lbrakk> (Some c, s)-*\<rightarrow>(co, t); s \<in> pre c; \<turnstile> c q \<rbrakk>
+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
@@ -247,9 +247,9 @@
apply simp_all
done
-subsection {* Soundness of the System for Parallel Programs *}
+subsection \<open>Soundness of the System for Parallel Programs\<close>
-lemma Parallel_length_post_P1: "(Parallel Ts,s) -P1\<rightarrow> (R', t) \<Longrightarrow>
+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)
@@ -259,8 +259,8 @@
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>
+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)
@@ -275,7 +275,7 @@
apply auto
done
-lemma interfree_aux1 [rule_format]:
+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)
@@ -285,13 +285,13 @@
apply force+
done
-lemma interfree_aux2 [rule_format]:
+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;
+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
@@ -302,10 +302,10 @@
apply(force elim: interfree_aux2 simp add:nth_list_update)
done
-text {* Strong Soundness Theorem for Parallel Programs:*}
+text \<open>Strong Soundness Theorem for Parallel Programs:\<close>
-lemma Parallel_Strong_Soundness_Seq_aux:
- "\<lbrakk>interfree Ts; i<length Ts; com(Ts ! i) = Some(AnnSeq c0 c1) \<rbrakk>
+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
@@ -317,14 +317,14 @@
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)))
+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
@@ -335,20 +335,20 @@
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)
+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 *}
+--\<open>Basic\<close>
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)
@@ -362,15 +362,15 @@
apply clarify
apply simp
apply(erule_tac x="pre y" in ballE)
- apply(force intro: converse_rtrancl_into_rtrancl
+ 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 *}
+--\<open>Seqs\<close>
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 *}
+--\<open>Await\<close>
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)
@@ -378,28 +378,28 @@
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
+ 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
+ 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]:
+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)
+ \<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 *}
+--\<open>Base\<close>
apply force
---{* Induction step *}
+--\<open>Induction step\<close>
apply clarify
apply(drule Parallel_length_post_PStar)
apply clarify
@@ -419,9 +419,9 @@
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>
+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+
@@ -431,7 +431,7 @@
apply (unfold com_validity_def)
apply(rule oghoare_induct)
apply(rule TrueI)+
---{* Parallel *}
+--\<open>Parallel\<close>
apply(simp add: SEM_def sem_def)
apply(clarify, rename_tac x y i Ts')
apply(frule Parallel_length_post_PStar)
@@ -445,19 +445,19 @@
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 *}
+--\<open>Basic\<close>
apply(simp add: SEM_def sem_def)
apply(force dest: rtrancl_imp_UN_relpow Basic_ntran)
---{* Seq *}
+--\<open>Seq\<close>
apply(rule subset_trans)
prefer 2 apply assumption
apply(simp add: L3_5ii L3_5i)
---{* Cond *}
+--\<open>Cond\<close>
apply(simp add: L3_5iv)
---{* While *}
+--\<open>While\<close>
apply(simp add: L3_5v)
- apply (blast dest: SEM_fwhile)
---{* Conseq *}
+ apply (blast dest: SEM_fwhile)
+--\<open>Conseq\<close>
apply(auto simp add: SEM_def sem_def)
done
--- a/src/HOL/Hoare_Parallel/OG_Syntax.thy Sat Dec 27 19:51:55 2014 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Syntax.thy Sat Dec 27 20:32:06 2014 +0100
@@ -2,8 +2,8 @@
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}. *}
+text\<open>Syntax for commands and for assertions and boolean expressions in
+ commands @{text com} and annotated commands @{text ann_com}.\<close>
abbreviation Skip :: "'a com" ("SKIP" 63)
where "SKIP \<equiv> Basic id"
@@ -28,14 +28,14 @@
("_ //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"
+ "_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)
- "_Cond" :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> 'a com \<Rightarrow> 'a com"
+ "_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"
@@ -55,7 +55,7 @@
"r AWAIT b THEN c END" \<rightharpoonup> "CONST AnnAwait r \<lbrace>b\<rbrace> c"
"r \<langle>c\<rangle>" \<rightleftharpoons> "r AWAIT CONST True THEN c END"
"r WAIT b END" \<rightleftharpoons> "r AWAIT b THEN SKIP END"
-
+
nonterminal prgs
syntax
@@ -63,7 +63,7 @@
"_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"
+ "_prg_scheme" :: "['a, 'a, 'a, 'a, 'a] \<Rightarrow> prgs"
("SCHEME [_ \<le> _ < _] _// _" [0,0,0,60, 90] 57)
translations
@@ -73,7 +73,7 @@
"_prg_scheme j i k c q" \<rightleftharpoons> "CONST map (\<lambda>i. (CONST Some c, q)) [j..<k]"
-print_translation {*
+print_translation \<open>
let
fun quote_tr' f (t :: ts) =
Term.list_comb (f $ Syntax_Trans.quote_tr' @{syntax_const "_antiquote"} t, ts)
@@ -123,6 +123,6 @@
(@{const_syntax AnnCond1}, K (annbexp_tr' @{syntax_const "_AnnCond1"})),
(@{const_syntax AnnCond2}, K (annbexp_tr' @{syntax_const "_AnnCond2"}))]
end;
-*}
+\<close>
end
\ No newline at end of file
--- a/src/HOL/Hoare_Parallel/OG_Tactics.thy Sat Dec 27 19:51:55 2014 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Tactics.thy Sat Dec 27 20:32:06 2014 +0100
@@ -1,4 +1,4 @@
-section {* Generation of Verification Conditions *}
+section \<open>Generation of Verification Conditions\<close>
theory OG_Tactics
imports OG_Hoare
@@ -7,15 +7,15 @@
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>
+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
+prefer 2
apply fast
apply assumption+
done
@@ -39,8 +39,8 @@
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>
+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)
@@ -49,17 +49,17 @@
apply force+
done
-lemma WhileRule: "\<lbrakk> p \<subseteq> i; \<parallel>- (i \<inter> b) c i ; (i \<inter> (-b)) \<subseteq> q \<rbrakk>
+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
+text \<open>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}"}: *}
+command where the boolean condition is @{text "{s. True}"}:\<close>
lemma AnnatomRule:
"\<lbrakk> atom_com(c); \<parallel>- r c q \<rbrakk> \<Longrightarrow> \<turnstile> (AnnAwait r {s. True} c) q"
@@ -81,93 +81,93 @@
apply simp
done
-text {* Lemmata to avoid using the definition of @{text
+text \<open>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: *}
+@{text interfree} by splitting it into different cases:\<close>
lemma interfree_aux_rule1: "interfree_aux(co, q, None)"
by(simp add:interfree_aux_def)
-lemma interfree_aux_rule2:
+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:
+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>
+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>
+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>
+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>
+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>
+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>
+
+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:
+
+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>
+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>
+ "\<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:
+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)
+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:
+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)
+by(simp add: interfree_aux_def oghoare_sound)
-lemma AnnAwait_atomics:
+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)
@@ -178,21 +178,21 @@
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>
+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)
+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:
+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
@@ -202,8 +202,8 @@
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))
+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)
@@ -219,14 +219,14 @@
apply(case_tac i,simp+)
done
-lemma MapAnnMap:
+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
+ \<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)
@@ -234,7 +234,7 @@
done
(*
lemma ParamParallelRule:
- "\<lbrakk> \<forall>k<n. \<turnstile> (c k) (Q k);
+ "\<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)
@@ -252,10 +252,10 @@
done
*)
-text {* The following are some useful lemmas and simplification
+text \<open>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. *}
+transformation.\<close>
lemma Compl_Collect: "-(Collect b) = {x. \<not>(b x)}"
by fast
@@ -273,7 +273,7 @@
Collect_mem_eq ball_simps option.simps primrecdef_list
lemmas ParallelConseq_list = INTER_eq Collect_conj_eq length_map length_upt length_append
-ML {*
+ML \<open>
fun before_interfree_simp_tac ctxt =
simp_tac (put_simpset HOL_basic_ss ctxt addsimps [@{thm com.simps}, @{thm post.simps}])
@@ -284,50 +284,50 @@
fun ParallelConseq ctxt =
simp_tac (put_simpset HOL_basic_ss ctxt
addsimps @{thms ParallelConseq_list} @ @{thms my_simp_list})
-*}
+\<close>
-text {* The following tactic applies @{text tac} to each conjunct in a
+text \<open>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: *}
+@{text n} subgoals, one for each conjunct:\<close>
-ML {*
+ML \<open>
fun conjI_Tac tac i st = st |>
( (EVERY [rtac conjI i,
conjI_Tac tac (i+1),
tac i]) ORELSE (tac i) )
-*}
+\<close>
-subsubsection {* Tactic for the generation of the verification conditions *}
+subsubsection \<open>Tactic for the generation of the verification conditions\<close>
-text {* The tactic basically uses two subtactics:
+text \<open>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
+\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
+\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
+ 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}
-*}
+\<close>
-ML {*
+ML \<open>
fun WlpTac ctxt i = (rtac (@{thm SeqRule}) i) THEN (HoareRuleTac ctxt false (i+1))
-and HoareRuleTac ctxt precond i st = st |>
+and HoareRuleTac ctxt precond i st = st |>
( (WlpTac ctxt i THEN HoareRuleTac ctxt precond i)
ORELSE
(FIRST[rtac (@{thm SkipRule}) i,
@@ -335,7 +335,7 @@
EVERY[rtac (@{thm ParallelConseqRule}) i,
ParallelConseq ctxt (i+2),
ParallelTac ctxt (i+1),
- ParallelConseq ctxt i],
+ ParallelConseq ctxt i],
EVERY[rtac (@{thm CondRule}) i,
HoareRuleTac ctxt false (i+2),
HoareRuleTac ctxt false (i+1)],
@@ -345,7 +345,7 @@
THEN (if precond then (K all_tac i) else (rtac (@{thm subset_refl}) i))))
and AnnWlpTac ctxt i = (rtac (@{thm AnnSeq}) i) THEN (AnnHoareRuleTac ctxt (i+1))
-and AnnHoareRuleTac ctxt i st = st |>
+and AnnHoareRuleTac ctxt i st = st |>
( (AnnWlpTac ctxt i THEN AnnHoareRuleTac ctxt i )
ORELSE
(FIRST[(rtac (@{thm AnnskipRule}) i),
@@ -387,7 +387,7 @@
rtac (@{thm allI}) i,rtac (@{thm impI}) i,
conjI_Tac (interfree_aux_Tac ctxt) i]])
-and interfree_Tac ctxt i st = st |>
+and interfree_Tac ctxt i st = st |>
(FIRST[rtac (@{thm interfree_Empty}) i,
EVERY[rtac (@{thm interfree_List}) i,
interfree_Tac ctxt (i+1),
@@ -396,7 +396,7 @@
rtac (@{thm allI}) i,rtac (@{thm allI}) i,rtac (@{thm impI}) i,
interfree_aux_Tac ctxt i ]])
-and interfree_aux_Tac ctxt i = (before_interfree_simp_tac ctxt i ) THEN
+and interfree_aux_Tac ctxt i = (before_interfree_simp_tac ctxt i ) THEN
(FIRST[rtac (@{thm interfree_aux_rule1}) i,
dest_assertions_Tac ctxt i])
@@ -441,59 +441,59 @@
EVERY[rtac (@{thm AnnAwait_atomics}) i,
HoareRuleTac ctxt true i],
K all_tac i])
-*}
+\<close>
-text {* The final tactic is given the name @{text oghoare}: *}
+text \<open>The final tactic is given the name @{text oghoare}:\<close>
-ML {*
+ML \<open>
fun oghoare_tac ctxt = SUBGOAL (fn (_, i) => HoareRuleTac ctxt true i)
-*}
+\<close>
-text {* Notice that the tactic for parallel programs @{text
+text \<open>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: *}
+generate verification conditions out of interference freedom tests:\<close>
-ML {*
+ML \<open>
fun annhoare_tac ctxt = SUBGOAL (fn (_, i) => AnnHoareRuleTac ctxt i)
fun interfree_aux_tac ctxt = SUBGOAL (fn (_, i) => interfree_aux_Tac ctxt i)
-*}
+\<close>
-text {* The so defined ML tactics are then ``exported'' to be used in
-Isabelle proofs. *}
+text \<open>The so defined ML tactics are then ``exported'' to be used in
+Isabelle proofs.\<close>
-method_setup oghoare = {*
- Scan.succeed (SIMPLE_METHOD' o oghoare_tac) *}
+method_setup oghoare = \<open>
+ Scan.succeed (SIMPLE_METHOD' o oghoare_tac)\<close>
"verification condition generator for the oghoare logic"
-method_setup annhoare = {*
- Scan.succeed (SIMPLE_METHOD' o annhoare_tac) *}
+method_setup annhoare = \<open>
+ Scan.succeed (SIMPLE_METHOD' o annhoare_tac)\<close>
"verification condition generator for the ann_hoare logic"
-method_setup interfree_aux = {*
- Scan.succeed (SIMPLE_METHOD' o interfree_aux_tac) *}
+method_setup interfree_aux = \<open>
+ Scan.succeed (SIMPLE_METHOD' o interfree_aux_tac)\<close>
"verification condition generator for interference freedom tests"
-text {* Tactics useful for dealing with the generated verification conditions: *}
+text \<open>Tactics useful for dealing with the generated verification conditions:\<close>
-method_setup conjI_tac = {*
- Scan.succeed (K (SIMPLE_METHOD' (conjI_Tac (K all_tac)))) *}
+method_setup conjI_tac = \<open>
+ Scan.succeed (K (SIMPLE_METHOD' (conjI_Tac (K all_tac))))\<close>
"verification condition generator for interference freedom tests"
-ML {*
+ML \<open>
fun disjE_Tac tac i st = st |>
( (EVERY [etac disjE i,
disjE_Tac tac (i+1),
tac i]) ORELSE (tac i) )
-*}
+\<close>
-method_setup disjE_tac = {*
- Scan.succeed (K (SIMPLE_METHOD' (disjE_Tac (K all_tac)))) *}
+method_setup disjE_tac = \<open>
+ Scan.succeed (K (SIMPLE_METHOD' (disjE_Tac (K all_tac))))\<close>
"verification condition generator for interference freedom tests"
end
--- a/src/HOL/Hoare_Parallel/OG_Tran.thy Sat Dec 27 19:51:55 2014 +0100
+++ b/src/HOL/Hoare_Parallel/OG_Tran.thy Sat Dec 27 20:32:06 2014 +0100
@@ -1,10 +1,10 @@
-section {* Operational Semantics *}
+section \<open>Operational Semantics\<close>
theory OG_Tran imports OG_Com begin
type_synonym 'a ann_com_op = "('a ann_com) option"
type_synonym 'a ann_triple_op = "('a ann_com_op \<times> 'a assn)"
-
+
primrec com :: "'a ann_triple_op \<Rightarrow> 'a ann_com_op" where
"com (c, q) = c"
@@ -14,10 +14,10 @@
definition All_None :: "'a ann_triple_op list \<Rightarrow> bool" where
"All_None Ts \<equiv> \<forall>(c, q) \<in> set Ts. c = None"
-subsection {* The Transition Relation *}
+subsection \<open>The Transition Relation\<close>
inductive_set
- ann_transition :: "(('a ann_com_op \<times> 'a) \<times> ('a ann_com_op \<times> 'a)) 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)
@@ -32,9 +32,9 @@
| AnnBasic: "(Some (AnnBasic r f), s) -1\<rightarrow> (None, f s)"
-| AnnSeq1: "(Some c0, s) -1\<rightarrow> (None, t) \<Longrightarrow>
+| 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>
+| 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)"
@@ -44,11 +44,11 @@
| 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>
+| 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)"
+ (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)"
@@ -66,10 +66,10 @@
monos "rtrancl_mono"
-text {* The corresponding abbreviations are: *}
+text \<open>The corresponding abbreviations are:\<close>
abbreviation
- ann_transition_n :: "('a ann_com_op \<times> 'a) \<Rightarrow> nat \<Rightarrow> ('a ann_com_op \<times> 'a)
+ 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"
@@ -79,11 +79,11 @@
"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"
+ 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 *}
+subsection \<open>Definition of Semantics\<close>
definition ann_sem :: "'a ann_com \<Rightarrow> 'a \<Rightarrow> 'a set" where
"ann_sem c \<equiv> \<lambda>s. {t. (Some c, s) -*\<rightarrow> (None, t)}"
@@ -104,17 +104,17 @@
"fwhile b c 0 = \<Omega>"
| "fwhile b c (Suc n) = Cond b (Seq c (fwhile b c n)) (Basic id)"
-subsubsection {* Proofs *}
+subsubsection \<open>Proofs\<close>
declare ann_transition_transition.intros [intro]
-inductive_cases transition_cases:
- "(Parallel T,s) -P1\<rightarrow> t"
+inductive_cases transition_cases:
+ "(Parallel T,s) -P1\<rightarrow> t"
"(Basic f, s) -P1\<rightarrow> t"
- "(Seq c1 c2, 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)]:
+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))
@@ -123,7 +123,7 @@
apply(force elim:transition_cases)
done
-lemma Parallel_AllNone_lemma [rule_format (no_asm)]:
+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))
@@ -148,25 +148,25 @@
apply(simp add:All_None_def)
done
-text {* Set of lemmas from Apt and Olderog "Verification of sequential
-and concurrent programs", page 63. *}
+text \<open>Set of lemmas from Apt and Olderog "Verification of sequential
+and concurrent programs", page 63.\<close>
-lemma L3_5i: "X\<subseteq>Y \<Longrightarrow> SEM c X \<subseteq> SEM c Y"
+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>
+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>
+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)
@@ -176,9 +176,9 @@
apply (fast intro!: le_SucI elim!: relpow_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
+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_relpow)
apply(fast dest: L3_5ii_lemma2 relpow_imp_rtrancl)
@@ -206,7 +206,7 @@
done
-lemma L3_5v_lemma1[rule_format]:
+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)
@@ -238,8 +238,8 @@
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>
+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
@@ -262,8 +262,8 @@
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>
+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)
@@ -292,7 +292,7 @@
apply(fast intro: L3_5v_lemma5)
done
-section {* Validity of Correctness Formulas *}
+section \<open>Validity of Correctness Formulas\<close>
definition com_validity :: "'a assn \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool" ("(3\<parallel>= _// _//_)" [90,55,90] 50) where
"\<parallel>= p c q \<equiv> SEM c p \<subseteq> q"
--- a/src/HOL/Hoare_Parallel/Quote_Antiquote.thy Sat Dec 27 19:51:55 2014 +0100
+++ b/src/HOL/Hoare_Parallel/Quote_Antiquote.thy Sat Dec 27 20:32:06 2014 +0100
@@ -1,4 +1,4 @@
-section {* Concrete Syntax *}
+section \<open>Concrete Syntax\<close>
theory Quote_Antiquote imports Main begin
@@ -10,11 +10,11 @@
translations
"\<lbrace>b\<rbrace>" \<rightharpoonup> "CONST Collect \<guillemotleft>b\<guillemotright>"
-parse_translation {*
+parse_translation \<open>
let
fun quote_tr [t] = Syntax_Trans.quote_tr @{syntax_const "_antiquote"} t
| quote_tr ts = raise TERM ("quote_tr", ts);
in [(@{syntax_const "_quote"}, K quote_tr)] end
-*}
+\<close>
end
\ No newline at end of file
--- a/src/HOL/Hoare_Parallel/RG_Com.thy Sat Dec 27 19:51:55 2014 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Com.thy Sat Dec 27 20:32:06 2014 +0100
@@ -1,21 +1,21 @@
-chapter {* The Rely-Guarantee Method *}
+chapter \<open>The Rely-Guarantee Method\<close>
-section {* Abstract Syntax *}
+section \<open>Abstract Syntax\<close>
theory RG_Com imports Main begin
-text {* Semantics of assertions and boolean expressions (bexp) as sets
+text \<open>Semantics of assertions and boolean expressions (bexp) as sets
of states. Syntax of commands @{text com} and parallel commands
-@{text par_com}. *}
+@{text par_com}.\<close>
type_synonym 'a bexp = "'a set"
-datatype 'a com =
+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"
+ | Cond "'a bexp" "'a com" "'a com"
+ | While "'a bexp" "'a com"
+ | Await "'a bexp" "'a com"
type_synonym 'a par_com = "'a com option list"
--- a/src/HOL/Hoare_Parallel/RG_Examples.thy Sat Dec 27 19:51:55 2014 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Examples.thy Sat Dec 27 20:32:06 2014 +0100
@@ -1,12 +1,12 @@
-section {* Examples *}
+section \<open>Examples\<close>
theory RG_Examples
imports RG_Syntax
begin
-lemmas definitions [simp]= stable_def Pre_def Rely_def Guar_def Post_def Com_def
+lemmas definitions [simp]= stable_def Pre_def Rely_def Guar_def Post_def Com_def
-subsection {* Set Elements of an Array to Zero *}
+subsection \<open>Set Elements of an Array to Zero\<close>
lemma le_less_trans2: "\<lbrakk>(j::nat)<k; i\<le> j\<rbrakk> \<Longrightarrow> i<k"
by simp
@@ -17,40 +17,40 @@
record Example1 =
A :: "nat list"
-lemma Example1:
+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>)
+ (\<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)
+apply (auto intro!: Basic)
done
-lemma Example1_parameterized:
+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>,
+ \<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(erule le_less_trans2)
apply(case_tac t,simp+)
apply (simp add:add.commute)
apply(simp add: add_le_mono)
@@ -66,38 +66,38 @@
done
-subsection {* Increment a Variable in Parallel *}
+subsection \<open>Increment a Variable in Parallel\<close>
-subsubsection {* Two components *}
+subsubsection \<open>Two components\<close>
record Example2 =
x :: nat
c_0 :: nat
c_1 :: nat
-lemma Example2:
+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
+ (\<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
+ (\<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>,
+ 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>]"
@@ -151,9 +151,9 @@
apply(auto intro!: Basic)
done
-subsubsection {* Parameterized *}
+subsubsection \<open>Parameterized\<close>
-lemma Example2_lemma2_aux: "j<n \<Longrightarrow>
+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)
@@ -165,11 +165,11 @@
apply arith
done
-lemma Example2_lemma2_aux2:
+lemma Example2_lemma2_aux2:
"j\<le> s \<Longrightarrow> (\<Sum>i::nat=0..<j. (b (s:=t)) i) = (\<Sum>i=0..<j. b i)"
by (induct j) simp_all
-lemma Example2_lemma2:
+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)
@@ -189,19 +189,19 @@
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 =
+record Example2_parameterized =
C :: "nat \<Rightarrow> nat"
y :: nat
-lemma Example2_parameterized: "0<n \<Longrightarrow>
+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>
+ (\<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>)
+ \<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)
@@ -229,9 +229,9 @@
apply simp_all
done
-subsection {* Find Least Element *}
+subsection \<open>Find Least Element\<close>
-text {* A previous lemma: *}
+text \<open>A previous lemma:\<close>
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" )
@@ -251,25 +251,25 @@
X :: "nat \<Rightarrow> nat"
Y :: "nat \<Rightarrow> nat"
-lemma Example3: "m mod n=0 \<Longrightarrow>
- \<turnstile> COBEGIN
+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
+ (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>
+ \<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>
+ \<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>)
+ \<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>
+ \<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 *}
+--\<open>5 subgoals left\<close>
apply force+
apply clarify
apply simp
@@ -298,29 +298,29 @@
apply auto
done
-text {* Same but with a list as auxiliary variable: *}
+text \<open>Same but with a list as auxiliary variable:\<close>
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
+ (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>
+ \<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>
+ \<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>
+ \<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)
-apply (auto cong del: strong_INF_cong strong_SUP_cong)
+apply (auto cong del: strong_INF_cong strong_SUP_cong)
apply force
apply (rule While)
apply force
--- a/src/HOL/Hoare_Parallel/RG_Hoare.thy Sat Dec 27 19:51:55 2014 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Hoare.thy Sat Dec 27 20:32:06 2014 +0100
@@ -1,23 +1,23 @@
-section {* The Proof System *}
+section \<open>The Proof System\<close>
theory RG_Hoare imports RG_Tran begin
-subsection {* Proof System for Component Programs *}
+subsection \<open>Proof System for Component Programs\<close>
declare Un_subset_iff [simp del] sup.bounded_iff [simp del]
-definition stable :: "'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> bool" where
- "stable \<equiv> \<lambda>f g. (\<forall>x y. x \<in> f \<longrightarrow> (x, y) \<in> g \<longrightarrow> y \<in> f)"
+definition stable :: "'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> bool" where
+ "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"
+ 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>
+ 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>
+| 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];
@@ -28,11 +28,11 @@
\<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},
+| 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]"
@@ -52,7 +52,7 @@
definition Com :: "'a rgformula \<Rightarrow> 'a com" where
"Com x \<equiv> fst x"
-subsection {* Proof System for Parallel Programs *}
+subsection \<open>Proof System for Parallel Programs\<close>
type_synonym 'a par_rgformula =
"('a rgformula) list \<times> 'a set \<times> ('a \<times> 'a) set \<times> ('a \<times> 'a) set \<times> 'a set"
@@ -61,20 +61,20 @@
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:
+ 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));
+ 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*}
+section \<open>Soundness\<close>
-subsubsection {* Some previous lemmas *}
+subsubsection \<open>Some previous lemmas\<close>
-lemma tl_of_assum_in_assum:
- "(P, s) # (P, t) # xs \<in> assum (pre, rely) \<Longrightarrow> stable pre rely
+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
@@ -88,51 +88,51 @@
apply simp
done
-lemma etran_in_comm:
+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>
+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!]:
+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(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!]:
+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(case_tac j,simp+)
apply(erule cptn.cases)
apply simp
apply force
apply force
done
-lemma takepar_cptn_is_par_cptn [rule_format,elim]:
+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(case_tac j,simp)
apply(rule ParCptnOne)
apply(force intro:par_cptn.intros elim:par_cptn.cases)
done
@@ -142,7 +142,7 @@
apply(induct "c")
apply(force elim: par_cptn.cases)
apply clarify
-apply(case_tac j,simp+)
+apply(case_tac j,simp+)
apply(erule par_cptn.cases)
apply simp
apply force
@@ -150,11 +150,11 @@
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(subgoal_tac "1 < length (x # xs)")
apply(drule dropcptn_is_cptn,simp+)
done
-lemma not_ctran_None [rule_format]:
+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
@@ -170,9 +170,9 @@
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
+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
@@ -189,7 +189,7 @@
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]:
+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)
@@ -202,15 +202,15 @@
apply simp
done
-lemma etran_or_ctran2_disjI1:
+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:
+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]:
+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)
@@ -226,11 +226,11 @@
apply(case_tac "\<forall>m. m<n \<longrightarrow> \<not> P m")
apply auto
done
-
-lemma stability [rule_format]:
+
+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. (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
@@ -239,7 +239,7 @@
apply(erule cptn.cases,simp)
apply simp
apply(case_tac k,simp,simp)
- apply(case_tac j,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")
@@ -252,7 +252,7 @@
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 simp
apply(erule mp)
apply(erule mp)
apply(rule Env)
@@ -277,13 +277,13 @@
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 *}
+subsection \<open>Soundness of the System for Component Programs\<close>
-subsubsection {* Soundness of the Basic rule *}
+subsubsection \<open>Soundness of the Basic rule\<close>
-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>
+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
@@ -306,8 +306,8 @@
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)
+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
@@ -322,8 +322,8 @@
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;
+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)
@@ -374,11 +374,11 @@
apply(simp add:last_length)+
done
-subsubsection{* Soundness of the Await rule *}
+subsubsection\<open>Soundness of the Await rule\<close>
-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>
+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
@@ -399,8 +399,8 @@
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)
+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
@@ -414,7 +414,7 @@
apply(rule_tac x=0 in exI,simp)
done
-lemma Star_imp_cptn:
+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)
@@ -438,12 +438,12 @@
apply(erule CptnComp)
apply clarify
done
-
-lemma Await_sound:
+
+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},
+ \<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},
+ \<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)
@@ -458,9 +458,9 @@
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. *}
+--\<open>here starts the different part.\<close>
apply(erule ctran.cases,simp_all)
- apply(drule Star_imp_cptn)
+ apply(drule Star_imp_cptn)
apply clarify
apply(erule_tac x=sa in allE)
apply clarify
@@ -491,7 +491,7 @@
apply(case_tac "x!Suc j",simp)
apply(rule ctran.cases,simp)
apply(simp_all)
-apply(drule Star_imp_cptn)
+apply(drule Star_imp_cptn)
apply clarify
apply(erule_tac x=sa in allE)
apply clarify
@@ -512,10 +512,10 @@
apply(simp add:last_length)+
done
-subsubsection{* Soundness of the Conditional rule *}
+subsubsection\<open>Soundness of the Conditional rule\<close>
-lemma Cond_sound:
- "\<lbrakk> stable pre rely; \<Turnstile> P1 sat [pre \<inter> b, rely, guar, post];
+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)
@@ -578,7 +578,7 @@
apply arith
done
-subsubsection{* Soundness of the Sequential rule *}
+subsubsection\<open>Soundness of the Sequential rule\<close>
inductive_cases Seq_cases [elim!]: "(Some (Seq P Q), s) -c\<rightarrow> t"
@@ -591,9 +591,9 @@
apply simp
done
-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>
+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)
@@ -619,11 +619,11 @@
apply simp
done
-lemma Seq_sound2 [rule_format]:
- "x \<in> cptn \<Longrightarrow> \<forall>s P i. x!0=(Some (Seq P Q), s) \<longrightarrow> i<length x
+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
+ (\<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)
@@ -668,7 +668,7 @@
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(subgoal_tac "lift Q (Some P2, ta) =(Some (Seq P2 Q), ta)")
apply(simp add:Cons_lift del:list.map)
apply(rule nth_tl_if)
apply force
@@ -697,7 +697,7 @@
apply simp
done
-lemma Seq_sound:
+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)
@@ -740,7 +740,7 @@
apply (simp del:list.map)
apply(simp only:last_lift_not_None)
apply simp
---{* @{text "\<exists>i<length x. fst (x ! i) = Some Q"} *}
+--\<open>@{text "\<exists>i<length x. fst (x ! i) = Some Q"}\<close>
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
@@ -756,7 +756,7 @@
apply(simp add:cp_def assum_def nth_append)
apply clarify
apply(erule_tac x=i in allE)
- back
+ back
apply(simp add:snd_lift)
apply(erule mp)
apply(force elim:etranE intro:Env simp add:lift_def)
@@ -780,7 +780,7 @@
apply(erule mp)
apply(case_tac "xs!m")
apply(force elim:etran.cases intro:Env simp add:lift_def)
- apply simp
+ apply simp
apply simp
apply clarify
apply(rule conjI,clarify)
@@ -788,12 +788,12 @@
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 "(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)
+ apply(erule_tac x="i-m" in allE)
back
back
apply(subgoal_tac "Suc (i - m) < length ys",simp)
@@ -812,7 +812,7 @@
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 simp
apply force
apply clarify
apply(rule tl_zero)
@@ -834,7 +834,7 @@
apply(case_tac ys,simp+)
done
-subsubsection{* Soundness of the While rule *}
+subsubsection\<open>Soundness of the While rule\<close>
lemma last_append[rule_format]:
"\<forall>xs. ys\<noteq>[] \<longrightarrow> ((xs@ys)!(length (xs@ys) - (Suc 0)))=(ys!(length ys - (Suc 0)))"
@@ -844,10 +844,10 @@
apply (simp add:nth_append)
done
-lemma assum_after_body:
- "\<lbrakk> \<Turnstile> P sat [pre \<inter> b, rely, guar, pre];
+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) #
+ (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)
@@ -875,20 +875,20 @@
apply(simp add:Cons_lift_append nth_append snd_lift)
done
-lemma While_sound_aux [rule_format]:
+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>
+ 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 *}
+--\<open>5 subgoals left\<close>
apply(simp add:comm_def)
---{* 4 subgoals left *}
+--\<open>4 subgoals left\<close>
apply(rule etran_in_comm)
apply(erule mp)
apply(erule tl_of_assum_in_assum,simp)
---{* While-None *}
+--\<open>While-None\<close>
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])
@@ -908,12 +908,12 @@
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(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 *}
+--\<open>WhileOne\<close>
apply(thin_tac "P = While b P \<longrightarrow> ?Q")
apply(rule ctran_in_comm,simp)
apply(simp add:Cons_lift del:list.map)
@@ -949,23 +949,23 @@
apply(case_tac "fst(xs!i)")
apply force
apply force
---{* last=None *}
+--\<open>last=None\<close>
apply clarify
apply(subgoal_tac "(map (lift (While b P)) ((Some P, sa) # xs))\<noteq>[]")
apply(drule last_conv_nth)
apply (simp del:list.map)
apply(simp only:last_lift_not_None)
apply simp
---{* WhileMore *}
+--\<open>WhileMore\<close>
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. *}
+--\<open>metiendo la hipotesis antes de dividir la conclusion.\<close>
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. *}
+--\<open>lo de antes.\<close>
apply(simp add:comm_def del:list.map last.simps)
apply(rule conjI)
apply clarify
@@ -1001,8 +1001,8 @@
apply(case_tac "fst(xs!i)")
apply force
apply force
---{* @{text "i \<ge> length xs"} *}
-apply(subgoal_tac "i-length xs <length ys")
+--\<open>@{text "i \<ge> length xs"}\<close>
+apply(subgoal_tac "i-length xs <length ys")
prefer 2
apply arith
apply(erule_tac x="i-length xs" in allE,clarify)
@@ -1012,7 +1012,7 @@
apply(erule mp)
apply(case_tac "last((Some P, sa) # xs)")
apply(simp add:lift_def del:last.simps)
---{* @{text "i>length xs"} *}
+--\<open>@{text "i>length xs"}\<close>
apply(case_tac "i-length xs")
apply arith
apply(simp add:nth_append del:list.map last.simps)
@@ -1021,7 +1021,7 @@
prefer 2
apply arith
apply simp
---{* last=None *}
+--\<open>last=None\<close>
apply clarify
apply(case_tac ys)
apply(simp add:Cons_lift del:list.map last.simps)
@@ -1042,7 +1042,7 @@
apply simp
done
-lemma While_sound:
+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]"
@@ -1063,10 +1063,10 @@
apply(case_tac x,simp+)
done
-subsubsection{* Soundness of the Rule of Consequence *}
+subsubsection\<open>Soundness of the Rule of Consequence\<close>
-lemma Conseq_sound:
- "\<lbrakk>pre \<subseteq> pre'; rely \<subseteq> rely'; guar' \<subseteq> guar; post' \<subseteq> post;
+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)
@@ -1077,9 +1077,9 @@
apply force
done
-subsubsection {* Soundness of the system for sequential component programs *}
+subsubsection \<open>Soundness of the system for sequential component programs\<close>
-theorem rgsound:
+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)
@@ -1088,42 +1088,42 @@
apply(force elim:While_sound)
apply(force elim:Await_sound)
apply(erule Conseq_sound,simp+)
-done
+done
-subsection {* Soundness of the System for Parallel Programs *}
+subsection \<open>Soundness of the System for Parallel Programs\<close>
definition ParallelCom :: "('a rgformula) list \<Rightarrow> 'a par_com" where
- "ParallelCom Ps \<equiv> map (Some \<circ> fst) Ps"
+ "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))
+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.
+ \<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> \<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 (rule ccontr)
+--\<open>By contradiction:\<close>
apply simp
apply(erule exE)
---{* the first c-tran that does not satisfy the guarantee-condition is from @{text "\<sigma>_i"} at step @{text "m"}. *}
+--\<open>the first c-tran that does not satisfy the guarantee-condition is from @{text "\<sigma>_i"} at step @{text "m"}.\<close>
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)"} *}
+--\<open>@{text "\<sigma>_i \<in> A(pre, rely_1)"}\<close>
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]"} *}
+--\<open>but this contradicts @{text "\<Turnstile> \<sigma>_i sat [pre_i,rely_i,guar_i,post_i]"}\<close>
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 (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)
@@ -1142,9 +1142,9 @@
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 *}
+--\<open>each etran in @{text "\<sigma>_1[0\<dots>m]"} corresponds to\<close>
apply(erule disjE)
---{* a c-tran in some @{text "\<sigma>_{ib}"} *}
+--\<open>a c-tran in some @{text "\<sigma>_{ib}"}\<close>
apply clarify
apply(case_tac "i=ib",simp)
apply(erule etranE,simp)
@@ -1160,21 +1160,21 @@
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}"} *}
+--\<open>or an e-tran in @{text "\<sigma>"},
+therefore it satisfies @{text "rely \<or> guar_{ib}"}\<close>
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))
+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.
+ \<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> \<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
@@ -1196,14 +1196,14 @@
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))
+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;
+ (\<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 \<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)
@@ -1247,14 +1247,14 @@
apply(force elim:par_cptn.cases)
done
-lemma five:
+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));
+ 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);
+ 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)
apply(subgoal_tac "(map (Some \<circ> fst) xs)\<noteq>[]")
@@ -1316,8 +1316,8 @@
apply(simp add:conjoin_def same_length_def)
done
-lemma ParallelEmpty [rule_format]:
- "\<forall>i s. x \<in> par_cp (ParallelCom []) s \<longrightarrow>
+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)
@@ -1335,8 +1335,8 @@
apply simp
done
-theorem par_rgsound:
- "\<turnstile> c SAT [pre, rely, guar, post] \<Longrightarrow>
+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)
@@ -1368,7 +1368,7 @@
apply(assumption+)
apply clarify
apply (erule allE, erule impE, assumption,erule rgsound)
- apply(assumption+)
+ apply(assumption+)
done
end
--- a/src/HOL/Hoare_Parallel/RG_Syntax.thy Sat Dec 27 19:51:55 2014 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Syntax.thy Sat Dec 27 20:32:06 2014 +0100
@@ -1,4 +1,4 @@
-section {* Concrete Syntax *}
+section \<open>Concrete Syntax\<close>
theory RG_Syntax
imports RG_Hoare Quote_Antiquote
@@ -45,17 +45,17 @@
translations
"_prg_scheme j i k c" \<rightleftharpoons> "(CONST map (\<lambda>i. c) [j..<k])"
-text {* Translations for variables before and after a transition: *}
+text \<open>Translations for variables before and after a transition:\<close>
-syntax
+syntax
"_before" :: "id \<Rightarrow> 'a" ("\<ordmasculine>_")
"_after" :: "id \<Rightarrow> 'a" ("\<ordfeminine>_")
-
+
translations
"\<ordmasculine>x" \<rightleftharpoons> "x \<acute>CONST fst"
"\<ordfeminine>x" \<rightleftharpoons> "x \<acute>CONST snd"
-print_translation {*
+print_translation \<open>
let
fun quote_tr' f (t :: ts) =
Term.list_comb (f $ Syntax_Trans.quote_tr' @{syntax_const "_antiquote"} t, ts)
@@ -77,6 +77,6 @@
(@{const_syntax Cond}, K (bexp_tr' @{syntax_const "_Cond"})),
(@{const_syntax While}, K (bexp_tr' @{syntax_const "_While"}))]
end
-*}
+\<close>
end
\ No newline at end of file
--- a/src/HOL/Hoare_Parallel/RG_Tran.thy Sat Dec 27 19:51:55 2014 +0100
+++ b/src/HOL/Hoare_Parallel/RG_Tran.thy Sat Dec 27 20:32:06 2014 +0100
@@ -1,12 +1,12 @@
-section {* Operational Semantics *}
+section \<open>Operational Semantics\<close>
theory RG_Tran
imports RG_Com
begin
-subsection {* Semantics of Component Programs *}
+subsection \<open>Semantics of Component Programs\<close>
-subsubsection {* Environment transitions *}
+subsubsection \<open>Environment transitions\<close>
type_synonym 'a conf = "(('a com) option) \<times> 'a"
@@ -20,7 +20,7 @@
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 *}
+subsubsection \<open>Component transitions\<close>
inductive_set
ctran :: "('a conf \<times> 'a conf) set"
@@ -46,7 +46,7 @@
monos "rtrancl_mono"
-subsection {* Semantics of Parallel Programs *}
+subsection \<open>Semantics of Parallel Programs\<close>
type_synonym 'a par_conf = "('a par_com) \<times> 'a"
@@ -69,9 +69,9 @@
(Ps ! i, s) -c\<rightarrow> (r, t) \<Longrightarrow> P) \<Longrightarrow> P"
by (induct c, induct c', erule par_ctran.cases, blast)
-subsection {* Computations *}
+subsection \<open>Computations\<close>
-subsubsection {* Sequential computations *}
+subsubsection \<open>Sequential computations\<close>
type_synonym 'a confs = "'a conf list"
@@ -84,7 +84,7 @@
definition cp :: "('a com) option \<Rightarrow> 'a \<Rightarrow> ('a confs) set" where
"cp P s \<equiv> {l. l!0=(P,s) \<and> l \<in> cptn}"
-subsubsection {* Parallel computations *}
+subsubsection \<open>Parallel computations\<close>
type_synonym 'a par_confs = "'a par_conf list"
@@ -97,7 +97,7 @@
definition par_cp :: "'a par_com \<Rightarrow> 'a \<Rightarrow> ('a par_confs) set" where
"par_cp P s \<equiv> {l. l!0=(P,s) \<and> l \<in> par_cptn}"
-subsection{* Modular Definition of Computation *}
+subsection\<open>Modular Definition of Computation\<close>
definition lift :: "'a com \<Rightarrow> 'a conf \<Rightarrow> 'a conf" where
"lift Q \<equiv> \<lambda>(P, s). (if P=None then (Some Q,s) else (Some(Seq (the P) Q), s))"
@@ -125,7 +125,7 @@
(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.*}
+subsection \<open>Equivalence of Both Definitions.\<close>
lemma last_length: "((a#xs)!(length xs))=last (a#xs)"
by (induct xs) auto
@@ -178,20 +178,20 @@
\<longrightarrow> (Some a, s) # (Q, t) # xs \<in> cptn_mod"
apply(induct a)
apply simp_all
---{* basic *}
+--\<open>basic\<close>
apply clarify
apply(erule ctran.cases,simp_all)
apply(rule CptnModNone,rule Basic,simp)
apply clarify
apply(erule ctran.cases,simp_all)
---{* Seq1 *}
+--\<open>Seq1\<close>
apply(rule_tac xs="[(None,ta)]" in CptnModSeq2)
apply(erule CptnModNone)
apply(rule CptnModOne)
apply simp
apply simp
apply(simp add:lift_def)
---{* Seq2 *}
+--\<open>Seq2\<close>
apply(erule_tac x=sa in allE)
apply(erule_tac x="Some P2" in allE)
apply(erule allE,erule impE, assumption)
@@ -208,12 +208,12 @@
apply (simp add:last_length)
apply (simp add:last_length)
apply(simp add:lift_def)
---{* Cond *}
+--\<open>Cond\<close>
apply clarify
apply(erule ctran.cases,simp_all)
apply(force elim: CptnModCondT)
apply(force elim: CptnModCondF)
---{* While *}
+--\<open>While\<close>
apply clarify
apply(erule ctran.cases,simp_all)
apply(rule CptnModNone,erule WhileF,simp)
@@ -223,7 +223,7 @@
apply(force elim:CptnModWhile1)
apply clarify
apply(force simp add:last_length elim:CptnModWhile2)
---{* await *}
+--\<open>await\<close>
apply clarify
apply(erule ctran.cases,simp_all)
apply(rule CptnModNone,erule Await,simp+)
@@ -295,7 +295,7 @@
apply(erule CondT,simp)
apply(rule CptnComp)
apply(erule CondF,simp)
---{* Seq1 *}
+--\<open>Seq1\<close>
apply(erule cptn.cases,simp_all)
apply(rule CptnOne)
apply clarify
@@ -315,7 +315,7 @@
apply(rule Seq2,simp)
apply(drule_tac P=P1 in lift_is_cptn)
apply(simp add:lift_def)
---{* Seq2 *}
+--\<open>Seq2\<close>
apply(rule cptn_append_is_cptn)
apply(drule_tac P=P1 in lift_is_cptn)
apply(simp add:lift_def)
@@ -325,12 +325,12 @@
apply(rule last_fst_esp)
apply (simp add:last_length)
apply(simp add:Cons_lift lift_def split_def last_conv_nth)
---{* While1 *}
+--\<open>While1\<close>
apply(rule CptnComp)
apply(rule WhileT,simp)
apply(drule_tac P="While b P" in lift_is_cptn)
apply(simp add:lift_def)
---{* While2 *}
+--\<open>While2\<close>
apply(rule CptnComp)
apply(rule WhileT,simp)
apply(rule cptn_append_is_cptn)
@@ -349,9 +349,9 @@
apply(erule cptn_if_cptn_mod)
done
-section {* Validity of Correctness Formulas*}
+section \<open>Validity of Correctness Formulas\<close>
-subsection {* Validity for Component Programs. *}
+subsection \<open>Validity for Component Programs.\<close>
type_synonym 'a rgformula =
"'a com \<times> 'a set \<times> ('a \<times> 'a) set \<times> ('a \<times> 'a) set \<times> 'a set"
@@ -370,7 +370,7 @@
"\<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. *}
+subsection \<open>Validity for Parallel Programs.\<close>
definition All_None :: "(('a com) option) list \<Rightarrow> bool" where
"All_None xs \<equiv> \<forall>c\<in>set xs. c=None"
@@ -389,9 +389,9 @@
"\<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 *}
+subsection \<open>Compositionality of the Semantics\<close>
-subsubsection {* Definition of the conjoin operator *}
+subsubsection \<open>Definition of the conjoin operator\<close>
definition same_length :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" where
"same_length c clist \<equiv> (\<forall>i<length clist. length(clist!i)=length c)"
@@ -411,7 +411,7 @@
definition conjoin :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" ("_ \<propto> _" [65,65] 64) where
"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 *}
+subsubsection \<open>Some previous lemmas\<close>
lemma list_eq_if [rule_format]:
"\<forall>ys. xs=ys \<longrightarrow> (length xs = length ys) \<longrightarrow> (\<forall>i<length xs. xs!i=ys!i)"
@@ -482,7 +482,7 @@
"P (ys!Suc j) \<longrightarrow> Suc j<length ys \<longrightarrow> ys\<noteq>[] \<longrightarrow> P (tl(ys)!j)"
by (induct ys) simp_all
-subsection {* The Semantics is Compositional *}
+subsection \<open>The Semantics is Compositional\<close>
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)
@@ -496,7 +496,7 @@
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 *}
+--\<open>first step is a Component step\<close>
apply clarify
apply simp
apply(subgoal_tac "a=(xs[i:=(fst(clist!i!0))])")
@@ -516,7 +516,7 @@
apply(erule etranE,simp)
apply(rule ParCptnComp)
apply(erule ParComp,simp)
---{* applying the induction hypothesis *}
+--\<open>applying the induction hypothesis\<close>
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)
@@ -630,7 +630,7 @@
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 *}
+--\<open>first step is an environmental step\<close>
apply clarify
apply(erule par_etran.cases)
apply simp