| author | wenzelm | 
| Tue, 28 Jul 2009 15:05:18 +0200 | |
| changeset 32252 | fd5e4a1a4ea6 | 
| parent 19769 | c40ce2de2020 | 
| child 35416 | d8d7d1b785af | 
| permissions | -rw-r--r-- | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/UNITY/Token  | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
4  | 
Copyright 1998 University of Cambridge  | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
5  | 
*)  | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
7  | 
|
| 15618 | 8  | 
header {*The Token Ring*}
 | 
9  | 
||
10  | 
theory Token  | 
|
11  | 
imports "../WFair"  | 
|
12  | 
||
13  | 
begin  | 
|
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
14  | 
|
| 15618 | 15  | 
text{*From Misra, "A Logic for Concurrent Programming" (1994), sections 5.2 and 13.2.*}
 | 
16  | 
||
17  | 
subsection{*Definitions*}
 | 
|
18  | 
||
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
19  | 
datatype pstate = Hungry | Eating | Thinking  | 
| 15618 | 20  | 
    --{*process states*}
 | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
21  | 
|
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
22  | 
record state =  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
23  | 
token :: "nat"  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
24  | 
proc :: "nat => pstate"  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
25  | 
|
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
26  | 
|
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
27  | 
constdefs  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
28  | 
HasTok :: "nat => state set"  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
29  | 
    "HasTok i == {s. token s = i}"
 | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
30  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
31  | 
H :: "nat => state set"  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
32  | 
    "H i == {s. proc s i = Hungry}"
 | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
33  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
34  | 
E :: "nat => state set"  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
35  | 
    "E i == {s. proc s i = Eating}"
 | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
36  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
37  | 
T :: "nat => state set"  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
38  | 
    "T i == {s. proc s i = Thinking}"
 | 
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
39  | 
|
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
40  | 
|
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
41  | 
locale Token =  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
42  | 
fixes N and F and nodeOrder and "next"  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
43  | 
defines nodeOrder_def:  | 
| 15618 | 44  | 
       "nodeOrder j == measure(%i. ((j+N)-i) mod N) \<inter> {..<N} \<times> {..<N}"
 | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
45  | 
and next_def:  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
46  | 
"next i == (Suc i) mod N"  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
47  | 
assumes N_positive [iff]: "0<N"  | 
| 13806 | 48  | 
and TR2: "F \<in> (T i) co (T i \<union> H i)"  | 
49  | 
and TR3: "F \<in> (H i) co (H i \<union> E i)"  | 
|
50  | 
and TR4: "F \<in> (H i - HasTok i) co (H i)"  | 
|
51  | 
and TR5: "F \<in> (HasTok i) co (HasTok i \<union> -(E i))"  | 
|
52  | 
and TR6: "F \<in> (H i \<inter> HasTok i) leadsTo (E i)"  | 
|
53  | 
and TR7: "F \<in> (HasTok i) leadsTo (HasTok (next i))"  | 
|
| 
13785
 
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  | 
|
| 13806 | 56  | 
lemma HasToK_partition: "[| s \<in> HasTok i; s \<in> HasTok j |] ==> i=j"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
57  | 
by (unfold HasTok_def, auto)  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
58  | 
|
| 13806 | 59  | 
lemma not_E_eq: "(s \<notin> E i) = (s \<in> H i | s \<in> T i)"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
60  | 
apply (simp add: H_def E_def T_def)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
61  | 
apply (case_tac "proc s i", auto)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
62  | 
done  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
63  | 
|
| 13806 | 64  | 
lemma (in Token) token_stable: "F \<in> stable (-(E i) \<union> (HasTok i))"  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
65  | 
apply (unfold stable_def)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
66  | 
apply (rule constrains_weaken)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
67  | 
apply (rule constrains_Un [OF constrains_Un [OF TR2 TR4] TR5])  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
68  | 
apply (auto simp add: not_E_eq)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
69  | 
apply (simp_all add: H_def E_def T_def)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
70  | 
done  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
71  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
72  | 
|
| 15618 | 73  | 
subsection{*Progress under Weak Fairness*}
 | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
