author | wenzelm |
Fri, 06 Oct 2000 17:35:58 +0200 | |
changeset 10168 | 50be659d4222 |
parent 9190 | b86ff604729f |
permissions | -rw-r--r-- |
5232 | 1 |
(* Title: HOL/UNITY/Deadlock |
5111 | 2 |
ID: $Id$ |
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
4776 | 5 |
|
5111 | 6 |
Deadlock examples from section 5.6 of |
7 |
Misra, "A Logic for Concurrent Programming", 1994 |
|
8 |
*) |
|
4776 | 9 |
|
10 |
(*Trivial, two-process case*) |
|
5069 | 11 |
Goalw [constrains_def, stable_def] |
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
6536
diff
changeset
|
12 |
"[| F : (A Int B) co A; F : (B Int A) co B |] ==> F : stable (A Int B)"; |
4776 | 13 |
by (Blast_tac 1); |
14 |
result(); |
|
15 |
||
16 |
||
17 |
(*a simplification step*) |
|
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
6536
diff
changeset
|
18 |
Goal "(INT i: atMost n. A(Suc i) Int A i) = (INT i: atMost (Suc n). A i)"; |
4776 | 19 |
by (induct_tac "n" 1); |
5648 | 20 |
by (ALLGOALS (asm_simp_tac (simpset() addsimps [atMost_Suc]))); |
9190 | 21 |
by Auto_tac; |
4776 | 22 |
qed "Collect_le_Int_equals"; |
23 |
||
24 |
(*Dual of the required property. Converse inclusion fails.*) |
|
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
6536
diff
changeset
|
25 |
Goal "(UN i: lessThan n. A i) Int (- A n) <= \ |
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
6536
diff
changeset
|
26 |
\ (UN i: lessThan n. (A i) Int (- A (Suc i)))"; |
4776 | 27 |
by (induct_tac "n" 1); |
28 |
by (Asm_simp_tac 1); |
|
5648 | 29 |
by (simp_tac (simpset() addsimps [lessThan_Suc]) 1); |
4776 | 30 |
by (Blast_tac 1); |
31 |
qed "UN_Int_Compl_subset"; |
|
32 |
||
33 |
||
34 |
(*Converse inclusion fails.*) |
|
5648 | 35 |
Goal "(INT i: lessThan n. -A i Un A (Suc i)) <= \ |
36 |
\ (INT i: lessThan n. -A i) Un A n"; |
|
4776 | 37 |
by (induct_tac "n" 1); |
38 |
by (Asm_simp_tac 1); |
|
5648 | 39 |
by (asm_simp_tac (simpset() addsimps [lessThan_Suc]) 1); |
4776 | 40 |
by (Blast_tac 1); |
41 |
qed "INT_Un_Compl_subset"; |
|
42 |
||
43 |
||
44 |
(*Specialized rewriting*) |
|
5648 | 45 |
Goal "A 0 Int (-(A n) Int (INT i: lessThan n. -A i Un A (Suc i))) = {}"; |
4776 | 46 |
by (blast_tac (claset() addIs [gr0I] |
47 |
addDs [impOfSubs INT_Un_Compl_subset]) 1); |
|
48 |
val lemma = result(); |
|
49 |
||
50 |
(*Reverse direction makes it harder to invoke the ind hyp*) |
|
5648 | 51 |
Goal "(INT i: atMost n. A i) = \ |
52 |
\ A 0 Int (INT i: lessThan n. -A i Un A(Suc i))"; |
|
4776 | 53 |
by (induct_tac "n" 1); |
54 |
by (Asm_simp_tac 1); |
|
55 |
by (asm_simp_tac |
|
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
6536
diff
changeset
|
56 |
(simpset() addsimps Int_ac @ [Int_Un_distrib, Int_Un_distrib2, lemma, |
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
6536
diff
changeset
|
57 |
lessThan_Suc, atMost_Suc]) 1); |
4776 | 58 |
qed "INT_le_equals_Int"; |
59 |
||
5648 | 60 |
Goal "(INT i: atMost (Suc n). A i) = \ |
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
6536
diff
changeset
|
61 |
\ A 0 Int (INT i: atMost n. -A i Un A(Suc i))"; |
5648 | 62 |
by (simp_tac (simpset() addsimps [lessThan_Suc_atMost, INT_le_equals_Int]) 1); |
4776 | 63 |
qed "INT_le_Suc_equals_Int"; |
64 |
||
65 |
||
66 |
(*The final deadlock example*) |
|
8334
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
7127
diff
changeset
|
67 |
val [zeroprem, allprem] = Goalw [stable_def] |
6536 | 68 |
"[| F : (A 0 Int A (Suc n)) co (A 0); \ |
8334
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
7127
diff
changeset
|
69 |
\ !!i. i: atMost n ==> F : (A(Suc i) Int A i) co (-A i Un A(Suc i)) |] \ |
5648 | 70 |
\ ==> F : stable (INT i: atMost (Suc n). A i)"; |
8334
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
7127
diff
changeset
|
71 |
by (rtac ([zeroprem, constrains_INT] MRS |
7127
48e235179ffb
added parentheses to cope with a possible reduction of the precedence of unary
paulson
parents:
6536
diff
changeset
|
72 |
constrains_Int RS constrains_weaken) 1); |
8334
7896bcbd8641
Added Tanja's Detects and Reachability theories. Also
paulson
parents:
7127
diff
changeset
|
73 |
by (etac allprem 1); |
4776 | 74 |
by (simp_tac (simpset() addsimps [Collect_le_Int_equals, |
75 |
Int_assoc, INT_absorb]) 1); |
|
5536 | 76 |
by (simp_tac (simpset() addsimps [INT_le_Suc_equals_Int]) 1); |
4776 | 77 |
result(); |
78 |