| author | nipkow | 
| Thu, 03 Nov 2022 14:20:07 +0100 | |
| changeset 76422 | 2612b3406b61 | 
| parent 71989 | bad75618fb82 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 59189 | 1  | 
section \<open>Operational Semantics\<close>  | 
| 13020 | 2  | 
|
| 15425 | 3  | 
theory RG_Tran  | 
4  | 
imports RG_Com  | 
|
5  | 
begin  | 
|
| 13020 | 6  | 
|
| 59189 | 7  | 
subsection \<open>Semantics of Component Programs\<close>  | 
| 13020 | 8  | 
|
| 59189 | 9  | 
subsubsection \<open>Environment transitions\<close>  | 
| 13020 | 10  | 
|
| 42174 | 11  | 
type_synonym 'a conf = "(('a com) option) \<times> 'a"
 | 
| 13020 | 12  | 
|
| 23746 | 13  | 
inductive_set  | 
14  | 
  etran :: "('a conf \<times> 'a conf) set" 
 | 
|
15  | 
  and etran' :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool"  ("_ -e\<rightarrow> _" [81,81] 80)
 | 
|
16  | 
where  | 
|
17  | 
"P -e\<rightarrow> Q \<equiv> (P,Q) \<in> etran"  | 
|
18  | 
| Env: "(P, s) -e\<rightarrow> (P, t)"  | 
|
19  | 
||
20  | 
lemma etranE: "c -e\<rightarrow> c' \<Longrightarrow> (\<And>P s t. c = (P, s) \<Longrightarrow> c' = (P, t) \<Longrightarrow> Q) \<Longrightarrow> Q"  | 
|
21  | 
by (induct c, induct c', erule etran.cases, blast)  | 
|
| 13020 | 22  | 
|
| 59189 | 23  | 
subsubsection \<open>Component transitions\<close>  | 
| 13020 | 24  | 
|
| 23746 | 25  | 
inductive_set  | 
26  | 
  ctran :: "('a conf \<times> 'a conf) set"
 | 
|
27  | 
  and ctran' :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool"   ("_ -c\<rightarrow> _" [81,81] 80)
 | 
|
28  | 
  and ctrans :: "'a conf \<Rightarrow> 'a conf \<Rightarrow> bool"   ("_ -c*\<rightarrow> _" [81,81] 80)
 | 
|
29  | 
where  | 
|
30  | 
"P -c\<rightarrow> Q \<equiv> (P,Q) \<in> ctran"  | 
|
| 67613 | 31  | 
| "P -c*\<rightarrow> Q \<equiv> (P,Q) \<in> ctran\<^sup>*"  | 
| 13020 | 32  | 
|
| 23746 | 33  | 
| Basic: "(Some(Basic f), s) -c\<rightarrow> (None, f s)"  | 
| 13020 | 34  | 
|
| 23746 | 35  | 
| Seq1: "(Some P0, s) -c\<rightarrow> (None, t) \<Longrightarrow> (Some(Seq P0 P1), s) -c\<rightarrow> (Some P1, t)"  | 
| 13020 | 36  | 
|
| 23746 | 37  | 
| Seq2: "(Some P0, s) -c\<rightarrow> (Some P2, t) \<Longrightarrow> (Some(Seq P0 P1), s) -c\<rightarrow> (Some(Seq P2 P1), t)"  | 
| 13020 | 38  | 
|
| 23746 | 39  | 
| CondT: "s\<in>b \<Longrightarrow> (Some(Cond b P1 P2), s) -c\<rightarrow> (Some P1, s)"  | 
40  | 
| CondF: "s\<notin>b \<Longrightarrow> (Some(Cond b P1 P2), s) -c\<rightarrow> (Some P2, s)"  | 
|
| 13020 | 41  | 
|
| 23746 | 42  | 
| WhileF: "s\<notin>b \<Longrightarrow> (Some(While b P), s) -c\<rightarrow> (None, s)"  | 
43  | 
| WhileT: "s\<in>b \<Longrightarrow> (Some(While b P), s) -c\<rightarrow> (Some(Seq P (While b P)), s)"  | 
|
| 13020 | 44  | 
|
| 23746 | 45  | 
| Await: "\<lbrakk>s\<in>b; (Some P, s) -c*\<rightarrow> (None, t)\<rbrakk> \<Longrightarrow> (Some(Await b P), s) -c\<rightarrow> (None, t)"  | 
| 13020 | 46  | 
|
47  | 
monos "rtrancl_mono"  | 
|
48  | 
||
| 59189 | 49  | 
subsection \<open>Semantics of Parallel Programs\<close>  | 
| 13020 | 50  | 
|
| 42174 | 51  | 
type_synonym 'a par_conf = "('a par_com) \<times> 'a"
 | 
| 23746 | 52  | 
|
53  | 
inductive_set  | 
|
| 13020 | 54  | 
  par_etran :: "('a par_conf \<times> 'a par_conf) set"
 | 
| 23746 | 55  | 
  and par_etran' :: "['a par_conf,'a par_conf] \<Rightarrow> bool" ("_ -pe\<rightarrow> _" [81,81] 80)
 | 
56  | 
where  | 
|
57  | 
"P -pe\<rightarrow> Q \<equiv> (P,Q) \<in> par_etran"  | 
|
58  | 
| ParEnv: "(Ps, s) -pe\<rightarrow> (Ps, t)"  | 
|
59  | 
||
60  | 
inductive_set  | 
|
| 13020 | 61  | 
  par_ctran :: "('a par_conf \<times> 'a par_conf) set"
 | 
| 23746 | 62  | 
  and par_ctran' :: "['a par_conf,'a par_conf] \<Rightarrow> bool" ("_ -pc\<rightarrow> _" [81,81] 80)
 | 
