| author | wenzelm | 
| Tue, 05 Oct 1999 15:30:14 +0200 | |
| changeset 7718 | 86755cc5b83c | 
| parent 7594 | 8a188ef6545e | 
| child 7878 | 43b03d412b82 | 
| permissions | -rw-r--r-- | 
| 5597 | 1  | 
(* Title: HOL/UNITY/Comp.thy  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
4  | 
Copyright 1998 University of Cambridge  | 
|
5  | 
||
6  | 
Composition  | 
|
7  | 
||
8  | 
From Chandy and Sanders, "Reasoning About Program Composition"  | 
|
9  | 
*)  | 
|
10  | 
||
11  | 
(*** component ***)  | 
|
12  | 
||
| 
6738
 
06189132c67b
component_eq_subset: a neat characterization of "component"
 
paulson 
parents: 
6703 
diff
changeset
 | 
13  | 
Goalw [component_def]  | 
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7386 
diff
changeset
 | 
14  | 
"H <= F | H <= G ==> H <= (F Join G)";  | 
| 7361 | 15  | 
by Auto_tac;  | 
16  | 
by (res_inst_tac [("x", "G Join Ga")] exI 1);
 | 
|
17  | 
by (res_inst_tac [("x", "G Join F")] exI 2);
 | 
|
18  | 
by (auto_tac (claset(), simpset() addsimps Join_ac));  | 
|
19  | 
qed "componentI";  | 
|
20  | 
||
21  | 
Goalw [component_def]  | 
|
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7386 
diff
changeset
 | 
22  | 
"(F <= G) = (Init G <= Init F & Acts F <= Acts G)";  | 
| 
6738
 
06189132c67b
component_eq_subset: a neat characterization of "component"
 
paulson 
parents: 
6703 
diff
changeset
 | 
23  | 
by (force_tac (claset() addSIs [exI, program_equalityI],  | 
| 7537 | 24  | 
simpset()) 1);  | 
| 
6738
 
06189132c67b
component_eq_subset: a neat characterization of "component"
 
paulson 
parents: 
6703 
diff
changeset
 | 
25  | 
qed "component_eq_subset";  | 
| 
 
06189132c67b
component_eq_subset: a neat characterization of "component"
 
paulson 
parents: 
6703 
diff
changeset
 | 
26  | 
|
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7386 
diff
changeset
 | 
27  | 
Goalw [component_def] "SKIP <= F";  | 
| 
6012
 
1894bfc4aee9
Addition of the States component; parts of Comp not working
 
paulson 
parents: 
5968 
diff
changeset
 | 
28  | 
by (force_tac (claset() addIs [Join_SKIP_left], simpset()) 1);  | 
| 
5612
 
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
 
paulson 
parents: 
5597 
diff
changeset
 | 
29  | 
qed "component_SKIP";  | 
| 
 
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
 
paulson 
parents: 
5597 
diff
changeset
 | 
30  | 
|
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7386 
diff
changeset
 | 
31  | 
Goalw [component_def] "F <= (F :: 'a program)";  | 
| 
6295
 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 
paulson 
parents: 
6018 
diff
changeset
 | 
32  | 
by (blast_tac (claset() addIs [Join_SKIP_right]) 1);  | 
| 5597 | 33  | 
qed "component_refl";  | 
34  | 
||
| 
5612
 
e981ca6f7332
Finished proofs to end of section 5.1 of Chandy and Sanders
 
paulson 
parents: 
5597 
diff
changeset
 | 
35  | 
AddIffs [component_SKIP, component_refl];  | 
| 5597 | 36  | 
|
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7386 
diff
changeset
 | 
37  | 
Goal "F <= SKIP ==> F = SKIP";  | 
| 6833 | 38  | 
by (auto_tac (claset() addSIs [program_equalityI],  | 
39  | 
simpset() addsimps [component_eq_subset]));  | 
|
40  | 
qed "SKIP_minimal";  | 
|
41  | 
||
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7386 
diff
changeset
 | 
42  | 
Goalw [component_def] "F <= (F Join G)";  | 
| 5968 | 43  | 
by (Blast_tac 1);  | 
44  | 
qed "component_Join1";  | 
|
45  | 
||
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7386 
diff
changeset
 | 
46  | 
Goalw [component_def] "G <= (F Join G)";  | 
| 5968 | 47  | 
by (simp_tac (simpset() addsimps [Join_commute]) 1);  | 
48  | 
by (Blast_tac 1);  | 
|
49  | 
qed "component_Join2";  | 
|
50  | 
||
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7386 
diff
changeset
 | 
