| author | wenzelm | 
| Mon, 04 Sep 2023 17:25:16 +0200 | |
| changeset 78646 | fff610f1a6f4 | 
| parent 63549 | b0d31c7def86 | 
| permissions | -rw-r--r-- | 
| 62008 | 1  | 
(* Title: HOL/HOLCF/IOA/Deadlock.thy  | 
| 40945 | 2  | 
Author: Olaf Müller  | 
| 17233 | 3  | 
*)  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents:  
diff
changeset
 | 
4  | 
|
| 62002 | 5  | 
section \<open>Deadlock freedom of I/O Automata\<close>  | 
| 
3433
 
2de17c994071
added deadlock freedom, polished definitions and proofs
 
mueller 
parents:  
diff
changeset
 | 
6  | 
|
| 17233 | 7  | 
theory Deadlock  | 
8  | 
imports RefCorrectness CompoScheds  | 
|
9  | 
begin  | 
|
10  | 
||
| 62156 | 11  | 
text \<open>Input actions may always be added to a schedule.\<close>  | 
| 19741 | 12  | 
|
13  | 
lemma scheds_input_enabled:  | 
|
| 63549 | 14  | 
"Filter (\<lambda>x. x \<in> act A) \<cdot> sch \<in> schedules A \<Longrightarrow> a \<in> inp A \<Longrightarrow> input_enabled A \<Longrightarrow> Finite sch  | 
15  | 
\<Longrightarrow> Filter (\<lambda>x. x \<in> act A) \<cdot> sch @@ a \<leadsto> nil \<in> schedules A"  | 
|
| 62156 | 16  | 
apply (simp add: schedules_def has_schedule_def)  | 
17  | 
apply auto  | 
|
18  | 
apply (frule inp_is_act)  | 
|
19  | 
apply (simp add: executions_def)  | 
|
| 62195 | 20  | 
apply (pair ex)  | 
| 62156 | 21  | 
apply (rename_tac s ex)  | 
22  | 
apply (subgoal_tac "Finite ex")  | 
|
23  | 
prefer 2  | 
|
24  | 
apply (simp add: filter_act_def)  | 
|
25  | 
defer  | 
|
26  | 
apply (rule_tac [2] Map2Finite [THEN iffD1])  | 
|
| 63549 | 27  | 
apply (rule_tac [2] t = "Map fst \<cdot> ex" in subst)  | 
| 62156 | 28  | 
prefer 2  | 
29  | 
apply assumption  | 
|
30  | 
apply (erule_tac [2] FiniteFilter)  | 
|
31  | 
text \<open>subgoal 1\<close>  | 
|
32  | 
apply (frule exists_laststate)  | 
|
33  | 
apply (erule allE)  | 
|
34  | 
apply (erule exE)  | 
|
35  | 
text \<open>using input-enabledness\<close>  | 
|
36  | 
apply (simp add: input_enabled_def)  | 
|
37  | 
apply (erule conjE)+  | 
|
38  | 
apply (erule_tac x = "a" in allE)  | 
|
39  | 
apply simp  | 
|
40  | 
apply (erule_tac x = "u" in allE)  | 
|
41  | 
apply (erule exE)  | 
|
42  | 
text \<open>instantiate execution\<close>  | 
|
43  | 
apply (rule_tac x = " (s, ex @@ (a, s2) \<leadsto> nil) " in exI)  | 
|
44  | 
apply (simp add: filter_act_def MapConc)  | 
|
45  | 
apply (erule_tac t = "u" in lemma_2_1)  | 
|
46  | 
apply simp  | 
|
47  | 
apply (rule sym)  | 
|
48  | 
apply assumption  | 
|
49  | 
done  | 
|
| 19741 | 50  | 
|
| 62002 | 51  | 
text \<open>  | 
| 62156 | 52  | 
Deadlock freedom: component B cannot block an out or int action of component  | 
53  | 
A in every schedule.  | 
|
54  | 
||
55  | 
Needs compositionality on schedule level, input-enabledness, compatibility  | 
|
56  | 
and distributivity of \<open>is_exec_frag\<close> over \<open>@@\<close>.  | 
|
| 62002 | 57  | 
\<close>  | 
| 19741 | 58  | 
|
| 62156 | 59  | 
lemma IOA_deadlock_free:  | 
60  | 
assumes "a \<in> local A"  | 
|
61  | 
and "Finite sch"  | 
|
62  | 
and "sch \<in> schedules (A \<parallel> B)"  | 
|
| 63549 | 63  | 
and "Filter (\<lambda>x. x \<in> act A) \<cdot> (sch @@ a \<leadsto> nil) \<in> schedules A"  | 
| 62156 | 64  | 
and "compatible A B"  | 
65  | 
and "input_enabled B"  | 
|
66  | 
shows "(sch @@ a \<leadsto> nil) \<in> schedules (A \<parallel> B)"  | 
|
| 62390 | 67  | 
supply if_split [split del]  | 
| 62156 | 68  | 
apply (insert assms)  | 
69  | 
apply (simp add: compositionality_sch locals_def)  | 
|
70  | 
apply (rule conjI)  | 
|
71  | 
text \<open>\<open>a \<in> act (A \<parallel> B)\<close>\<close>  | 
|
72  | 
prefer 2  | 
|
73  | 
apply (simp add: actions_of_par)  | 
|
74  | 
apply (blast dest: int_is_act out_is_act)  | 
|
| 19741 | 75  | 
|
| 62156 | 76  | 
text \<open>\<open>Filter B (sch @@ [a]) \<in> schedules B\<close>\<close>  | 
77  | 
apply (case_tac "a \<in> int A")  | 
|
78  | 
apply (drule intA_is_not_actB)  | 
|
79  | 
apply (assumption) (* \<longrightarrow> a \<notin> act B *)  | 
|
80  | 
apply simp  | 
|
| 19741 | 81  | 
|
| 62156 | 82  | 
text \<open>case \<open>a \<notin> int A\<close>, i.e. \<open>a \<in> out A\<close>\<close>  | 
83  | 
apply (case_tac "a \<notin> act B")  | 
|
84  | 
apply simp  | 
|
85  | 
text \<open>case \<open>a \<in> act B\<close>\<close>  | 
|
86  | 
apply simp  | 
|
87  | 
apply (subgoal_tac "a \<in> out A")  | 
|
88  | 
prefer 2  | 
|
89  | 
apply blast  | 
|
90  | 
apply (drule outAactB_is_inpB)  | 
|
91  | 
apply assumption  | 
|
92  | 
apply assumption  | 
|
93  | 
apply (rule scheds_input_enabled)  | 
|
94  | 
apply simp  | 
|
95  | 
apply assumption+  | 
|
96  | 
done  | 
|
| 19741 | 97  | 
|
| 17233 | 98  | 
end  |