author | kleing |
Tue, 13 May 2003 08:59:21 +0200 | |
changeset 14024 | 213dcc39358f |
parent 13806 | fd40c9d9076b |
child 15618 | 05bad476e0f0 |
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 |
The Token Ring. |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
7 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
8 |
From Misra, "A Logic for Concurrent Programming" (1994), sections 5.2 and 13.2. |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
9 |
*) |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
10 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
11 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
12 |
theory Token = WFair: |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
13 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
14 |
(*process states*) |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
15 |
datatype pstate = Hungry | Eating | Thinking |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
16 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
17 |
record state = |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
18 |
token :: "nat" |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
19 |
proc :: "nat => pstate" |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
20 |
|
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 |
constdefs |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
23 |
HasTok :: "nat => state set" |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
24 |
"HasTok i == {s. token s = i}" |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
25 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
26 |
H :: "nat => state set" |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
27 |
"H i == {s. proc s i = Hungry}" |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
28 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
29 |
E :: "nat => state set" |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
30 |
"E i == {s. proc s i = Eating}" |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
31 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
32 |
T :: "nat => state set" |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
33 |
"T i == {s. proc s i = Thinking}" |
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
34 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
35 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
36 |
locale Token = |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
37 |
fixes N and F and nodeOrder and "next" |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
38 |
defines nodeOrder_def: |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
39 |
"nodeOrder j == (inv_image less_than (%i. ((j+N)-i) mod N)) Int |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
40 |
(lessThan N <*> lessThan N)" |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
41 |
and next_def: |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
42 |
"next i == (Suc i) mod N" |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
43 |
assumes N_positive [iff]: "0<N" |
13806 | 44 |
and TR2: "F \<in> (T i) co (T i \<union> H i)" |
45 |
and TR3: "F \<in> (H i) co (H i \<union> E i)" |
|
46 |
and TR4: "F \<in> (H i - HasTok i) co (H i)" |
|
47 |
and TR5: "F \<in> (HasTok i) co (HasTok i \<union> -(E i))" |
|
48 |
and TR6: "F \<in> (H i \<inter> HasTok i) leadsTo (E i)" |
|
49 |
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
|
50 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
51 |
|
13806 | 52 |
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
|
53 |
by (unfold HasTok_def, auto) |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
54 |
|
13806 | 55 |
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
|
56 |
apply (simp add: H_def E_def T_def) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
57 |
apply (case_tac "proc s i", auto) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
58 |
done |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
59 |
|
13806 | 60 |
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
|
61 |
apply (unfold stable_def) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
62 |
apply (rule constrains_weaken) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
63 |
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
|
64 |
apply (auto simp add: not_E_eq) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
65 |
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
|
66 |
done |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
67 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
68 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
69 |
(*** Progress under weak fairness ***) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
70 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
71 |
lemma (in Token) wf_nodeOrder: "wf(nodeOrder j)" |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
72 |
apply (unfold nodeOrder_def) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
73 |
apply (rule wf_less_than [THEN wf_inv_image, THEN wf_subset], blast) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
74 |
done |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
75 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
76 |
lemma (in Token) nodeOrder_eq: |
13806 | 77 |
"[| i<N; j<N |] ==> ((next i, i) \<in> nodeOrder j) = (i \<noteq> j)" |
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
78 |
apply (unfold nodeOrder_def next_def inv_image_def) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
79 |
apply (auto simp add: mod_Suc mod_geq) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
80 |
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
|
81 |
done |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
82 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
83 |
(*From "A Logic for Concurrent Programming", but not used in Chapter 4. |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
84 |
Note the use of case_tac. Reasoning about leadsTo takes practice!*) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
85 |
lemma (in Token) TR7_nodeOrder: |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
86 |
"[| i<N; j<N |] ==> |
13806 | 87 |
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
|
88 |
apply (case_tac "i=j") |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
89 |
apply (blast intro: subset_imp_leadsTo) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
90 |
apply (rule TR7 [THEN leadsTo_weaken_R]) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
91 |
apply (auto simp add: HasTok_def nodeOrder_eq) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
92 |
done |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
93 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
94 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
95 |
(*Chapter 4 variant, the one actually used below.*) |
13806 | 96 |
lemma (in Token) TR7_aux: "[| i<N; j<N; i\<noteq>j |] |
97 |
==> 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
|
98 |
apply (rule TR7 [THEN leadsTo_weaken_R]) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
99 |
apply (auto simp add: HasTok_def nodeOrder_eq) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
100 |
done |
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
101 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
102 |
lemma (in Token) token_lemma: |
13806 | 103 |
"({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
|
104 |
by auto |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
105 |
|
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
106 |
|
13785
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
107 |
(*Misra's TR9: the token reaches an arbitrary node*) |
13806 | 108 |
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
|
109 |
apply (rule leadsTo_weaken_R) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
110 |
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
|
111 |
in wf_nodeOrder [THEN bounded_induct]) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
112 |
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
|
113 |
prefer 2 apply blast |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
114 |
apply clarify |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
115 |
apply (rule TR7_aux [THEN leadsTo_weaken]) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
116 |
apply (auto simp add: HasTok_def nodeOrder_def) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
117 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
118 |
|
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
119 |
(*Misra's TR8: a hungry process eventually eats*) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
120 |
lemma (in Token) token_progress: |
13806 | 121 |
"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
|
122 |
apply (rule leadsTo_cancel1 [THEN leadsTo_Un_duplicate]) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
123 |
apply (rule_tac [2] TR6) |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
124 |
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
|
125 |
done |
e2fcd88be55d
Partial conversion of UNITY to Isar new-style theories
paulson
parents:
11195
diff
changeset
|
126 |
|
11195
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
127 |
|
65ede8dfe304
reorganization of HOL/UNITY, moving examples to subdirectories Simple and Comp
paulson
parents:
diff
changeset
|
128 |
end |