74  | 
|
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
75  | 
lemma (in Token) wf_nodeOrder: "wf(nodeOrder j)"  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
76  | 
apply (unfold nodeOrder_def)  | 
| 15618 | 77  | 
apply (rule wf_measure [THEN wf_subset], blast)  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
78  | 
done  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
79  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
80  | 
lemma (in Token) nodeOrder_eq:  | 
| 13806 | 81  | 
"[| i<N; j<N |] ==> ((next i, i) \<in> nodeOrder j) = (i \<noteq> j)"  | 
| 
19769
 
c40ce2de2020
Added [simp]-lemmas "in_inv_image" and "in_lex_prod" in the spirit of "in_measure".
 
krauss 
parents: 
15618 
diff
changeset
 | 
82  | 
apply (unfold nodeOrder_def next_def)  | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
83  | 
apply (auto simp add: mod_Suc mod_geq)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
84  | 
apply (auto split add: nat_diff_split simp add: linorder_neq_iff mod_geq)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
85  | 
done  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
86  | 
|
| 15618 | 87  | 
text{*From "A Logic for Concurrent Programming", but not used in Chapter 4.
 | 
88  | 
  Note the use of @{text case_tac}.  Reasoning about leadsTo takes practice!*}
 | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
89  | 
lemma (in Token) TR7_nodeOrder:  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
90  | 
"[| i<N; j<N |] ==>  | 
| 13806 | 91  | 
      F \<in> (HasTok i) leadsTo ({s. (token s, i) \<in> nodeOrder j} \<union> HasTok j)"
 | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
92  | 
apply (case_tac "i=j")  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
93  | 
apply (blast intro: subset_imp_leadsTo)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
94  | 
apply (rule TR7 [THEN leadsTo_weaken_R])  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
95  | 
apply (auto simp add: HasTok_def nodeOrder_eq)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
96  | 
done  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
97  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
98  | 
|
| 15618 | 99  | 
text{*Chapter 4 variant, the one actually used below.*}
 | 
| 13806 | 100  | 
lemma (in Token) TR7_aux: "[| i<N; j<N; i\<noteq>j |]  | 
101  | 
      ==> F \<in> (HasTok i) leadsTo {s. (token s, i) \<in> nodeOrder j}"
 | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
102  | 
apply (rule TR7 [THEN leadsTo_weaken_R])  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
103  | 
apply (auto simp add: HasTok_def nodeOrder_eq)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
104  | 
done  | 
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
105  | 
|
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
106  | 
lemma (in Token) token_lemma:  | 
| 13806 | 107  | 
     "({s. token s < N} \<inter> token -` {m}) = (if m<N then token -` {m} else {})"
 | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
108  | 
by auto  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
109  | 
|
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
110  | 
|
| 15618 | 111  | 
text{*Misra's TR9: the token reaches an arbitrary node*}
 | 
| 13806 | 112  | 
lemma  (in Token) leadsTo_j: "j<N ==> F \<in> {s. token s < N} leadsTo (HasTok j)"
 | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
113  | 
apply (rule leadsTo_weaken_R)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
114  | 
apply (rule_tac I = "-{j}" and f = token and B = "{}" 
 | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
115  | 
in wf_nodeOrder [THEN bounded_induct])  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
116  | 
apply (simp_all (no_asm_simp) add: token_lemma vimage_Diff HasTok_def)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
117  | 
prefer 2 apply blast  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
118  | 
apply clarify  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
119  | 
apply (rule TR7_aux [THEN leadsTo_weaken])  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
120  | 
apply (auto simp add: HasTok_def nodeOrder_def)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
121  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
122  | 
|
| 15618 | 123  | 
text{*Misra's TR8: a hungry process eventually eats*}
 | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
124  | 
lemma (in Token) token_progress:  | 
| 13806 | 125  | 
     "j<N ==> F \<in> ({s. token s < N} \<inter> H j) leadsTo (E j)"
 | 
| 
13785
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
126  | 
apply (rule leadsTo_cancel1 [THEN leadsTo_Un_duplicate])  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
127  | 
apply (rule_tac [2] TR6)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
128  | 
apply (rule psp [OF leadsTo_j TR3, THEN leadsTo_weaken], blast+)  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
129  | 
done  | 
| 
 
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
 
paulson 
parents: 
11195 
diff
changeset
 | 
130  | 
|
| 
11195
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
131  | 
|
| 
 
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
 
paulson 
parents:  
diff
changeset
 | 
132  | 
end  |