| author | wenzelm | 
| Mon, 21 Dec 2015 17:20:57 +0100 | |
| changeset 61889 | 42d902e074e8 | 
| parent 59189 | ad8e0a789af6 | 
| child 62390 | 842917225d56 | 
| permissions | -rw-r--r-- | 
| 59189 | 1  | 
section \<open>Operational Semantics\<close>  | 
| 13020 | 2  | 
|
| 16417 | 3  | 
theory OG_Tran imports OG_Com begin  | 
| 13020 | 4  | 
|
| 42174 | 5  | 
type_synonym 'a ann_com_op = "('a ann_com) option"
 | 
6  | 
type_synonym 'a ann_triple_op = "('a ann_com_op \<times> 'a assn)"
 | 
|
| 59189 | 7  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35107 
diff
changeset
 | 
8  | 
primrec com :: "'a ann_triple_op \<Rightarrow> 'a ann_com_op" where  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35107 
diff
changeset
 | 
9  | 
"com (c, q) = c"  | 
| 13020 | 10  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35107 
diff
changeset
 | 
11  | 
primrec post :: "'a ann_triple_op \<Rightarrow> 'a assn" where  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35107 
diff
changeset
 | 
12  | 
"post (c, q) = q"  | 
| 13020 | 13  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35107 
diff
changeset
 | 
14  | 
definition All_None :: "'a ann_triple_op list \<Rightarrow> bool" where  | 
| 13020 | 15  | 
"All_None Ts \<equiv> \<forall>(c, q) \<in> set Ts. c = None"  | 
16  | 
||
| 59189 | 17  | 
subsection \<open>The Transition Relation\<close>  | 
| 13020 | 18  | 
|
| 23746 | 19  | 
inductive_set  | 
| 59189 | 20  | 
  ann_transition :: "(('a ann_com_op \<times> 'a) \<times> ('a ann_com_op \<times> 'a)) set"
 | 
| 23746 | 21  | 
  and transition :: "(('a com \<times> 'a) \<times> ('a com \<times> 'a)) set"
 | 
22  | 
  and ann_transition' :: "('a ann_com_op \<times> 'a) \<Rightarrow> ('a ann_com_op \<times> 'a) \<Rightarrow> bool"
 | 
|
23  | 
    ("_ -1\<rightarrow> _"[81,81] 100)
 | 
|
24  | 
  and transition' :: "('a com \<times> 'a) \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"
 | 
|
25  | 
    ("_ -P1\<rightarrow> _"[81,81] 100)
 | 
|
26  | 
  and transitions :: "('a com \<times> 'a) \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"
 | 
|
27  | 
    ("_ -P*\<rightarrow> _"[81,81] 100)
 | 
