| author | paulson | 
| Thu, 17 Jun 1999 10:35:01 +0200 | |
| changeset 6833 | 15d6c121d75f | 
| parent 6678 | d83f27b03529 | 
| child 6909 | 21601bc4f3c2 | 
| 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 | ||
| 6678 | 9 | (*record_split_name causes a BIG blow-up (doubles the run-time)*) | 
| 5706 | 10 | claset_ref() := claset() delSWrapper record_split_name; | 
| 4776 | 11 | |
| 6565 | 12 | Addsimps [Mutex_def RS def_prg_Init]; | 
| 13 | program_defs_ref := [Mutex_def]; | |
| 4776 | 14 | |
| 5426 
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
 paulson parents: 
5424diff
changeset | 15 | Addsimps (map simp_of_act | 
| 6565 | 16 | [U0_def, U1_def, U2_def, U3_def, U4_def, | 
| 17 | V0_def, V1_def, V2_def, V3_def, V4_def]); | |
| 4776 | 18 | |
| 6570 | 19 | Addsimps (map simp_of_set [IU_def, IV_def, bad_IU_def]); | 
| 4776 | 20 | |
| 5232 | 21 | (*Simplification for records*) | 
| 5426 
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
 paulson parents: 
5424diff
changeset | 22 | Addsimps (thms"state.update_defs"); | 
| 5232 | 23 | |
| 6570 | 24 | Goal "Mutex : Always IU"; | 
| 25 | by (rtac AlwaysI 1); | |
| 5426 
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
 paulson parents: 
5424diff
changeset | 26 | by (constrains_tac 2); | 
| 5232 | 27 | by Auto_tac; | 
| 6570 | 28 | qed "IU"; | 
| 4776 | 29 | |
| 6570 | 30 | Goal "Mutex : Always IV"; | 
| 31 | by (rtac AlwaysI 1); | |
| 5426 
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
 paulson parents: 
5424diff
changeset | 32 | by (constrains_tac 2); | 
| 4776 | 33 | by Auto_tac; | 
| 6570 | 34 | qed "IV"; | 
| 4776 | 35 | |
| 5240 
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
 paulson parents: 
5232diff
changeset | 36 | (*The safety property: mutual exclusion*) | 
| 6570 | 37 | Goal "Mutex : Always {s. ~ (m s = #3 & n s = #3)}";
 | 
| 38 | by (rtac Always_weaken 1); | |
| 39 | by (rtac ([IU, IV] MRS Always_Int) 1); | |
| 4776 | 40 | by Auto_tac; | 
| 41 | qed "mutual_exclusion"; | |
| 42 | ||
| 43 | ||
| 6565 | 44 | (*The bad invariant FAILS in V1*) | 
| 6570 | 45 | Goal "Mutex : Always bad_IU"; | 
| 46 | by (rtac AlwaysI 1); | |
| 5426 
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
 paulson parents: 
5424diff
changeset | 47 | by (constrains_tac 2); | 
| 6615 | 48 | by Auto_tac; | 
| 4776 | 49 | (*Resulting state: n=1, p=false, m=4, u=false. | 
| 6565 | 50 | Execution of V1 (the command of process v guarded by n=1) sets p:=true, | 
| 4776 | 51 | violating the invariant!*) | 
| 52 | (*Check that subgoals remain: proof failed.*) | |
| 53 | getgoal 1; | |
| 54 | ||
| 55 | ||
| 6565 | 56 | Goal "(#1 <= i & i <= #3) = (i = #1 | i = #2 | i = #3)"; | 
| 6615 | 57 | by (arith_tac 1); | 
| 5596 | 58 | qed "eq_123"; | 
| 59 | ||
| 60 | ||
| 4776 | 61 | (*** Progress for U ***) | 
| 62 | ||
| 6565 | 63 | Goalw [Unless_def] "Mutex : {s. m s=#2} Unless {s. m s=#3}";
 | 
| 5426 
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
 paulson parents: 
5424diff
changeset | 64 | by (constrains_tac 1); | 
| 4776 | 65 | qed "U_F0"; | 
| 66 | ||
| 6565 | 67 | Goal "Mutex : {s. m s=#1} LeadsTo {s. p s = v s & m s = #2}";
 | 
| 68 | by (ensures_tac "U1" 1); | |
| 4776 | 69 | qed "U_F1"; | 
| 70 | ||
| 6565 | 71 | Goal "Mutex : {s. ~ p s & m s = #2} LeadsTo {s. m s = #3}";
 | 
