152 |
159 |
153 fun |
160 fun |
154 holds_for :: "(tape \<Rightarrow> bool) \<Rightarrow> config \<Rightarrow> bool" ("_ holds'_for _" [100, 99] 100) |
161 holds_for :: "(tape \<Rightarrow> bool) \<Rightarrow> config \<Rightarrow> bool" ("_ holds'_for _" [100, 99] 100) |
155 where |
162 where |
156 "P holds_for (s, l, r) = P (l, r)" |
163 "P holds_for (s, l, r) = P (l, r)" |
157 |
|
158 (* |
|
159 lemma step_0 [simp]: |
|
160 shows "step (0, (l, r)) p = (0, (l, r))" |
|
161 by simp |
|
162 |
|
163 lemma steps_0 [simp]: |
|
164 shows "steps (0, (l, r)) p n = (0, (l, r))" |
|
165 by (induct n) (simp_all) |
|
166 *) |
|
167 |
164 |
168 lemma is_final_holds[simp]: |
165 lemma is_final_holds[simp]: |
169 assumes "is_final c" |
166 assumes "is_final c" |
170 shows "Q holds_for (steps c (p, off) n) = Q holds_for c" |
167 shows "Q holds_for (steps c (p, off) n) = Q holds_for c" |
171 using assms |
168 using assms |
172 apply(induct n) |
169 apply(induct n) |
|
170 apply(auto) |
173 apply(case_tac [!] c) |
171 apply(case_tac [!] c) |
174 apply(auto) |
172 apply(auto) |
175 done |
173 done |
176 |
174 |
177 type_synonym assert = "tape \<Rightarrow> bool" |
175 type_synonym assert = "tape \<Rightarrow> bool" |
178 |
176 |
179 definition |
177 definition |
180 assert_imp :: "assert \<Rightarrow> assert \<Rightarrow> bool" ("_ \<mapsto> _" [0, 0] 100) |
178 assert_imp :: "assert \<Rightarrow> assert \<Rightarrow> bool" ("_ \<mapsto> _" [0, 0] 100) |
181 where |
179 where |
182 "P \<mapsto> Q = (\<forall>l r. P (l, r) \<longrightarrow> Q (l, r))" |
180 "P \<mapsto> Q \<equiv> \<forall>l r. P (l, r) \<longrightarrow> Q (l, r)" |
183 |
181 |
184 lemma holds_for_imp: |
182 lemma holds_for_imp: |
185 assumes "P holds_for c" |
183 assumes "P holds_for c" |
186 and "P \<mapsto> Q" |
184 and "P \<mapsto> Q" |
187 shows "Q holds_for c" |
185 shows "Q holds_for c" |
188 using assms unfolding assert_imp_def by (case_tac c, auto) |
186 using assms unfolding assert_imp_def |
|
187 by (case_tac c) (auto) |
189 |
188 |
190 definition |
189 definition |
191 Hoare :: "assert \<Rightarrow> tprog \<Rightarrow> assert \<Rightarrow> bool" ("({(1_)}/ (_)/ {(1_)})" 50) |
190 Hoare :: "assert \<Rightarrow> tprog0 \<Rightarrow> assert \<Rightarrow> bool" ("({(1_)}/ (_)/ {(1_)})" 50) |
192 where |
191 where |
193 "{P} p {Q} \<equiv> |
192 "{P} p {Q} \<equiv> |
194 (\<forall>l r. P (l, r) \<longrightarrow> (\<exists>n. is_final (steps (1, (l, r)) p n) \<and> Q holds_for (steps (1, (l, r)) p n)))" |
193 (\<forall>l r. P (l, r) \<longrightarrow> (\<exists>n. is_final (steps0 (1, (l, r)) p n) \<and> Q holds_for (steps0 (1, (l, r)) p n)))" |
195 |
194 |
196 lemma HoareI: |
195 lemma HoareI: |
197 assumes "\<And>l r. P (l, r) \<Longrightarrow> \<exists>n. is_final (steps (1, (l, r)) p n) \<and> Q holds_for (steps (1, (l, r)) p n)" |
196 assumes "\<And>l r. P (l, r) \<Longrightarrow> \<exists>n. is_final (steps0 (1, (l, r)) p n) \<and> Q holds_for (steps0 (1, (l, r)) p n)" |
198 shows "{P} p {Q}" |
197 shows "{P} p {Q}" |
199 unfolding Hoare_def using assms by auto |
198 unfolding Hoare_def using assms by auto |
200 |
199 |
201 |
200 |
202 lemma step_0 [simp]: |
201 lemma step_0 [simp]: |
205 |
204 |
206 lemma steps_0 [simp]: |
205 lemma steps_0 [simp]: |
207 shows "steps (0, (l, r)) p n = (0, (l, r))" |
206 shows "steps (0, (l, r)) p n = (0, (l, r))" |
208 by (induct n) (simp_all) |
207 by (induct n) (simp_all) |
209 |
208 |
210 declare steps.simps[simp del] |
209 (* if the machine is in the halting state, there must have |
211 |
210 been a state just before the halting state *) |
212 lemma before_final_old: |
|
213 assumes "steps (1, tp) A n = (0, tp')" |
|
214 obtains n' where "\<not> is_final (steps (1, tp) A n')" |
|
215 and "steps (1, tp) A (Suc n') = (0, tp')" |
|
216 proof - |
|
217 from assms have "\<exists> n'. \<not> is_final (steps (1, tp) A n') \<and> |
|
218 steps (1, tp) A (Suc n') = (0, tp')" |
|
219 proof(induct n arbitrary: tp', simp add: steps.simps) |
|
220 fix n tp' |
|
221 assume ind: |
|
222 "\<And>tp'. steps (1, tp) A n = (0, tp') \<Longrightarrow> |
|
223 \<exists>n'. \<not> is_final (steps (1, tp) A n') \<and> steps (1, tp) A (Suc n') = (0, tp')" |
|
224 and h: " steps (1, tp) A (Suc n) = (0, tp')" |
|
225 from h show "\<exists>n'. \<not> is_final (steps (1, tp) A n') \<and> steps (1, tp) A (Suc n') = (0, tp')" |
|
226 proof(simp add: step_red del: steps.simps, |
|
227 case_tac "(steps (Suc 0, tp) A n)", case_tac "a = 0", simp) |
|
228 fix a b c |
|
229 assume " steps (Suc 0, tp) A n = (0, tp')" |
|
230 hence "\<exists>n'. \<not> is_final (steps (1, tp) A n') \<and> steps (1, tp) A (Suc n') = (0, tp')" |
|
231 apply(rule_tac ind, simp) |
|
232 done |
|
233 thus "\<exists>n'. \<not> is_final (steps (Suc 0, tp) A n') \<and> step (steps (Suc 0, tp) A n') A = (0, tp')" |
|
234 apply(simp) |
|
235 done |
|
236 next |
|
237 fix a b c |
|
238 assume "steps (Suc 0, tp) A n = (a, b, c)" |
|
239 "step (steps (Suc 0, tp) A n) A = (0, tp')" |
|
240 "a \<noteq> 0" |
|
241 thus "\<exists>n'. \<not> is_final (steps (Suc 0, tp) A n') \<and> |
|
242 step (steps (Suc 0, tp) A n') A = (0, tp')" |
|
243 apply(rule_tac x = n in exI) |
|
244 apply(simp) |
|
245 done |
|
246 qed |
|
247 qed |
|
248 thus "(\<And>n'. \<lbrakk>\<not> is_final (steps (1, tp) A n'); |
|
249 steps (1, tp) A (Suc n') = (0, tp')\<rbrakk> \<Longrightarrow> thesis) \<Longrightarrow> thesis" |
|
250 by auto |
|
251 qed |
|
252 |
|
253 lemma before_final: |
211 lemma before_final: |
254 assumes "steps (1, tp) A n = (0, tp')" |
212 assumes "steps0 (1, tp) A n = (0, tp')" |
255 shows "\<exists> n'. \<not> is_final (steps (1, tp) A n') \<and> steps (1, tp) A (Suc n') = (0, tp')" |
213 shows "\<exists> n'. \<not> is_final (steps0 (1, tp) A n') \<and> steps0 (1, tp) A (Suc n') = (0, tp')" |
256 using assms |
214 using assms |
257 proof(induct n arbitrary: tp') |
215 proof(induct n arbitrary: tp') |
258 case (0 tp') |
216 case (0 tp') |
259 have asm: "steps (1, tp) A 0 = (0, tp')" by fact |
217 have asm: "steps0 (1, tp) A 0 = (0, tp')" by fact |
260 then show "\<exists>n'. \<not> is_final (steps (1, tp) A n') \<and> steps (1, tp) A (Suc n') = (0, tp')" |
218 then show "\<exists>n'. \<not> is_final (steps0 (1, tp) A n') \<and> steps0 (1, tp) A (Suc n') = (0, tp')" |
261 by (simp add: steps.simps) |
219 by simp |
262 next |
220 next |
263 case (Suc n tp') |
221 case (Suc n tp') |
264 have ih: "\<And>tp'. steps (1, tp) A n = (0, tp') \<Longrightarrow> |
222 have ih: "\<And>tp'. steps0 (1, tp) A n = (0, tp') \<Longrightarrow> |
265 \<exists>n'. \<not> is_final (steps (1, tp) A n') \<and> steps (1, tp) A (Suc n') = (0, tp')" by fact |
223 \<exists>n'. \<not> is_final (steps0 (1, tp) A n') \<and> steps0 (1, tp) A (Suc n') = (0, tp')" by fact |
266 have asm: "steps (1, tp) A (Suc n) = (0, tp')" by fact |
224 have asm: "steps0 (1, tp) A (Suc n) = (0, tp')" by fact |
267 obtain s l r where cases: "steps (1, tp) A n = (s, l, r)" |
225 obtain s l r where cases: "steps0 (1, tp) A n = (s, l, r)" |
268 by (auto intro: is_final.cases) |
226 by (auto intro: is_final.cases) |
269 then show "\<exists>n'. \<not> is_final (steps (1, tp) A n') \<and> steps (1, tp) A (Suc n') = (0, tp')" |
227 then show "\<exists>n'. \<not> is_final (steps0 (1, tp) A n') \<and> steps0 (1, tp) A (Suc n') = (0, tp')" |
270 proof (cases "s = 0") |
228 proof (cases "s = 0") |
271 case True (* in halting state *) |
229 case True (* in halting state *) |
272 then have "steps (1, tp) A n = (0, tp')" |
230 then have "steps0 (1, tp) A n = (0, tp')" |
273 using asm cases by simp |
231 using asm cases by (simp del: steps.simps) |
274 then show ?thesis using ih by simp |
232 then show ?thesis using ih by simp |
275 next |
233 next |
276 case False (* not in halting state *) |
234 case False (* not in halting state *) |
277 then have "\<not> is_final (steps (1, tp) A n) \<and> steps (1, tp) A (Suc n) = (0, tp')" |
235 then have "\<not> is_final (steps0 (1, tp) A n) \<and> steps0 (1, tp) A (Suc n) = (0, tp')" |
278 using asm cases by simp |
236 using asm cases by simp |
279 then show ?thesis by auto |
237 then show ?thesis by auto |
280 qed |
238 qed |
281 qed |
239 qed |
282 |
240 |
|
241 |
|
242 lemma length_comp: |
|
243 shows "length (A |+| B) = length A + length B" |
|
244 by auto |
|
245 |
|
246 declare steps.simps[simp del] |
283 declare tm_comp.simps [simp del] adjust.simps[simp del] shift.simps[simp del] |
247 declare tm_comp.simps [simp del] adjust.simps[simp del] shift.simps[simp del] |
284 |
248 |
285 lemma length_comp: |
|
286 "length (A |+| B) = length A + length B" |
|
287 apply(auto simp: tm_comp.simps) |
|
288 done |
|
289 |
249 |
290 lemma tmcomp_fetch_in_first: |
250 lemma tmcomp_fetch_in_first: |
291 assumes "case (fetch A a x) of (ac, ns) \<Rightarrow> ns \<noteq> 0" |
251 assumes "case (fetch A a x) of (ac, ns) \<Rightarrow> ns \<noteq> 0" |
292 shows "fetch (A |+| B) a x = fetch A a x" |
252 shows "fetch (A |+| B) a x = fetch A a x" |
293 using assms |
253 using assms |
530 |
490 |
531 lemma Hoare_plus_halt: |
491 lemma Hoare_plus_halt: |
532 assumes aimpb: "Q1 \<mapsto> P2" |
492 assumes aimpb: "Q1 \<mapsto> P2" |
533 and A_wf : "tm_wf (A, 0)" |
493 and A_wf : "tm_wf (A, 0)" |
534 and B_wf : "tm_wf (B, 0)" |
494 and B_wf : "tm_wf (B, 0)" |
535 and A_halt : "{P1} (A, 0) {Q1}" |
495 and A_halt : "{P1} A {Q1}" |
536 and B_halt : "{P2} (B, 0) {Q2}" |
496 and B_halt : "{P2} B {Q2}" |
537 shows "{P1} (A |+| B, 0) {Q2}" |
497 shows "{P1} A |+| B {Q2}" |
538 proof(rule HoareI) |
498 proof(rule HoareI) |
539 fix l r |
499 fix l r |
540 assume h: "P1 (l, r)" |
500 assume h: "P1 (l, r)" |
541 then obtain n1 |
501 then obtain n1 |
542 where "is_final (steps (1, l, r) (A, 0) n1)" and "Q1 holds_for (steps (1, l, r) (A, 0) n1)" |
502 where "is_final (steps0 (1, l, r) A n1)" and "Q1 holds_for (steps0 (1, l, r) A n1)" |
543 using A_halt unfolding Hoare_def by auto |
503 using A_halt unfolding Hoare_def by auto |
544 then obtain l' r' where "steps (1, l, r) (A, 0) n1 = (0, l', r')" and c: "Q1 holds_for (0, l', r')" |
504 then obtain l' r' where "steps0 (1, l, r) A n1 = (0, l', r')" and c: "Q1 holds_for (0, l', r')" |
545 by(case_tac "steps (1, l, r) (A, 0) n1") (auto) |
505 by(case_tac "steps0 (1, l, r) A n1") (auto) |
546 then obtain stpa where d: "steps (1, l, r) (A |+| B, 0) stpa = (Suc (length A div 2), l', r')" |
506 then obtain stpa where d: "steps0 (1, l, r) (A |+| B) stpa = (Suc (length A div 2), l', r')" |
547 using A_wf by(rule_tac t_merge_pre_halt_same) (auto) |
507 using A_wf by(rule_tac t_merge_pre_halt_same) (auto) |
548 moreover |
508 moreover |
549 from c aimpb have "P2 holds_for (0, l', r')" |
509 from c aimpb have "P2 holds_for (0, l', r')" |
550 by (rule holds_for_imp) |
510 by (rule holds_for_imp) |
551 then have "P2 (l', r')" by auto |
511 then have "P2 (l', r')" by auto |
552 then obtain n2 |
512 then obtain n2 |
553 where "is_final (steps (1, l', r') (B, 0) n2)" and "Q2 holds_for (steps (1, l', r') (B, 0) n2)" |
513 where "is_final (steps0 (1, l', r') B n2)" and "Q2 holds_for (steps0 (1, l', r') B n2)" |
554 using B_halt unfolding Hoare_def by auto |
514 using B_halt unfolding Hoare_def by auto |
555 then obtain l'' r'' where "steps (1, l', r') (B, 0) n2 = (0, l'', r'')" and g: "Q2 holds_for (0, l'', r'')" |
515 then obtain l'' r'' where "steps0 (1, l', r') B n2 = (0, l'', r'')" and g: "Q2 holds_for (0, l'', r'')" |
556 by (case_tac "steps (1, l', r') (B, 0) n2", auto) |
516 by (case_tac "steps0 (1, l', r') B n2", auto) |
557 then have "steps (Suc (length A div 2), l', r') (A |+| B, 0) n2 = (0, l'', r'')" |
517 then have "steps0 (Suc (length A div 2), l', r') (A |+| B) n2 = (0, l'', r'')" |
558 by (rule_tac t_merge_second_halt_same) (auto simp: A_wf B_wf) |
518 by (rule_tac t_merge_second_halt_same) (auto simp: A_wf B_wf) |
559 ultimately show |
519 ultimately show |
560 "\<exists>n. is_final (steps (1, l, r) (A |+| B, 0) n) \<and> Q2 holds_for (steps (1, l, r) (A |+| B, 0) n)" |
520 "\<exists>n. is_final (steps0 (1, l, r) (A |+| B) n) \<and> Q2 holds_for (steps0 (1, l, r) (A |+| B) n)" |
561 using g |
521 using g |
562 apply(rule_tac x = "stpa + n2" in exI) |
522 apply(rule_tac x = "stpa + n2" in exI) |
563 apply(simp add: steps_add) |
523 apply(simp add: steps_add) |
564 done |
524 done |
565 qed |
525 qed |
566 |
526 |
567 definition |
527 definition |
568 Hoare_unhalt :: "assert \<Rightarrow> tprog \<Rightarrow> bool" ("({(1_)}/ (_))" 50) |
528 Hoare_unhalt :: "assert \<Rightarrow> tprog0 \<Rightarrow> bool" ("({(1_)}/ (_))" 50) |
569 where |
529 where |
570 "{P} p \<equiv> |
530 "{P} p \<equiv> |
571 (\<forall>l r. P (l, r) \<longrightarrow> (\<forall> n . \<not> (is_final (steps (1, (l, r)) p n))))" |
531 (\<forall>l r. P (l, r) \<longrightarrow> (\<forall> n . \<not> (is_final (steps0 (1, (l, r)) p n))))" |
572 |
532 |
573 lemma Hoare_unhalt_I: |
533 lemma Hoare_unhalt_I: |
574 assumes "\<And>l r. P (l, r) \<Longrightarrow> \<forall> n. \<not> is_final (steps (1, (l, r)) p n)" |
534 assumes "\<And>l r. P (l, r) \<Longrightarrow> \<forall> n. \<not> is_final (steps0 (1, (l, r)) p n)" |
575 shows "{P} p" |
535 shows "{P} p" |
576 unfolding Hoare_unhalt_def using assms by auto |
536 unfolding Hoare_unhalt_def using assms by auto |
577 |
537 |
578 lemma Hoare_plus_unhalt: |
538 lemma Hoare_plus_unhalt: |
|
539 fixes A B :: tprog0 |
579 assumes aimpb: "Q1 \<mapsto> P2" |
540 assumes aimpb: "Q1 \<mapsto> P2" |
580 and A_wf : "tm_wf (A, 0)" |
541 and A_wf : "tm_wf (A, 0)" |
581 and B_wf : "tm_wf (B, 0)" |
542 and B_wf : "tm_wf (B, 0)" |
582 and A_halt : "{P1} (A, 0) {Q1}" |
543 and A_halt : "{P1} A {Q1}" |
583 and B_uhalt : "{P2} (B, 0)" |
544 and B_uhalt : "{P2} B" |
584 shows "{P1} (A |+| B, 0)" |
545 shows "{P1} (A |+| B)" |
585 proof(rule_tac Hoare_unhalt_I) |
546 proof(rule_tac Hoare_unhalt_I) |
586 fix l r |
547 fix l r |
587 assume h: "P1 (l, r)" |
548 assume h: "P1 (l, r)" |
588 then obtain n1 where a: "is_final (steps (1, l, r) (A, 0) n1)" and b: "Q1 holds_for (steps (1, l, r) (A, 0) n1)" |
549 then obtain n1 where a: "is_final (steps0 (1, l, r) A n1)" and b: "Q1 holds_for (steps0 (1, l, r) A n1)" |
589 using A_halt unfolding Hoare_def by auto |
550 using A_halt unfolding Hoare_def by auto |
590 then obtain l' r' where "steps (1, l, r) (A, 0) n1 = (0, l', r')" and c: "Q1 holds_for (0, l', r')" |
551 then obtain l' r' where "steps0 (1, l, r) A n1 = (0, l', r')" and c: "Q1 holds_for (0, l', r')" |
591 by(case_tac "steps (1, l, r) (A, 0) n1", auto) |
552 by(case_tac "steps0 (1, l, r) A n1", auto) |
592 then obtain stpa where d: "steps (1, l, r) (A |+| B, 0) stpa = (Suc (length A div 2), l', r')" |
553 then obtain stpa where d: "steps0 (1, l, r) (A |+| B) stpa = (Suc (length A div 2), l', r')" |
593 using A_wf |
554 using A_wf |
594 by(rule_tac t_merge_pre_halt_same, auto) |
555 by(rule_tac t_merge_pre_halt_same, auto) |
595 from c aimpb have "P2 holds_for (0, l', r')" |
556 from c aimpb have "P2 holds_for (0, l', r')" |
596 by(rule holds_for_imp) |
557 by(rule holds_for_imp) |
597 from this have "P2 (l', r')" by auto |
558 from this have "P2 (l', r')" by auto |
598 from this have e: "\<forall> n. \<not> is_final (steps (Suc 0, l', r') (B, 0) n) " |
559 from this have e: "\<forall> n. \<not> is_final (steps0 (Suc 0, l', r') B n) " |
599 using B_uhalt unfolding Hoare_unhalt_def |
560 using B_uhalt unfolding Hoare_unhalt_def |
600 by auto |
561 by auto |
601 from e show "\<forall>n. \<not> is_final (steps (1, l, r) (A |+| B, 0) n)" |
562 from e show "\<forall>n. \<not> is_final (steps0 (1, l, r) (A |+| B) n)" |
602 proof(rule_tac allI, case_tac "n > stpa") |
563 proof(rule_tac allI, case_tac "n > stpa") |
603 fix n |
564 fix n |
604 assume h2: "stpa < n" |
565 assume h2: "stpa < n" |
605 hence "\<not> is_final (steps (Suc 0, l', r') (B, 0) (n - stpa))" |
566 hence "\<not> is_final (steps0 (Suc 0, l', r') B (n - stpa))" |
606 using e |
567 using e |
607 apply(erule_tac x = "n - stpa" in allE) by simp |
568 apply(erule_tac x = "n - stpa" in allE) by simp |
608 then obtain s'' l'' r'' where f: "steps (Suc 0, l', r') (B, 0) (n - stpa) = (s'', l'', r'')" and g: "s'' \<noteq> 0" |
569 then obtain s'' l'' r'' where f: "steps0 (Suc 0, l', r') B (n - stpa) = (s'', l'', r'')" and g: "s'' \<noteq> 0" |
609 apply(case_tac "steps (Suc 0, l', r') (B, 0) (n - stpa)", auto) |
570 apply(case_tac "steps0 (Suc 0, l', r') B (n - stpa)", auto) |
610 done |
571 done |
611 have k: "steps (Suc (length A div 2), l', r') (A |+| B, 0) (n - stpa) = (s''+ length A div 2, l'', r'') " |
572 have k: "steps0 (Suc (length A div 2), l', r') (A |+| B) (n - stpa) = (s''+ length A div 2, l'', r'') " |
612 using A_wf B_wf f g |
573 using A_wf B_wf f g |
613 apply(drule_tac t_merge_second_same, auto) |
574 apply(drule_tac t_merge_second_same, auto) |
614 done |
575 done |
615 show "\<not> is_final (steps (1, l, r) (A |+| B, 0) n)" |
576 show "\<not> is_final (steps0 (1, l, r) (A |+| B) n)" |
616 proof - |
577 proof - |
617 have "\<not> is_final (steps (1, l, r) (A |+| B, 0) (stpa + (n - stpa)))" |
578 have "\<not> is_final (steps0 (1, l, r) (A |+| B) (stpa + (n - stpa)))" |
618 using d k A_wf |
579 using d k A_wf |
619 apply(simp only: steps_add d, simp add: tm_wf.simps) |
580 apply(simp only: steps_add d, simp add: tm_wf.simps) |
620 done |
581 done |
621 thus "\<not> is_final (steps (1, l, r) (A |+| B, 0) n)" |
582 thus "\<not> is_final (steps0 (1, l, r) (A |+| B) n)" |
622 using h2 by simp |
583 using h2 by simp |
623 qed |
584 qed |
624 next |
585 next |
625 fix n |
586 fix n |
626 assume h2: "\<not> stpa < n" |
587 assume h2: "\<not> stpa < n" |
627 with d show "\<not> is_final (steps (1, l, r) (A |+| B, 0) n)" |
588 with d show "\<not> is_final (steps0 (1, l, r) (A |+| B) n)" |
628 apply(auto) |
589 apply(auto) |
629 apply(subgoal_tac "\<exists> d. stpa = n + d", auto) |
590 apply(subgoal_tac "\<exists> d. stpa = n + d", auto) |
630 apply(case_tac "(steps (Suc 0, l, r) (A |+| B, 0) n)", simp) |
591 apply(case_tac "(steps0 (Suc 0, l, r) (A |+| B) n)", simp) |
631 apply(rule_tac x = "stpa - n" in exI, simp) |
592 apply(rule_tac x = "stpa - n" in exI, simp) |
632 done |
593 done |
633 qed |
594 qed |
634 qed |
595 qed |
635 |
596 |