| author | wenzelm | 
| Fri, 02 May 2008 16:39:44 +0200 | |
| changeset 26770 | d688166808c0 | 
| parent 26359 | 6d437bde2f1d | 
| child 27208 | 5fe899199f85 | 
| permissions | -rw-r--r-- | 
| 3433 
2de17c994071
added deadlock freedom, polished definitions and proofs
 mueller parents: diff
changeset | 1 | (* Title: HOLCF/IOA/meta_theory/Deadlock.thy | 
| 
2de17c994071
added deadlock freedom, polished definitions and proofs
 mueller parents: diff
changeset | 2 | ID: $Id$ | 
| 12218 | 3 | Author: Olaf Müller | 
| 17233 | 4 | *) | 
| 3433 
2de17c994071
added deadlock freedom, polished definitions and proofs
 mueller parents: diff
changeset | 5 | |
| 17233 | 6 | header {* Deadlock freedom of I/O Automata *}
 | 
| 3433 
2de17c994071
added deadlock freedom, polished definitions and proofs
 mueller parents: diff
changeset | 7 | |
| 17233 | 8 | theory Deadlock | 
| 9 | imports RefCorrectness CompoScheds | |
| 10 | begin | |
| 11 | ||
| 19741 | 12 | text {* input actions may always be added to a schedule *}
 | 
| 13 | ||
| 14 | lemma scheds_input_enabled: | |
| 15 | "[| Filter (%x. x:act A)$sch : schedules A; a:inp A; input_enabled A; Finite sch|] | |
| 16 | ==> Filter (%x. x:act A)$sch @@ a>>nil : schedules A" | |
| 17 | apply (simp add: schedules_def has_schedule_def) | |
| 26359 | 18 | apply auto | 
| 19741 | 19 | apply (frule inp_is_act) | 
| 20 | apply (simp add: executions_def) | |
| 21 | apply (tactic {* pair_tac "ex" 1 *})
 | |
| 22 | apply (rename_tac s ex) | |
| 23 | apply (subgoal_tac "Finite ex") | |
| 24 | prefer 2 | |
| 25 | apply (simp add: filter_act_def) | |
| 26 | defer | |
| 27 | apply (rule_tac [2] Map2Finite [THEN iffD1]) | |
| 28 | apply (rule_tac [2] t = "Map fst$ex" in subst) | |
| 29 | prefer 2 apply (assumption) | |
| 30 | apply (erule_tac [2] FiniteFilter) | |
| 31 | (* subgoal 1 *) | |
| 32 | apply (frule exists_laststate) | |
| 33 | apply (erule allE) | |
| 34 | apply (erule exE) | |
| 35 | (* using input-enabledness *) | |
| 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 | (* instantiate execution *) | |
| 43 | apply (rule_tac x = " (s,ex @@ (a,s2) >>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 | |
| 50 | ||
| 51 | text {*
 | |
| 52 | Deadlock freedom: component B cannot block an out or int action | |
| 53 | of component A in every schedule. | |
| 54 | Needs compositionality on schedule level, input-enabledness, compatibility | |
| 55 | and distributivity of is_exec_frag over @@ | |
| 56 | *} | |
| 57 | ||
| 58 | declare split_if [split del] | |
| 59 | lemma IOA_deadlock_free: "[| a : local A; Finite sch; sch : schedules (A||B); | |
| 60 | Filter (%x. x:act A)$(sch @@ a>>nil) : schedules A; compatible A B; input_enabled B |] | |
| 61 | ==> (sch @@ a>>nil) : schedules (A||B)" | |
| 62 | apply (simp add: compositionality_sch locals_def) | |
| 63 | apply (rule conjI) | |
| 64 | (* a : act (A||B) *) | |
| 65 | prefer 2 | |
| 66 | apply (simp add: actions_of_par) | |
| 67 | apply (blast dest: int_is_act out_is_act) | |
| 68 | ||
| 69 | (* Filter B (sch@@[a]) : schedules B *) | |
| 70 | ||
| 71 | apply (case_tac "a:int A") | |
| 72 | apply (drule intA_is_not_actB) | |
| 73 | apply (assumption) (* --> a~:act B *) | |
| 74 | apply simp | |
| 75 | ||
| 76 | (* case a~:int A , i.e. a:out A *) | |
| 77 | apply (case_tac "a~:act B") | |
| 78 | apply simp | |
| 79 | (* case a:act B *) | |
| 80 | apply simp | |
| 81 | apply (subgoal_tac "a:out A") | |
| 82 | prefer 2 apply (blast) | |
| 83 | apply (drule outAactB_is_inpB) | |
| 84 | apply assumption | |
| 85 | apply assumption | |
| 86 | apply (rule scheds_input_enabled) | |
| 87 | apply simp | |
| 88 | apply assumption+ | |
| 89 | done | |
| 90 | ||
| 91 | declare split_if [split] | |
| 92 | ||
| 17233 | 93 | end |