| 6570 | 72 | by (cut_facts_tac [IU] 1); | 
| 6565 | 73 | by (ensures_tac "U2" 1); | 
| 4776 | 74 | qed "U_F2"; | 
| 75 | ||
| 6565 | 76 | Goal "Mutex : {s. m s = #3} LeadsTo {s. p s}";
 | 
| 77 | by (res_inst_tac [("B", "{s. m s = #4}")] LeadsTo_Trans 1);
 | |
| 78 | by (ensures_tac "U4" 2); | |
| 79 | by (ensures_tac "U3" 1); | |
| 4776 | 80 | qed "U_F3"; | 
| 81 | ||
| 6565 | 82 | Goal "Mutex : {s. m s = #2} LeadsTo {s. p s}";
 | 
| 5340 | 83 | by (rtac ([LeadsTo_weaken_L, Int_lower2 RS subset_imp_LeadsTo] | 
| 4776 | 84 | MRS LeadsTo_Diff) 1); | 
| 85 | by (rtac ([U_F2, U_F3] MRS LeadsTo_Trans) 1); | |
| 86 | by (auto_tac (claset() addSEs [less_SucE], simpset())); | |
| 5240 
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
 paulson parents: 
5232diff
changeset | 87 | val U_lemma2 = result(); | 
| 4776 | 88 | |
| 6565 | 89 | Goal "Mutex : {s. m s = #1} LeadsTo {s. p s}";
 | 
| 5240 
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
 paulson parents: 
5232diff
changeset | 90 | by (rtac ([U_F1 RS LeadsTo_weaken_R, U_lemma2] MRS LeadsTo_Trans) 1); | 
| 4776 | 91 | by (Blast_tac 1); | 
| 5240 
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
 paulson parents: 
5232diff
changeset | 92 | val U_lemma1 = result(); | 
| 4776 | 93 | |
| 6565 | 94 | Goal "Mutex : {s. #1 <= m s & m s <= #3} LeadsTo {s. p s}";
 | 
| 5596 | 95 | 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: 
5232diff
changeset | 96 | U_lemma1, U_lemma2, U_F3] ) 1); | 
| 
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
 paulson parents: 
5232diff
changeset | 97 | val U_lemma123 = result(); | 
| 4776 | 98 | |
| 99 | (*Misra's F4*) | |
| 6565 | 100 | Goal "Mutex : {s. u s} LeadsTo {s. p s}";
 | 
| 6570 | 101 | by (rtac ([IU, U_lemma123] MRS Always_LeadsTo_weaken) 1); | 
| 4776 | 102 | by Auto_tac; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5277diff
changeset | 103 | qed "u_Leadsto_p"; | 
| 4776 | 104 | |
| 105 | ||
| 106 | (*** Progress for V ***) | |
| 107 | ||
| 108 | ||
| 6565 | 109 | Goalw [Unless_def] "Mutex : {s. n s=#2} Unless {s. n s=#3}";
 | 
| 5426 
566f47250bd0
A new approach, using simp_of_act and simp_of_set to activate definitions when
 paulson parents: 
5424diff
changeset | 110 | by (constrains_tac 1); | 
| 4776 | 111 | qed "V_F0"; | 
| 112 | ||
| 6565 | 113 | Goal "Mutex : {s. n s=#1} LeadsTo {s. p s = (~ u s) & n s = #2}";
 | 
| 114 | by (ensures_tac "V1" 1); | |
| 4776 | 115 | qed "V_F1"; | 
| 116 | ||
| 6565 | 117 | Goal "Mutex : {s. p s & n s = #2} LeadsTo {s. n s = #3}";
 | 
| 6570 | 118 | by (cut_facts_tac [IV] 1); | 
| 6565 | 119 | by (ensures_tac "V2" 1); | 
| 4776 | 120 | qed "V_F2"; | 
| 121 | ||
| 6565 | 122 | Goal "Mutex : {s. n s = #3} LeadsTo {s. ~ p s}";
 | 
| 123 | by (res_inst_tac [("B", "{s. n s = #4}")] LeadsTo_Trans 1);
 | |
| 124 | by (ensures_tac "V4" 2); | |
| 125 | by (ensures_tac "V3" 1); | |
| 4776 | 126 | qed "V_F3"; | 
| 127 | ||
| 6565 | 128 | Goal "Mutex : {s. n s = #2} LeadsTo {s. ~ p s}";
 | 
| 5340 | 129 | by (rtac ([LeadsTo_weaken_L, Int_lower2 RS subset_imp_LeadsTo] | 
| 4776 | 130 | MRS LeadsTo_Diff) 1); | 
| 131 | by (rtac ([V_F2, V_F3] MRS LeadsTo_Trans) 1); | |
| 132 | by (auto_tac (claset() addSEs [less_SucE], simpset())); | |
| 5240 
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
 paulson parents: 
5232diff
changeset | 133 | val V_lemma2 = result(); | 
| 4776 | 134 | |
| 6565 | 135 | Goal "Mutex : {s. n s = #1} LeadsTo {s. ~ p s}";
 | 
| 5240 
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
 paulson parents: 
5232diff
changeset | 136 | by (rtac ([V_F1 RS LeadsTo_weaken_R, V_lemma2] MRS LeadsTo_Trans) 1); | 
| 4776 | 137 | by (Blast_tac 1); | 
| 5240 
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
 paulson parents: 
5232diff
changeset | 138 | val V_lemma1 = result(); | 
| 4776 | 139 | |
| 6565 | 140 | Goal "Mutex : {s. #1 <= n s & n s <= #3} LeadsTo {s. ~ p s}";
 | 
| 5596 | 141 | 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: 
5232diff
changeset | 142 | V_lemma1, V_lemma2, V_F3] ) 1); | 
| 
bbcd79ef7cf2
Constant "invariant" and new constrains_tac, ensures_tac
 paulson parents: 