51  | 
Goalw [component_def] "i : I ==> (F i) <= (JN i:I. (F i))";  | 
| 
6295
 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 
paulson 
parents: 
6018 
diff
changeset
 | 
52  | 
by (blast_tac (claset() addIs [JN_absorb]) 1);  | 
| 5968 | 53  | 
qed "component_JN";  | 
54  | 
||
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7386 
diff
changeset
 | 
55  | 
Goalw [component_def] "[| F <= G; G <= H |] ==> F <= (H :: 'a program)";  | 
| 
6295
 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 
paulson 
parents: 
6018 
diff
changeset
 | 
56  | 
by (blast_tac (claset() addIs [Join_assoc RS sym]) 1);  | 
| 5597 | 57  | 
qed "component_trans";  | 
58  | 
||
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7386 
diff
changeset
 | 
59  | 
Goal "[| F <= G; G <= F |] ==> F = (G :: 'a program)";  | 
| 
6738
 
06189132c67b
component_eq_subset: a neat characterization of "component"
 
paulson 
parents: 
6703 
diff
changeset
 | 
60  | 
by (full_simp_tac (simpset() addsimps [component_eq_subset]) 1);  | 
| 
 
06189132c67b
component_eq_subset: a neat characterization of "component"
 
paulson 
parents: 
6703 
diff
changeset
 | 
61  | 
by (blast_tac (claset() addSIs [program_equalityI]) 1);  | 
| 6703 | 62  | 
qed "component_antisym";  | 
| 5597 | 63  | 
|
| 
5804
 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 
paulson 
parents: 
5706 
diff
changeset
 | 
64  | 
Goalw [component_def]  | 
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7386 
diff
changeset
 | 
65  | 
"F <= H = (EX G. F Join G = H & Disjoint F G)";  | 
| 
6295
 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 
paulson 
parents: 
6018 
diff
changeset
 | 
66  | 
by (blast_tac (claset() addSIs [Diff_Disjoint, Join_Diff2]) 1);  | 
| 
5804
 
8e0a4c4fd67b
Revising the Client proof as suggested by Michel Charpentier.  New lemmas
 
paulson 
parents: 
5706 
diff
changeset
 | 
67  | 
qed "component_eq";  | 
| 5597 | 68  | 
|
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7386 
diff
changeset
 | 
69  | 
Goal "((F Join G) <= H) = (F <= H & G <= H)";  | 
| 7537 | 70  | 
by (simp_tac (simpset() addsimps [component_eq_subset]) 1);  | 
| 6833 | 71  | 
by (Blast_tac 1);  | 
72  | 
qed "Join_component_iff";  | 
|
73  | 
||
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7386 
diff
changeset
 | 
74  | 
Goal "[| F <= G; G : A co B |] ==> F : A co B";  | 
| 7386 | 75  | 
by (auto_tac (claset(),  | 
76  | 
simpset() addsimps [constrains_def, component_eq_subset]));  | 
|
77  | 
qed "component_constrains";  | 
|
| 
6295
 
351b3c2b0d83
removed the infernal States, eqStates, compatible, etc.
 
paulson 
parents: 
6018 
diff
changeset
 | 
78  | 
|
| 
7399
 
cf780c2bcccf
changed "component" infix in HOL/UNITY/Comp.thy to be overloaded <
 
paulson 
parents: 
7386 
diff
changeset
 | 
79  | 
bind_thm ("program_less_le", strict_component_def RS meta_eq_to_obj_eq);
 | 
| 
7594
 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 
paulson 
parents: 
7537 
diff
changeset
 | 
80  | 
|
| 
 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 
paulson 
parents: 
7537 
diff
changeset
 | 
81  | 
Goal "Diff F (Acts G) <= F";  | 
| 
 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 
paulson 
parents: 
7537 
diff
changeset
 | 
82  | 
by (auto_tac (claset() addSIs [program_equalityI],  | 
| 
 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 
paulson 
parents: 
7537 
diff
changeset
 | 
83  | 
simpset() addsimps [Diff_def, component_eq_subset]));  | 
| 
 
8a188ef6545e
working version with  co-guarantees-leadsto  results
 
paulson 
parents: 
7537 
diff
changeset
 | 
84  | 
qed "Diff_component";  |