| author | haftmann | 
| Mon, 12 Jun 2006 09:14:41 +0200 | |
| changeset 19852 | b06db8e4476b | 
| parent 19741 | f65265d71426 | 
| child 25135 | 4f8176c940cf | 
| permissions | -rw-r--r-- | 
| 3071 | 1 | (* Title: HOLCF/IOA/meta_theory/RefMappings.thy | 
| 3275 | 2 | ID: $Id$ | 
| 12218 | 3 | Author: Olaf Müller | 
| 3071 | 4 | *) | 
| 5 | ||
| 17233 | 6 | header {* Refinement Mappings in HOLCF/IOA *}
 | 
| 3071 | 7 | |
| 17233 | 8 | theory RefMappings | 
| 9 | imports Traces | |
| 10 | begin | |
| 11 | ||
| 12 | defaultsort type | |
| 3071 | 13 | |
| 14 | consts | |
| 3433 
2de17c994071
added deadlock freedom, polished definitions and proofs
 mueller parents: 
3275diff
changeset | 15 |   move         ::"[('a,'s)ioa,('a,'s)pairs,'s,'a,'s] => bool"
 | 
| 3071 | 16 |   is_ref_map   ::"[('s1=>'s2),('a,'s1)ioa,('a,'s2)ioa] => bool"
 | 
| 17 |   is_weak_ref_map ::"[('s1=>'s2),('a,'s1)ioa,('a,'s2)ioa] => bool"
 | |
| 17233 | 18 | |
| 4559 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3433diff
changeset | 19 | |
| 3071 | 20 | defs | 
| 21 | ||
| 17233 | 22 | move_def: | 
| 23 | "move ioa ex s a t == | |
| 24 | (is_exec_frag ioa (s,ex) & Finite ex & | |
| 25 | laststate (s,ex)=t & | |
| 10835 | 26 | mk_trace ioa$ex = (if a:ext(ioa) then a>>nil else nil))" | 
| 3071 | 27 | |
| 17233 | 28 | is_ref_map_def: | 
| 29 | "is_ref_map f C A == | |
| 30 | (!s:starts_of(C). f(s):starts_of(A)) & | |
| 31 | (!s t a. reachable C s & | |
| 32 | s -a--C-> t | |
| 3071 | 33 | --> (? ex. move A ex (f s) a (f t)))" | 
| 17233 | 34 | |
| 35 | is_weak_ref_map_def: | |
| 36 | "is_weak_ref_map f C A == | |
| 37 | (!s:starts_of(C). f(s):starts_of(A)) & | |
| 38 | (!s t a. reachable C s & | |
| 39 | s -a--C-> t | |
| 40 | --> (if a:ext(C) | |
| 3071 | 41 | then (f s) -a--A-> (f t) | 
| 17233 | 42 | else (f s)=(f t)))" | 
| 3071 | 43 | |
| 19741 | 44 | |
| 45 | subsection "transitions and moves" | |
| 46 | ||
| 47 | ||
| 48 | lemma transition_is_ex: "s -a--A-> t ==> ? ex. move A ex s a t" | |
| 49 | apply (rule_tac x = " (a,t) >>nil" in exI) | |
| 50 | apply (simp add: move_def) | |
| 51 | done | |
| 52 | ||
| 53 | ||
| 54 | lemma nothing_is_ex: "(~a:ext A) & s=t ==> ? ex. move A ex s a t" | |
| 55 | apply (rule_tac x = "nil" in exI) | |
| 56 | apply (simp add: move_def) | |
| 57 | done | |
| 58 | ||
| 59 | ||
| 60 | lemma ei_transitions_are_ex: "(s -a--A-> s') & (s' -a'--A-> s'') & (~a':ext A) | |
| 61 | ==> ? ex. move A ex s a s''" | |
| 62 | apply (rule_tac x = " (a,s') >> (a',s'') >>nil" in exI) | |
| 63 | apply (simp add: move_def) | |
| 64 | done | |
| 65 | ||
| 66 | ||
| 67 | lemma eii_transitions_are_ex: "(s1 -a1--A-> s2) & (s2 -a2--A-> s3) & (s3 -a3--A-> s4) & | |
| 68 | (~a2:ext A) & (~a3:ext A) ==> | |
| 69 | ? ex. move A ex s1 a1 s4" | |
| 70 | apply (rule_tac x = " (a1,s2) >> (a2,s3) >> (a3,s4) >>nil" in exI) | |
| 71 | apply (simp add: move_def) | |
| 72 | done | |
| 73 | ||
| 74 | ||
| 75 | subsection "weak_ref_map and ref_map" | |
| 76 | ||
| 77 | lemma imp_conj_lemma: | |
| 78 | "[| ext C = ext A; | |
| 79 | is_weak_ref_map f C A |] ==> is_ref_map f C A" | |
| 80 | apply (unfold is_weak_ref_map_def is_ref_map_def) | |
| 81 | apply (tactic "safe_tac set_cs") | |
| 82 | apply (case_tac "a:ext A") | |
| 83 | apply (rule transition_is_ex) | |
| 84 | apply (simp (no_asm_simp)) | |
| 85 | apply (rule nothing_is_ex) | |
| 86 | apply simp | |
| 87 | done | |
| 88 | ||
| 89 | ||
| 90 | lemma imp_conj_lemma: "(P ==> Q-->R) ==> P&Q --> R" | |
| 91 | by blast | |
| 92 | ||
| 93 | declare split_if [split del] | |
| 94 | declare if_weak_cong [cong del] | |
| 95 | ||
| 96 | lemma rename_through_pmap: "[| is_weak_ref_map f C A |] | |
| 97 | ==> (is_weak_ref_map f (rename C g) (rename A g))" | |
| 98 | apply (simp add: is_weak_ref_map_def) | |
| 99 | apply (rule conjI) | |
| 100 | (* 1: start states *) | |
| 101 | apply (simp add: rename_def rename_set_def starts_of_def) | |
| 102 | (* 2: reachable transitions *) | |
| 103 | apply (rule allI)+ | |
| 104 | apply (rule imp_conj_lemma) | |
| 105 | apply (simp (no_asm) add: rename_def rename_set_def) | |
| 106 | apply (simp add: externals_def asig_inputs_def asig_outputs_def asig_of_def trans_of_def) | |
| 107 | apply safe | |
| 108 | apply (simplesubst split_if) | |
| 109 | apply (rule conjI) | |
| 110 | apply (rule impI) | |
| 111 | apply (erule disjE) | |
| 112 | apply (erule exE) | |
| 113 | apply (erule conjE) | |
| 114 | (* x is input *) | |
| 115 | apply (drule sym) | |
| 116 | apply (drule sym) | |
| 117 | apply simp | |
| 118 | apply hypsubst+ | |
| 119 | apply (frule reachable_rename) | |
| 120 | apply simp | |
| 121 | (* x is output *) | |
| 122 | apply (erule exE) | |
| 123 | apply (erule conjE) | |
| 124 | apply (drule sym) | |
| 125 | apply (drule sym) | |
| 126 | apply simp | |
| 127 | apply hypsubst+ | |
| 128 | apply (frule reachable_rename) | |
| 129 | apply simp | |
| 130 | (* x is internal *) | |
| 131 | apply (frule reachable_rename) | |
| 132 | apply auto | |
| 133 | done | |
| 134 | ||
| 135 | declare split_if [split] | |
| 136 | declare if_weak_cong [cong] | |
| 4559 
8e604d885b54
added files containing temproal logic and abstraction;
 mueller parents: 
3433diff
changeset | 137 | |
| 3071 | 138 | end |