author | wenzelm |
Wed, 03 Feb 1999 16:42:40 +0100 | |
changeset 6188 | c40e5ac04e3e |
parent 5983 | 79e301a6a51b |
child 6536 | 281d44905cab |
permissions | -rw-r--r-- |
5111 | 1 |
(* Title: HOL/UNITY/Mutex |
4776 | 2 |
ID: $Id$ |
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1998 University of Cambridge |
|
5 |
||
6 |
Based on "A Family of 2-Process Mutual Exclusion Algorithms" by J Misra |
|
7 |
*) |
|
8 |
||
5232 | 9 |
(*split_all_tac causes a big blow-up*) |
5706 | 10 |
claset_ref() := claset() delSWrapper record_split_name; |
4776 | 11 |
|
5648 | 12 |
Addsimps [Mprg_def RS def_prg_Init]; |
13 |
program_defs_ref := [Mprg_def]; |
|
4776 | 14 |
|
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
15 |
Addsimps (map simp_of_act |
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
16 |
[cmd0U_def, cmd1U_def, cmd2U_def, cmd3U_def, cmd4U_def, |
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
17 |
cmd0V_def, cmd1V_def, cmd2V_def, cmd3V_def, cmd4V_def]); |
4776 | 18 |
|
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
19 |
Addsimps (map simp_of_set |
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
20 |
[invariantU_def, invariantV_def, bad_invariantU_def]); |
4776 | 21 |
|
5232 | 22 |
(*Simplification for records*) |
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
23 |
Addsimps (thms"state.update_defs"); |
5232 | 24 |
|
5648 | 25 |
Goal "Mprg : Invariant invariantU"; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
26 |
by (rtac InvariantI 1); |
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
27 |
by (constrains_tac 2); |
5232 | 28 |
by Auto_tac; |
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
29 |
qed "invariantU"; |
4776 | 30 |
|
5648 | 31 |
Goal "Mprg : Invariant invariantV"; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
32 |
by (rtac InvariantI 1); |
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
33 |
by (constrains_tac 2); |
4776 | 34 |
by Auto_tac; |
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
35 |
qed "invariantV"; |
4776 | 36 |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
37 |
val invariantUV = Invariant_Int_rule [invariantU, invariantV]; |
4776 | 38 |
|
39 |
||
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
40 |
(*The safety property: mutual exclusion*) |
5596 | 41 |
Goal "(reachable Mprg) Int {s. MM s = #3 & NN s = #3} = {}"; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
42 |
by (cut_facts_tac [invariantUV RS Invariant_includes_reachable] 1); |
4776 | 43 |
by Auto_tac; |
44 |
qed "mutual_exclusion"; |
|
45 |
||
46 |
||
5232 | 47 |
(*The bad invariant FAILS in cmd1V*) |
5648 | 48 |
Goal "Mprg : Invariant bad_invariantU"; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
49 |
by (rtac InvariantI 1); |
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
50 |
by (constrains_tac 2); |
5596 | 51 |
by (Force_tac 1); |
52 |
(*Needs a decision procedure to simplify the resulting state*) |
|
53 |
by (auto_tac (claset(), |
|
54 |
simpset_of Int.thy addsimps |
|
55 |
[zadd_int, integ_of_Pls, integ_of_Min, |
|
56 |
integ_of_BIT, le_int_Suc_eq])); |
|
57 |
by (dtac zle_trans 1 THEN assume_tac 1); |
|
58 |
by (full_simp_tac (simpset_of Int.thy) 1); |
|
59 |
by (asm_full_simp_tac (simpset() addsimps int_simps) 1); |
|
4776 | 60 |
(*Resulting state: n=1, p=false, m=4, u=false. |
5232 | 61 |
Execution of cmd1V (the command of process v guarded by n=1) sets p:=true, |
4776 | 62 |
violating the invariant!*) |
63 |
(*Check that subgoals remain: proof failed.*) |
|
64 |
getgoal 1; |
|
65 |
||
66 |
||
5596 | 67 |
Goal "(#1 <= m & m <= #3) = (m = #1 | m = #2 | m = #3)"; |
68 |
by (auto_tac (claset(), |
|
69 |
simpset_of Int.thy addsimps |
|
70 |
[zle_iff_zadd, zadd_int, integ_of_Pls, integ_of_Min, |
|
71 |
integ_of_BIT])); |
|
72 |
qed "eq_123"; |
|
73 |
||
74 |
||
4776 | 75 |
(*** Progress for U ***) |
76 |
||
5648 | 77 |
Goalw [Unless_def] "Mprg : Unless {s. MM s=#2} {s. MM s=#3}"; |
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
78 |
by (constrains_tac 1); |
4776 | 79 |
qed "U_F0"; |
80 |
||
5648 | 81 |
Goal "Mprg : LeadsTo {s. MM s=#1} {s. PP s = VV s & MM s = #2}"; |
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
82 |
by (ensures_tac "cmd1U" 1); |
4776 | 83 |
qed "U_F1"; |
84 |
||
5648 | 85 |
Goal "Mprg : LeadsTo {s. ~ PP s & MM s = #2} {s. MM s = #3}"; |
5340 | 86 |
by (cut_facts_tac [invariantU] 1); |
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
87 |
by (ensures_tac "cmd2U" 1); |
4776 | 88 |
qed "U_F2"; |
89 |
||
5648 | 90 |
Goal "Mprg : LeadsTo {s. MM s = #3} {s. PP s}"; |
5596 | 91 |
by (res_inst_tac [("B", "{s. MM s = #4}")] LeadsTo_Trans 1); |
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
92 |
by (ensures_tac "cmd4U" 2); |
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
93 |
by (ensures_tac "cmd3U" 1); |
4776 | 94 |
qed "U_F3"; |
95 |
||
5648 | 96 |
Goal "Mprg : LeadsTo {s. MM s = #2} {s. PP s}"; |
5340 | 97 |
by (rtac ([LeadsTo_weaken_L, Int_lower2 RS subset_imp_LeadsTo] |
4776 | 98 |
MRS LeadsTo_Diff) 1); |
99 |
by (rtac ([U_F2, U_F3] MRS LeadsTo_Trans) 1); |
|
100 |
by (auto_tac (claset() addSEs [less_SucE], simpset())); |
|
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
101 |
val U_lemma2 = result(); |
4776 | 102 |
|
5648 | 103 |
Goal "Mprg : LeadsTo {s. MM s = #1} {s. PP s}"; |
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
104 |
by (rtac ([U_F1 RS LeadsTo_weaken_R, U_lemma2] MRS LeadsTo_Trans) 1); |
4776 | 105 |
by (Blast_tac 1); |
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
106 |
val U_lemma1 = result(); |
4776 | 107 |
|
5648 | 108 |
Goal "Mprg : LeadsTo {s. #1 <= MM s & MM s <= #3} {s. PP s}"; |
5596 | 109 |
by (simp_tac (simpset() addsimps [eq_123, Collect_disj_eq, LeadsTo_Un_distrib, |
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
110 |
U_lemma1, U_lemma2, U_F3] ) 1); |
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
111 |
val U_lemma123 = result(); |
4776 | 112 |
|
113 |
(*Misra's F4*) |
|
5648 | 114 |
Goal "Mprg : LeadsTo {s. UU s} {s. PP s}"; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
115 |
by (rtac ([invariantU, U_lemma123] MRS Invariant_LeadsTo_weaken) 1); |
4776 | 116 |
by Auto_tac; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
117 |
qed "u_Leadsto_p"; |
4776 | 118 |
|
119 |
||
120 |
(*** Progress for V ***) |
|
121 |
||
122 |
||
5648 | 123 |
Goalw [Unless_def] "Mprg : Unless {s. NN s=#2} {s. NN s=#3}"; |
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
124 |
by (constrains_tac 1); |
4776 | 125 |
qed "V_F0"; |
126 |
||
5648 | 127 |
Goal "Mprg : LeadsTo {s. NN s=#1} {s. PP s = (~ UU s) & NN s = #2}"; |
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
128 |
by (ensures_tac "cmd1V" 1); |
4776 | 129 |
qed "V_F1"; |
130 |
||
5648 | 131 |
Goal "Mprg : LeadsTo {s. PP s & NN s = #2} {s. NN s = #3}"; |
5340 | 132 |
by (cut_facts_tac [invariantV] 1); |
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
133 |
by (ensures_tac "cmd2V" 1); |
4776 | 134 |
qed "V_F2"; |
135 |
||
5648 | 136 |
Goal "Mprg : LeadsTo {s. NN s = #3} {s. ~ PP s}"; |
5596 | 137 |
by (res_inst_tac [("B", "{s. NN s = #4}")] LeadsTo_Trans 1); |
5426
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
138 |
by (ensures_tac "cmd4V" 2); |
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents:
5424
diff
changeset
|
139 |
by (ensures_tac "cmd3V" 1); |
4776 | 140 |
qed "V_F3"; |
141 |
||
5648 | 142 |
Goal "Mprg : LeadsTo {s. NN s = #2} {s. ~ PP s}"; |
5340 | 143 |
by (rtac ([LeadsTo_weaken_L, Int_lower2 RS subset_imp_LeadsTo] |
4776 | 144 |
MRS LeadsTo_Diff) 1); |
145 |
by (rtac ([V_F2, V_F3] MRS LeadsTo_Trans) 1); |
|
146 |
by (auto_tac (claset() addSEs [less_SucE], simpset())); |
|
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
147 |
val V_lemma2 = result(); |
4776 | 148 |
|
5648 | 149 |
Goal "Mprg : LeadsTo {s. NN s = #1} {s. ~ PP s}"; |
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
150 |
by (rtac ([V_F1 RS LeadsTo_weaken_R, V_lemma2] MRS LeadsTo_Trans) 1); |
4776 | 151 |
by (Blast_tac 1); |
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
152 |
val V_lemma1 = result(); |
4776 | 153 |
|
5648 | 154 |
Goal "Mprg : LeadsTo {s. #1 <= NN s & NN s <= #3} {s. ~ PP s}"; |
5596 | 155 |
by (simp_tac (simpset() addsimps [eq_123, Collect_disj_eq, LeadsTo_Un_distrib, |
5240
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
156 |
V_lemma1, V_lemma2, V_F3] ) 1); |
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
paulson
parents:
5232
diff
changeset
|
157 |
val V_lemma123 = result(); |
4776 | 158 |
|
159 |
||
160 |
(*Misra's F4*) |
|
5648 | 161 |
Goal "Mprg : LeadsTo {s. VV s} {s. ~ PP s}"; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
162 |
by (rtac ([invariantV, V_lemma123] MRS Invariant_LeadsTo_weaken) 1); |
4776 | 163 |
by Auto_tac; |
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
164 |
qed "v_Leadsto_not_p"; |
4776 | 165 |
|
166 |
||
167 |
(** Absence of starvation **) |
|
168 |
||
169 |
(*Misra's F6*) |
|
5648 | 170 |
Goal "Mprg : LeadsTo {s. MM s = #1} {s. MM s = #3}"; |
4776 | 171 |
by (rtac LeadsTo_Un_duplicate 1); |
172 |
by (rtac LeadsTo_cancel2 1); |
|
173 |
by (rtac U_F2 2); |
|
174 |
by (simp_tac (simpset() addsimps [Collect_conj_eq] ) 1); |
|
175 |
by (stac Un_commute 1); |
|
176 |
by (rtac LeadsTo_Un_duplicate 1); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
177 |
by (rtac ([v_Leadsto_not_p, U_F0] MRS PSP_Unless RSN(2, LeadsTo_cancel2)) 1); |
4776 | 178 |
by (rtac (U_F1 RS LeadsTo_weaken_R) 1); |
179 |
by (auto_tac (claset() addSEs [less_SucE], simpset())); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
180 |
qed "m1_Leadsto_3"; |
4776 | 181 |
|
182 |
||
183 |
(*The same for V*) |
|
5648 | 184 |
Goal "Mprg : LeadsTo {s. NN s = #1} {s. NN s = #3}"; |
4776 | 185 |
by (rtac LeadsTo_Un_duplicate 1); |
186 |
by (rtac LeadsTo_cancel2 1); |
|
187 |
by (rtac V_F2 2); |
|
188 |
by (simp_tac (simpset() addsimps [Collect_conj_eq] ) 1); |
|
189 |
by (stac Un_commute 1); |
|
190 |
by (rtac LeadsTo_Un_duplicate 1); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
191 |
by (rtac ([u_Leadsto_p, V_F0] MRS PSP_Unless RSN(2, LeadsTo_cancel2)) 1); |
4776 | 192 |
by (rtac (V_F1 RS LeadsTo_weaken_R) 1); |
193 |
by (auto_tac (claset() addSEs [less_SucE], simpset())); |
|
5313
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
5277
diff
changeset
|
194 |
qed "n1_Leadsto_3"; |