21 by (exhaust_tac "proc s i" 1); |
21 by (exhaust_tac "proc s i" 1); |
22 by Auto_tac; |
22 by Auto_tac; |
23 qed "not_E_eq"; |
23 qed "not_E_eq"; |
24 |
24 |
25 (*This proof is in the "massaging" style and is much faster! *) |
25 (*This proof is in the "massaging" style and is much faster! *) |
26 Goalw [stable_def] "stable Acts (Compl(E i) Un (HasTok i))"; |
26 Goalw [stable_def] "stable acts (Compl(E i) Un (HasTok i))"; |
27 by (rtac constrains_weaken 1); |
27 by (rtac constrains_weaken 1); |
28 by (rtac ([[TR2, TR4] MRS constrains_Un, TR5] MRS constrains_Un) 1); |
28 by (rtac ([[TR2, TR4] MRS constrains_Un, TR5] MRS constrains_Un) 1); |
29 by (auto_tac (claset(), simpset() addsimps [not_E_eq])); |
29 by (auto_tac (claset(), simpset() addsimps [not_E_eq])); |
30 by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [H_def, E_def, T_def]))); |
30 by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [H_def, E_def, T_def]))); |
31 qed "token_stable"; |
31 qed "token_stable"; |
75 |
75 |
76 |
76 |
77 (*From "A Logic for Concurrent Programming", but not used in Chapter 4. |
77 (*From "A Logic for Concurrent Programming", but not used in Chapter 4. |
78 Note the use of case_tac. Reasoning about leadsTo takes practice!*) |
78 Note the use of case_tac. Reasoning about leadsTo takes practice!*) |
79 Goal "[| i<N; j<N |] ==> \ |
79 Goal "[| i<N; j<N |] ==> \ |
80 \ leadsTo Acts (HasTok i) ({s. (token s, i) : nodeOrder j} Un HasTok j)"; |
80 \ leadsTo acts (HasTok i) ({s. (token s, i) : nodeOrder j} Un HasTok j)"; |
81 by (case_tac "i=j" 1); |
81 by (case_tac "i=j" 1); |
82 by (blast_tac (claset() addIs [subset_imp_leadsTo]) 1); |
82 by (blast_tac (claset() addIs [subset_imp_leadsTo]) 1); |
83 by (rtac (TR7 RS leadsTo_weaken_R) 1); |
83 by (rtac (TR7 RS leadsTo_weaken_R) 1); |
84 by (Clarify_tac 1); |
84 by (Clarify_tac 1); |
85 by (asm_full_simp_tac (simpset() addsimps [HasTok_def, nodeOrder_eq]) 1); |
85 by (asm_full_simp_tac (simpset() addsimps [HasTok_def, nodeOrder_eq]) 1); |
86 qed "TR7_nodeOrder"; |
86 qed "TR7_nodeOrder"; |
87 |
87 |
88 |
88 |
89 (*Chapter 4 variant, the one actually used below.*) |
89 (*Chapter 4 variant, the one actually used below.*) |
90 Goal "[| i<N; j<N; i~=j |] \ |
90 Goal "[| i<N; j<N; i~=j |] \ |
91 \ ==> leadsTo Acts (HasTok i) {s. (token s, i) : nodeOrder j}"; |
91 \ ==> leadsTo acts (HasTok i) {s. (token s, i) : nodeOrder j}"; |
92 by (rtac (TR7 RS leadsTo_weaken_R) 1); |
92 by (rtac (TR7 RS leadsTo_weaken_R) 1); |
93 by (Clarify_tac 1); |
93 by (Clarify_tac 1); |
94 by (asm_full_simp_tac (simpset() addsimps [HasTok_def, nodeOrder_eq]) 1); |
94 by (asm_full_simp_tac (simpset() addsimps [HasTok_def, nodeOrder_eq]) 1); |
95 qed "TR7_aux"; |
95 qed "TR7_aux"; |
96 |
96 |
99 by Auto_tac; |
99 by Auto_tac; |
100 val token_lemma = result(); |
100 val token_lemma = result(); |
101 |
101 |
102 |
102 |
103 (*Misra's TR9: the token reaches an arbitrary node*) |
103 (*Misra's TR9: the token reaches an arbitrary node*) |
104 Goal "j<N ==> leadsTo Acts {s. token s < N} (HasTok j)"; |
104 Goal "j<N ==> leadsTo acts {s. token s < N} (HasTok j)"; |
105 by (rtac leadsTo_weaken_R 1); |
105 by (rtac leadsTo_weaken_R 1); |
106 by (res_inst_tac [("I", "Compl{j}"), ("f", "token"), ("B", "{}")] |
106 by (res_inst_tac [("I", "Compl{j}"), ("f", "token"), ("B", "{}")] |
107 (wf_nodeOrder RS bounded_induct) 1); |
107 (wf_nodeOrder RS bounded_induct) 1); |
108 by (ALLGOALS (asm_simp_tac (simpset() addsimps [token_lemma, vimage_Diff, |
108 by (ALLGOALS (asm_simp_tac (simpset() addsimps [token_lemma, vimage_Diff, |
109 HasTok_def]))); |
109 HasTok_def]))); |
113 by (ALLGOALS (asm_simp_tac (simpset() addsimps [nodeOrder_def, HasTok_def]))); |
113 by (ALLGOALS (asm_simp_tac (simpset() addsimps [nodeOrder_def, HasTok_def]))); |
114 by (ALLGOALS Blast_tac); |
114 by (ALLGOALS Blast_tac); |
115 qed "leadsTo_j"; |
115 qed "leadsTo_j"; |
116 |
116 |
117 (*Misra's TR8: a hungry process eventually eats*) |
117 (*Misra's TR8: a hungry process eventually eats*) |
118 Goal "j<N ==> leadsTo Acts ({s. token s < N} Int H j) (E j)"; |
118 Goal "j<N ==> leadsTo acts ({s. token s < N} Int H j) (E j)"; |
119 by (rtac (leadsTo_cancel1 RS leadsTo_Un_duplicate) 1); |
119 by (rtac (leadsTo_cancel1 RS leadsTo_Un_duplicate) 1); |
120 by (rtac TR6 2); |
120 by (rtac TR6 2); |
121 by (rtac leadsTo_weaken_R 1); |
121 by (rtac leadsTo_weaken_R 1); |
122 by (rtac ([leadsTo_j, TR3] MRS PSP) 1); |
122 by (rtac ([leadsTo_j, TR3] MRS PSP) 1); |
123 by (ALLGOALS Blast_tac); |
123 by (ALLGOALS Blast_tac); |