| author | wenzelm | 
| Fri, 27 Jan 2006 19:03:19 +0100 | |
| changeset 18812 | a4554848b59e | 
| parent 17288 | aa3833fb7bee | 
| permissions | -rw-r--r-- | 
| 4530 | 1  | 
(* Title: HOL/IOA/IOA.ML  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
3  | 
Author: Tobias Nipkow & Konrad Slind  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
4  | 
Copyright 1994 TU Muenchen  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
7  | 
Addsimps [Let_def];  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
8  | 
|
| 17288 | 9  | 
bind_thms ("ioa_projections", [asig_of_def, starts_of_def, trans_of_def]);
 | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
10  | 
|
| 17288 | 11  | 
bind_thms ("exec_rws", [executions_def,is_execution_fragment_def]);
 | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
12  | 
|
| 5069 | 13  | 
Goal  | 
| 8114 | 14  | 
"asig_of(x,y,z) = x & starts_of(x,y,z) = y & trans_of(x,y,z) = z";  | 
| 4089 | 15  | 
by (simp_tac (simpset() addsimps ioa_projections) 1);  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
16  | 
qed "ioa_triple_proj";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
17  | 
|
| 5069 | 18  | 
Goalw [ioa_def,state_trans_def,actions_def, is_asig_def]  | 
| 
5148
 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5143 
diff
changeset
 | 
19  | 
"[| IOA(A); (s1,a,s2):trans_of(A) |] ==> a:actions(asig_of(A))";  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
20  | 
by (REPEAT(etac conjE 1));  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
21  | 
by (EVERY1[etac allE, etac impE, atac]);  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
22  | 
by (Asm_full_simp_tac 1);  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
23  | 
qed "trans_in_actions";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
24  | 
|
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
25  | 
|
| 5069 | 26  | 
Goal "filter_oseq p (filter_oseq p s) = filter_oseq p s";  | 
| 4089 | 27  | 
by (simp_tac (simpset() addsimps [filter_oseq_def]) 1);  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
28  | 
by (rtac ext 1);  | 
| 
8442
 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 
wenzelm 
parents: 
8423 
diff
changeset
 | 
29  | 
by (case_tac "s(i)" 1);  | 
| 3346 | 30  | 
by (Asm_simp_tac 1);  | 
| 4831 | 31  | 
by (Asm_simp_tac 1);  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
32  | 
qed "filter_oseq_idemp";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
33  | 
|
| 5069 | 34  | 
Goalw [mk_trace_def,filter_oseq_def]  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
35  | 
"(mk_trace A s n = None) = \  | 
| 3346 | 36  | 
\ (s(n)=None | (? a. s(n)=Some(a) & a ~: externals(asig_of(A)))) \  | 
37  | 
\ & \  | 
|
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
38  | 
\ (mk_trace A s n = Some(a)) = \  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
39  | 
\ (s(n)=Some(a) & a : externals(asig_of(A)))";  | 
| 
8442
 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 
wenzelm 
parents: 
8423 
diff
changeset
 | 
40  | 
by (case_tac "s(n)" 1);  | 
| 4831 | 41  | 
by (ALLGOALS Asm_simp_tac);  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
42  | 
by (Fast_tac 1);  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
43  | 
qed "mk_trace_thm";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
44  | 
|
| 
5143
 
b94cd208f073
Removal of leading "\!\!..." from most Goal commands
 
paulson 
parents: 
5132 
diff
changeset
 | 
45  | 
Goalw [reachable_def] "s:starts_of(A) ==> reachable A s";  | 
| 3842 | 46  | 
  by (res_inst_tac [("x","(%i. None,%i. s)")] bexI 1);
 | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
47  | 
by (Simp_tac 1);  | 
| 4089 | 48  | 
by (asm_simp_tac (simpset() addsimps exec_rws) 1);  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
49  | 
qed "reachable_0";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
50  | 
|
| 5069 | 51  | 
Goalw (reachable_def::exec_rws)  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
52  | 
"!!A. [| reachable A s; (s,a,t) : trans_of(A) |] ==> reachable A t";  | 
| 4089 | 53  | 
by (asm_full_simp_tac (simpset() delsimps bex_simps) 1);  | 
| 4828 | 54  | 
by (split_all_tac 1);  | 
| 5132 | 55  | 
by Safe_tac;  | 
| 4828 | 56  | 
by (rename_tac "ex1 ex2 n" 1);  | 
| 
4772
 
8c7e7eaffbdf
split_all_tac now fails if there is nothing to split
 
oheimb 
parents: 
4530 
diff
changeset
 | 
57  | 
  by (res_inst_tac [("x","(%i. if i<n then ex1 i                    \
 | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
58  | 
\ else (if i=n then Some a else None), \  | 
| 
4772
 
8c7e7eaffbdf
split_all_tac now fails if there is nothing to split
 
oheimb 
parents: 
4530 
diff
changeset
 | 
59  | 
\ %i. if i<Suc n then ex2 i else t)")] bexI 1);  | 
| 4828 | 60  | 
   by (res_inst_tac [("x","Suc n")] exI 1);
 | 
| 
4772
 
8c7e7eaffbdf
split_all_tac now fails if there is nothing to split
 
oheimb 
parents: 
4530 
diff
changeset
 | 
61  | 
by (Simp_tac 1);  | 
| 
 
8c7e7eaffbdf
split_all_tac now fails if there is nothing to split
 
oheimb 
parents: 
4530 
diff
changeset
 | 
62  | 
by (Asm_full_simp_tac 1);  | 
| 4828 | 63  | 
by (rtac allI 1);  | 
| 9166 | 64  | 
  by (cut_inst_tac [("m","na"),("n","n")] less_linear 1);
 | 
65  | 
by Auto_tac;  | 
|
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
66  | 
qed "reachable_n";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
67  | 
|
| 17288 | 68  | 
val [p1,p2] = goalw (the_context ()) [invariant_def]  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
69  | 
"[| !!s. s:starts_of(A) ==> P(s); \  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
70  | 
\ !!s t a. [|reachable A s; P(s)|] ==> (s,a,t): trans_of(A) --> P(t) |] \  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
71  | 
\ ==> invariant A P";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
72  | 
by (rewrite_goals_tac(reachable_def::Let_def::exec_rws));  | 
| 5132 | 73  | 
by Safe_tac;  | 
| 4828 | 74  | 
by (rename_tac "ex1 ex2 n" 1);  | 
| 
4772
 
8c7e7eaffbdf
split_all_tac now fails if there is nothing to split
 
oheimb 
parents: 
4530 
diff
changeset
 | 
75  | 
  by (res_inst_tac [("Q","reachable A (ex2 n)")] conjunct1 1);
 | 
| 
 
8c7e7eaffbdf
split_all_tac now fails if there is nothing to split
 
oheimb 
parents: 
4530 
diff
changeset
 | 
76  | 
by (Asm_full_simp_tac 1);  | 
| 5184 | 77  | 
by (induct_tac "n" 1);  | 
| 
4772
 
8c7e7eaffbdf
split_all_tac now fails if there is nothing to split
 
oheimb 
parents: 
4530 
diff
changeset
 | 
78  | 
by (fast_tac (claset() addIs [p1,reachable_0]) 1);  | 
| 5184 | 79  | 
  by (eres_inst_tac[("x","na")] allE 1);
 | 
| 
8442
 
96023903c2df
case_tac now subsumes both boolean and datatype cases;
 
wenzelm 
parents: 
8423 
diff
changeset
 | 
80  | 
by (case_tac "ex1 na" 1 THEN ALLGOALS Asm_full_simp_tac);  | 
| 4153 | 81  | 
by Safe_tac;  | 
| 
4772
 
8c7e7eaffbdf
split_all_tac now fails if there is nothing to split
 
oheimb 
parents: 
4530 
diff
changeset
 | 
82  | 
by (etac (p2 RS mp) 1);  | 
| 
 
8c7e7eaffbdf
split_all_tac now fails if there is nothing to split
 
oheimb 
parents: 
4530 
diff
changeset
 | 
83  | 
by (ALLGOALS(fast_tac(claset() addDs [reachable_n])));  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
84  | 
qed "invariantI";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
85  | 
|
| 17288 | 86  | 
val [p1,p2] = goal (the_context ())  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
87  | 
"[| !!s. s : starts_of(A) ==> P(s); \  | 
| 9166 | 88  | 
\ !!s t a. reachable A s ==> P(s) --> (s,a,t):trans_of(A) --> P(t) \  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
89  | 
\ |] ==> invariant A P";  | 
| 4089 | 90  | 
by (fast_tac (claset() addSIs [invariantI] addSDs [p1,p2]) 1);  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
91  | 
qed "invariantI1";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
92  | 
|
| 17288 | 93  | 
val [p1,p2] = goalw (the_context ()) [invariant_def]  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
94  | 
"[| invariant A P; reachable A s |] ==> P(s)";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
95  | 
by (rtac (p2 RS (p1 RS spec RS mp)) 1);  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
96  | 
qed "invariantE";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
97  | 
|
| 17288 | 98  | 
Goal  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
99  | 
"actions(asig_comp a b) = actions(a) Un actions(b)";  | 
| 4089 | 100  | 
by (simp_tac (simpset() addsimps  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
101  | 
([actions_def,asig_comp_def]@asig_projections)) 1);  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
102  | 
by (Fast_tac 1);  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
103  | 
qed "actions_asig_comp";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
104  | 
|
| 5069 | 105  | 
Goal  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
106  | 
"starts_of(A || B) = {p. fst(p):starts_of(A) & snd(p):starts_of(B)}";
 | 
| 4089 | 107  | 
by (simp_tac (simpset() addsimps (par_def::ioa_projections)) 1);  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
108  | 
qed "starts_of_par";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
109  | 
|
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
110  | 
(* Every state in an execution is reachable *)  | 
| 17288 | 111  | 
Goalw [reachable_def]  | 
| 
5148
 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5143 
diff
changeset
 | 
112  | 
"ex:executions(A) ==> !n. reachable A (snd ex n)";  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
113  | 
by (Fast_tac 1);  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
114  | 
qed "states_of_exec_reachable";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
115  | 
|
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
116  | 
|
| 17288 | 117  | 
Goal  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
118  | 
"(s,a,t) : trans_of(A || B || C || D) = \  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
119  | 
\ ((a:actions(asig_of(A)) | a:actions(asig_of(B)) | a:actions(asig_of(C)) | \  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
120  | 
\ a:actions(asig_of(D))) & \  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
121  | 
\ (if a:actions(asig_of(A)) then (fst(s),a,fst(t)):trans_of(A) \  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
122  | 
\ else fst t=fst s) & \  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
123  | 
\ (if a:actions(asig_of(B)) then (fst(snd(s)),a,fst(snd(t))):trans_of(B) \  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
124  | 
\ else fst(snd(t))=fst(snd(s))) & \  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
125  | 
\ (if a:actions(asig_of(C)) then \  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
126  | 
\ (fst(snd(snd(s))),a,fst(snd(snd(t)))):trans_of(C) \  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
127  | 
\ else fst(snd(snd(t)))=fst(snd(snd(s)))) & \  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
128  | 
\ (if a:actions(asig_of(D)) then \  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
129  | 
\ (snd(snd(snd(s))),a,snd(snd(snd(t)))):trans_of(D) \  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
130  | 
\ else snd(snd(snd(t)))=snd(snd(snd(s)))))";  | 
| 9166 | 131  | 
(*SLOW*)  | 
132  | 
by (simp_tac (simpset() addsimps ([par_def,actions_asig_comp,Pair_fst_snd_eq]  | 
|
133  | 
@ ioa_projections)) 1);  | 
|
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
134  | 
qed "trans_of_par4";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
135  | 
|
| 5069 | 136  | 
Goal "starts_of(restrict ioa acts) = starts_of(ioa) & \  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
137  | 
\ trans_of(restrict ioa acts) = trans_of(ioa) & \  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
138  | 
\ reachable (restrict ioa acts) s = reachable ioa s";  | 
| 4089 | 139  | 
by (simp_tac (simpset() addsimps ([is_execution_fragment_def,executions_def,  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
140  | 
reachable_def,restrict_def]@ioa_projections)) 1);  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
141  | 
qed "cancel_restrict";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
142  | 
|
| 5069 | 143  | 
Goal "asig_of(A || B) = asig_comp (asig_of A) (asig_of B)";  | 
| 4089 | 144  | 
by (simp_tac (simpset() addsimps (par_def::ioa_projections)) 1);  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
145  | 
qed "asig_of_par";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
146  | 
|
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
147  | 
|
| 5069 | 148  | 
Goal "externals(asig_of(A1||A2)) = \  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
149  | 
\ (externals(asig_of(A1)) Un externals(asig_of(A2)))";  | 
| 4089 | 150  | 
by (asm_full_simp_tac (simpset() addsimps [externals_def,asig_of_par,asig_comp_def,asig_inputs_def,asig_outputs_def,Un_def,set_diff_def]) 1);  | 
| 17288 | 151  | 
by (rtac set_ext 1);  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
152  | 
by (Fast_tac 1);  | 
| 17288 | 153  | 
qed"externals_of_par";  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
154  | 
|
| 5069 | 155  | 
Goalw [externals_def,actions_def,compat_ioas_def,compat_asigs_def]  | 
| 
5148
 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5143 
diff
changeset
 | 
156  | 
"[| compat_ioas A1 A2; a:externals(asig_of(A1))|] ==> a~:internals(asig_of(A2))";  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
157  | 
by (Asm_full_simp_tac 1);  | 
| 12886 | 158  | 
by (Best_tac 1);  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
159  | 
qed"ext1_is_not_int2";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
160  | 
|
| 5069 | 161  | 
Goalw [externals_def,actions_def,compat_ioas_def,compat_asigs_def]  | 
| 
5148
 
74919e8f221c
More tidying and removal of "\!\!... from Goal commands
 
paulson 
parents: 
5143 
diff
changeset
 | 
162  | 
"[| compat_ioas A2 A1 ; a:externals(asig_of(A1))|] ==> a~:internals(asig_of(A2))";  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
163  | 
by (Asm_full_simp_tac 1);  | 
| 12886 | 164  | 
by (Best_tac 1);  | 
| 
3078
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
165  | 
qed"ext2_is_not_int1";  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
166  | 
|
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
167  | 
val ext1_ext2_is_not_act2 = ext1_is_not_int2 RS int_and_ext_is_act;  | 
| 
 
984866a8f905
old IOA meta theory (see also new version in HOLCF/IOA/meta_theory);
 
mueller 
parents:  
diff
changeset
 | 
168  | 
val ext1_ext2_is_not_act1 = ext2_is_not_int1 RS int_and_ext_is_act;  |