65 lemmas hom_ioas = |
65 lemmas hom_ioas = |
66 env_ioas [simp] impl_ioas [simp] impl_trans [simp] impl_asigs [simp] |
66 env_ioas [simp] impl_ioas [simp] impl_trans [simp] impl_asigs [simp] |
67 asig_projections set_lemmas |
67 asig_projections set_lemmas |
68 |
68 |
69 |
69 |
70 subsection {* lemmas about reduce *} |
70 subsection \<open>lemmas about reduce\<close> |
71 |
71 |
72 lemma l_iff_red_nil: "(reduce l = []) = (l = [])" |
72 lemma l_iff_red_nil: "(reduce l = []) = (l = [])" |
73 by (induct l) (auto split: list.split) |
73 by (induct l) (auto split: list.split) |
74 |
74 |
75 lemma hd_is_reduce_hd: "s ~= [] --> hd s = hd (reduce s)" |
75 lemma hd_is_reduce_hd: "s ~= [] --> hd s = hd (reduce s)" |
76 by (induct s) (auto split: list.split) |
76 by (induct s) (auto split: list.split) |
77 |
77 |
78 text {* to be used in the following Lemma *} |
78 text \<open>to be used in the following Lemma\<close> |
79 lemma rev_red_not_nil [rule_format]: |
79 lemma rev_red_not_nil [rule_format]: |
80 "l ~= [] --> reverse (reduce l) ~= []" |
80 "l ~= [] --> reverse (reduce l) ~= []" |
81 by (induct l) (auto split: list.split) |
81 by (induct l) (auto split: list.split) |
82 |
82 |
83 text {* shows applicability of the induction hypothesis of the following Lemma 1 *} |
83 text \<open>shows applicability of the induction hypothesis of the following Lemma 1\<close> |
84 lemma last_ind_on_first: |
84 lemma last_ind_on_first: |
85 "l ~= [] ==> hd (reverse (reduce (a # l))) = hd (reverse (reduce l))" |
85 "l ~= [] ==> hd (reverse (reduce (a # l))) = hd (reverse (reduce l))" |
86 apply simp |
86 apply simp |
87 apply (tactic {* auto_tac (put_simpset HOL_ss @{context} |
87 apply (tactic \<open>auto_tac (put_simpset HOL_ss @{context} |
88 addsimps (@{thms reverse.simps} @ [@{thm hd_append}, @{thm rev_red_not_nil}]) |
88 addsimps (@{thms reverse.simps} @ [@{thm hd_append}, @{thm rev_red_not_nil}]) |
89 |> Splitter.add_split @{thm list.split}) *}) |
89 |> Splitter.add_split @{thm list.split})\<close>) |
90 done |
90 done |
91 |
91 |
92 text {* Main Lemma 1 for @{text "S_pkt"} in showing that reduce is refinement. *} |
92 text \<open>Main Lemma 1 for \<open>S_pkt\<close> in showing that reduce is refinement.\<close> |
93 lemma reduce_hd: |
93 lemma reduce_hd: |
94 "if x=hd(reverse(reduce(l))) & reduce(l)~=[] then |
94 "if x=hd(reverse(reduce(l))) & reduce(l)~=[] then |
95 reduce(l@[x])=reduce(l) else |
95 reduce(l@[x])=reduce(l) else |
96 reduce(l@[x])=reduce(l)@[x]" |
96 reduce(l@[x])=reduce(l)@[x]" |
97 apply (simplesubst split_if) |
97 apply (simplesubst split_if) |
98 apply (rule conjI) |
98 apply (rule conjI) |
99 txt {* @{text "-->"} *} |
99 txt \<open>\<open>-->\<close>\<close> |
100 apply (induct_tac "l") |
100 apply (induct_tac "l") |
101 apply (simp (no_asm)) |
101 apply (simp (no_asm)) |
102 apply (case_tac "list=[]") |
102 apply (case_tac "list=[]") |
103 apply simp |
103 apply simp |
104 apply (rule impI) |
104 apply (rule impI) |
106 apply (cut_tac l = "list" in cons_not_nil) |
106 apply (cut_tac l = "list" in cons_not_nil) |
107 apply (simp del: reduce_Cons) |
107 apply (simp del: reduce_Cons) |
108 apply (erule exE)+ |
108 apply (erule exE)+ |
109 apply hypsubst |
109 apply hypsubst |
110 apply (simp del: reduce_Cons add: last_ind_on_first l_iff_red_nil) |
110 apply (simp del: reduce_Cons add: last_ind_on_first l_iff_red_nil) |
111 txt {* @{text "<--"} *} |
111 txt \<open>\<open><--\<close>\<close> |
112 apply (simp (no_asm) add: and_de_morgan_and_absorbe l_iff_red_nil) |
112 apply (simp (no_asm) add: and_de_morgan_and_absorbe l_iff_red_nil) |
113 apply (induct_tac "l") |
113 apply (induct_tac "l") |
114 apply (simp (no_asm)) |
114 apply (simp (no_asm)) |
115 apply (case_tac "list=[]") |
115 apply (case_tac "list=[]") |
116 apply (cut_tac [2] l = "list" in cons_not_nil) |
116 apply (cut_tac [2] l = "list" in cons_not_nil) |
118 apply (auto simp del: reduce_Cons simp add: last_ind_on_first l_iff_red_nil split: split_if) |
118 apply (auto simp del: reduce_Cons simp add: last_ind_on_first l_iff_red_nil split: split_if) |
119 apply simp |
119 apply simp |
120 done |
120 done |
121 |
121 |
122 |
122 |
123 text {* Main Lemma 2 for R_pkt in showing that reduce is refinement. *} |
123 text \<open>Main Lemma 2 for R_pkt in showing that reduce is refinement.\<close> |
124 lemma reduce_tl: "s~=[] ==> |
124 lemma reduce_tl: "s~=[] ==> |
125 if hd(s)=hd(tl(s)) & tl(s)~=[] then |
125 if hd(s)=hd(tl(s)) & tl(s)~=[] then |
126 reduce(tl(s))=reduce(s) else |
126 reduce(tl(s))=reduce(s) else |
127 reduce(tl(s))=tl(reduce(s))" |
127 reduce(tl(s))=tl(reduce(s))" |
128 apply (cut_tac l = "s" in cons_not_nil) |
128 apply (cut_tac l = "s" in cons_not_nil) |
130 apply (erule exE)+ |
130 apply (erule exE)+ |
131 apply (auto split: list.split) |
131 apply (auto split: list.split) |
132 done |
132 done |
133 |
133 |
134 |
134 |
135 subsection {* Channel Abstraction *} |
135 subsection \<open>Channel Abstraction\<close> |
136 |
136 |
137 declare split_if [split del] |
137 declare split_if [split del] |
138 |
138 |
139 lemma channel_abstraction: "is_weak_ref_map reduce ch_ioa ch_fin_ioa" |
139 lemma channel_abstraction: "is_weak_ref_map reduce ch_ioa ch_fin_ioa" |
140 apply (simp (no_asm) add: is_weak_ref_map_def) |
140 apply (simp (no_asm) add: is_weak_ref_map_def) |
141 txt {* main-part *} |
141 txt \<open>main-part\<close> |
142 apply (rule allI)+ |
142 apply (rule allI)+ |
143 apply (rule imp_conj_lemma) |
143 apply (rule imp_conj_lemma) |
144 apply (induct_tac "a") |
144 apply (induct_tac "a") |
145 txt {* 2 cases *} |
145 txt \<open>2 cases\<close> |
146 apply (simp_all (no_asm) cong del: if_weak_cong add: externals_def) |
146 apply (simp_all (no_asm) cong del: if_weak_cong add: externals_def) |
147 txt {* fst case *} |
147 txt \<open>fst case\<close> |
148 apply (rule impI) |
148 apply (rule impI) |
149 apply (rule disjI2) |
149 apply (rule disjI2) |
150 apply (rule reduce_hd) |
150 apply (rule reduce_hd) |
151 txt {* snd case *} |
151 txt \<open>snd case\<close> |
152 apply (rule impI) |
152 apply (rule impI) |
153 apply (erule conjE)+ |
153 apply (erule conjE)+ |
154 apply (erule disjE) |
154 apply (erule disjE) |
155 apply (simp add: l_iff_red_nil) |
155 apply (simp add: l_iff_red_nil) |
156 apply (erule hd_is_reduce_hd [THEN mp]) |
156 apply (erule hd_is_reduce_hd [THEN mp]) |
162 done |
162 done |
163 |
163 |
164 declare split_if [split] |
164 declare split_if [split] |
165 |
165 |
166 lemma sender_abstraction: "is_weak_ref_map reduce srch_ioa srch_fin_ioa" |
166 lemma sender_abstraction: "is_weak_ref_map reduce srch_ioa srch_fin_ioa" |
167 apply (tactic {* |
167 apply (tactic \<open> |
168 simp_tac (put_simpset HOL_ss @{context} |
168 simp_tac (put_simpset HOL_ss @{context} |
169 addsimps [@{thm srch_fin_ioa_def}, @{thm rsch_fin_ioa_def}, |
169 addsimps [@{thm srch_fin_ioa_def}, @{thm rsch_fin_ioa_def}, |
170 @{thm srch_ioa_def}, @{thm rsch_ioa_def}, @{thm rename_through_pmap}, |
170 @{thm srch_ioa_def}, @{thm rsch_ioa_def}, @{thm rename_through_pmap}, |
171 @{thm channel_abstraction}]) 1 *}) |
171 @{thm channel_abstraction}]) 1\<close>) |
172 done |
172 done |
173 |
173 |
174 lemma receiver_abstraction: "is_weak_ref_map reduce rsch_ioa rsch_fin_ioa" |
174 lemma receiver_abstraction: "is_weak_ref_map reduce rsch_ioa rsch_fin_ioa" |
175 apply (tactic {* |
175 apply (tactic \<open> |
176 simp_tac (put_simpset HOL_ss @{context} |
176 simp_tac (put_simpset HOL_ss @{context} |
177 addsimps [@{thm srch_fin_ioa_def}, @{thm rsch_fin_ioa_def}, |
177 addsimps [@{thm srch_fin_ioa_def}, @{thm rsch_fin_ioa_def}, |
178 @{thm srch_ioa_def}, @{thm rsch_ioa_def}, @{thm rename_through_pmap}, |
178 @{thm srch_ioa_def}, @{thm rsch_ioa_def}, @{thm rename_through_pmap}, |
179 @{thm channel_abstraction}]) 1 *}) |
179 @{thm channel_abstraction}]) 1\<close>) |
180 done |
180 done |
181 |
181 |
182 |
182 |
183 text {* 3 thms that do not hold generally! The lucky restriction here is |
183 text \<open>3 thms that do not hold generally! The lucky restriction here is |
184 the absence of internal actions. *} |
184 the absence of internal actions.\<close> |
185 lemma sender_unchanged: "is_weak_ref_map (%id. id) sender_ioa sender_ioa" |
185 lemma sender_unchanged: "is_weak_ref_map (%id. id) sender_ioa sender_ioa" |
186 apply (simp (no_asm) add: is_weak_ref_map_def) |
186 apply (simp (no_asm) add: is_weak_ref_map_def) |
187 txt {* main-part *} |
187 txt \<open>main-part\<close> |
188 apply (rule allI)+ |
188 apply (rule allI)+ |
189 apply (induct_tac a) |
189 apply (induct_tac a) |
190 txt {* 7 cases *} |
190 txt \<open>7 cases\<close> |
191 apply (simp_all (no_asm) add: externals_def) |
191 apply (simp_all (no_asm) add: externals_def) |
192 done |
192 done |
193 |
193 |
194 text {* 2 copies of before *} |
194 text \<open>2 copies of before\<close> |
195 lemma receiver_unchanged: "is_weak_ref_map (%id. id) receiver_ioa receiver_ioa" |
195 lemma receiver_unchanged: "is_weak_ref_map (%id. id) receiver_ioa receiver_ioa" |
196 apply (simp (no_asm) add: is_weak_ref_map_def) |
196 apply (simp (no_asm) add: is_weak_ref_map_def) |
197 txt {* main-part *} |
197 txt \<open>main-part\<close> |
198 apply (rule allI)+ |
198 apply (rule allI)+ |
199 apply (induct_tac a) |
199 apply (induct_tac a) |
200 txt {* 7 cases *} |
200 txt \<open>7 cases\<close> |
201 apply (simp_all (no_asm) add: externals_def) |
201 apply (simp_all (no_asm) add: externals_def) |
202 done |
202 done |
203 |
203 |
204 lemma env_unchanged: "is_weak_ref_map (%id. id) env_ioa env_ioa" |
204 lemma env_unchanged: "is_weak_ref_map (%id. id) env_ioa env_ioa" |
205 apply (simp (no_asm) add: is_weak_ref_map_def) |
205 apply (simp (no_asm) add: is_weak_ref_map_def) |
206 txt {* main-part *} |
206 txt \<open>main-part\<close> |
207 apply (rule allI)+ |
207 apply (rule allI)+ |
208 apply (induct_tac a) |
208 apply (induct_tac a) |
209 txt {* 7 cases *} |
209 txt \<open>7 cases\<close> |
210 apply (simp_all (no_asm) add: externals_def) |
210 apply (simp_all (no_asm) add: externals_def) |
211 done |
211 done |
212 |
212 |
213 |
213 |
214 lemma compat_single_ch: "compatible srch_ioa rsch_ioa" |
214 lemma compat_single_ch: "compatible srch_ioa rsch_ioa" |
289 apply (induct_tac x) |
289 apply (induct_tac x) |
290 apply simp_all |
290 apply simp_all |
291 done |
291 done |
292 |
292 |
293 |
293 |
294 text {* lemmata about externals of channels *} |
294 text \<open>lemmata about externals of channels\<close> |
295 lemma ext_single_ch: "externals(asig_of(srch_fin_ioa)) = externals(asig_of(srch_ioa)) & |
295 lemma ext_single_ch: "externals(asig_of(srch_fin_ioa)) = externals(asig_of(srch_ioa)) & |
296 externals(asig_of(rsch_fin_ioa)) = externals(asig_of(rsch_ioa))" |
296 externals(asig_of(rsch_fin_ioa)) = externals(asig_of(rsch_ioa))" |
297 by (simp add: externals_def) |
297 by (simp add: externals_def) |
298 |
298 |
299 |
299 |
300 subsection {* Soundness of Abstraction *} |
300 subsection \<open>Soundness of Abstraction\<close> |
301 |
301 |
302 lemmas ext_simps = externals_of_par ext_single_ch |
302 lemmas ext_simps = externals_of_par ext_single_ch |
303 and compat_simps = compat_single_ch compat_single_fin_ch compat_rec |
303 and compat_simps = compat_single_ch compat_single_fin_ch compat_rec |
304 compat_rec_fin compat_sen compat_sen_fin compat_env compat_env_fin |
304 compat_rec_fin compat_sen compat_sen_fin compat_env compat_env_fin |
305 and abstractions = env_unchanged sender_unchanged |
305 and abstractions = env_unchanged sender_unchanged |