|
28  | 
where  | 
|
29  | 
"con_0 -1\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> ann_transition"  | 
|
30  | 
| "con_0 -P1\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> transition"  | 
|
31  | 
| "con_0 -P*\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> transition\<^sup>*"  | 
|
32  | 
||
33  | 
| AnnBasic: "(Some (AnnBasic r f), s) -1\<rightarrow> (None, f s)"  | 
|
34  | 
||
| 59189 | 35  | 
| AnnSeq1: "(Some c0, s) -1\<rightarrow> (None, t) \<Longrightarrow>  | 
| 23746 | 36  | 
(Some (AnnSeq c0 c1), s) -1\<rightarrow> (Some c1, t)"  | 
| 59189 | 37  | 
| AnnSeq2: "(Some c0, s) -1\<rightarrow> (Some c2, t) \<Longrightarrow>  | 
| 23746 | 38  | 
(Some (AnnSeq c0 c1), s) -1\<rightarrow> (Some (AnnSeq c2 c1), t)"  | 
39  | 
||
40  | 
| AnnCond1T: "s \<in> b \<Longrightarrow> (Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> (Some c1, s)"  | 
|
41  | 
| AnnCond1F: "s \<notin> b \<Longrightarrow> (Some (AnnCond1 r b c1 c2), s) -1\<rightarrow> (Some c2, s)"  | 
|
| 13020 | 42  | 
|
| 23746 | 43  | 
| AnnCond2T: "s \<in> b \<Longrightarrow> (Some (AnnCond2 r b c), s) -1\<rightarrow> (Some c, s)"  | 
44  | 
| AnnCond2F: "s \<notin> b \<Longrightarrow> (Some (AnnCond2 r b c), s) -1\<rightarrow> (None, s)"  | 
|
45  | 
||
46  | 
| AnnWhileF: "s \<notin> b \<Longrightarrow> (Some (AnnWhile r b i c), s) -1\<rightarrow> (None, s)"  | 
|
| 59189 | 47  | 
| AnnWhileT: "s \<in> b \<Longrightarrow> (Some (AnnWhile r b i c), s) -1\<rightarrow>  | 
| 23746 | 48  | 
(Some (AnnSeq c (AnnWhile i b i c)), s)"  | 
49  | 
||
50  | 
| AnnAwait: "\<lbrakk> s \<in> b; atom_com c; (c, s) -P*\<rightarrow> (Parallel [], t) \<rbrakk> \<Longrightarrow>  | 
|
| 59189 | 51  | 
(Some (AnnAwait r b c), s) -1\<rightarrow> (None, t)"  | 
| 23746 | 52  | 
|
53  | 
| Parallel: "\<lbrakk> i<length Ts; Ts!i = (Some c, q); (Some c, s) -1\<rightarrow> (r, t) \<rbrakk>  | 
|
54  | 
\<Longrightarrow> (Parallel Ts, s) -P1\<rightarrow> (Parallel (Ts [i:=(r, q)]), t)"  | 
|
55  | 
||
56  | 
| Basic: "(Basic f, s) -P1\<rightarrow> (Parallel [], f s)"  | 
|
57  | 
||
58  | 
| Seq1: "All_None Ts \<Longrightarrow> (Seq (Parallel Ts) c, s) -P1\<rightarrow> (c, s)"  | 
|
59  | 
| Seq2: "(c0, s) -P1\<rightarrow> (c2, t) \<Longrightarrow> (Seq c0 c1, s) -P1\<rightarrow> (Seq c2 c1, t)"  | 
|
60  | 
||
61  | 
| CondT: "s \<in> b \<Longrightarrow> (Cond b c1 c2, s) -P1\<rightarrow> (c1, s)"  | 
|
62  | 
| CondF: "s \<notin> b \<Longrightarrow> (Cond b c1 c2, s) -P1\<rightarrow> (c2, s)"  | 
|
63  | 
||
64  | 
| WhileF: "s \<notin> b \<Longrightarrow> (While b i c, s) -P1\<rightarrow> (Parallel [], s)"  | 
|
65  | 
| WhileT: "s \<in> b \<Longrightarrow> (While b i c, s) -P1\<rightarrow> (Seq c (While b i c), s)"  | 
|
66  | 
||
67  | 
monos "rtrancl_mono"  | 
|
| 13020 | 68  | 
|
| 59189 | 69  | 
text \<open>The corresponding abbreviations are:\<close>  | 
| 13020 | 70  | 
|
| 23746 | 71  | 
abbreviation  | 
| 59189 | 72  | 
  ann_transition_n :: "('a ann_com_op \<times> 'a) \<Rightarrow> nat \<Rightarrow> ('a ann_com_op \<times> 'a)
 | 
| 23746 | 73  | 
                           \<Rightarrow> bool"  ("_ -_\<rightarrow> _"[81,81] 100)  where
 | 
| 
30952
 
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
 
haftmann 
parents: 
30304 
diff
changeset
 | 
74  | 
"con_0 -n\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> ann_transition ^^ n"  | 
| 13020 | 75  | 
|
| 23746 | 76  | 
abbreviation  | 
77  | 
  ann_transitions :: "('a ann_com_op \<times> 'a) \<Rightarrow> ('a ann_com_op \<times> 'a) \<Rightarrow> bool"
 | 
|
78  | 
                           ("_ -*\<rightarrow> _"[81,81] 100)  where
 | 
|
79  | 
"con_0 -*\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> ann_transition\<^sup>*"  | 
|
| 13020 | 80  | 
|
| 23746 | 81  | 
abbreviation  | 
| 59189 | 82  | 
  transition_n :: "('a com \<times> 'a) \<Rightarrow> nat \<Rightarrow> ('a com \<times> 'a) \<Rightarrow> bool"
 | 
| 23746 | 83  | 
                          ("_ -P_\<rightarrow> _"[81,81,81] 100)  where
 | 
| 
30952
 
7ab2716dd93b
power operation on functions with syntax o^; power operation on relations with syntax ^^
 
