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