63  | 
where  | 
|
64  | 
"P -pc\<rightarrow> Q \<equiv> (P,Q) \<in> par_ctran"  | 
|
65  | 
| ParComp: "\<lbrakk>i<length Ps; (Ps!i, s) -c\<rightarrow> (r, t)\<rbrakk> \<Longrightarrow> (Ps, s) -pc\<rightarrow> (Ps[i:=r], t)"  | 
|
| 13020 | 66  | 
|
| 23746 | 67  | 
lemma par_ctranE: "c -pc\<rightarrow> c' \<Longrightarrow>  | 
68  | 
(\<And>i Ps s r t. c = (Ps, s) \<Longrightarrow> c' = (Ps[i := r], t) \<Longrightarrow> i < length Ps \<Longrightarrow>  | 
|
69  | 
(Ps ! i, s) -c\<rightarrow> (r, t) \<Longrightarrow> P) \<Longrightarrow> P"  | 
|
70  | 
by (induct c, induct c', erule par_ctran.cases, blast)  | 
|
| 13020 | 71  | 
|
| 59189 | 72  | 
subsection \<open>Computations\<close>  | 
| 13020 | 73  | 
|
| 59189 | 74  | 
subsubsection \<open>Sequential computations\<close>  | 
| 13020 | 75  | 
|
| 42174 | 76  | 
type_synonym 'a confs = "'a conf list"  | 
| 23746 | 77  | 
|
| 42174 | 78  | 
inductive_set cptn :: "'a confs set"  | 
| 23746 | 79  | 
where  | 
| 13020 | 80  | 
CptnOne: "[(P,s)] \<in> cptn"  | 
| 23746 | 81  | 
| CptnEnv: "(P, t)#xs \<in> cptn \<Longrightarrow> (P,s)#(P,t)#xs \<in> cptn"  | 
82  | 
| CptnComp: "\<lbrakk>(P,s) -c\<rightarrow> (Q,t); (Q, t)#xs \<in> cptn \<rbrakk> \<Longrightarrow> (P,s)#(Q,t)#xs \<in> cptn"  | 
|
| 13020 | 83  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
84  | 
definition cp :: "('a com) option \<Rightarrow> 'a \<Rightarrow> ('a confs) set" where
 | 
| 13020 | 85  | 
  "cp P s \<equiv> {l. l!0=(P,s) \<and> l \<in> cptn}"  
 | 
86  | 
||
| 59189 | 87  | 
subsubsection \<open>Parallel computations\<close>  | 
| 13020 | 88  | 
|
| 42174 | 89  | 
type_synonym 'a par_confs = "'a par_conf list"  | 
| 23746 | 90  | 
|
| 42174 | 91  | 
inductive_set par_cptn :: "'a par_confs set"  | 
| 23746 | 92  | 
where  | 
| 13020 | 93  | 
ParCptnOne: "[(P,s)] \<in> par_cptn"  | 
| 23746 | 94  | 
| ParCptnEnv: "(P,t)#xs \<in> par_cptn \<Longrightarrow> (P,s)#(P,t)#xs \<in> par_cptn"  | 
95  | 
| ParCptnComp: "\<lbrakk> (P,s) -pc\<rightarrow> (Q,t); (Q,t)#xs \<in> par_cptn \<rbrakk> \<Longrightarrow> (P,s)#(Q,t)#xs \<in> par_cptn"  | 
|
| 13020 | 96  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
97  | 
definition par_cp :: "'a par_com \<Rightarrow> 'a \<Rightarrow> ('a par_confs) set" where
 | 
| 13020 | 98  | 
  "par_cp P s \<equiv> {l. l!0=(P,s) \<and> l \<in> par_cptn}"  
 | 
99  | 
||
| 59189 | 100  | 
subsection\<open>Modular Definition of Computation\<close>  | 
| 13020 | 101  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
102  | 
definition lift :: "'a com \<Rightarrow> 'a conf \<Rightarrow> 'a conf" where  | 
| 13020 | 103  | 
"lift Q \<equiv> \<lambda>(P, s). (if P=None then (Some Q,s) else (Some(Seq (the P) Q), s))"  | 
104  | 
||
| 23746 | 105  | 
inductive_set cptn_mod :: "('a confs) set"
 | 
106  | 
where  | 
|
| 13020 | 107  | 
CptnModOne: "[(P, s)] \<in> cptn_mod"  | 
| 23746 | 108  | 
| CptnModEnv: "(P, t)#xs \<in> cptn_mod \<Longrightarrow> (P, s)#(P, t)#xs \<in> cptn_mod"  | 
109  | 
| CptnModNone: "\<lbrakk>(Some P, s) -c\<rightarrow> (None, t); (None, t)#xs \<in> cptn_mod \<rbrakk> \<Longrightarrow> (Some P,s)#(None, t)#xs \<in>cptn_mod"  | 
|
110  | 
| CptnModCondT: "\<lbrakk>(Some P0, s)#ys \<in> cptn_mod; s \<in> b \<rbrakk> \<Longrightarrow> (Some(Cond b P0 P1), s)#(Some P0, s)#ys \<in> cptn_mod"  | 
|
111  | 
| CptnModCondF: "\<lbrakk>(Some P1, s)#ys \<in> cptn_mod; s \<notin> b \<rbrakk> \<Longrightarrow> (Some(Cond b P0 P1), s)#(Some P1, s)#ys \<in> cptn_mod"  | 
|
112  | 
| CptnModSeq1: "\<lbrakk>(Some P0, s)#xs \<in> cptn_mod; zs=map (lift P1) xs \<rbrakk>  | 
|
| 13020 | 113  | 
\<Longrightarrow> (Some(Seq P0 P1), s)#zs \<in> cptn_mod"  | 
| 23746 | 114  | 
| CptnModSeq2:  | 
| 13020 | 115  | 
"\<lbrakk>(Some P0, s)#xs \<in> cptn_mod; fst(last ((Some P0, s)#xs)) = None;  | 
116  | 
(Some P1, snd(last ((Some P0, s)#xs)))#ys \<in> cptn_mod;  | 
|
117  | 
zs=(map (lift P1) xs)@ys \<rbrakk> \<Longrightarrow> (Some(Seq P0 P1), s)#zs \<in> cptn_mod"  | 
|
118  | 
||
| 23746 | 119  | 
| CptnModWhile1:  | 
| 13020 | 120  | 
"\<lbrakk> (Some P, s)#xs \<in> cptn_mod; s \<in> b; zs=map (lift (While b P)) xs \<rbrakk>  | 
121  | 
\<Longrightarrow> (Some(While b P), s)#(Some(Seq P (While b P)), s)#zs \<in> cptn_mod"  | 
|
| 23746 | 122  | 
| CptnModWhile2:  | 
| 13020 | 123  | 
"\<lbrakk> (Some P, s)#xs \<in> cptn_mod; fst(last ((Some P, s)#xs))=None; s \<in> b;  | 
124  | 
zs=(map (lift (While b P)) xs)@ys;  | 
|
125  | 
(Some(While b P), snd(last ((Some P, s)#xs)))#ys \<in> cptn_mod\<rbrakk>  | 
|
126  | 
\<Longrightarrow> (Some(While b P), s)#(Some(Seq P (While b P)), s)#zs \<in> cptn_mod"  | 
|
127  | 
||
| 59189 | 128  | 
subsection \<open>Equivalence of Both Definitions.\<close>  | 
| 13020 | 129  | 
|
130  | 
lemma last_length: "((a#xs)!(length xs))=last (a#xs)"  | 
|
| 51119 | 131  | 
by (induct xs) auto  | 
| 13020 | 132  | 
|
133  | 
lemma div_seq [rule_format]: "list \<in> cptn_mod \<Longrightarrow>  | 
|
134  | 
(\<forall>s P Q zs. list=(Some (Seq P Q), s)#zs \<longrightarrow>  | 
|
135  | 
(\<exists>xs. (Some P, s)#xs \<in> cptn_mod \<and> (zs=(map (lift Q) xs) \<or>  | 
|
136  | 
( fst(((Some P, s)#xs)!length xs)=None \<and>  | 
|
137  | 
(\<exists>ys. (Some Q, snd(((Some P, s)#xs)!length xs))#ys \<in> cptn_mod  | 
|
138  | 
\<and> zs=(map (lift (Q)) xs)@ys)))))"  | 
|
139  | 
apply(erule cptn_mod.induct)  | 
|
140  | 
apply simp_all  | 
|
141  | 
apply clarify  | 
|
142  | 
apply(force intro:CptnModOne)  | 
|
143  | 
apply clarify  | 
|
144  | 
apply(erule_tac x=Pa in allE)  | 
|
145  | 
apply(erule_tac x=Q in allE)  | 
|
146  | 
apply simp  | 
|
147  | 
apply clarify  | 
|
148  | 
apply(erule disjE)  | 
|
149  | 
apply(rule_tac x="(Some Pa,t)#xsa" in exI)  | 
|
150  | 
apply(rule conjI)  | 
|
151  | 
apply clarify  | 
|
152  | 
apply(erule CptnModEnv)  | 
|
153  | 
apply(rule disjI1)  | 
|
154  | 
apply(simp add:lift_def)  | 
|
155  | 
apply clarify  | 
|
156  | 
apply(rule_tac x="(Some Pa,t)#xsa" in exI)  | 
|
157  | 
apply(rule conjI)  | 
|
158  | 
apply(erule CptnModEnv)  | 
|
159  | 
apply(rule disjI2)  | 
|
160  | 
apply(rule conjI)  | 
|
161  | 
apply(case_tac xsa,simp,simp)  | 
|
162  | 
apply(rule_tac x="ys" in exI)  | 
|
163  | 
apply(rule conjI)  | 
|
164  | 
apply simp  | 
|
165  | 
apply(simp add:lift_def)  | 
|
166  | 
apply clarify  | 
|
| 23746 | 167  | 
apply(erule ctran.cases,simp_all)  | 
| 13020 | 168  | 
apply clarify  | 
169  | 
apply(rule_tac x="xs" in exI)  | 
|
170  | 
apply simp  | 
|
171  | 
apply clarify  | 
|
172  | 
apply(rule_tac x="xs" in exI)  | 
|
173  | 
apply(simp add: last_length)  | 
|
174  | 
done  | 
|
175  | 
||
176  | 
lemma cptn_onlyif_cptn_mod_aux [rule_format]:  | 
|
177  | 
"\<forall>s Q t xs.((Some a, s), Q, t) \<in> ctran \<longrightarrow> (Q, t) # xs \<in> cptn_mod  | 
|
178  | 
\<longrightarrow> (Some a, s) # (Q, t) # xs \<in> cptn_mod"  | 
|
| 
71989
 
bad75618fb82
extraction of equations x = t from premises beneath meta-all
 
haftmann 
parents: 
67613 
diff
changeset
 | 
179  | 
supply [[simproc del: defined_all]]  | 
| 13020 | 180  | 
apply(induct a)  | 
181  | 
apply simp_all  | 
|
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
62390 
diff
changeset
 | 
182  | 
\<comment> \<open>basic\<close>  | 
| 13020 | 183  | 
apply clarify  | 
| 23746 | 184  | 
apply(erule ctran.cases,simp_all)  | 
| 13020 | 185  | 
apply(rule CptnModNone,rule Basic,simp)  | 
186  | 
apply clarify  | 
|
| 23746 | 187  | 
apply(erule ctran.cases,simp_all)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
62390 
diff
changeset
 | 
188  | 
\<comment> \<open>Seq1\<close>  | 
| 13020 | 189  | 
apply(rule_tac xs="[(None,ta)]" in CptnModSeq2)  | 
190  | 
apply(erule CptnModNone)  | 
|
191  | 
apply(rule CptnModOne)  | 
|
192  | 
apply simp  | 
|
193  | 
apply simp  | 
|
194  | 
apply(simp add:lift_def)  | 
|
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
62390 
diff
changeset
 | 
195  | 
\<comment> \<open>Seq2\<close>  | 
| 13020 | 196  | 
apply(erule_tac x=sa in allE)  | 
197  | 
apply(erule_tac x="Some P2" in allE)  | 
|
198  | 
apply(erule allE,erule impE, assumption)  | 
|
199  | 
apply(drule div_seq,simp)  | 
|
200  | 
apply clarify  | 
|
201  | 
apply(erule disjE)  | 
|
202  | 
apply clarify  | 
|
203  | 
apply(erule allE,erule impE, assumption)  | 
|
204  | 
apply(erule_tac CptnModSeq1)  | 
|
205  | 
apply(simp add:lift_def)  | 
|
206  | 
apply clarify  | 
|
207  | 
apply(erule allE,erule impE, assumption)  | 
|
208  | 
apply(erule_tac CptnModSeq2)  | 
|
209  | 
apply (simp add:last_length)  | 
|
210  | 
apply (simp add:last_length)  | 
|
211  | 
apply(simp add:lift_def)  | 
|
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
62390 
diff
changeset
 | 
212  | 
\<comment> \<open>Cond\<close>  | 
| 13020 | 213  | 
apply clarify  | 
| 23746 | 214  | 
apply(erule ctran.cases,simp_all)  | 
| 13020 | 215  | 
apply(force elim: CptnModCondT)  | 
216  | 
apply(force elim: CptnModCondF)  | 
|
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
62390 
diff
changeset
 | 
217  | 
\<comment> \<open>While\<close>  | 
| 13020 | 218  | 
apply clarify  | 
| 23746 | 219  | 
apply(erule ctran.cases,simp_all)  | 
| 13020 | 220  | 
apply(rule CptnModNone,erule WhileF,simp)  | 
221  | 
apply(drule div_seq,force)  | 
|
222  | 
apply clarify  | 
|
223  | 
apply (erule disjE)  | 
|
224  | 
apply(force elim:CptnModWhile1)  | 
|
225  | 
apply clarify  | 
|
226  | 
apply(force simp add:last_length elim:CptnModWhile2)  | 
|
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
62390 
diff
changeset
 | 
227  | 
\<comment> \<open>await\<close>  | 
| 13020 | 228  | 
apply clarify  | 
| 23746 | 229  | 
apply(erule ctran.cases,simp_all)  | 
| 13020 | 230  | 
apply(rule CptnModNone,erule Await,simp+)  | 
231  | 
done  | 
|
232  | 
||
233  | 
lemma cptn_onlyif_cptn_mod [rule_format]: "c \<in> cptn \<Longrightarrow> c \<in> cptn_mod"  | 
|
234  | 
apply(erule cptn.induct)  | 
|
235  | 
apply(rule CptnModOne)  | 
|
236  | 
apply(erule CptnModEnv)  | 
|
237  | 
apply(case_tac P)  | 
|
238  | 
apply simp  | 
|
| 23746 | 239  | 
apply(erule ctran.cases,simp_all)  | 
| 13020 | 240  | 
apply(force elim:cptn_onlyif_cptn_mod_aux)  | 
241  | 
done  | 
|
242  | 
||
243  | 
lemma lift_is_cptn: "c\<in>cptn \<Longrightarrow> map (lift P) c \<in> cptn"  | 
|
244  | 
apply(erule cptn.induct)  | 
|
245  | 
apply(force simp add:lift_def CptnOne)  | 
|
246  | 
apply(force intro:CptnEnv simp add:lift_def)  | 
|
| 23746 | 247  | 
apply(force simp add:lift_def intro:CptnComp Seq2 Seq1 elim:ctran.cases)  | 
| 13020 | 248  | 
done  | 
249  | 
||
250  | 
lemma cptn_append_is_cptn [rule_format]:  | 
|
251  | 
"\<forall>b a. b#c1\<in>cptn \<longrightarrow> a#c2\<in>cptn \<longrightarrow> (b#c1)!length c1=a \<longrightarrow> b#c1@c2\<in>cptn"  | 
|
252  | 
apply(induct c1)  | 
|
253  | 
apply simp  | 
|
254  | 
apply clarify  | 
|
| 23746 | 255  | 
apply(erule cptn.cases,simp_all)  | 
| 13020 | 256  | 
apply(force intro:CptnEnv)  | 
257  | 
apply(force elim:CptnComp)  | 
|
258  | 
done  | 
|
259  | 
||
260  | 
lemma last_lift: "\<lbrakk>xs\<noteq>[]; fst(xs!(length xs - (Suc 0)))=None\<rbrakk>  | 
|
261  | 
\<Longrightarrow> fst((map (lift P) xs)!(length (map (lift P) xs)- (Suc 0)))=(Some P)"  | 
|
| 51119 | 262  | 
by (cases "(xs ! (length xs - (Suc 0)))") (simp add:lift_def)  | 
| 13020 | 263  | 
|
264  | 
lemma last_fst [rule_format]: "P((a#x)!length x) \<longrightarrow> \<not>P a \<longrightarrow> P (x!(length x - (Suc 0)))"  | 
|
| 51119 | 265  | 
by (induct x) simp_all  | 
| 13020 | 266  | 
|
267  | 
lemma last_fst_esp:  | 
|
268  | 
"fst(((Some a,s)#xs)!(length xs))=None \<Longrightarrow> fst(xs!(length xs - (Suc 0)))=None"  | 
|
269  | 
apply(erule last_fst)  | 
|
270  | 
apply simp  | 
|
271  | 
done  | 
|
272  | 
||
273  | 
lemma last_snd: "xs\<noteq>[] \<Longrightarrow>  | 
|
274  | 
snd(((map (lift P) xs))!(length (map (lift P) xs) - (Suc 0)))=snd(xs!(length xs - (Suc 0)))"  | 
|
| 51119 | 275  | 
by (cases "(xs ! (length xs - (Suc 0)))") (simp_all add:lift_def)  | 
| 13020 | 276  | 
|
277  | 
lemma Cons_lift: "(Some (Seq P Q), s) # (map (lift Q) xs) = map (lift Q) ((Some P, s) # xs)"  | 
|
| 51119 | 278  | 
by (simp add:lift_def)  | 
| 13020 | 279  | 
|
280  | 
lemma Cons_lift_append:  | 
|
281  | 
"(Some (Seq P Q), s) # (map (lift Q) xs) @ ys = map (lift Q) ((Some P, s) # xs)@ ys "  | 
|
| 51119 | 282  | 
by (simp add:lift_def)  | 
| 13020 | 283  | 
|
284  | 
lemma lift_nth: "i<length xs \<Longrightarrow> map (lift Q) xs ! i = lift Q (xs! i)"  | 
|
| 51119 | 285  | 
by (simp add:lift_def)  | 
| 13020 | 286  | 
|
287  | 
lemma snd_lift: "i< length xs \<Longrightarrow> snd(lift Q (xs ! i))= snd (xs ! i)"  | 
|
| 51119 | 288  | 
by (cases "xs!i") (simp add:lift_def)  | 
| 13020 | 289  | 
|
290  | 
lemma cptn_if_cptn_mod: "c \<in> cptn_mod \<Longrightarrow> c \<in> cptn"  | 
|
291  | 
apply(erule cptn_mod.induct)  | 
|
292  | 
apply(rule CptnOne)  | 
|
293  | 
apply(erule CptnEnv)  | 
|
294  | 
apply(erule CptnComp,simp)  | 
|
295  | 
apply(rule CptnComp)  | 
|
| 41842 | 296  | 
apply(erule CondT,simp)  | 
| 13020 | 297  | 
apply(rule CptnComp)  | 
| 41842 | 298  | 
apply(erule CondF,simp)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
62390 
diff
changeset
 | 
299  | 
\<comment> \<open>Seq1\<close>  | 
| 23746 | 300  | 
apply(erule cptn.cases,simp_all)  | 
| 13020 | 301  | 
apply(rule CptnOne)  | 
302  | 
apply clarify  | 
|
303  | 
apply(drule_tac P=P1 in lift_is_cptn)  | 
|
304  | 
apply(simp add:lift_def)  | 
|
305  | 
apply(rule CptnEnv,simp)  | 
|
306  | 
apply clarify  | 
|
307  | 
apply(simp add:lift_def)  | 
|
308  | 
apply(rule conjI)  | 
|
309  | 
apply clarify  | 
|
310  | 
apply(rule CptnComp)  | 
|
311  | 
apply(rule Seq1,simp)  | 
|
312  | 
apply(drule_tac P=P1 in lift_is_cptn)  | 
|
313  | 
apply(simp add:lift_def)  | 
|
314  | 
apply clarify  | 
|
315  | 
apply(rule CptnComp)  | 
|
316  | 
apply(rule Seq2,simp)  | 
|
317  | 
apply(drule_tac P=P1 in lift_is_cptn)  | 
|
318  | 
apply(simp add:lift_def)  | 
|
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
62390 
diff
changeset
 | 
319  | 
\<comment> \<open>Seq2\<close>  | 
| 13020 | 320  | 
apply(rule cptn_append_is_cptn)  | 
321  | 
apply(drule_tac P=P1 in lift_is_cptn)  | 
|
322  | 
apply(simp add:lift_def)  | 
|
323  | 
apply simp  | 
|
| 62390 | 324  | 
apply(simp split: if_split_asm)  | 
| 41842 | 325  | 
apply(frule_tac P=P1 in last_lift)  | 
326  | 
apply(rule last_fst_esp)  | 
|
327  | 
apply (simp add:last_length)  | 
|
328  | 
apply(simp add:Cons_lift lift_def split_def last_conv_nth)  | 
|
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
62390 
diff
changeset
 | 
329  | 
\<comment> \<open>While1\<close>  | 
| 13020 | 330  | 
apply(rule CptnComp)  | 
| 41842 | 331  | 
apply(rule WhileT,simp)  | 
| 13020 | 332  | 
apply(drule_tac P="While b P" in lift_is_cptn)  | 
333  | 
apply(simp add:lift_def)  | 
|
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
62390 
diff
changeset
 | 
334  | 
\<comment> \<open>While2\<close>  | 
| 13020 | 335  | 
apply(rule CptnComp)  | 
| 41842 | 336  | 
apply(rule WhileT,simp)  | 
| 13020 | 337  | 
apply(rule cptn_append_is_cptn)  | 
| 41842 | 338  | 
apply(drule_tac P="While b P" in lift_is_cptn)  | 
| 13020 | 339  | 
apply(simp add:lift_def)  | 
340  | 
apply simp  | 
|
| 62390 | 341  | 
apply(simp split: if_split_asm)  | 
| 41842 | 342  | 
apply(frule_tac P="While b P" in last_lift)  | 
343  | 
apply(rule last_fst_esp,simp add:last_length)  | 
|
344  | 
apply(simp add:Cons_lift lift_def split_def last_conv_nth)  | 
|
| 13020 | 345  | 
done  | 
346  | 
||
347  | 
theorem cptn_iff_cptn_mod: "(c \<in> cptn) = (c \<in> cptn_mod)"  | 
|
348  | 
apply(rule iffI)  | 
|
349  | 
apply(erule cptn_onlyif_cptn_mod)  | 
|
350  | 
apply(erule cptn_if_cptn_mod)  | 
|
351  | 
done  | 
|
352  | 
||
| 59189 | 353  | 
section \<open>Validity of Correctness Formulas\<close>  | 
| 13020 | 354  | 
|
| 59189 | 355  | 
subsection \<open>Validity for Component Programs.\<close>  | 
| 13020 | 356  | 
|
| 42174 | 357  | 
type_synonym 'a rgformula =  | 
358  | 
  "'a com \<times> 'a set \<times> ('a \<times> 'a) set \<times> ('a \<times> 'a) set \<times> 'a set"
 | 
|
| 13020 | 359  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
360  | 
definition assum :: "('a set \<times> ('a \<times> 'a) set) \<Rightarrow> ('a confs) set" where
 | 
| 13020 | 361  | 
  "assum \<equiv> \<lambda>(pre, rely). {c. snd(c!0) \<in> pre \<and> (\<forall>i. Suc i<length c \<longrightarrow> 
 | 
362  | 
c!i -e\<rightarrow> c!(Suc i) \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> rely)}"  | 
|
363  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
364  | 
definition comm :: "(('a \<times> 'a) set \<times> 'a set) \<Rightarrow> ('a confs) set" where
 | 
| 13020 | 365  | 
  "comm \<equiv> \<lambda>(guar, post). {c. (\<forall>i. Suc i<length c \<longrightarrow> 
 | 
366  | 
c!i -c\<rightarrow> c!(Suc i) \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> guar) \<and>  | 
|
367  | 
(fst (last c) = None \<longrightarrow> snd (last c) \<in> post)}"  | 
|
368  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
369  | 
definition com_validity :: "'a com \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> 'a set \<Rightarrow> bool" 
 | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
370  | 
                 ("\<Turnstile> _ sat [_, _, _, _]" [60,0,0,0,0] 45) where
 | 
| 13020 | 371  | 
"\<Turnstile> P sat [pre, rely, guar, post] \<equiv>  | 
372  | 
\<forall>s. cp (Some P) s \<inter> assum(pre, rely) \<subseteq> comm(guar, post)"  | 
|
373  | 
||
| 59189 | 374  | 
subsection \<open>Validity for Parallel Programs.\<close>  | 
| 13020 | 375  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
376  | 
definition All_None :: "(('a com) option) list \<Rightarrow> bool" where
 | 
| 13020 | 377  | 
"All_None xs \<equiv> \<forall>c\<in>set xs. c=None"  | 
378  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
379  | 
definition par_assum :: "('a set \<times> ('a \<times> 'a) set) \<Rightarrow> ('a par_confs) set" where
 | 
| 13020 | 380  | 
  "par_assum \<equiv> \<lambda>(pre, rely). {c. snd(c!0) \<in> pre \<and> (\<forall>i. Suc i<length c \<longrightarrow> 
 | 
381  | 
c!i -pe\<rightarrow> c!Suc i \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> rely)}"  | 
|
382  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
383  | 
definition par_comm :: "(('a \<times> 'a) set \<times> 'a set) \<Rightarrow> ('a par_confs) set" where
 | 
| 13020 | 384  | 
  "par_comm \<equiv> \<lambda>(guar, post). {c. (\<forall>i. Suc i<length c \<longrightarrow>   
 | 
385  | 
c!i -pc\<rightarrow> c!Suc i \<longrightarrow> (snd(c!i), snd(c!Suc i)) \<in> guar) \<and>  | 
|
386  | 
(All_None (fst (last c)) \<longrightarrow> snd( last c) \<in> post)}"  | 
|
387  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
388  | 
definition par_com_validity :: "'a  par_com \<Rightarrow> 'a set \<Rightarrow> ('a \<times> 'a) set \<Rightarrow> ('a \<times> 'a) set 
 | 
| 
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
389  | 
\<Rightarrow> 'a set \<Rightarrow> bool"  ("\<Turnstile> _ SAT [_, _, _, _]" [60,0,0,0,0] 45) where
 | 
| 13020 | 390  | 
"\<Turnstile> Ps SAT [pre, rely, guar, post] \<equiv>  | 
391  | 
\<forall>s. par_cp Ps s \<inter> par_assum(pre, rely) \<subseteq> par_comm(guar, post)"  | 
|
392  | 
||
| 59189 | 393  | 
subsection \<open>Compositionality of the Semantics\<close>  | 
| 13020 | 394  | 
|
| 59189 | 395  | 
subsubsection \<open>Definition of the conjoin operator\<close>  | 
| 13020 | 396  | 
|
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
397  | 
definition same_length :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" where
 | 
| 13020 | 398  | 
"same_length c clist \<equiv> (\<forall>i<length clist. length(clist!i)=length c)"  | 
399  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
400  | 
definition same_state :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" where
 | 
| 13020 | 401  | 
"same_state c clist \<equiv> (\<forall>i <length clist. \<forall>j<length c. snd(c!j) = snd((clist!i)!j))"  | 
402  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
403  | 
definition same_program :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" where
 | 
| 13020 | 404  | 
"same_program c clist \<equiv> (\<forall>j<length c. fst(c!j) = map (\<lambda>x. fst(nth x j)) clist)"  | 
405  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
406  | 
definition compat_label :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool" where
 | 
| 13020 | 407  | 
"compat_label c clist \<equiv> (\<forall>j. Suc j<length c \<longrightarrow>  | 
408  | 
(c!j -pc\<rightarrow> c!Suc j \<and> (\<exists>i<length clist. (clist!i)!j -c\<rightarrow> (clist!i)! Suc j \<and>  | 
|
| 
13022
 
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
 
prensani 
parents: 
13020 
diff
changeset
 | 
409  | 
(\<forall>l<length clist. l\<noteq>i \<longrightarrow> (clist!l)!j -e\<rightarrow> (clist!l)! Suc j))) \<or>  | 
| 13020 | 410  | 
(c!j -pe\<rightarrow> c!Suc j \<and> (\<forall>i<length clist. (clist!i)!j -e\<rightarrow> (clist!i)! Suc j)))"  | 
411  | 
||
| 
35416
 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 
haftmann 
parents: 
32621 
diff
changeset
 | 
412  | 
definition conjoin :: "'a par_confs \<Rightarrow> ('a confs) list \<Rightarrow> bool"  ("_ \<propto> _" [65,65] 64) where
 | 
| 13020 | 413  | 
"c \<propto> clist \<equiv> (same_length c clist) \<and> (same_state c clist) \<and> (same_program c clist) \<and> (compat_label c clist)"  | 
414  | 
||
| 59189 | 415  | 
subsubsection \<open>Some previous lemmas\<close>  | 
| 13020 | 416  | 
|
| 
13022
 
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
 
prensani 
parents: 
13020 
diff
changeset
 | 
417  | 
lemma list_eq_if [rule_format]:  | 
| 
 
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
 
prensani 
parents: 
13020 
diff
changeset
 | 
418  | 
"\<forall>ys. xs=ys \<longrightarrow> (length xs = length ys) \<longrightarrow> (\<forall>i<length xs. xs!i=ys!i)"  | 
| 51119 | 419  | 
by (induct xs) auto  | 
| 13020 | 420  | 
|
421  | 
lemma list_eq: "(length xs = length ys \<and> (\<forall>i<length xs. xs!i=ys!i)) = (xs=ys)"  | 
|
422  | 
apply(rule iffI)  | 
|
423  | 
apply clarify  | 
|
424  | 
apply(erule nth_equalityI)  | 
|
425  | 
apply simp+  | 
|
426  | 
done  | 
|
427  | 
||
428  | 
lemma nth_tl: "\<lbrakk> ys!0=a; ys\<noteq>[] \<rbrakk> \<Longrightarrow> ys=(a#(tl ys))"  | 
|
| 51119 | 429  | 
by (cases ys) simp_all  | 
| 13020 | 430  | 
|
431  | 
lemma nth_tl_if [rule_format]: "ys\<noteq>[] \<longrightarrow> ys!0=a \<longrightarrow> P ys \<longrightarrow> P (a#(tl ys))"  | 
|
| 51119 | 432  | 
by (induct ys) simp_all  | 
| 13020 | 433  | 
|
434  | 
lemma nth_tl_onlyif [rule_format]: "ys\<noteq>[] \<longrightarrow> ys!0=a \<longrightarrow> P (a#(tl ys)) \<longrightarrow> P ys"  | 
|
| 51119 | 435  | 
by (induct ys) simp_all  | 
| 13020 | 436  | 
|
437  | 
lemma seq_not_eq1: "Seq c1 c2\<noteq>c1"  | 
|
| 51119 | 438  | 
by (induct c1) auto  | 
| 13020 | 439  | 
|
440  | 
lemma seq_not_eq2: "Seq c1 c2\<noteq>c2"  | 
|
| 51119 | 441  | 
by (induct c2) auto  | 
| 13020 | 442  | 
|
443  | 
lemma if_not_eq1: "Cond b c1 c2 \<noteq>c1"  | 
|
| 51119 | 444  | 
by (induct c1) auto  | 
| 13020 | 445  | 
|
446  | 
lemma if_not_eq2: "Cond b c1 c2\<noteq>c2"  | 
|
| 51119 | 447  | 
by (induct c2) auto  | 
| 13020 | 448  | 
|
449  | 
lemmas seq_and_if_not_eq [simp] = seq_not_eq1 seq_not_eq2  | 
|
450  | 
seq_not_eq1 [THEN not_sym] seq_not_eq2 [THEN not_sym]  | 
|
451  | 
if_not_eq1 if_not_eq2 if_not_eq1 [THEN not_sym] if_not_eq2 [THEN not_sym]  | 
|
452  | 
||
| 23746 | 453  | 
lemma prog_not_eq_in_ctran_aux:  | 
454  | 
assumes c: "(P,s) -c\<rightarrow> (Q,t)"  | 
|
455  | 
shows "P\<noteq>Q" using c  | 
|
456  | 
by (induct x1 \<equiv> "(P,s)" x2 \<equiv> "(Q,t)" arbitrary: P s Q t) auto  | 
|
| 13020 | 457  | 
|
458  | 
lemma prog_not_eq_in_ctran [simp]: "\<not> (P,s) -c\<rightarrow> (P,t)"  | 
|
459  | 
apply clarify  | 
|
460  | 
apply(drule prog_not_eq_in_ctran_aux)  | 
|
461  | 
apply simp  | 
|
462  | 
done  | 
|
463  | 
||
464  | 
lemma prog_not_eq_in_par_ctran_aux [rule_format]: "(P,s) -pc\<rightarrow> (Q,t) \<Longrightarrow> (P\<noteq>Q)"  | 
|
465  | 
apply(erule par_ctran.induct)  | 
|
466  | 
apply(drule prog_not_eq_in_ctran_aux)  | 
|
467  | 
apply clarify  | 
|
468  | 
apply(drule list_eq_if)  | 
|
469  | 
apply simp_all  | 
|
470  | 
apply force  | 
|
471  | 
done  | 
|
472  | 
||
473  | 
lemma prog_not_eq_in_par_ctran [simp]: "\<not> (P,s) -pc\<rightarrow> (P,t)"  | 
|
474  | 
apply clarify  | 
|
475  | 
apply(drule prog_not_eq_in_par_ctran_aux)  | 
|
476  | 
apply simp  | 
|
477  | 
done  | 
|
478  | 
||
479  | 
lemma tl_in_cptn: "\<lbrakk> a#xs \<in>cptn; xs\<noteq>[] \<rbrakk> \<Longrightarrow> xs\<in>cptn"  | 
|
| 51119 | 480  | 
by (force elim: cptn.cases)  | 
| 13020 | 481  | 
|
| 
13022
 
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
 
prensani 
parents: 
13020 
diff
changeset
 | 
482  | 
lemma tl_zero[rule_format]:  | 
| 
 
b115b305612f
New order in the loading of theories (Quote-antiquote right before the OG_Syntax and RG_Syntax respectively)
 
prensani 
parents: 
13020 
diff
changeset
 | 
483  | 
"P (ys!Suc j) \<longrightarrow> Suc j<length ys \<longrightarrow> ys\<noteq>[] \<longrightarrow> P (tl(ys)!j)"  | 
| 51119 | 484  | 
by (induct ys) simp_all  | 
| 13020 | 485  | 
|
| 59189 | 486  | 
subsection \<open>The Semantics is Compositional\<close>  | 
| 13020 | 487  | 
|
488  | 
lemma aux_if [rule_format]:  | 
|
489  | 
"\<forall>xs s clist. (length clist = length xs \<and> (\<forall>i<length xs. (xs!i,s)#clist!i \<in> cptn)  | 
|
490  | 
\<and> ((xs, s)#ys \<propto> map (\<lambda>i. (fst i,s)#snd i) (zip xs clist))  | 
|
491  | 
\<longrightarrow> (xs, s)#ys \<in> par_cptn)"  | 
|
492  | 
apply(induct ys)  | 
|
493  | 
apply(clarify)  | 
|
494  | 
apply(rule ParCptnOne)  | 
|
495  | 
apply(clarify)  | 
|
496  | 
apply(simp add:conjoin_def compat_label_def)  | 
|
497  | 
apply clarify  | 
|
| 59807 | 498  | 
apply(erule_tac x="0" and P="\<lambda>j. H j \<longrightarrow> (P j \<or> Q j)" for H P Q in all_dupE, simp)  | 
| 13020 | 499  | 
apply(erule disjE)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
62390 
diff
changeset
 | 
500  | 
\<comment> \<open>first step is a Component step\<close>  | 
| 13020 | 501  | 
apply clarify  | 
502  | 
apply simp  | 
|
503  | 
apply(subgoal_tac "a=(xs[i:=(fst(clist!i!0))])")  | 
|
504  | 
apply(subgoal_tac "b=snd(clist!i!0)",simp)  | 
|
505  | 
prefer 2  | 
|
506  | 
apply(simp add: same_state_def)  | 
|
507  | 
apply(erule_tac x=i in allE,erule impE,assumption,  | 
|
| 59807 | 508  | 
erule_tac x=1 and P="\<lambda>j. (H j) \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE, simp)  | 
| 13020 | 509  | 
prefer 2  | 
510  | 
apply(simp add:same_program_def)  | 
|
| 59807 | 511  | 
apply(erule_tac x=1 and P="\<lambda>j. H j \<longrightarrow> (fst (s j))=(t j)" for H s t in allE,simp)  | 
| 13020 | 512  | 
apply(rule nth_equalityI,simp)  | 
513  | 
apply clarify  | 
|
514  | 
apply(case_tac "i=ia",simp,simp)  | 
|
| 59807 | 515  | 
apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE)  | 
| 13020 | 516  | 
apply(drule_tac t=i in not_sym,simp)  | 
| 23746 | 517  | 
apply(erule etranE,simp)  | 
| 13020 | 518  | 
apply(rule ParCptnComp)  | 
519  | 
apply(erule ParComp,simp)  | 
|
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
62390 
diff
changeset
 | 
520  | 
\<comment> \<open>applying the induction hypothesis\<close>  | 
| 13020 | 521  | 
apply(erule_tac x="xs[i := fst (clist ! i ! 0)]" in allE)  | 
522  | 
apply(erule_tac x="snd (clist ! i ! 0)" in allE)  | 
|
523  | 
apply(erule mp)  | 
|
524  | 
apply(rule_tac x="map tl clist" in exI,simp)  | 
|
525  | 
apply(rule conjI,clarify)  | 
|
526  | 
apply(case_tac "i=ia",simp)  | 
|
527  | 
apply(rule nth_tl_if)  | 
|
528  | 
apply(force simp add:same_length_def length_Suc_conv)  | 
|
529  | 
apply simp  | 
|
530  | 
apply(erule allE,erule impE,assumption,erule tl_in_cptn)  | 
|
531  | 
apply(force simp add:same_length_def length_Suc_conv)  | 
|
532  | 
apply(rule nth_tl_if)  | 
|
533  | 
apply(force simp add:same_length_def length_Suc_conv)  | 
|
534  | 
apply(simp add:same_state_def)  | 
|
535  | 
apply(erule_tac x=ia in allE, erule impE, assumption,  | 
|
| 59807 | 536  | 
erule_tac x=1 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE)  | 
537  | 
apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE)  | 
|
| 13020 | 538  | 
apply(drule_tac t=i in not_sym,simp)  | 
| 23746 | 539  | 
apply(erule etranE,simp)  | 
| 13020 | 540  | 
apply(erule allE,erule impE,assumption,erule tl_in_cptn)  | 
541  | 
apply(force simp add:same_length_def length_Suc_conv)  | 
|
542  | 
apply(simp add:same_length_def same_state_def)  | 
|
543  | 
apply(rule conjI)  | 
|
544  | 
apply clarify  | 
|
545  | 
apply(case_tac j,simp,simp)  | 
|
546  | 
apply(erule_tac x=ia in allE, erule impE, assumption,  | 
|
| 59807 | 547  | 
erule_tac x="Suc(Suc nat)" and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)  | 
| 13020 | 548  | 
apply(force simp add:same_length_def length_Suc_conv)  | 
549  | 
apply(rule conjI)  | 
|
550  | 
apply(simp add:same_program_def)  | 
|
551  | 
apply clarify  | 
|
552  | 
apply(case_tac j,simp)  | 
|
553  | 
apply(rule nth_equalityI,simp)  | 
|
554  | 
apply clarify  | 
|
555  | 
apply(case_tac "i=ia",simp,simp)  | 
|
| 59807 | 556  | 
apply(erule_tac x="Suc(Suc nat)" and P="\<lambda>j. H j \<longrightarrow> (fst (s j))=(t j)" for H s t in allE,simp)  | 
| 13020 | 557  | 
apply(rule nth_equalityI,simp,simp)  | 
558  | 
apply(force simp add:length_Suc_conv)  | 
|
559  | 
apply(rule allI,rule impI)  | 
|
| 59807 | 560  | 
apply(erule_tac x="Suc j" and P="\<lambda>j. H j \<longrightarrow> (I j \<or> J j)" for H I J in allE,simp)  | 
| 13020 | 561  | 
apply(erule disjE)  | 
562  | 
apply clarify  | 
|
563  | 
apply(rule_tac x=ia in exI,simp)  | 
|
564  | 
apply(case_tac "i=ia",simp)  | 
|
565  | 
apply(rule conjI)  | 
|
566  | 
apply(force simp add: length_Suc_conv)  | 
|
567  | 
apply clarify  | 
|
| 59807 | 568  | 
apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE,erule impE,assumption)  | 
569  | 
apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE,erule impE,assumption)  | 
|
| 13020 | 570  | 
apply simp  | 
571  | 
apply(case_tac j,simp)  | 
|
572  | 
apply(rule tl_zero)  | 
|
573  | 
apply(erule_tac x=l in allE, erule impE, assumption,  | 
|
| 59807 | 574  | 
erule_tac x=1 and P="\<lambda>j. (H j) \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)  | 
| 23746 | 575  | 
apply(force elim:etranE intro:Env)  | 
| 13020 | 576  | 
apply force  | 
577  | 
apply force  | 
|
578  | 
apply simp  | 
|
579  | 
apply(rule tl_zero)  | 
|
580  | 
apply(erule tl_zero)  | 
|
581  | 
apply force  | 
|
582  | 
apply force  | 
|
583  | 
apply force  | 
|
584  | 
apply force  | 
|
585  | 
apply(rule conjI,simp)  | 
|
586  | 
apply(rule nth_tl_if)  | 
|
587  | 
apply force  | 
|
588  | 
apply(erule_tac x=ia in allE, erule impE, assumption,  | 
|
| 59807 | 589  | 
erule_tac x=1 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE)  | 
590  | 
apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE)  | 
|
| 13020 | 591  | 
apply(drule_tac t=i in not_sym,simp)  | 
| 23746 | 592  | 
apply(erule etranE,simp)  | 
| 13020 | 593  | 
apply(erule tl_zero)  | 
594  | 
apply force  | 
|
595  | 
apply force  | 
|
596  | 
apply clarify  | 
|
597  | 
apply(case_tac "i=l",simp)  | 
|
598  | 
apply(rule nth_tl_if)  | 
|
| 59807 | 599  | 
apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 600  | 
apply simp  | 
| 59807 | 601  | 
apply(erule_tac P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE,erule impE,assumption,erule impE,assumption)  | 
| 13020 | 602  | 
apply(erule tl_zero,force)  | 
| 59807 | 603  | 
apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 604  | 
apply(rule nth_tl_if)  | 
| 59807 | 605  | 
apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 606  | 
apply(erule_tac x=l in allE, erule impE, assumption,  | 
| 59807 | 607  | 
erule_tac x=1 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE)  | 
608  | 
apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE,erule impE, assumption,simp)  | 
|
| 23746 | 609  | 
apply(erule etranE,simp)  | 
| 13020 | 610  | 
apply(rule tl_zero)  | 
611  | 
apply force  | 
|
612  | 
apply force  | 
|
| 59807 | 613  | 
apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 614  | 
apply(rule disjI2)  | 
615  | 
apply(case_tac j,simp)  | 
|
616  | 
apply clarify  | 
|
617  | 
apply(rule tl_zero)  | 
|
| 59807 | 618  | 
apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> I j\<in>etran" for H I in allE,erule impE, assumption)  | 
| 13020 | 619  | 
apply(case_tac "i=ia",simp,simp)  | 
620  | 
apply(erule_tac x=ia in allE, erule impE, assumption,  | 
|
| 59807 | 621  | 
erule_tac x=1 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE)  | 
622  | 
apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE,erule impE, assumption,simp)  | 
|
| 23746 | 623  | 
apply(force elim:etranE intro:Env)  | 
| 13020 | 624  | 
apply force  | 
| 59807 | 625  | 
apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 626  | 
apply simp  | 
627  | 
apply clarify  | 
|
628  | 
apply(rule tl_zero)  | 
|
629  | 
apply(rule tl_zero,force)  | 
|
630  | 
apply force  | 
|
| 59807 | 631  | 
apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 632  | 
apply force  | 
| 59807 | 633  | 
apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 
67443
 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 