haftmann 
parents: 
30304 
diff
changeset
 | 
84  | 
"con_0 -Pn\<rightarrow> con_1 \<equiv> (con_0, con_1) \<in> transition ^^ n"  | 
| 13020 | 85  | 
|
| 59189 | 86  | 
subsection \<open>Definition of Semantics\<close>  | 
| 13020 | 87  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35107 
diff
changeset
 | 
88  | 
definition ann_sem :: "'a ann_com \<Rightarrow> 'a \<Rightarrow> 'a set" where  | 
| 13020 | 89  | 
  "ann_sem c \<equiv> \<lambda>s. {t. (Some c, s) -*\<rightarrow> (None, t)}"
 | 
90  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35107 
diff
changeset
 | 
91  | 
definition ann_SEM :: "'a ann_com \<Rightarrow> 'a set \<Rightarrow> 'a set" where  | 
| 
52141
 
eff000cab70f
weaker precendence of syntax for big intersection and union on sets
 
haftmann 
parents: 
46362 
diff
changeset
 | 
92  | 
"ann_SEM c S \<equiv> \<Union>(ann_sem c ` S)"  | 
| 13020 | 93  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35107 
diff
changeset
 | 
94  | 
definition sem :: "'a com \<Rightarrow> 'a \<Rightarrow> 'a set" where  | 
| 13020 | 95  | 
  "sem c \<equiv> \<lambda>s. {t. \<exists>Ts. (c, s) -P*\<rightarrow> (Parallel Ts, t) \<and> All_None Ts}"
 | 
96  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35107 
diff
changeset
 | 
97  | 
definition SEM :: "'a com \<Rightarrow> 'a set \<Rightarrow> 'a set" where  | 
| 
52141
 
eff000cab70f
weaker precendence of syntax for big intersection and union on sets
 
haftmann 
parents: 
46362 
diff
changeset
 | 
98  | 
"SEM c S \<equiv> \<Union>(sem c ` S)"  | 
| 13020 | 99  | 
|
| 35107 | 100  | 
abbreviation Omega :: "'a com"    ("\<Omega>" 63)
 | 
101  | 
where "\<Omega> \<equiv> While UNIV UNIV (Basic id)"  | 
|
| 13020 | 102  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35107 
diff
changeset
 | 
103  | 
primrec fwhile :: "'a bexp \<Rightarrow> 'a com \<Rightarrow> nat \<Rightarrow> 'a com" where  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35107 
diff
changeset
 | 
104  | 
"fwhile b c 0 = \<Omega>"  | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35107 
diff
changeset
 | 
