| author | blanchet | 
| Tue, 21 Sep 2010 10:02:50 +0200 | |
| changeset 39598 | 57413334669d | 
| parent 37936 | 1e4c5015a72e | 
| child 44871 | fbfdc5ac86be | 
| permissions | -rw-r--r-- | 
| 37936 | 1  | 
(* Title: HOL/UNITY/Simple/Deadlock.thy  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
3  | 
Copyright 1998 University of Cambridge  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
4  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
5  | 
Deadlock examples from section 5.6 of  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
6  | 
Misra, "A Logic for Concurrent Programming", 1994  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
7  | 
*)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
8  | 
|
| 16417 | 9  | 
theory Deadlock imports UNITY begin  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
10  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
11  | 
(*Trivial, two-process case*)  | 
| 13806 | 12  | 
lemma "[| F \<in> (A \<inter> B) co A; F \<in> (B \<inter> A) co B |] ==> F \<in> stable (A \<inter> B)"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
13  | 
by (unfold constrains_def stable_def, blast)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
14  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
15  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
16  | 
(*a simplification step*)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
17  | 
lemma Collect_le_Int_equals:  | 
| 13806 | 18  | 
"(\<Inter>i \<in> atMost n. A(Suc i) \<inter> A i) = (\<Inter>i \<in> atMost (Suc n). A i)"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
19  | 
apply (induct_tac "n")  | 
| 13806 | 20  | 
apply (auto simp add: atMost_Suc)  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
21  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
22  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
23  | 
(*Dual of the required property. Converse inclusion fails.*)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
24  | 
lemma UN_Int_Compl_subset:  | 
| 13806 | 25  | 
"(\<Union>i \<in> lessThan n. A i) \<inter> (- A n) \<subseteq>  | 
26  | 
(\<Union>i \<in> lessThan n. (A i) \<inter> (- A (Suc i)))"  | 
|
27  | 
apply (induct_tac "n", simp)  | 
|
28  | 
apply (simp add: lessThan_Suc, blast)  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
29  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
30  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
31  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
32  | 
(*Converse inclusion fails.*)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
33  | 
lemma INT_Un_Compl_subset:  | 
| 13806 | 34  | 
"(\<Inter>i \<in> lessThan n. -A i \<union> A (Suc i)) \<subseteq>  | 
35  | 
(\<Inter>i \<in> lessThan n. -A i) \<union> A n"  | 
|
36  | 
apply (induct_tac "n", simp)  | 
|
| 
26824
 
32e612e77edb
Replaced blast by fast in proof of INT_Un_Compl_subset, since blast looped
 
berghofe 
parents: 
16417 
diff
changeset
 | 
37  | 
apply (simp add: lessThan_Suc, fast)  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
38  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
39  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
40  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
41  | 
(*Specialized rewriting*)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
42  | 
lemma INT_le_equals_Int_lemma:  | 
| 13806 | 43  | 
     "A 0 \<inter> (-(A n) \<inter> (\<Inter>i \<in> lessThan n. -A i \<union> A (Suc i))) = {}"
 | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
44  | 
by (blast intro: gr0I dest: INT_Un_Compl_subset [THEN subsetD])  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
45  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
46  | 
(*Reverse direction makes it harder to invoke the ind hyp*)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
47  | 
lemma INT_le_equals_Int:  | 
| 13806 | 48  | 
"(\<Inter>i \<in> atMost n. A i) =  | 
49  | 
A 0 \<inter> (\<Inter>i \<in> lessThan n. -A i \<union> A(Suc i))"  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
50  | 
apply (induct_tac "n", simp)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
51  | 
apply (simp add: Int_ac Int_Un_distrib Int_Un_distrib2  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
52  | 
INT_le_equals_Int_lemma lessThan_Suc atMost_Suc)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
53  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
54  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
55  | 
lemma INT_le_Suc_equals_Int:  | 
| 13806 | 56  | 
"(\<Inter>i \<in> atMost (Suc n). A i) =  | 
57  | 
A 0 \<inter> (\<Inter>i \<in> atMost n. -A i \<union> A(Suc i))"  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
58  | 
by (simp add: lessThan_Suc_atMost INT_le_equals_Int)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
59  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
60  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
61  | 
(*The final deadlock example*)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
62  | 
lemma  | 
| 13806 | 63  | 
assumes zeroprem: "F \<in> (A 0 \<inter> A (Suc n)) co (A 0)"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
64  | 
and allprem:  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
26824 
diff
changeset
 | 
65  | 
"!!i. i \<in> atMost n ==> F \<in> (A(Suc i) \<inter> A i) co (-A i \<union> A(Suc i))"  | 
| 13806 | 66  | 
shows "F \<in> stable (\<Inter>i \<in> atMost (Suc n). A i)"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
67  | 
apply (unfold stable_def)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
68  | 
apply (rule constrains_Int [THEN constrains_weaken])  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
69  | 
apply (rule zeroprem)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
70  | 
apply (rule constrains_INT)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
71  | 
apply (erule allprem)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
72  | 
apply (simp add: Collect_le_Int_equals Int_assoc INT_absorb)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
73  | 
apply (simp add: INT_le_Suc_equals_Int)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
74  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
75  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
76  | 
end  |