| author | wenzelm | 
| Thu, 22 Jul 2010 18:08:39 +0200 | |
| changeset 37936 | 1e4c5015a72e | 
| parent 16417 | 9bc16273c2d4 | 
| child 57488 | 58db442609ac | 
| permissions | -rw-r--r-- | 
| 37936 | 1  | 
(* Title: HOL/UNITY/Detects.thy  | 
| 8334 | 2  | 
Author: Tanja Vos, Cambridge University Computer Laboratory  | 
3  | 
Copyright 2000 University of Cambridge  | 
|
4  | 
||
5  | 
Detects definition (Section 3.8 of Chandy & Misra) using LeadsTo  | 
|
6  | 
*)  | 
|
7  | 
||
| 13798 | 8  | 
header{*The Detects Relation*}
 | 
9  | 
||
| 16417 | 10  | 
theory Detects imports FP SubstAx begin  | 
| 8334 | 11  | 
|
12  | 
consts  | 
|
13  | 
op_Detects :: "['a set, 'a set] => 'a program set" (infixl "Detects" 60)  | 
|
14  | 
op_Equality :: "['a set, 'a set] => 'a set" (infixl "<==>" 60)  | 
|
15  | 
||
16  | 
defs  | 
|
| 13805 | 17  | 
Detects_def: "A Detects B == (Always (-A \<union> B)) \<inter> (B LeadsTo A)"  | 
18  | 
Equality_def: "A <==> B == (-A \<union> B) \<inter> (A \<union> -B)"  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
19  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
20  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
21  | 
(* Corollary from Sectiom 3.6.4 *)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
22  | 
|
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13805 
diff
changeset
 | 
23  | 
lemma Always_at_FP:  | 
| 
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13805 
diff
changeset
 | 
24  | 
"[|F \<in> A LeadsTo B; all_total F|] ==> F \<in> Always (-((FP F) \<inter> A \<inter> -B))"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
25  | 
apply (rule LeadsTo_empty)  | 
| 13805 | 26  | 
apply (subgoal_tac "F \<in> (FP F \<inter> A \<inter> - B) LeadsTo (B \<inter> (FP F \<inter> -B))")  | 
27  | 
apply (subgoal_tac [2] " (FP F \<inter> A \<inter> - B) = (A \<inter> (FP F \<inter> -B))")  | 
|
28  | 
apply (subgoal_tac "(B \<inter> (FP F \<inter> -B)) = {}")
 | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
29  | 
apply auto  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
30  | 
apply (blast intro: PSP_Stable stable_imp_Stable stable_FP_Int)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
31  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
32  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
33  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
34  | 
lemma Detects_Trans:  | 
| 13805 | 35  | 
"[| F \<in> A Detects B; F \<in> B Detects C |] ==> F \<in> A Detects C"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
36  | 
apply (unfold Detects_def Int_def)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
37  | 
apply (simp (no_asm))  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
38  | 
apply safe  | 
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13805 
diff
changeset
 | 
39  | 
apply (rule_tac [2] LeadsTo_Trans, auto)  | 
| 13805 | 40  | 
apply (subgoal_tac "F \<in> Always ((-A \<union> B) \<inter> (-B \<union> C))")  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
41  | 
apply (blast intro: Always_weaken)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
42  | 
apply (simp add: Always_Int_distrib)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
43  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
44  | 
|
| 13805 | 45  | 
lemma Detects_refl: "F \<in> A Detects A"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
46  | 
apply (unfold Detects_def)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
47  | 
apply (simp (no_asm) add: Un_commute Compl_partition subset_imp_LeadsTo)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
48  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
49  | 
|
| 13805 | 50  | 
lemma Detects_eq_Un: "(A<==>B) = (A \<inter> B) \<union> (-A \<inter> -B)"  | 
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13805 
diff
changeset
 | 
51  | 
by (unfold Equality_def, blast)  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
52  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
53  | 
(*Not quite antisymmetry: sets A and B agree in all reachable states *)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
54  | 
lemma Detects_antisym:  | 
| 13805 | 55  | 
"[| F \<in> A Detects B; F \<in> B Detects A|] ==> F \<in> Always (A <==> B)"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
56  | 
apply (unfold Detects_def Equality_def)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
57  | 
apply (simp add: Always_Int_I Un_commute)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
58  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
59  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
60  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
61  | 
(* Theorem from Section 3.8 *)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
62  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
63  | 
lemma Detects_Always:  | 
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13805 
diff
changeset
 | 
64  | 
"[|F \<in> A Detects B; all_total F|] ==> F \<in> Always (-(FP F) \<union> (A <==> B))"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
65  | 
apply (unfold Detects_def Equality_def)  | 
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13805 
diff
changeset
 | 
66  | 
apply (simp add: Un_Int_distrib Always_Int_distrib)  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
67  | 
apply (blast dest: Always_at_FP intro: Always_weaken)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
68  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
69  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
70  | 
(* Theorem from exercise 11.1 Section 11.3.1 *)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
71  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
72  | 
lemma Detects_Imp_LeadstoEQ:  | 
| 13805 | 73  | 
"F \<in> A Detects B ==> F \<in> UNIV LeadsTo (A <==> B)"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
74  | 
apply (unfold Detects_def Equality_def)  | 
| 
13812
 
91713a1915ee
converting HOL/UNITY to use unconditional fairness
 
paulson 
parents: 
13805 
diff
changeset
 | 
75  | 
apply (rule_tac B = B in LeadsTo_Diff)  | 
| 13805 | 76  | 
apply (blast intro: Always_LeadsToI subset_imp_LeadsTo)  | 
77  | 
apply (blast intro: Always_LeadsTo_weaken)  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
78  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
8334 
diff
changeset
 | 
79  | 
|
| 8334 | 80  | 
|
81  | 
end  | 
|
82  |