105  | 
| "fwhile b c (Suc n) = Cond b (Seq c (fwhile b c n)) (Basic id)"  | 
| 13020 | 106  | 
|
| 59189 | 107  | 
subsubsection \<open>Proofs\<close>  | 
| 13020 | 108  | 
|
109  | 
declare ann_transition_transition.intros [intro]  | 
|
| 59189 | 110  | 
inductive_cases transition_cases:  | 
111  | 
"(Parallel T,s) -P1\<rightarrow> t"  | 
|
| 13020 | 112  | 
"(Basic f, s) -P1\<rightarrow> t"  | 
| 59189 | 113  | 
"(Seq c1 c2, s) -P1\<rightarrow> t"  | 
| 13020 | 114  | 
"(Cond b c1 c2, s) -P1\<rightarrow> t"  | 
115  | 
"(While b i c, s) -P1\<rightarrow> t"  | 
|
116  | 
||
| 59189 | 117  | 
lemma Parallel_empty_lemma [rule_format (no_asm)]:  | 
| 13020 | 118  | 
"(Parallel [],s) -Pn\<rightarrow> (Parallel Ts,t) \<longrightarrow> Ts=[] \<and> n=0 \<and> s=t"  | 
119  | 
apply(induct n)  | 
|
120  | 
apply(simp (no_asm))  | 
|
121  | 
apply clarify  | 
|
| 46362 | 122  | 
apply(drule relpow_Suc_D2)  | 
| 13020 | 123  | 
apply(force elim:transition_cases)  | 
124  | 
done  | 
|
125  | 
||
| 59189 | 126  | 
lemma Parallel_AllNone_lemma [rule_format (no_asm)]:  | 
| 13020 | 127  | 
"All_None Ss \<longrightarrow> (Parallel Ss,s) -Pn\<rightarrow> (Parallel Ts,t) \<longrightarrow> Ts=Ss \<and> n=0 \<and> s=t"  | 
128  | 
apply(induct "n")  | 
|
129  | 
apply(simp (no_asm))  | 
|
130  | 
apply clarify  | 
|
| 46362 | 131  | 
apply(drule relpow_Suc_D2)  | 
| 13020 | 132  | 
apply clarify  | 
133  | 
apply(erule transition_cases,simp_all)  | 
|
134  | 
apply(force dest:nth_mem simp add:All_None_def)  | 
|
135  | 
done  | 
|
136  | 
||
137  | 
lemma Parallel_AllNone: "All_None Ts \<Longrightarrow> (SEM (Parallel Ts) X) = X"  | 
|
138  | 
apply (unfold SEM_def sem_def)  | 
|
139  | 
apply auto  | 
|
| 46362 | 140  | 
apply(drule rtrancl_imp_UN_relpow)  | 
| 13020 | 141  | 
apply clarify  | 
142  | 
apply(drule Parallel_AllNone_lemma)  | 
|
143  | 
apply auto  | 
|
144  | 
done  | 
|
145  | 
||
146  | 
lemma Parallel_empty: "Ts=[] \<Longrightarrow> (SEM (Parallel Ts) X) = X"  | 
|
147  | 
apply(rule Parallel_AllNone)  | 
|
148  | 
apply(simp add:All_None_def)  | 
|
149  | 
done  | 
|
150  | 
||
| 59189 | 151  | 
text \<open>Set of lemmas from Apt and Olderog "Verification of sequential  | 
152  | 
and concurrent programs", page 63.\<close>  | 
|
| 13020 | 153  | 
|
| 59189 | 154  | 
lemma L3_5i: "X\<subseteq>Y \<Longrightarrow> SEM c X \<subseteq> SEM c Y"  | 
| 13020 | 155  | 
apply (unfold SEM_def)  | 
156  | 
apply force  | 
|
157  | 
done  | 
|
158  | 
||
| 59189 | 159  | 
lemma L3_5ii_lemma1:  | 
160  | 
"\<lbrakk> (c1, s1) -P*\<rightarrow> (Parallel Ts, s2); All_None Ts;  | 
|
161  | 
(c2, s2) -P*\<rightarrow> (Parallel Ss, s3); All_None Ss \<rbrakk>  | 
|
| 13020 | 162  | 
\<Longrightarrow> (Seq c1 c2, s1) -P*\<rightarrow> (Parallel Ss, s3)"  | 
163  | 
apply(erule converse_rtrancl_induct2)  | 
|
164  | 
apply(force intro:converse_rtrancl_into_rtrancl)+  | 
|
165  | 
done  | 
|
166  | 
||
| 59189 | 167  | 
lemma L3_5ii_lemma2 [rule_format (no_asm)]:  | 
168  | 
"\<forall>c1 c2 s t. (Seq c1 c2, s) -Pn\<rightarrow> (Parallel Ts, t) \<longrightarrow>  | 
|
169  | 
(All_None Ts) \<longrightarrow> (\<exists>y m Rs. (c1,s) -P*\<rightarrow> (Parallel Rs, y) \<and>  | 
|
| 13020 | 170  | 
(All_None Rs) \<and> (c2, y) -Pm\<rightarrow> (Parallel Ts, t) \<and> m \<le> n)"  | 
171  | 
apply(induct "n")  | 
|
172  | 
apply(force)  | 
|
| 46362 | 173  | 
apply(safe dest!: relpow_Suc_D2)  | 
| 13020 | 174  | 
apply(erule transition_cases,simp_all)  | 
175  | 
apply (fast intro!: le_SucI)  | 
|
| 46362 | 176  | 
apply (fast intro!: le_SucI elim!: relpow_imp_rtrancl converse_rtrancl_into_rtrancl)  | 
| 13020 | 177  | 
done  | 
178  | 
||
| 59189 | 179  | 
lemma L3_5ii_lemma3:  | 
180  | 
"\<lbrakk>(Seq c1 c2,s) -P*\<rightarrow> (Parallel Ts,t); All_None Ts\<rbrakk> \<Longrightarrow>  | 
|
181  | 
(\<exists>y Rs. (c1,s) -P*\<rightarrow> (Parallel Rs,y) \<and> All_None Rs  | 
|
| 13020 | 182  | 
\<and> (c2,y) -P*\<rightarrow> (Parallel Ts,t))"  | 
| 46362 | 183  | 
apply(drule rtrancl_imp_UN_relpow)  | 
184  | 
apply(fast dest: L3_5ii_lemma2 relpow_imp_rtrancl)  | 
|
| 13020 | 185  | 
done  | 
186  | 
||
187  | 
lemma L3_5ii: "SEM (Seq c1 c2) X = SEM c2 (SEM c1 X)"  | 
|
188  | 
apply (unfold SEM_def sem_def)  | 
|
189  | 
apply auto  | 
|
190  | 
apply(fast dest: L3_5ii_lemma3)  | 
|
191  | 
apply(fast elim: L3_5ii_lemma1)  | 
|
192  | 
done  | 
|
193  | 
||
194  | 
lemma L3_5iii: "SEM (Seq (Seq c1 c2) c3) X = SEM (Seq c1 (Seq c2 c3)) X"  | 
|
195  | 
apply (simp (no_asm) add: L3_5ii)  | 
|
196  | 
done  | 
|
197  | 
||
198  | 
lemma L3_5iv:  | 
|
199  | 
"SEM (Cond b c1 c2) X = (SEM c1 (X \<inter> b)) Un (SEM c2 (X \<inter> (-b)))"  | 
|
200  | 
apply (unfold SEM_def sem_def)  | 
|
201  | 
apply auto  | 
|
202  | 
apply(erule converse_rtranclE)  | 
|
203  | 
prefer 2  | 
|
204  | 
apply (erule transition_cases,simp_all)  | 
|
205  | 
apply(fast intro: converse_rtrancl_into_rtrancl elim: transition_cases)+  | 
|
206  | 
done  | 
|
207  | 
||
208  | 
||
| 59189 | 209  | 
lemma L3_5v_lemma1[rule_format]:  | 
| 13020 | 210  | 
"(S,s) -Pn\<rightarrow> (T,t) \<longrightarrow> S=\<Omega> \<longrightarrow> (\<not>(\<exists>Rs. T=(Parallel Rs) \<and> All_None Rs))"  | 
211  | 
apply (unfold UNIV_def)  | 
|
212  | 
apply(rule nat_less_induct)  | 
|
213  | 
apply safe  | 
|
| 46362 | 214  | 
apply(erule relpow_E2)  | 
| 13020 | 215  | 
apply simp_all  | 
216  | 
apply(erule transition_cases)  | 
|
217  | 
apply simp_all  | 
|
| 46362 | 218  | 
apply(erule relpow_E2)  | 
| 13020 | 219  | 
apply(simp add: Id_def)  | 
220  | 
apply(erule transition_cases,simp_all)  | 
|
221  | 
apply clarify  | 
|
222  | 
apply(erule transition_cases,simp_all)  | 
|
| 46362 | 223  | 
apply(erule relpow_E2,simp)  | 
| 13020 | 224  | 
apply clarify  | 
225  | 
apply(erule transition_cases)  | 
|
226  | 
apply simp+  | 
|
227  | 
apply clarify  | 
|
228  | 
apply(erule transition_cases)  | 
|
229  | 
apply simp_all  | 
|
230  | 
done  | 
|
231  | 
||
232  | 
lemma L3_5v_lemma2: "\<lbrakk>(\<Omega>, s) -P*\<rightarrow> (Parallel Ts, t); All_None Ts \<rbrakk> \<Longrightarrow> False"  | 
|
| 46362 | 233  | 
apply(fast dest: rtrancl_imp_UN_relpow L3_5v_lemma1)  | 
| 13020 | 234  | 
done  | 
235  | 
||
236  | 
lemma L3_5v_lemma3: "SEM (\<Omega>) S = {}"
 | 