5232diff
changeset | 143 | val V_lemma123 = result(); | 
| 4776 | 144 | |
| 145 | ||
| 146 | (*Misra's F4*) | |
| 6565 | 147 | Goal "Mutex : {s. v s} LeadsTo {s. ~ p s}";
 | 
| 6570 | 148 | by (rtac ([IV, V_lemma123] MRS Always_LeadsTo_weaken) 1); | 
| 4776 | 149 | by Auto_tac; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5277diff
changeset | 150 | qed "v_Leadsto_not_p"; | 
| 4776 | 151 | |
| 152 | ||
| 153 | (** Absence of starvation **) | |
| 154 | ||
| 155 | (*Misra's F6*) | |
| 6565 | 156 | Goal "Mutex : {s. m s = #1} LeadsTo {s. m s = #3}";
 | 
| 6615 | 157 | by (rtac (LeadsTo_cancel2 RS LeadsTo_Un_duplicate) 1); | 
| 4776 | 158 | by (rtac U_F2 2); | 
| 159 | by (simp_tac (simpset() addsimps [Collect_conj_eq] ) 1); | |
| 160 | by (stac Un_commute 1); | |
| 6615 | 161 | by (rtac (LeadsTo_cancel2 RS LeadsTo_Un_duplicate) 1); | 
| 162 | by (rtac ([v_Leadsto_not_p, U_F0] MRS PSP_Unless) 2); | |
| 4776 | 163 | by (rtac (U_F1 RS LeadsTo_weaken_R) 1); | 
| 6615 | 164 | by Auto_tac; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5277diff
changeset | 165 | qed "m1_Leadsto_3"; | 
| 4776 | 166 | |
| 167 | (*The same for V*) | |
| 6565 | 168 | Goal "Mutex : {s. n s = #1} LeadsTo {s. n s = #3}";
 | 
| 6615 | 169 | by (rtac (LeadsTo_cancel2 RS LeadsTo_Un_duplicate) 1); | 
| 4776 | 170 | by (rtac V_F2 2); | 
| 171 | by (simp_tac (simpset() addsimps [Collect_conj_eq] ) 1); | |
| 172 | by (stac Un_commute 1); | |
| 6615 | 173 | by (rtac (LeadsTo_cancel2 RS LeadsTo_Un_duplicate) 1); | 
| 174 | by (rtac ([u_Leadsto_p, V_F0] MRS PSP_Unless) 2); | |
| 4776 | 175 | by (rtac (V_F1 RS LeadsTo_weaken_R) 1); | 
| 6615 | 176 | by Auto_tac; | 
| 5313 
1861a564d7e2
Constrains, Stable, Invariant...more of the substitution axiom, but Union
 paulson parents: 
5277diff
changeset | 177 | qed "n1_Leadsto_3"; |