author | wenzelm |
Sat, 16 Jan 2016 23:24:50 +0100 | |
changeset 62192 | bdaa0eb0fc74 |
parent 62008 | cbedaddc9351 |
child 62193 | 0826f6b6ba09 |
permissions | -rw-r--r-- |
62008 | 1 |
(* Title: HOL/HOLCF/IOA/Abstraction.thy |
40945 | 2 |
Author: Olaf Müller |
17233 | 3 |
*) |
4559 | 4 |
|
62002 | 5 |
section \<open>Abstraction Theory -- tailored for I/O automata\<close> |
4559 | 6 |
|
17233 | 7 |
theory Abstraction |
8 |
imports LiveIOA |
|
9 |
begin |
|
4559 | 10 |
|
36452 | 11 |
default_sort type |
4559 | 12 |
|
62192 | 13 |
definition cex_abs :: "('s1 \<Rightarrow> 's2) \<Rightarrow> ('a, 's1) execution \<Rightarrow> ('a, 's2) execution" |
14 |
where "cex_abs f ex = (f (fst ex), Map (\<lambda>(a, t). (a, f t)) $ (snd ex))" |
|
15 |
||
16 |
text \<open>equals cex_abs on Sequences -- after ex2seq application\<close> |
|
17 |
definition cex_absSeq :: |
|
18 |
"('s1 \<Rightarrow> 's2) \<Rightarrow> ('a option, 's1) transition Seq \<Rightarrow> ('a option, 's2)transition Seq" |
|
19 |
where "cex_absSeq f = (\<lambda>s. Map (\<lambda>(s, a, t). (f s, a, f t)) $ s)" |
|
4559 | 20 |
|
62192 | 21 |
definition is_abstraction :: "('s1 \<Rightarrow> 's2) \<Rightarrow> ('a, 's1) ioa \<Rightarrow> ('a, 's2) ioa \<Rightarrow> bool" |
22 |
where "is_abstraction f C A \<longleftrightarrow> |
|
23 |
((\<forall>s \<in> starts_of C. f s \<in> starts_of A) \<and> |
|
24 |
(\<forall>s t a. reachable C s \<and> s \<midarrow>a\<midarrow>C\<rightarrow> t \<longrightarrow> f s \<midarrow>a\<midarrow>A\<rightarrow> f t))" |
|
25 |
||
26 |
definition weakeningIOA :: "('a, 's2) ioa \<Rightarrow> ('a, 's1) ioa \<Rightarrow> ('s1 \<Rightarrow> 's2) \<Rightarrow> bool" |
|
27 |
where "weakeningIOA A C h \<longleftrightarrow> (\<forall>ex. ex \<in> executions C \<longrightarrow> cex_abs h ex \<in> executions A)" |
|
4559 | 28 |
|
62192 | 29 |
definition temp_strengthening :: "('a, 's2) ioa_temp \<Rightarrow> ('a, 's1) ioa_temp \<Rightarrow> ('s1 \<Rightarrow> 's2) \<Rightarrow> bool" |
30 |
where "temp_strengthening Q P h \<longleftrightarrow> (\<forall>ex. (cex_abs h ex \<TTurnstile> Q) \<longrightarrow> (ex \<TTurnstile> P))" |
|
31 |
||
32 |
definition temp_weakening :: "('a, 's2) ioa_temp \<Rightarrow> ('a, 's1) ioa_temp \<Rightarrow> ('s1 \<Rightarrow> 's2) \<Rightarrow> bool" |
|
33 |
where "temp_weakening Q P h \<longleftrightarrow> temp_strengthening (\<^bold>\<not> Q) (\<^bold>\<not> P) h" |
|
4559 | 34 |
|
62192 | 35 |
definition state_strengthening :: "('a, 's2) step_pred \<Rightarrow> ('a, 's1) step_pred \<Rightarrow> ('s1 \<Rightarrow> 's2) \<Rightarrow> bool" |
36 |
where "state_strengthening Q P h \<longleftrightarrow> (\<forall>s t a. Q (h s, a, h t) \<longrightarrow> P (s, a, t))" |
|
4559 | 37 |
|
62192 | 38 |
definition state_weakening :: "('a, 's2) step_pred \<Rightarrow> ('a, 's1) step_pred \<Rightarrow> ('s1 \<Rightarrow> 's2) \<Rightarrow> bool" |
39 |
where "state_weakening Q P h \<longleftrightarrow> state_strengthening (\<^bold>\<not> Q) (\<^bold>\<not> P) h" |
|
40 |
||
41 |
definition is_live_abstraction :: "('s1 \<Rightarrow> 's2) \<Rightarrow> ('a, 's1) live_ioa \<Rightarrow> ('a, 's2) live_ioa \<Rightarrow> bool" |
|
42 |
where "is_live_abstraction h CL AM \<longleftrightarrow> |
|
43 |
is_abstraction h (fst CL) (fst AM) \<and> temp_weakening (snd AM) (snd CL) h" |
|
4559 | 44 |
|
45 |
||
62192 | 46 |
(* thm about ex2seq which is not provable by induction as ex2seq is not continous *) |
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
47 |
axiomatization where |
62192 | 48 |
ex2seq_abs_cex: "ex2seq (cex_abs h ex) = cex_absSeq h (ex2seq ex)" |
4559 | 49 |
|
5976 | 50 |
(* analog to the proved thm strength_Box - proof skipped as trivial *) |
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
51 |
axiomatization where |
62192 | 52 |
weak_Box: "temp_weakening P Q h \<Longrightarrow> temp_weakening (\<box>P) (\<box>Q) h" |
53 |
||
5976 | 54 |
(* analog to the proved thm strength_Next - proof skipped as trivial *) |
62192 | 55 |
axiomatization where |
56 |
weak_Next: "temp_weakening P Q h \<Longrightarrow> temp_weakening (Next P) (Next Q) h" |
|
4559 | 57 |
|
19741 | 58 |
|
62192 | 59 |
subsection \<open>\<open>cex_abs\<close>\<close> |
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
60 |
|
62192 | 61 |
lemma cex_abs_UU [simp]: "cex_abs f (s, UU) = (f s, UU)" |
19741 | 62 |
by (simp add: cex_abs_def) |
63 |
||
62192 | 64 |
lemma cex_abs_nil [simp]: "cex_abs f (s,nil) = (f s, nil)" |
19741 | 65 |
by (simp add: cex_abs_def) |
66 |
||
62192 | 67 |
lemma cex_abs_cons [simp]: |
68 |
"cex_abs f (s, (a, t) \<leadsto> ex) = (f s, (a, f t) \<leadsto> (snd (cex_abs f (t, ex))))" |
|
19741 | 69 |
by (simp add: cex_abs_def) |
70 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
71 |
|
62192 | 72 |
subsection \<open>Lemmas\<close> |
19741 | 73 |
|
62192 | 74 |
lemma temp_weakening_def2: "temp_weakening Q P h \<longleftrightarrow> (\<forall>ex. (ex \<TTurnstile> P) \<longrightarrow> (cex_abs h ex \<TTurnstile> Q))" |
19741 | 75 |
apply (simp add: temp_weakening_def temp_strengthening_def NOT_def temp_sat_def satisfies_def) |
76 |
apply auto |
|
77 |
done |
|
78 |
||
62192 | 79 |
lemma state_weakening_def2: "state_weakening Q P h \<longleftrightarrow> (\<forall>s t a. P (s, a, t) \<longrightarrow> Q (h s, a, h t))" |
19741 | 80 |
apply (simp add: state_weakening_def state_strengthening_def NOT_def) |
81 |
apply auto |
|
82 |
done |
|
83 |
||
84 |
||
62192 | 85 |
subsection \<open>Abstraction Rules for Properties\<close> |
19741 | 86 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
87 |
lemma exec_frag_abstraction [rule_format]: |
62192 | 88 |
"is_abstraction h C A \<Longrightarrow> |
89 |
\<forall>s. reachable C s \<and> is_exec_frag C (s, xs) \<longrightarrow> is_exec_frag A (cex_abs h (s, xs))" |
|
90 |
apply (simp add: cex_abs_def) |
|
91 |
apply (tactic \<open>pair_induct_tac @{context} "xs" [@{thm is_exec_frag_def}] 1\<close>) |
|
92 |
txt \<open>main case\<close> |
|
93 |
apply (auto dest: reachable.reachable_n simp add: is_abstraction_def) |
|
94 |
done |
|
95 |
||
96 |
lemma abs_is_weakening: "is_abstraction h C A \<Longrightarrow> weakeningIOA A C h" |
|
97 |
apply (simp add: weakeningIOA_def) |
|
98 |
apply auto |
|
99 |
apply (simp add: executions_def) |
|
100 |
txt \<open>start state\<close> |
|
101 |
apply (rule conjI) |
|
102 |
apply (simp add: is_abstraction_def cex_abs_def) |
|
103 |
txt \<open>is-execution-fragment\<close> |
|
104 |
apply (erule exec_frag_abstraction) |
|
105 |
apply (simp add: reachable.reachable_0) |
|
106 |
done |
|
19741 | 107 |
|
108 |
||
62192 | 109 |
lemma AbsRuleT1: |
110 |
"is_abstraction h C A \<Longrightarrow> validIOA A Q \<Longrightarrow> temp_strengthening Q P h \<Longrightarrow> validIOA C P" |
|
111 |
apply (drule abs_is_weakening) |
|
112 |
apply (simp add: weakeningIOA_def validIOA_def temp_strengthening_def) |
|
113 |
apply (auto simp add: split_paired_all) |
|
114 |
done |
|
19741 | 115 |
|
116 |
||
62192 | 117 |
(* FIXME: to TLS.thy *) |
19741 | 118 |
|
62192 | 119 |
lemma IMPLIES_temp_sat [simp]: "(ex \<TTurnstile> P \<^bold>\<longrightarrow> Q) \<longleftrightarrow> ((ex \<TTurnstile> P) \<longrightarrow> (ex \<TTurnstile> Q))" |
19741 | 120 |
by (simp add: IMPLIES_def temp_sat_def satisfies_def) |
121 |
||
62192 | 122 |
lemma AND_temp_sat [simp]: "(ex \<TTurnstile> P \<^bold>\<and> Q) \<longleftrightarrow> ((ex \<TTurnstile> P) \<and> (ex \<TTurnstile> Q))" |
19741 | 123 |
by (simp add: AND_def temp_sat_def satisfies_def) |
124 |
||
62192 | 125 |
lemma OR_temp_sat [simp]: "(ex \<TTurnstile> P \<^bold>\<or> Q) \<longleftrightarrow> ((ex \<TTurnstile> P) \<or> (ex \<TTurnstile> Q))" |
19741 | 126 |
by (simp add: OR_def temp_sat_def satisfies_def) |
127 |
||
62192 | 128 |
lemma NOT_temp_sat [simp]: "(ex \<TTurnstile> \<^bold>\<not> P) \<longleftrightarrow> (\<not> (ex \<TTurnstile> P))" |
19741 | 129 |
by (simp add: NOT_def temp_sat_def satisfies_def) |
130 |
||
131 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
132 |
lemma AbsRuleT2: |
62192 | 133 |
"is_live_abstraction h (C, L) (A, M) \<Longrightarrow> validLIOA (A, M) Q \<Longrightarrow> temp_strengthening Q P h |
134 |
\<Longrightarrow> validLIOA (C, L) P" |
|
135 |
apply (unfold is_live_abstraction_def) |
|
136 |
apply auto |
|
137 |
apply (drule abs_is_weakening) |
|
138 |
apply (simp add: weakeningIOA_def temp_weakening_def2 validLIOA_def validIOA_def temp_strengthening_def) |
|
139 |
apply (auto simp add: split_paired_all) |
|
140 |
done |
|
19741 | 141 |
|
142 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
143 |
lemma AbsRuleTImprove: |
62192 | 144 |
"is_live_abstraction h (C, L) (A, M) \<Longrightarrow> validLIOA (A,M) (H1 \<^bold>\<longrightarrow> Q) \<Longrightarrow> |
145 |
temp_strengthening Q P h \<Longrightarrow> temp_weakening H1 H2 h \<Longrightarrow> validLIOA (C, L) H2 \<Longrightarrow> |
|
146 |
validLIOA (C, L) P" |
|
147 |
apply (unfold is_live_abstraction_def) |
|
148 |
apply auto |
|
149 |
apply (drule abs_is_weakening) |
|
150 |
apply (simp add: weakeningIOA_def temp_weakening_def2 validLIOA_def validIOA_def temp_strengthening_def) |
|
151 |
apply (auto simp add: split_paired_all) |
|
152 |
done |
|
19741 | 153 |
|
154 |
||
62192 | 155 |
subsection \<open>Correctness of safe abstraction\<close> |
156 |
||
157 |
lemma abstraction_is_ref_map: "is_abstraction h C A \<Longrightarrow> is_ref_map h C A" |
|
158 |
apply (auto simp: is_abstraction_def is_ref_map_def) |
|
159 |
apply (rule_tac x = "(a,h t) \<leadsto>nil" in exI) |
|
160 |
apply (simp add: move_def) |
|
161 |
done |
|
19741 | 162 |
|
62192 | 163 |
lemma abs_safety: "inp C = inp A \<Longrightarrow> out C = out A \<Longrightarrow> is_abstraction h C A \<Longrightarrow> C =<| A" |
164 |
apply (simp add: ioa_implements_def) |
|
165 |
apply (rule trace_inclusion) |
|
166 |
apply (simp (no_asm) add: externals_def) |
|
167 |
apply (auto)[1] |
|
168 |
apply (erule abstraction_is_ref_map) |
|
169 |
done |
|
19741 | 170 |
|
171 |
||
62192 | 172 |
subsection \<open>Correctness of life abstraction\<close> |
173 |
||
174 |
text \<open> |
|
175 |
Reduces to \<open>Filter (Map fst x) = Filter (Map fst (Map (\<lambda>(a, t). (a, x)) x)\<close>, |
|
176 |
that is to special Map Lemma. |
|
177 |
\<close> |
|
178 |
lemma traces_coincide_abs: |
|
179 |
"ext C = ext A \<Longrightarrow> mk_trace C $ xs = mk_trace A $ (snd (cex_abs f (s, xs)))" |
|
180 |
apply (unfold cex_abs_def mk_trace_def filter_act_def) |
|
181 |
apply simp |
|
182 |
apply (tactic \<open>pair_induct_tac @{context} "xs" [] 1\<close>) |
|
183 |
done |
|
19741 | 184 |
|
185 |
||
62192 | 186 |
text \<open> |
187 |
Does not work with \<open>abstraction_is_ref_map\<close> as proof of \<open>abs_safety\<close>, because |
|
188 |
\<open>is_live_abstraction\<close> includes \<open>temp_strengthening\<close> which is necessarily based |
|
189 |
on \<open>cex_abs\<close> and not on \<open>corresp_ex\<close>. Thus, the proof is redone in a more specific |
|
190 |
way for \<open>cex_abs\<close>. |
|
191 |
\<close> |
|
192 |
lemma abs_liveness: |
|
193 |
"inp C = inp A \<Longrightarrow> out C = out A \<Longrightarrow> is_live_abstraction h (C, M) (A, L) \<Longrightarrow> |
|
194 |
live_implements (C, M) (A, L)" |
|
195 |
apply (simp add: is_live_abstraction_def live_implements_def livetraces_def liveexecutions_def) |
|
196 |
apply auto |
|
197 |
apply (rule_tac x = "cex_abs h ex" in exI) |
|
198 |
apply auto |
|
199 |
text \<open>Traces coincide\<close> |
|
62002 | 200 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
19741 | 201 |
apply (rule traces_coincide_abs) |
202 |
apply (simp (no_asm) add: externals_def) |
|
203 |
apply (auto)[1] |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
204 |
|
62192 | 205 |
text \<open>\<open>cex_abs\<close> is execution\<close> |
62002 | 206 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
19741 | 207 |
apply (simp add: executions_def) |
62192 | 208 |
text \<open>start state\<close> |
19741 | 209 |
apply (rule conjI) |
210 |
apply (simp add: is_abstraction_def cex_abs_def) |
|
62192 | 211 |
text \<open>\<open>is_execution_fragment\<close>\<close> |
19741 | 212 |
apply (erule exec_frag_abstraction) |
213 |
apply (simp add: reachable.reachable_0) |
|
214 |
||
62192 | 215 |
text \<open>Liveness\<close> |
216 |
apply (simp add: temp_weakening_def2) |
|
217 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
|
218 |
done |
|
19741 | 219 |
|
62192 | 220 |
(* FIXME: to Traces.thy *) |
19741 | 221 |
|
62192 | 222 |
lemma implements_trans: "A =<| B \<Longrightarrow> B =<| C \<Longrightarrow> A =<| C" |
223 |
by (auto simp add: ioa_implements_def) |
|
19741 | 224 |
|
225 |
||
62192 | 226 |
subsection \<open>Abstraction Rules for Automata\<close> |
19741 | 227 |
|
62192 | 228 |
lemma AbsRuleA1: |
229 |
"inp C = inp A \<Longrightarrow> out C = out A \<Longrightarrow> inp Q = inp P \<Longrightarrow> out Q = out P \<Longrightarrow> |
|
230 |
is_abstraction h1 C A \<Longrightarrow> A =<| Q \<Longrightarrow> is_abstraction h2 Q P \<Longrightarrow> C =<| P" |
|
231 |
apply (drule abs_safety) |
|
232 |
apply assumption+ |
|
233 |
apply (drule abs_safety) |
|
234 |
apply assumption+ |
|
235 |
apply (erule implements_trans) |
|
236 |
apply (erule implements_trans) |
|
237 |
apply assumption |
|
238 |
done |
|
19741 | 239 |
|
62192 | 240 |
lemma AbsRuleA2: |
241 |
"inp C = inp A \<Longrightarrow> out C = out A \<Longrightarrow> inp Q = inp P \<Longrightarrow> out Q = out P \<Longrightarrow> |
|
242 |
is_live_abstraction h1 (C, LC) (A, LA) \<Longrightarrow> live_implements (A, LA) (Q, LQ) \<Longrightarrow> |
|
243 |
is_live_abstraction h2 (Q, LQ) (P, LP) \<Longrightarrow> live_implements (C, LC) (P, LP)" |
|
244 |
apply (drule abs_liveness) |
|
245 |
apply assumption+ |
|
246 |
apply (drule abs_liveness) |
|
247 |
apply assumption+ |
|
248 |
apply (erule live_implements_trans) |
|
249 |
apply (erule live_implements_trans) |
|
250 |
apply assumption |
|
251 |
done |
|
19741 | 252 |
|
253 |
||
254 |
declare split_paired_All [simp del] |
|
255 |
||
256 |
||
62192 | 257 |
subsection \<open>Localizing Temporal Strengthenings and Weakenings\<close> |
19741 | 258 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
259 |
lemma strength_AND: |
62192 | 260 |
"temp_strengthening P1 Q1 h \<Longrightarrow> temp_strengthening P2 Q2 h \<Longrightarrow> |
261 |
temp_strengthening (P1 \<^bold>\<and> P2) (Q1 \<^bold>\<and> Q2) h" |
|
262 |
by (auto simp: temp_strengthening_def) |
|
19741 | 263 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
264 |
lemma strength_OR: |
62192 | 265 |
"temp_strengthening P1 Q1 h \<Longrightarrow> temp_strengthening P2 Q2 h \<Longrightarrow> |
266 |
temp_strengthening (P1 \<^bold>\<or> P2) (Q1 \<^bold>\<or> Q2) h" |
|
267 |
by (auto simp: temp_strengthening_def) |
|
19741 | 268 |
|
62192 | 269 |
lemma strength_NOT: "temp_weakening P Q h \<Longrightarrow> temp_strengthening (\<^bold>\<not> P) (\<^bold>\<not> Q) h" |
270 |
by (auto simp: temp_weakening_def2 temp_strengthening_def) |
|
19741 | 271 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
272 |
lemma strength_IMPLIES: |
62192 | 273 |
"temp_weakening P1 Q1 h \<Longrightarrow> temp_strengthening P2 Q2 h \<Longrightarrow> |
274 |
temp_strengthening (P1 \<^bold>\<longrightarrow> P2) (Q1 \<^bold>\<longrightarrow> Q2) h" |
|
275 |
by (simp add: temp_weakening_def2 temp_strengthening_def) |
|
19741 | 276 |
|
277 |
||
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
278 |
lemma weak_AND: |
62192 | 279 |
"temp_weakening P1 Q1 h \<Longrightarrow> temp_weakening P2 Q2 h \<Longrightarrow> |
280 |
temp_weakening (P1 \<^bold>\<and> P2) (Q1 \<^bold>\<and> Q2) h" |
|
281 |
by (simp add: temp_weakening_def2) |
|
19741 | 282 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
283 |
lemma weak_OR: |
62192 | 284 |
"temp_weakening P1 Q1 h \<Longrightarrow> temp_weakening P2 Q2 h \<Longrightarrow> |
285 |
temp_weakening (P1 \<^bold>\<or> P2) (Q1 \<^bold>\<or> Q2) h" |
|
286 |
by (simp add: temp_weakening_def2) |
|
19741 | 287 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
288 |
lemma weak_NOT: |
62192 | 289 |
"temp_strengthening P Q h \<Longrightarrow> temp_weakening (\<^bold>\<not> P) (\<^bold>\<not> Q) h" |
290 |
by (auto simp add: temp_weakening_def2 temp_strengthening_def) |
|
19741 | 291 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
292 |
lemma weak_IMPLIES: |
62192 | 293 |
"temp_strengthening P1 Q1 h \<Longrightarrow> temp_weakening P2 Q2 h \<Longrightarrow> |
294 |
temp_weakening (P1 \<^bold>\<longrightarrow> P2) (Q1 \<^bold>\<longrightarrow> Q2) h" |
|
295 |
by (simp add: temp_weakening_def2 temp_strengthening_def) |
|
19741 | 296 |
|
297 |
||
62002 | 298 |
subsubsection \<open>Box\<close> |
19741 | 299 |
|
62192 | 300 |
(* FIXME: should be same as nil_is_Conc2 when all nils are turned to right side! *) |
301 |
lemma UU_is_Conc: "(UU = x @@ y) = (((x::'a Seq)= UU) | (x = nil \<and> y = UU))" |
|
302 |
by (tactic \<open>Seq_case_simp_tac @{context} "x" 1\<close>) |
|
19741 | 303 |
|
304 |
lemma ex2seqConc [rule_format]: |
|
62192 | 305 |
"Finite s1 \<longrightarrow> (\<forall>ex. (s \<noteq> nil \<and> s \<noteq> UU \<and> ex2seq ex = s1 @@ s) \<longrightarrow> (\<exists>ex'. s = ex2seq ex'))" |
306 |
apply (rule impI) |
|
307 |
apply (tactic \<open>Seq_Finite_induct_tac @{context} 1\<close>) |
|
308 |
apply blast |
|
309 |
text \<open>main case\<close> |
|
310 |
apply clarify |
|
311 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
|
312 |
apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>) |
|
313 |
text \<open>\<open>UU\<close> case\<close> |
|
314 |
apply (simp add: nil_is_Conc) |
|
315 |
text \<open>\<open>nil\<close> case\<close> |
|
316 |
apply (simp add: nil_is_Conc) |
|
317 |
text \<open>cons case\<close> |
|
318 |
apply (tactic \<open>pair_tac @{context} "aa" 1\<close>) |
|
319 |
apply auto |
|
320 |
done |
|
19741 | 321 |
|
322 |
(* important property of ex2seq: can be shiftet, as defined "pointwise" *) |
|
62192 | 323 |
lemma ex2seq_tsuffix: "tsuffix s (ex2seq ex) \<Longrightarrow> \<exists>ex'. s = (ex2seq ex')" |
324 |
apply (unfold tsuffix_def suffix_def) |
|
325 |
apply auto |
|
326 |
apply (drule ex2seqConc) |
|
327 |
apply auto |
|
328 |
done |
|
19741 | 329 |
|
330 |
||
62192 | 331 |
(* FIXME: to Sequence.thy *) |
19741 | 332 |
|
62192 | 333 |
lemma Mapnil: "Map f $ s = nil \<longleftrightarrow> s = nil" |
334 |
by (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>) |
|
19741 | 335 |
|
62192 | 336 |
lemma MapUU: "Map f $ s = UU \<longleftrightarrow> s = UU" |
337 |
by (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>) |
|
19741 | 338 |
|
339 |
||
62192 | 340 |
(*important property of cex_absSeq: As it is a 1to1 correspondence, |
19741 | 341 |
properties carry over *) |
62192 | 342 |
lemma cex_absSeq_tsuffix: "tsuffix s t \<Longrightarrow> tsuffix (cex_absSeq h s) (cex_absSeq h t)" |
343 |
apply (unfold tsuffix_def suffix_def cex_absSeq_def) |
|
344 |
apply auto |
|
345 |
apply (simp add: Mapnil) |
|
346 |
apply (simp add: MapUU) |
|
347 |
apply (rule_tac x = "Map (% (s,a,t) . (h s,a, h t))$s1" in exI) |
|
348 |
apply (simp add: Map2Finite MapConc) |
|
349 |
done |
|
19741 | 350 |
|
351 |
||
62192 | 352 |
lemma strength_Box: "temp_strengthening P Q h \<Longrightarrow> temp_strengthening (\<box>P) (\<box>Q) h" |
353 |
apply (unfold temp_strengthening_def state_strengthening_def temp_sat_def satisfies_def Box_def) |
|
354 |
apply clarify |
|
355 |
apply (frule ex2seq_tsuffix) |
|
356 |
apply clarify |
|
357 |
apply (drule_tac h = "h" in cex_absSeq_tsuffix) |
|
358 |
apply (simp add: ex2seq_abs_cex) |
|
359 |
done |
|
19741 | 360 |
|
361 |
||
62002 | 362 |
subsubsection \<open>Init\<close> |
19741 | 363 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
364 |
lemma strength_Init: |
62192 | 365 |
"state_strengthening P Q h \<Longrightarrow> temp_strengthening (Init P) (Init Q) h" |
366 |
apply (unfold temp_strengthening_def state_strengthening_def |
|
367 |
temp_sat_def satisfies_def Init_def unlift_def) |
|
368 |
apply auto |
|
369 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
|
370 |
apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>) |
|
371 |
apply (tactic \<open>pair_tac @{context} "a" 1\<close>) |
|
372 |
done |
|
19741 | 373 |
|
374 |
||
62002 | 375 |
subsubsection \<open>Next\<close> |
19741 | 376 |
|
62192 | 377 |
lemma TL_ex2seq_UU: "TL $ (ex2seq (cex_abs h ex)) = UU \<longleftrightarrow> TL $ (ex2seq ex) = UU" |
378 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
|
379 |
apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>) |
|
380 |
apply (tactic \<open>pair_tac @{context} "a" 1\<close>) |
|
381 |
apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>) |
|
382 |
apply (tactic \<open>pair_tac @{context} "a" 1\<close>) |
|
383 |
done |
|
19741 | 384 |
|
62192 | 385 |
lemma TL_ex2seq_nil: "TL $ (ex2seq (cex_abs h ex)) = nil \<longleftrightarrow> TL $ (ex2seq ex) = nil" |
386 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
|
387 |
apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>) |
|
388 |
apply (tactic \<open>pair_tac @{context} "a" 1\<close>) |
|
389 |
apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>) |
|
390 |
apply (tactic \<open>pair_tac @{context} "a" 1\<close>) |
|
391 |
done |
|
19741 | 392 |
|
62192 | 393 |
(* FIXME: put to Sequence Lemmas *) |
394 |
lemma MapTL: "Map f $ (TL $ s) = TL $ (Map f $ s)" |
|
395 |
by (tactic \<open>Seq_induct_tac @{context} "s" [] 1\<close>) |
|
19741 | 396 |
|
62192 | 397 |
(*important property of cex_absSeq: As it is a 1to1 correspondence, |
398 |
properties carry over*) |
|
399 |
lemma cex_absSeq_TL: "cex_absSeq h (TL $ s) = TL $ (cex_absSeq h s)" |
|
400 |
by (simp add: MapTL cex_absSeq_def) |
|
19741 | 401 |
|
402 |
(* important property of ex2seq: can be shiftet, as defined "pointwise" *) |
|
62192 | 403 |
lemma TLex2seq: "snd ex \<noteq> UU \<Longrightarrow> snd ex \<noteq> nil \<Longrightarrow> \<exists>ex'. TL$(ex2seq ex) = ex2seq ex'" |
404 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
|
405 |
apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>) |
|
406 |
apply (tactic \<open>pair_tac @{context} "a" 1\<close>) |
|
407 |
apply auto |
|
408 |
done |
|
19741 | 409 |
|
62192 | 410 |
lemma ex2seqnilTL: "TL $ (ex2seq ex) \<noteq> nil \<longleftrightarrow> snd ex \<noteq> nil \<and> snd ex \<noteq> UU" |
411 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
|
412 |
apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>) |
|
413 |
apply (tactic \<open>pair_tac @{context} "a" 1\<close>) |
|
414 |
apply (tactic \<open>Seq_case_simp_tac @{context} "s" 1\<close>) |
|
415 |
apply (tactic \<open>pair_tac @{context} "a" 1\<close>) |
|
416 |
done |
|
19741 | 417 |
|
62192 | 418 |
lemma strength_Next: "temp_strengthening P Q h \<Longrightarrow> temp_strengthening (Next P) (Next Q) h" |
419 |
apply (unfold temp_strengthening_def state_strengthening_def temp_sat_def satisfies_def Next_def) |
|
420 |
apply simp |
|
421 |
apply auto |
|
422 |
apply (simp add: TL_ex2seq_nil TL_ex2seq_UU) |
|
423 |
apply (simp add: TL_ex2seq_nil TL_ex2seq_UU) |
|
424 |
apply (simp add: TL_ex2seq_nil TL_ex2seq_UU) |
|
425 |
apply (simp add: TL_ex2seq_nil TL_ex2seq_UU) |
|
426 |
text \<open>cons case\<close> |
|
427 |
apply (simp add: TL_ex2seq_nil TL_ex2seq_UU ex2seq_abs_cex cex_absSeq_TL [symmetric] ex2seqnilTL) |
|
428 |
apply (erule conjE) |
|
429 |
apply (drule TLex2seq) |
|
430 |
apply assumption |
|
431 |
apply auto |
|
432 |
done |
|
19741 | 433 |
|
434 |
||
62192 | 435 |
text \<open>Localizing Temporal Weakenings - 2\<close> |
19741 | 436 |
|
62192 | 437 |
lemma weak_Init: "state_weakening P Q h \<Longrightarrow> temp_weakening (Init P) (Init Q) h" |
438 |
apply (simp add: temp_weakening_def2 state_weakening_def2 |
|
439 |
temp_sat_def satisfies_def Init_def unlift_def) |
|
440 |
apply auto |
|
441 |
apply (tactic \<open>pair_tac @{context} "ex" 1\<close>) |
|
442 |
apply (tactic \<open>Seq_case_simp_tac @{context} "x2" 1\<close>) |
|
443 |
apply (tactic \<open>pair_tac @{context} "a" 1\<close>) |
|
444 |
done |
|
19741 | 445 |
|
446 |
||
62002 | 447 |
text \<open>Localizing Temproal Strengthenings - 3\<close> |
19741 | 448 |
|
62192 | 449 |
lemma strength_Diamond: "temp_strengthening P Q h \<Longrightarrow> temp_strengthening (\<diamond>P) (\<diamond>Q) h" |
450 |
apply (unfold Diamond_def) |
|
451 |
apply (rule strength_NOT) |
|
452 |
apply (rule weak_Box) |
|
453 |
apply (erule weak_NOT) |
|
454 |
done |
|
19741 | 455 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
456 |
lemma strength_Leadsto: |
62192 | 457 |
"temp_weakening P1 P2 h \<Longrightarrow> temp_strengthening Q1 Q2 h \<Longrightarrow> |
458 |
temp_strengthening (P1 \<leadsto> Q1) (P2 \<leadsto> Q2) h" |
|
459 |
apply (unfold Leadsto_def) |
|
460 |
apply (rule strength_Box) |
|
461 |
apply (erule strength_IMPLIES) |
|
462 |
apply (erule strength_Diamond) |
|
463 |
done |
|
19741 | 464 |
|
465 |
||
62002 | 466 |
text \<open>Localizing Temporal Weakenings - 3\<close> |
19741 | 467 |
|
62192 | 468 |
lemma weak_Diamond: "temp_weakening P Q h \<Longrightarrow> temp_weakening (\<diamond>P) (\<diamond>Q) h" |
469 |
apply (unfold Diamond_def) |
|
470 |
apply (rule weak_NOT) |
|
471 |
apply (rule strength_Box) |
|
472 |
apply (erule strength_NOT) |
|
473 |
done |
|
19741 | 474 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
475 |
lemma weak_Leadsto: |
62192 | 476 |
"temp_strengthening P1 P2 h \<Longrightarrow> temp_weakening Q1 Q2 h \<Longrightarrow> |
477 |
temp_weakening (P1 \<leadsto> Q1) (P2 \<leadsto> Q2) h" |
|
478 |
apply (unfold Leadsto_def) |
|
479 |
apply (rule weak_Box) |
|
480 |
apply (erule weak_IMPLIES) |
|
481 |
apply (erule weak_Diamond) |
|
482 |
done |
|
19741 | 483 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
484 |
lemma weak_WF: |
62192 | 485 |
"(\<And>s. Enabled A acts (h s) \<Longrightarrow> Enabled C acts s) |
486 |
\<Longrightarrow> temp_weakening (WF A acts) (WF C acts) h" |
|
487 |
apply (unfold WF_def) |
|
488 |
apply (rule weak_IMPLIES) |
|
489 |
apply (rule strength_Diamond) |
|
490 |
apply (rule strength_Box) |
|
491 |
apply (rule strength_Init) |
|
492 |
apply (rule_tac [2] weak_Box) |
|
493 |
apply (rule_tac [2] weak_Diamond) |
|
494 |
apply (rule_tac [2] weak_Init) |
|
495 |
apply (auto simp add: state_weakening_def state_strengthening_def |
|
496 |
xt2_def plift_def option_lift_def NOT_def) |
|
497 |
done |
|
19741 | 498 |
|
25135
4f8176c940cf
modernized specifications ('definition', 'axiomatization');
wenzelm
parents:
23560
diff
changeset
|
499 |
lemma weak_SF: |
62192 | 500 |
"(\<And>s. Enabled A acts (h s) \<Longrightarrow> Enabled C acts s) |
501 |
\<Longrightarrow> temp_weakening (SF A acts) (SF C acts) h" |
|
502 |
apply (unfold SF_def) |
|
503 |
apply (rule weak_IMPLIES) |
|
504 |
apply (rule strength_Box) |
|
505 |
apply (rule strength_Diamond) |
|
506 |
apply (rule strength_Init) |
|
507 |
apply (rule_tac [2] weak_Box) |
|
508 |
apply (rule_tac [2] weak_Diamond) |
|
509 |
apply (rule_tac [2] weak_Init) |
|
510 |
apply (auto simp add: state_weakening_def state_strengthening_def |
|
511 |
xt2_def plift_def option_lift_def NOT_def) |
|
512 |
done |
|
19741 | 513 |
|
514 |
||
515 |
lemmas weak_strength_lemmas = |
|
516 |
weak_OR weak_AND weak_NOT weak_IMPLIES weak_Box weak_Next weak_Init |
|
517 |
weak_Diamond weak_Leadsto strength_OR strength_AND strength_NOT |
|
518 |
strength_IMPLIES strength_Box strength_Next strength_Init |
|
519 |
strength_Diamond strength_Leadsto weak_WF weak_SF |
|
520 |
||
62002 | 521 |
ML \<open> |
30607
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents:
27208
diff
changeset
|
522 |
fun abstraction_tac ctxt = |
42793 | 523 |
SELECT_GOAL (auto_tac |
524 |
(ctxt addSIs @{thms weak_strength_lemmas} |
|
51717
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents:
42814
diff
changeset
|
525 |
addsimps [@{thm state_strengthening_def}, @{thm state_weakening_def}])) |
62002 | 526 |
\<close> |
17233 | 527 |
|
62002 | 528 |
method_setup abstraction = \<open>Scan.succeed (SIMPLE_METHOD' o abstraction_tac)\<close> |
30607
c3d1590debd8
eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents:
27208
diff
changeset
|
529 |
|
17233 | 530 |
end |