|
237  | 
apply (unfold SEM_def sem_def)  | 
|
238  | 
apply(fast dest: L3_5v_lemma2)  | 
|
239  | 
done  | 
|
240  | 
||
| 59189 | 241  | 
lemma L3_5v_lemma4 [rule_format]:  | 
242  | 
"\<forall>s. (While b i c, s) -Pn\<rightarrow> (Parallel Ts, t) \<longrightarrow> All_None Ts \<longrightarrow>  | 
|
| 13020 | 243  | 
(\<exists>k. (fwhile b c k, s) -P*\<rightarrow> (Parallel Ts, t))"  | 
244  | 
apply(rule nat_less_induct)  | 
|
245  | 
apply safe  | 
|
| 46362 | 246  | 
apply(erule relpow_E2)  | 
| 13020 | 247  | 
apply safe  | 
248  | 
apply(erule transition_cases,simp_all)  | 
|
249  | 
apply (rule_tac x = "1" in exI)  | 
|
250  | 
apply(force dest: Parallel_empty_lemma intro: converse_rtrancl_into_rtrancl simp add: Id_def)  | 
|
251  | 
apply safe  | 
|
252  | 
apply(drule L3_5ii_lemma2)  | 
|
253  | 
apply safe  | 
|
254  | 
apply(drule le_imp_less_Suc)  | 
|
255  | 
apply (erule allE , erule impE,assumption)  | 
|
256  | 
apply (erule allE , erule impE, assumption)  | 
|
257  | 
apply safe  | 
|
258  | 
apply (rule_tac x = "k+1" in exI)  | 
|
259  | 
apply(simp (no_asm))  | 
|
260  | 
apply(rule converse_rtrancl_into_rtrancl)  | 
|
261  | 
apply fast  | 
|
262  | 
apply(fast elim: L3_5ii_lemma1)  | 
|
263  | 
done  | 
|
264  | 
||
| 59189 | 265  | 
lemma L3_5v_lemma5 [rule_format]:  | 
266  | 
"\<forall>s. (fwhile b c k, s) -P*\<rightarrow> (Parallel Ts, t) \<longrightarrow> All_None Ts \<longrightarrow>  | 
|
| 13020 | 267  | 
(While b i c, s) -P*\<rightarrow> (Parallel Ts,t)"  | 
268  | 
apply(induct "k")  | 
|
269  | 
apply(force dest: L3_5v_lemma2)  | 
|
270  | 
apply safe  | 
|
271  | 
apply(erule converse_rtranclE)  | 
|
272  | 
apply simp_all  | 
|
273  | 
apply(erule transition_cases,simp_all)  | 
|
274  | 
apply(rule converse_rtrancl_into_rtrancl)  | 
|
275  | 
apply(fast)  | 
|
276  | 
apply(fast elim!: L3_5ii_lemma1 dest: L3_5ii_lemma3)  | 
|
| 46362 | 277  | 
apply(drule rtrancl_imp_UN_relpow)  | 
| 13020 | 278  | 
apply clarify  | 
| 46362 | 279  | 
apply(erule relpow_E2)  | 
| 13020 | 280  | 
apply simp_all  | 
281  | 
apply(erule transition_cases,simp_all)  | 
|
282  | 
apply(fast dest: Parallel_empty_lemma)  | 
|
283  | 
done  | 
|
284  | 
||
285  | 
lemma L3_5v: "SEM (While b i c) = (\<lambda>x. (\<Union>k. SEM (fwhile b c k) x))"  | 
|
286  | 
apply(rule ext)  | 
|
287  | 
apply (simp add: SEM_def sem_def)  | 
|
288  | 
apply safe  | 
|
| 46362 | 289  | 
apply(drule rtrancl_imp_UN_relpow,simp)  | 
| 13020 | 290  | 
apply clarify  | 
291  | 
apply(fast dest:L3_5v_lemma4)  | 
|
292  | 
apply(fast intro: L3_5v_lemma5)  | 
|
293  | 
done  | 
|
294  | 
||
| 59189 | 295  | 
section \<open>Validity of Correctness Formulas\<close>  | 
| 13020 | 296  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35107 
diff
changeset
 | 
297  | 
definition com_validity :: "'a assn \<Rightarrow> 'a com \<Rightarrow> 'a assn \<Rightarrow> bool" ("(3\<parallel>= _// _//_)" [90,55,90] 50) where
 | 
| 13020 | 298  | 
"\<parallel>= p c q \<equiv> SEM c p \<subseteq> q"  | 
299  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
35107 
diff
changeset
 | 
300  | 
definition ann_com_validity :: "'a ann_com \<Rightarrow> 'a assn \<Rightarrow> bool" ("\<Turnstile> _ _" [60,90] 45) where
 | 
| 13020 | 301  | 
"\<Turnstile> c q \<equiv> ann_SEM c (pre c) \<subseteq> q"  | 
302  | 
||
303  | 
end  |