wenzelm 
parents: 
62390 
diff
changeset
 | 
634  | 
\<comment> \<open>first step is an environmental step\<close>  | 
| 13020 | 635  | 
apply clarify  | 
| 23746 | 636  | 
apply(erule par_etran.cases)  | 
| 13020 | 637  | 
apply simp  | 
638  | 
apply(rule ParCptnEnv)  | 
|
639  | 
apply(erule_tac x="Ps" in allE)  | 
|
640  | 
apply(erule_tac x="t" in allE)  | 
|
641  | 
apply(erule mp)  | 
|
642  | 
apply(rule_tac x="map tl clist" in exI,simp)  | 
|
643  | 
apply(rule conjI)  | 
|
644  | 
apply clarify  | 
|
| 59807 | 645  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> I j \<in> cptn" for H I in allE,simp)  | 
| 23746 | 646  | 
apply(erule cptn.cases)  | 
| 13020 | 647  | 
apply(simp add:same_length_def)  | 
| 59807 | 648  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 649  | 
apply(simp add:same_state_def)  | 
650  | 
apply(erule_tac x=i in allE, erule impE, assumption,  | 
|
| 59807 | 651  | 
erule_tac x=1 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)  | 
652  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> J j \<in>etran" for H J in allE,simp)  | 
|
| 23746 | 653  | 
apply(erule etranE,simp)  | 
| 13020 | 654  | 
apply(simp add:same_state_def same_length_def)  | 
655  | 
apply(rule conjI,clarify)  | 
|
656  | 
apply(case_tac j,simp,simp)  | 
|
657  | 
apply(erule_tac x=i in allE, erule impE, assumption,  | 
|
| 59807 | 658  | 
erule_tac x="Suc(Suc nat)" and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)  | 
| 13020 | 659  | 
apply(rule tl_zero)  | 
660  | 
apply(simp)  | 
|
661  | 
apply force  | 
|
| 59807 | 662  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 663  | 
apply(rule conjI)  | 
664  | 
apply(simp add:same_program_def)  | 
|
665  | 
apply clarify  | 
|
666  | 
apply(case_tac j,simp)  | 
|
667  | 
apply(rule nth_equalityI,simp)  | 
|
668  | 
apply clarify  | 
|
669  | 
apply simp  | 
|
| 59807 | 670  | 
apply(erule_tac x="Suc(Suc nat)" and P="\<lambda>j. H j \<longrightarrow> (fst (s j))=(t j)" for H s t in allE,simp)  | 
| 13020 | 671  | 
apply(rule nth_equalityI,simp,simp)  | 
672  | 
apply(force simp add:length_Suc_conv)  | 
|
673  | 
apply(rule allI,rule impI)  | 
|
| 59807 | 674  | 
apply(erule_tac x="Suc j" and P="\<lambda>j. H j \<longrightarrow> (I j \<or> J j)" for H I J in allE,simp)  | 
| 13020 | 675  | 
apply(erule disjE)  | 
676  | 
apply clarify  | 
|
677  | 
apply(rule_tac x=i in exI,simp)  | 
|
678  | 
apply(rule conjI)  | 
|
| 59807 | 679  | 
apply(erule_tac x=i and P="\<lambda>i. H i \<longrightarrow> J i \<in>etran" for H J in allE, erule impE, assumption)  | 
| 23746 | 680  | 
apply(erule etranE,simp)  | 
| 13020 | 681  | 
apply(erule_tac x=i in allE, erule impE, assumption,  | 
| 59807 | 682  | 
erule_tac x=1 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)  | 
| 13020 | 683  | 
apply(rule nth_tl_if)  | 
| 59807 | 684  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 685  | 
apply simp  | 
686  | 
apply(erule tl_zero,force)  | 
|
| 59807 | 687  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 688  | 
apply clarify  | 
| 59807 | 689  | 
apply(erule_tac x=l and P="\<lambda>i. H i \<longrightarrow> J i \<in>etran" for H J in allE, erule impE, assumption)  | 
| 23746 | 690  | 
apply(erule etranE,simp)  | 
| 13020 | 691  | 
apply(erule_tac x=l in allE, erule impE, assumption,  | 
| 59807 | 692  | 
erule_tac x=1 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)  | 
| 13020 | 693  | 
apply(rule nth_tl_if)  | 
| 59807 | 694  | 
apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 695  | 
apply simp  | 
696  | 
apply(rule tl_zero,force)  | 
|
697  | 
apply force  | 
|
| 59807 | 698  | 
apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 699  | 
apply(rule disjI2)  | 
700  | 
apply simp  | 
|
701  | 
apply clarify  | 
|
702  | 
apply(case_tac j,simp)  | 
|
703  | 
apply(rule tl_zero)  | 
|
| 59807 | 704  | 
apply(erule_tac x=i and P="\<lambda>i. H i \<longrightarrow> J i \<in>etran" for H J in allE, erule impE, assumption)  | 
705  | 
apply(erule_tac x=i and P="\<lambda>i. H i \<longrightarrow> J i \<in>etran" for H J in allE, erule impE, assumption)  | 
|
| 23746 | 706  | 
apply(force elim:etranE intro:Env)  | 
| 13020 | 707  | 
apply force  | 
| 59807 | 708  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 709  | 
apply simp  | 
710  | 
apply(rule tl_zero)  | 
|
711  | 
apply(rule tl_zero,force)  | 
|
712  | 
apply force  | 
|
| 59807 | 713  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 714  | 
apply force  | 
| 59807 | 715  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 716  | 
done  | 
717  | 
||
718  | 
lemma aux_onlyif [rule_format]: "\<forall>xs s. (xs, s)#ys \<in> par_cptn \<longrightarrow>  | 
|
719  | 
(\<exists>clist. (length clist = length xs) \<and>  | 
|
720  | 
(xs, s)#ys \<propto> map (\<lambda>i. (fst i,s)#(snd i)) (zip xs clist) \<and>  | 
|
721  | 
(\<forall>i<length xs. (xs!i,s)#(clist!i) \<in> cptn))"  | 
|
| 
71989
 
bad75618fb82
extraction of equations x = t from premises beneath meta-all
 
haftmann 
parents: 
67613 
diff
changeset
 | 
722  | 
supply [[simproc del: defined_all]]  | 
| 13020 | 723  | 
apply(induct ys)  | 
724  | 
apply(clarify)  | 
|
| 15425 | 725  | 
apply(rule_tac x="map (\<lambda>i. []) [0..<length xs]" in exI)  | 
| 13020 | 726  | 
apply(simp add: conjoin_def same_length_def same_state_def same_program_def compat_label_def)  | 
727  | 
apply(rule conjI)  | 
|
728  | 
apply(rule nth_equalityI,simp,simp)  | 
|
729  | 
apply(force intro: cptn.intros)  | 
|
730  | 
apply(clarify)  | 
|
| 23746 | 731  | 
apply(erule par_cptn.cases,simp)  | 
| 13020 | 732  | 
apply simp  | 
733  | 
apply(erule_tac x="xs" in allE)  | 
|
734  | 
apply(erule_tac x="t" in allE,simp)  | 
|
735  | 
apply clarify  | 
|
| 15425 | 736  | 
apply(rule_tac x="(map (\<lambda>j. (P!j, t)#(clist!j)) [0..<length P])" in exI,simp)  | 
| 13020 | 737  | 
apply(rule conjI)  | 
738  | 
prefer 2  | 
|
739  | 
apply clarify  | 
|
740  | 
apply(rule CptnEnv,simp)  | 
|
741  | 
apply(simp add:conjoin_def same_length_def same_state_def)  | 
|
742  | 
apply (rule conjI)  | 
|
743  | 
apply clarify  | 
|
744  | 
apply(case_tac j,simp,simp)  | 
|
745  | 
apply(rule conjI)  | 
|
746  | 
apply(simp add:same_program_def)  | 
|
747  | 
apply clarify  | 
|
748  | 
apply(case_tac j,simp)  | 
|
749  | 
apply(rule nth_equalityI,simp,simp)  | 
|
750  | 
apply simp  | 
|
751  | 
apply(rule nth_equalityI,simp,simp)  | 
|
752  | 
apply(simp add:compat_label_def)  | 
|
753  | 
apply clarify  | 
|
754  | 
apply(case_tac j,simp)  | 
|
755  | 
apply(simp add:ParEnv)  | 
|
756  | 
apply clarify  | 
|
757  | 
apply(simp add:Env)  | 
|
758  | 
apply simp  | 
|
759  | 
apply(erule_tac x=nat in allE,erule impE, assumption)  | 
|
760  | 
apply(erule disjE,simp)  | 
|
761  | 
apply clarify  | 
|
762  | 
apply(rule_tac x=i in exI,simp)  | 
|
763  | 
apply force  | 
|
| 23746 | 764  | 
apply(erule par_ctran.cases,simp)  | 
| 13020 | 765  | 
apply(erule_tac x="Ps[i:=r]" in allE)  | 
766  | 
apply(erule_tac x="ta" in allE,simp)  | 
|
767  | 
apply clarify  | 
|
| 15425 | 768  | 
apply(rule_tac x="(map (\<lambda>j. (Ps!j, ta)#(clist!j)) [0..<length Ps]) [i:=((r, ta)#(clist!i))]" in exI,simp)  | 
| 13020 | 769  | 
apply(rule conjI)  | 
770  | 
prefer 2  | 
|
771  | 
apply clarify  | 
|
772  | 
apply(case_tac "i=ia",simp)  | 
|
773  | 
apply(erule CptnComp)  | 
|
| 59807 | 774  | 
apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> (I j \<in> cptn)" for H I in allE,simp)  | 
| 13020 | 775  | 
apply simp  | 
776  | 
apply(erule_tac x=ia in allE)  | 
|
777  | 
apply(rule CptnEnv,simp)  | 
|
778  | 
apply(simp add:conjoin_def)  | 
|
779  | 
apply (rule conjI)  | 
|
780  | 
apply(simp add:same_length_def)  | 
|
781  | 
apply clarify  | 
|
782  | 
apply(case_tac "i=ia",simp,simp)  | 
|
783  | 
apply(rule conjI)  | 
|
784  | 
apply(simp add:same_state_def)  | 
|
785  | 
apply clarify  | 
|
| 13601 | 786  | 
apply(case_tac j, simp, simp (no_asm_simp))  | 
| 13020 | 787  | 
apply(case_tac "i=ia",simp,simp)  | 
788  | 
apply(rule conjI)  | 
|
789  | 
apply(simp add:same_program_def)  | 
|
790  | 
apply clarify  | 
|
791  | 
apply(case_tac j,simp)  | 
|
792  | 
apply(rule nth_equalityI,simp,simp)  | 
|
793  | 
apply simp  | 
|
794  | 
apply(rule nth_equalityI,simp,simp)  | 
|
| 59807 | 795  | 
apply(erule_tac x=nat and P="\<lambda>j. H j \<longrightarrow> (fst (a j))=((b j))" for H a b in allE)  | 
| 13020 | 796  | 
apply(case_tac nat)  | 
797  | 
apply clarify  | 
|
798  | 
apply(case_tac "i=ia",simp,simp)  | 
|
799  | 
apply clarify  | 
|
800  | 
apply(case_tac "i=ia",simp,simp)  | 
|
801  | 
apply(simp add:compat_label_def)  | 
|
802  | 
apply clarify  | 
|
803  | 
apply(case_tac j)  | 
|
804  | 
apply(rule conjI,simp)  | 
|
805  | 
apply(erule ParComp,assumption)  | 
|
806  | 
apply clarify  | 
|
807  | 
apply(rule_tac x=i in exI,simp)  | 
|
808  | 
apply clarify  | 
|
809  | 
apply(rule Env)  | 
|
810  | 
apply simp  | 
|
| 59807 | 811  | 
apply(erule_tac x=nat and P="\<lambda>j. H j \<longrightarrow> (P j \<or> Q j)" for H P Q in allE,simp)  | 
| 13020 | 812  | 
apply(erule disjE)  | 
813  | 
apply clarify  | 
|
814  | 
apply(rule_tac x=ia in exI,simp)  | 
|
815  | 
apply(rule conjI)  | 
|
816  | 
apply(case_tac "i=ia",simp,simp)  | 
|
817  | 
apply clarify  | 
|
818  | 
apply(case_tac "i=l",simp)  | 
|
819  | 
apply(case_tac "l=ia",simp,simp)  | 
|
820  | 
apply(erule_tac x=l in allE,erule impE,assumption,erule impE, assumption,simp)  | 
|
821  | 
apply simp  | 
|
822  | 
apply(erule_tac x=l in allE,erule impE,assumption,erule impE, assumption,simp)  | 
|
823  | 
apply clarify  | 
|
| 59807 | 824  | 
apply(erule_tac x=ia and P="\<lambda>j. H j \<longrightarrow> (P j)\<in>etran" for H P in allE, erule impE, assumption)  | 
| 13601 | 825  | 
apply(case_tac "i=ia",simp,simp)  | 
| 13020 | 826  | 
done  | 
827  | 
||
828  | 
lemma one_iff_aux: "xs\<noteq>[] \<Longrightarrow> (\<forall>ys. ((xs, s)#ys \<in> par_cptn) =  | 
|
829  | 
(\<exists>clist. length clist= length xs \<and>  | 
|
830  | 
((xs, s)#ys \<propto> map (\<lambda>i. (fst i,s)#(snd i)) (zip xs clist)) \<and>  | 
|
831  | 
(\<forall>i<length xs. (xs!i,s)#(clist!i) \<in> cptn))) =  | 
|
832  | 
 (par_cp (xs) s = {c. \<exists>clist. (length clist)=(length xs) \<and>
 | 
|
833  | 
(\<forall>i<length clist. (clist!i) \<in> cp(xs!i) s) \<and> c \<propto> clist})"  | 
|
834  | 
apply (rule iffI)  | 
|
835  | 
apply(rule subset_antisym)  | 
|
836  | 
apply(rule subsetI)  | 
|
837  | 
apply(clarify)  | 
|
838  | 
apply(simp add:par_cp_def cp_def)  | 
|
839  | 
apply(case_tac x)  | 
|
| 23746 | 840  | 
apply(force elim:par_cptn.cases)  | 
| 13020 | 841  | 
apply simp  | 
| 
55417
 
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
 
blanchet 
parents: 
51120 
diff
changeset
 | 
842  | 
apply(rename_tac a list)  | 
| 13020 | 843  | 
apply(erule_tac x="list" in allE)  | 
844  | 
apply clarify  | 
|
845  | 
apply simp  | 
|
846  | 
apply(rule_tac x="map (\<lambda>i. (fst i, s) # snd i) (zip xs clist)" in exI,simp)  | 
|
847  | 
apply(rule subsetI)  | 
|
848  | 
apply(clarify)  | 
|
849  | 
apply(case_tac x)  | 
|
850  | 
apply(erule_tac x=0 in allE)  | 
|
851  | 
apply(simp add:cp_def conjoin_def same_length_def same_program_def same_state_def compat_label_def)  | 
|
852  | 
apply clarify  | 
|
| 23746 | 853  | 
apply(erule cptn.cases,force,force,force)  | 
| 13020 | 854  | 
apply(simp add:par_cp_def conjoin_def same_length_def same_program_def same_state_def compat_label_def)  | 
855  | 
apply clarify  | 
|
| 59807 | 856  | 
apply(erule_tac x=0 and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in all_dupE)  | 
| 13020 | 857  | 
apply(subgoal_tac "a = xs")  | 
858  | 
apply(subgoal_tac "b = s",simp)  | 
|
859  | 
prefer 3  | 
|
| 59807 | 860  | 
apply(erule_tac x=0 and P="\<lambda>j. H j \<longrightarrow> (fst (s j))=((t j))" for H s t in allE)  | 
| 13020 | 861  | 
apply (simp add:cp_def)  | 
862  | 
apply(rule nth_equalityI,simp,simp)  | 
|
863  | 
prefer 2  | 
|
864  | 
apply(erule_tac x=0 in allE)  | 
|
865  | 
apply (simp add:cp_def)  | 
|
| 59807 | 866  | 
apply(erule_tac x=0 and P="\<lambda>j. H j \<longrightarrow> (\<forall>i. T i \<longrightarrow> (snd (d j i))=(snd (e j i)))" for H T d e in allE,simp)  | 
867  | 
apply(erule_tac x=0 and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)  | 
|
| 13020 | 868  | 
apply(erule_tac x=list in allE)  | 
869  | 
apply(rule_tac x="map tl clist" in exI,simp)  | 
|
870  | 
apply(rule conjI)  | 
|
871  | 
apply clarify  | 
|
872  | 
apply(case_tac j,simp)  | 
|
873  | 
apply(erule_tac x=i in allE, erule impE, assumption,  | 
|
| 59807 | 874  | 
erule_tac x="0" and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE,simp)  | 
| 13020 | 875  | 
apply(erule_tac x=i in allE, erule impE, assumption,  | 
| 59807 | 876  | 
erule_tac x="Suc nat" and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE)  | 
877  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)  | 
|
| 13020 | 878  | 
apply(case_tac "clist!i",simp,simp)  | 
879  | 
apply(rule conjI)  | 
|
880  | 
apply clarify  | 
|
881  | 
apply(rule nth_equalityI,simp,simp)  | 
|
882  | 
apply(case_tac j)  | 
|
883  | 
apply clarify  | 
|
884  | 
apply(erule_tac x=i in allE)  | 
|
885  | 
apply(simp add:cp_def)  | 
|
886  | 
apply clarify  | 
|
887  | 
apply simp  | 
|
| 59807 | 888  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)  | 
| 13020 | 889  | 
apply(case_tac "clist!i",simp,simp)  | 
| 59807 | 890  | 
apply(thin_tac "H = (\<exists>i. J i)" for H J)  | 
| 13020 | 891  | 
apply(rule conjI)  | 
892  | 
apply clarify  | 
|
893  | 
apply(erule_tac x=j in allE,erule impE, assumption,erule disjE)  | 
|
894  | 
apply clarify  | 
|
895  | 
apply(rule_tac x=i in exI,simp)  | 
|
896  | 
apply(case_tac j,simp)  | 
|
897  | 
apply(rule conjI)  | 
|
898  | 
apply(erule_tac x=i in allE)  | 
|
899  | 
apply(simp add:cp_def)  | 
|
| 59807 | 900  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)  | 
| 13020 | 901  | 
apply(case_tac "clist!i",simp,simp)  | 
902  | 
apply clarify  | 
|
903  | 
apply(erule_tac x=l in allE)  | 
|
| 59807 | 904  | 
apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE)  | 
| 13020 | 905  | 
apply clarify  | 
906  | 
apply(simp add:cp_def)  | 
|
| 59807 | 907  | 
apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)  | 
| 13020 | 908  | 
apply(case_tac "clist!l",simp,simp)  | 
909  | 
apply simp  | 
|
910  | 
apply(rule conjI)  | 
|
| 59807 | 911  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)  | 
| 13020 | 912  | 
apply(case_tac "clist!i",simp,simp)  | 
913  | 
apply clarify  | 
|
| 59807 | 914  | 
apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE)  | 
915  | 
apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)  | 
|
| 13020 | 916  | 
apply(case_tac "clist!l",simp,simp)  | 
917  | 
apply clarify  | 
|
918  | 
apply(erule_tac x=i in allE)  | 
|
919  | 
apply(simp add:cp_def)  | 
|
| 59807 | 920  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)  | 
| 13020 | 921  | 
apply(case_tac "clist!i",simp)  | 
922  | 
apply(rule nth_tl_if,simp,simp)  | 
|
| 59807 | 923  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (P j)\<in>etran" for H P in allE, erule impE, assumption,simp)  | 
| 13020 | 924  | 
apply(simp add:cp_def)  | 
925  | 
apply clarify  | 
|
926  | 
apply(rule nth_tl_if)  | 
|
| 59807 | 927  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)  | 
| 13020 | 928  | 
apply(case_tac "clist!i",simp,simp)  | 
929  | 
apply force  | 
|
930  | 
apply force  | 
|
931  | 
apply clarify  | 
|
932  | 
apply(rule iffI)  | 
|
933  | 
apply(simp add:par_cp_def)  | 
|
934  | 
apply(erule_tac c="(xs, s) # ys" in equalityCE)  | 
|
935  | 
apply simp  | 
|
936  | 
apply clarify  | 
|
937  | 
apply(rule_tac x="map tl clist" in exI)  | 
|
938  | 
apply simp  | 
|
939  | 
apply (rule conjI)  | 
|
940  | 
apply(simp add:conjoin_def cp_def)  | 
|
941  | 
apply(rule conjI)  | 
|
942  | 
apply clarify  | 
|
943  | 
apply(unfold same_length_def)  | 
|
944  | 
apply clarify  | 
|
| 59807 | 945  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,simp)  | 
| 13020 | 946  | 
apply(rule conjI)  | 
947  | 
apply(simp add:same_state_def)  | 
|
948  | 
apply clarify  | 
|
949  | 
apply(erule_tac x=i in allE, erule impE, assumption,  | 
|
| 59807 | 950  | 
erule_tac x=j and P="\<lambda>j. H j \<longrightarrow> (snd (d j))=(snd (e j))" for H d e in allE)  | 
| 13020 | 951  | 
apply(case_tac j,simp)  | 
| 59807 | 952  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)  | 
| 13020 | 953  | 
apply(case_tac "clist!i",simp,simp)  | 
954  | 
apply(rule conjI)  | 
|
955  | 
apply(simp add:same_program_def)  | 
|
956  | 
apply clarify  | 
|
957  | 
apply(rule nth_equalityI,simp,simp)  | 
|
958  | 
apply(case_tac j,simp)  | 
|
959  | 
apply clarify  | 
|
| 59807 | 960  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)  | 
| 13020 | 961  | 
apply(case_tac "clist!i",simp,simp)  | 
962  | 
apply clarify  | 
|
963  | 
apply(simp add:compat_label_def)  | 
|
964  | 
apply(rule allI,rule impI)  | 
|
965  | 
apply(erule_tac x=j in allE,erule impE, assumption)  | 
|
966  | 
apply(erule disjE)  | 
|
967  | 
apply clarify  | 
|
968  | 
apply(rule_tac x=i in exI,simp)  | 
|
969  | 
apply(rule conjI)  | 
|
970  | 
apply(erule_tac x=i in allE)  | 
|
971  | 
apply(case_tac j,simp)  | 
|
| 59807 | 972  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)  | 
| 13020 | 973  | 
apply(case_tac "clist!i",simp,simp)  | 
| 59807 | 974  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)  | 
| 13020 | 975  | 
apply(case_tac "clist!i",simp,simp)  | 
976  | 
apply clarify  | 
|
| 59807 | 977  | 
apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> I j \<longrightarrow> J j" for H I J in allE)  | 
978  | 
apply(erule_tac x=l and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE)  | 
|
| 13020 | 979  | 
apply(case_tac "clist!l",simp,simp)  | 
980  | 
apply(erule_tac x=l in allE,simp)  | 
|
981  | 
apply(rule disjI2)  | 
|
982  | 
apply clarify  | 
|
983  | 
apply(rule tl_zero)  | 
|
984  | 
apply(case_tac j,simp,simp)  | 
|
985  | 
apply(rule tl_zero,force)  | 
|
986  | 
apply force  | 
|
| 59807 | 987  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 988  | 
apply force  | 
| 59807 | 989  | 
apply(erule_tac x=i and P="\<lambda>j. H j \<longrightarrow> (length (s j) = t)" for H s t in allE,force)  | 
| 13020 | 990  | 
apply clarify  | 
991  | 
apply(erule_tac x=i in allE)  | 
|
992  | 
apply(simp add:cp_def)  | 
|
993  | 
apply(rule nth_tl_if)  | 
|
994  | 
apply(simp add:conjoin_def)  | 
|
995  | 
apply clarify  | 
|
996  | 
apply(simp add:same_length_def)  | 
|
997  | 
apply(erule_tac x=i in allE,simp)  | 
|
998  | 
apply simp  | 
|
999  | 
apply simp  | 
|
1000  | 
apply simp  | 
|
1001  | 
apply clarify  | 
|
1002  | 
apply(erule_tac c="(xs, s) # ys" in equalityCE)  | 
|
1003  | 
apply(simp add:par_cp_def)  | 
|
1004  | 
apply simp  | 
|
1005  | 
apply(erule_tac x="map (\<lambda>i. (fst i, s) # snd i) (zip xs clist)" in allE)  | 
|
1006  | 
apply simp  | 
|
1007  | 
apply clarify  | 
|
1008  | 
apply(simp add:cp_def)  | 
|
1009  | 
done  | 
|
1010  | 
||
1011  | 
theorem one: "xs\<noteq>[] \<Longrightarrow>  | 
|
1012  | 
 par_cp xs s = {c. \<exists>clist. (length clist)=(length xs) \<and> 
 | 
|
1013  | 
(\<forall>i<length clist. (clist!i) \<in> cp(xs!i) s) \<and> c \<propto> clist}"  | 
|
1014  | 
apply(frule one_iff_aux)  | 
|
1015  | 
apply(drule sym)  | 
|
1016  | 
apply(erule iffD2)  | 
|
1017  | 
apply clarify  | 
|
1018  | 
apply(rule iffI)  | 
|
1019  | 
apply(erule aux_onlyif)  | 
|
1020  | 
apply clarify  | 
|
1021  | 
apply(force intro:aux_if)  | 
|
1022  | 
done  | 
|
1023  | 
||
| 13187 | 1024  | 
end  |