388 mkeps :: "rexp \<Rightarrow> val" |
388 mkeps :: "rexp \<Rightarrow> val" |
389 where |
389 where |
390 "mkeps(EMPTY) = Void" |
390 "mkeps(EMPTY) = Void" |
391 | "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)" |
391 | "mkeps(SEQ r1 r2) = Seq (mkeps r1) (mkeps r2)" |
392 | "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))" |
392 | "mkeps(ALT r1 r2) = (if nullable(r1) then Left (mkeps r1) else Right (mkeps r2))" |
393 |
|
394 lemma mkeps_nullable: |
|
395 assumes "nullable(r)" |
|
396 shows "\<turnstile> mkeps r : r" |
|
397 using assms |
|
398 apply(induct rule: nullable.induct) |
|
399 apply(auto intro: Prf.intros) |
|
400 done |
|
401 |
|
402 lemma mkeps_flat: |
|
403 assumes "nullable(r)" |
|
404 shows "flat (mkeps r) = []" |
|
405 using assms |
|
406 apply(induct rule: nullable.induct) |
|
407 apply(auto) |
|
408 done |
|
409 |
|
410 text {* |
|
411 The value mkeps returns is always the correct POSIX |
|
412 value. |
|
413 *} |
|
414 |
|
415 section {* Sulzmann's Ordering of values *} |
|
416 |
|
417 inductive ValOrd :: "val \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<succ>_ _" [100, 100, 100] 100) |
|
418 where |
|
419 "v2 \<succ>r2 v2' \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1 v2')" |
|
420 | "\<lbrakk>v1 \<succ>r1 v1'; v1 \<noteq> v1'\<rbrakk> \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" |
|
421 | "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Right v2)" |
|
422 | "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Left v1)" |
|
423 | "v2 \<succ>r2 v2' \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Right v2')" |
|
424 | "v1 \<succ>r1 v1' \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Left v1')" |
|
425 | "Void \<succ>EMPTY Void" |
|
426 | "(Char c) \<succ>(CHAR c) (Char c)" |
|
427 |
|
428 inductive ValOrd2 :: "val \<Rightarrow> val \<Rightarrow> bool" ("_ 2\<succ> _" [100, 100] 100) |
|
429 where |
|
430 "v2 2\<succ> v2' \<Longrightarrow> (Seq v1 v2) 2\<succ> (Seq v1 v2')" |
|
431 | "\<lbrakk>v1 2\<succ> v1'; v1 \<noteq> v1'\<rbrakk> \<Longrightarrow> (Seq v1 v2) 2\<succ> (Seq v1' v2')" |
|
432 | "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) 2\<succ> (Right v2)" |
|
433 | "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) 2\<succ> (Left v1)" |
|
434 | "v2 2\<succ> v2' \<Longrightarrow> (Right v2) 2\<succ> (Right v2')" |
|
435 | "v1 2\<succ> v1' \<Longrightarrow> (Left v1) 2\<succ> (Left v1')" |
|
436 | "Void 2\<succ> Void" |
|
437 | "(Char c) 2\<succ> (Char c)" |
|
438 |
|
439 lemma Ord1: |
|
440 "v1 \<succ>r v2 \<Longrightarrow> v1 2\<succ> v2" |
|
441 apply(induct rule: ValOrd.induct) |
|
442 apply(auto intro: ValOrd2.intros) |
|
443 done |
|
444 |
|
445 lemma Ord2: |
|
446 "v1 2\<succ> v2 \<Longrightarrow> \<exists>r. v1 \<succ>r v2" |
|
447 apply(induct v1 v2 rule: ValOrd2.induct) |
|
448 apply(auto intro: ValOrd.intros) |
|
449 done |
|
450 |
|
451 lemma Ord3: |
|
452 "\<lbrakk>v1 2\<succ> v2; \<turnstile> v1 : r\<rbrakk> \<Longrightarrow> v1 \<succ>r v2" |
|
453 apply(induct v1 v2 arbitrary: r rule: ValOrd2.induct) |
|
454 apply(auto intro: ValOrd.intros elim: Prf.cases) |
|
455 done |
|
456 |
|
457 lemma ValOrd_refl: |
|
458 assumes "\<turnstile> v : r" |
|
459 shows "v \<succ>r v" |
|
460 using assms |
|
461 apply(induct) |
|
462 apply(auto intro: ValOrd.intros) |
|
463 done |
|
464 |
|
465 lemma ValOrd_total: |
|
466 shows "\<lbrakk>\<turnstile> v1 : r; \<turnstile> v2 : r\<rbrakk> \<Longrightarrow> v1 \<succ>r v2 \<or> v2 \<succ>r v1" |
|
467 apply(induct r arbitrary: v1 v2) |
|
468 apply(auto) |
|
469 apply(erule Prf.cases) |
|
470 apply(simp_all)[5] |
|
471 apply(erule Prf.cases) |
|
472 apply(simp_all)[5] |
|
473 apply(erule Prf.cases) |
|
474 apply(simp_all)[5] |
|
475 apply (metis ValOrd.intros(7)) |
|
476 apply(erule Prf.cases) |
|
477 apply(simp_all)[5] |
|
478 apply(erule Prf.cases) |
|
479 apply(simp_all)[5] |
|
480 apply (metis ValOrd.intros(8)) |
|
481 apply(erule Prf.cases) |
|
482 apply(simp_all)[5] |
|
483 apply(erule Prf.cases) |
|
484 apply(simp_all)[5] |
|
485 apply(clarify) |
|
486 apply(case_tac "v1a = v1b") |
|
487 apply(simp) |
|
488 apply(rule ValOrd.intros(1)) |
|
489 apply (metis ValOrd.intros(1)) |
|
490 apply(rule ValOrd.intros(2)) |
|
491 apply(auto)[2] |
|
492 apply(erule contrapos_np) |
|
493 apply(rule ValOrd.intros(2)) |
|
494 apply(auto) |
|
495 apply(erule Prf.cases) |
|
496 apply(simp_all)[5] |
|
497 apply(erule Prf.cases) |
|
498 apply(simp_all)[5] |
|
499 apply (metis Ord1 Ord3 Prf.intros(2) ValOrd2.intros(6)) |
|
500 apply(rule ValOrd.intros) |
|
501 apply(erule contrapos_np) |
|
502 apply(rule ValOrd.intros) |
|
503 apply (metis le_eq_less_or_eq neq_iff) |
|
504 apply(erule Prf.cases) |
|
505 apply(simp_all)[5] |
|
506 apply(rule ValOrd.intros) |
|
507 apply(erule contrapos_np) |
|
508 apply(rule ValOrd.intros) |
|
509 apply (metis le_eq_less_or_eq neq_iff) |
|
510 apply(rule ValOrd.intros) |
|
511 apply(erule contrapos_np) |
|
512 apply(rule ValOrd.intros) |
|
513 by metis |
|
514 |
|
515 lemma ValOrd_anti: |
|
516 shows "\<lbrakk>\<turnstile> v1 : r; \<turnstile> v2 : r; v1 \<succ>r v2; v2 \<succ>r v1\<rbrakk> \<Longrightarrow> v1 = v2" |
|
517 apply(induct r arbitrary: v1 v2) |
|
518 apply(erule Prf.cases) |
|
519 apply(simp_all)[5] |
|
520 apply(erule Prf.cases) |
|
521 apply(simp_all)[5] |
|
522 apply(erule Prf.cases) |
|
523 apply(simp_all)[5] |
|
524 apply(erule Prf.cases) |
|
525 apply(simp_all)[5] |
|
526 apply(erule Prf.cases) |
|
527 apply(simp_all)[5] |
|
528 apply(erule Prf.cases) |
|
529 apply(simp_all)[5] |
|
530 apply(erule Prf.cases) |
|
531 apply(simp_all)[5] |
|
532 apply(erule ValOrd.cases) |
|
533 apply(simp_all)[8] |
|
534 apply(erule ValOrd.cases) |
|
535 apply(simp_all)[8] |
|
536 apply(erule ValOrd.cases) |
|
537 apply(simp_all)[8] |
|
538 apply(erule Prf.cases) |
|
539 apply(simp_all)[5] |
|
540 apply(erule Prf.cases) |
|
541 apply(simp_all)[5] |
|
542 apply(erule ValOrd.cases) |
|
543 apply(simp_all)[8] |
|
544 apply(erule ValOrd.cases) |
|
545 apply(simp_all)[8] |
|
546 apply(erule ValOrd.cases) |
|
547 apply(simp_all)[8] |
|
548 apply(erule ValOrd.cases) |
|
549 apply(simp_all)[8] |
|
550 apply(erule Prf.cases) |
|
551 apply(simp_all)[5] |
|
552 apply(erule ValOrd.cases) |
|
553 apply(simp_all)[8] |
|
554 apply(erule ValOrd.cases) |
|
555 apply(simp_all)[8] |
|
556 apply(erule ValOrd.cases) |
|
557 apply(simp_all)[8] |
|
558 apply(erule ValOrd.cases) |
|
559 apply(simp_all)[8] |
|
560 done |
|
561 |
|
562 lemma refl_on_ValOrd: |
|
563 "refl_on (Values r s) {(v1, v2). v1 \<succ>r v2 \<and> v1 \<in> Values r s \<and> v2 \<in> Values r s}" |
|
564 unfolding refl_on_def |
|
565 apply(auto) |
|
566 apply(rule ValOrd_refl) |
|
567 apply(simp add: Values_def) |
|
568 done |
|
569 |
|
570 |
|
571 section {* Posix definition *} |
|
572 |
|
573 definition POSIX :: "val \<Rightarrow> rexp \<Rightarrow> bool" |
|
574 where |
|
575 "POSIX v r \<equiv> (\<turnstile> v : r \<and> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v \<succ>r v'))" |
|
576 |
|
577 definition POSIX2 :: "val \<Rightarrow> rexp \<Rightarrow> bool" |
|
578 where |
|
579 "POSIX2 v r \<equiv> (\<turnstile> v : r \<and> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v 2\<succ> v'))" |
|
580 |
|
581 lemma "POSIX v r = POSIX2 v r" |
|
582 unfolding POSIX_def POSIX2_def |
|
583 apply(auto) |
|
584 apply(rule Ord1) |
|
585 apply(auto) |
|
586 apply(rule Ord3) |
|
587 apply(auto) |
|
588 done |
|
589 |
|
590 section {* POSIX for some constructors *} |
|
591 |
|
592 lemma POSIX_SEQ1: |
|
593 assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2" |
|
594 shows "POSIX v1 r1" |
|
595 using assms |
|
596 unfolding POSIX_def |
|
597 apply(auto) |
|
598 apply(drule_tac x="Seq v' v2" in spec) |
|
599 apply(simp) |
|
600 apply(erule impE) |
|
601 apply(rule Prf.intros) |
|
602 apply(simp) |
|
603 apply(simp) |
|
604 apply(erule ValOrd.cases) |
|
605 apply(simp_all) |
|
606 apply(clarify) |
|
607 by (metis ValOrd_refl) |
|
608 |
|
609 lemma POSIX_SEQ2: |
|
610 assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2" |
|
611 shows "POSIX v2 r2" |
|
612 using assms |
|
613 unfolding POSIX_def |
|
614 apply(auto) |
|
615 apply(drule_tac x="Seq v1 v'" in spec) |
|
616 apply(simp) |
|
617 apply(erule impE) |
|
618 apply(rule Prf.intros) |
|
619 apply(simp) |
|
620 apply(simp) |
|
621 apply(erule ValOrd.cases) |
|
622 apply(simp_all) |
|
623 done |
|
624 |
|
625 lemma POSIX_ALT2: |
|
626 assumes "POSIX (Left v1) (ALT r1 r2)" |
|
627 shows "POSIX v1 r1" |
|
628 using assms |
|
629 unfolding POSIX_def |
|
630 apply(auto) |
|
631 apply(erule Prf.cases) |
|
632 apply(simp_all)[5] |
|
633 apply(drule_tac x="Left v'" in spec) |
|
634 apply(simp) |
|
635 apply(drule mp) |
|
636 apply(rule Prf.intros) |
|
637 apply(auto) |
|
638 apply(erule ValOrd.cases) |
|
639 apply(simp_all) |
|
640 done |
|
641 |
|
642 lemma POSIX_ALT1a: |
|
643 assumes "POSIX (Right v2) (ALT r1 r2)" |
|
644 shows "POSIX v2 r2" |
|
645 using assms |
|
646 unfolding POSIX_def |
|
647 apply(auto) |
|
648 apply(erule Prf.cases) |
|
649 apply(simp_all)[5] |
|
650 apply(drule_tac x="Right v'" in spec) |
|
651 apply(simp) |
|
652 apply(drule mp) |
|
653 apply(rule Prf.intros) |
|
654 apply(auto) |
|
655 apply(erule ValOrd.cases) |
|
656 apply(simp_all) |
|
657 done |
|
658 |
|
659 lemma POSIX_ALT1b: |
|
660 assumes "POSIX (Right v2) (ALT r1 r2)" |
|
661 shows "(\<forall>v'. (\<turnstile> v' : r2 \<and> flat v' = flat v2) \<longrightarrow> v2 \<succ>r2 v')" |
|
662 using assms |
|
663 apply(drule_tac POSIX_ALT1a) |
|
664 unfolding POSIX_def |
|
665 apply(auto) |
|
666 done |
|
667 |
|
668 lemma POSIX_ALT_I1: |
|
669 assumes "POSIX v1 r1" |
|
670 shows "POSIX (Left v1) (ALT r1 r2)" |
|
671 using assms |
|
672 unfolding POSIX_def |
|
673 apply(auto) |
|
674 apply (metis Prf.intros(2)) |
|
675 apply(rotate_tac 2) |
|
676 apply(erule Prf.cases) |
|
677 apply(simp_all)[5] |
|
678 apply(auto) |
|
679 apply(rule ValOrd.intros) |
|
680 apply(auto) |
|
681 apply(rule ValOrd.intros) |
|
682 by simp |
|
683 |
|
684 lemma POSIX_ALT_I2: |
|
685 assumes "POSIX v2 r2" "\<forall>v'. \<turnstile> v' : r1 \<longrightarrow> length (flat v2) > length (flat v')" |
|
686 shows "POSIX (Right v2) (ALT r1 r2)" |
|
687 using assms |
|
688 unfolding POSIX_def |
|
689 apply(auto) |
|
690 apply (metis Prf.intros) |
|
691 apply(rotate_tac 3) |
|
692 apply(erule Prf.cases) |
|
693 apply(simp_all)[5] |
|
694 apply(auto) |
|
695 apply(rule ValOrd.intros) |
|
696 apply metis |
|
697 done |
|
698 |
|
699 lemma mkeps_POSIX: |
|
700 assumes "nullable r" |
|
701 shows "POSIX (mkeps r) r" |
|
702 using assms |
|
703 apply(induct r) |
|
704 apply(auto)[1] |
|
705 apply(simp add: POSIX_def) |
|
706 apply(auto)[1] |
|
707 apply (metis Prf.intros(4)) |
|
708 apply(erule Prf.cases) |
|
709 apply(simp_all)[5] |
|
710 apply (metis ValOrd.intros) |
|
711 apply(simp) |
|
712 apply(auto)[1] |
|
713 apply(simp add: POSIX_def) |
|
714 apply(auto)[1] |
|
715 apply (metis mkeps.simps(2) mkeps_nullable nullable.simps(5)) |
|
716 apply(rotate_tac 6) |
|
717 apply(erule Prf.cases) |
|
718 apply(simp_all)[5] |
|
719 apply (simp add: mkeps_flat) |
|
720 apply(case_tac "mkeps r1a = v1") |
|
721 apply(simp) |
|
722 apply (metis ValOrd.intros(1)) |
|
723 apply (rule ValOrd.intros(2)) |
|
724 apply metis |
|
725 apply(simp) |
|
726 (* ALT case *) |
|
727 thm mkeps.simps |
|
728 apply(simp) |
|
729 apply(erule disjE) |
|
730 apply(simp) |
|
731 apply (metis POSIX_ALT_I1) |
|
732 (* *) |
|
733 apply(auto)[1] |
|
734 thm POSIX_ALT_I1 |
|
735 apply (metis POSIX_ALT_I1) |
|
736 apply(simp (no_asm) add: POSIX_def) |
|
737 apply(auto)[1] |
|
738 apply(rule Prf.intros(3)) |
|
739 apply(simp only: POSIX_def) |
|
740 apply(rotate_tac 4) |
|
741 apply(erule Prf.cases) |
|
742 apply(simp_all)[5] |
|
743 thm mkeps_flat |
|
744 apply(simp add: mkeps_flat) |
|
745 apply(auto)[1] |
|
746 thm Prf_flat_L nullable_correctness |
|
747 apply (metis Prf_flat_L nullable_correctness) |
|
748 apply(rule ValOrd.intros) |
|
749 apply(subst (asm) POSIX_def) |
|
750 apply(clarify) |
|
751 apply(drule_tac x="v2" in spec) |
|
752 by simp |
|
753 |
|
754 |
393 |
755 section {* Derivatives *} |
394 section {* Derivatives *} |
756 |
395 |
757 fun |
396 fun |
758 der :: "char \<Rightarrow> rexp \<Rightarrow> rexp" |
397 der :: "char \<Rightarrow> rexp \<Rightarrow> rexp" |
875 apply(rule Prf.intros) |
528 apply(rule Prf.intros) |
876 apply (metis Cons_eq_append_conv) |
529 apply (metis Cons_eq_append_conv) |
877 apply(simp) |
530 apply(simp) |
878 done |
531 done |
879 |
532 |
|
533 lemma v4: |
|
534 assumes "\<turnstile> v : der c r" |
|
535 shows "flat (injval r c v) = c # (flat v)" |
|
536 using assms |
|
537 apply(induct arbitrary: v rule: der.induct) |
|
538 apply(simp) |
|
539 apply(erule Prf.cases) |
|
540 apply(simp_all)[5] |
|
541 apply(simp) |
|
542 apply(erule Prf.cases) |
|
543 apply(simp_all)[5] |
|
544 apply(simp) |
|
545 apply(case_tac "c = c'") |
|
546 apply(simp) |
|
547 apply(auto)[1] |
|
548 apply(erule Prf.cases) |
|
549 apply(simp_all)[5] |
|
550 apply(simp) |
|
551 apply(erule Prf.cases) |
|
552 apply(simp_all)[5] |
|
553 apply(simp) |
|
554 apply(erule Prf.cases) |
|
555 apply(simp_all)[5] |
|
556 apply(simp) |
|
557 apply(case_tac "nullable r1") |
|
558 apply(simp) |
|
559 apply(erule Prf.cases) |
|
560 apply(simp_all (no_asm_use))[5] |
|
561 apply(auto)[1] |
|
562 apply(erule Prf.cases) |
|
563 apply(simp_all)[5] |
|
564 apply(clarify) |
|
565 apply(simp only: injval.simps flat.simps) |
|
566 apply(auto)[1] |
|
567 apply (metis mkeps_flat) |
|
568 apply(simp) |
|
569 apply(erule Prf.cases) |
|
570 apply(simp_all)[5] |
|
571 done |
|
572 |
|
573 lemma v4_proj: |
|
574 assumes "\<turnstile> v : r" and "\<exists>s. (flat v) = c # s" |
|
575 shows "c # flat (projval r c v) = flat v" |
|
576 using assms |
|
577 apply(induct rule: Prf.induct) |
|
578 prefer 4 |
|
579 apply(simp) |
|
580 prefer 4 |
|
581 apply(simp) |
|
582 prefer 2 |
|
583 apply(simp) |
|
584 prefer 2 |
|
585 apply(simp) |
|
586 apply(auto) |
|
587 by (metis Cons_eq_append_conv) |
|
588 |
|
589 lemma v4_proj2: |
|
590 assumes "\<turnstile> v : r" and "(flat v) = c # s" |
|
591 shows "flat (projval r c v) = s" |
|
592 using assms |
|
593 by (metis list.inject v4_proj) |
|
594 |
|
595 |
|
596 section {* Alternative Posix definition *} |
|
597 |
|
598 inductive |
|
599 PMatch :: "string \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<in> _ \<rightarrow> _" [100, 100, 100] 100) |
|
600 where |
|
601 "[] \<in> EMPTY \<rightarrow> Void" |
|
602 | "[c] \<in> (CHAR c) \<rightarrow> (Char c)" |
|
603 | "s \<in> r1 \<rightarrow> v \<Longrightarrow> s \<in> (ALT r1 r2) \<rightarrow> (Left v)" |
|
604 | "\<lbrakk>s \<in> r2 \<rightarrow> v; s \<notin> L(r1)\<rbrakk> \<Longrightarrow> s \<in> (ALT r1 r2) \<rightarrow> (Right v)" |
|
605 | "\<lbrakk>s1 \<in> r1 \<rightarrow> v1; s2 \<in> r2 \<rightarrow> v2; |
|
606 \<not>(\<exists>s3 s4. s3 \<noteq> [] \<and> s3 @ s4 = s2 \<and> (s1 @ s2) \<in> L r1 \<and> s4 \<in> L r2)\<rbrakk> \<Longrightarrow> |
|
607 (s1 @ s2) \<in> (SEQ r1 r2) \<rightarrow> (Seq v1 v2)" |
|
608 |
|
609 lemma PMatch_mkeps: |
|
610 assumes "nullable r" |
|
611 shows " [] \<in> r \<rightarrow> mkeps r" |
|
612 using assms |
|
613 apply(induct r) |
|
614 apply(auto) |
|
615 apply (metis PMatch.intros(1)) |
|
616 apply(subst append.simps(1)[symmetric]) |
|
617 apply (rule PMatch.intros) |
|
618 apply(simp) |
|
619 apply(simp) |
|
620 apply(auto)[1] |
|
621 apply (rule PMatch.intros) |
|
622 apply(simp) |
|
623 apply (rule PMatch.intros) |
|
624 apply(simp) |
|
625 apply (rule PMatch.intros) |
|
626 apply(simp) |
|
627 by (metis nullable_correctness) |
|
628 |
|
629 |
|
630 lemma PMatch1: |
|
631 assumes "s \<in> r \<rightarrow> v" |
|
632 shows "\<turnstile> v : r" "flat v = s" |
|
633 using assms |
|
634 apply(induct s r v rule: PMatch.induct) |
|
635 apply(auto) |
|
636 apply (metis Prf.intros(4)) |
|
637 apply (metis Prf.intros(5)) |
|
638 apply (metis Prf.intros(2)) |
|
639 apply (metis Prf.intros(3)) |
|
640 apply (metis Prf.intros(1)) |
|
641 by (metis Prf.intros(1)) |
|
642 |
|
643 lemma PMAtch2: |
|
644 assumes "s \<in> (der c r) \<rightarrow> v" |
|
645 shows "(c#s) \<in> r \<rightarrow> (injval r c v)" |
|
646 using assms |
|
647 apply(induct c r arbitrary: s v rule: der.induct) |
|
648 apply(auto) |
|
649 apply(erule PMatch.cases) |
|
650 apply(simp_all)[5] |
|
651 apply(erule PMatch.cases) |
|
652 apply(simp_all)[5] |
|
653 apply(case_tac "c = c'") |
|
654 apply(simp) |
|
655 apply(erule PMatch.cases) |
|
656 apply(simp_all)[5] |
|
657 apply (metis PMatch.intros(2)) |
|
658 apply(simp) |
|
659 apply(erule PMatch.cases) |
|
660 apply(simp_all)[5] |
|
661 apply(erule PMatch.cases) |
|
662 apply(simp_all)[5] |
|
663 apply (metis PMatch.intros(3)) |
|
664 apply(clarify) |
|
665 apply(rule PMatch.intros) |
|
666 apply metis |
|
667 apply(simp add: L_flat_Prf) |
|
668 apply(auto)[1] |
|
669 thm v3_proj |
|
670 apply(frule_tac c="c" in v3_proj) |
|
671 apply metis |
|
672 apply(drule_tac x="projval r1 c v" in spec) |
|
673 apply(drule mp) |
|
674 apply (metis v4_proj2) |
|
675 apply(simp) |
|
676 apply(case_tac "nullable r1") |
|
677 apply(simp) |
|
678 defer |
|
679 apply(simp) |
|
680 apply(erule PMatch.cases) |
|
681 apply(simp_all)[5] |
|
682 apply(clarify) |
|
683 apply(subst append.simps(2)[symmetric]) |
|
684 apply(rule PMatch.intros) |
|
685 apply metis |
|
686 apply metis |
|
687 apply(auto)[1] |
|
688 apply(simp add: L_flat_Prf) |
|
689 apply(auto)[1] |
|
690 apply(frule_tac c="c" in v3_proj) |
|
691 apply metis |
|
692 apply(drule_tac x="projval r1 c v" in spec) |
|
693 apply(drule mp) |
|
694 apply (metis v4_proj2) |
|
695 apply(simp) |
|
696 (* nullable case *) |
|
697 apply(erule PMatch.cases) |
|
698 apply(simp_all)[5] |
|
699 apply(clarify) |
|
700 apply(erule PMatch.cases) |
|
701 apply(simp_all)[5] |
|
702 apply(clarify) |
|
703 apply(subst append.simps(2)[symmetric]) |
|
704 apply(rule PMatch.intros) |
|
705 apply metis |
|
706 apply metis |
|
707 apply(auto)[1] |
|
708 apply(simp add: L_flat_Prf) |
|
709 apply(auto)[1] |
|
710 apply(frule_tac c="c" in v3_proj) |
|
711 apply metis |
|
712 apply(drule_tac x="projval r1 c v" in spec) |
|
713 apply(drule mp) |
|
714 apply (metis v4_proj2) |
|
715 apply(simp) |
|
716 (* interesting case *) |
|
717 apply(clarify) |
|
718 apply(simp) |
|
719 thm L.simps |
|
720 apply(subst (asm) L.simps(4)[symmetric]) |
|
721 apply(simp only: L_flat_Prf) |
|
722 apply(simp) |
|
723 apply(subst append.simps(1)[symmetric]) |
|
724 apply(rule PMatch.intros) |
|
725 apply (metis PMatch_mkeps) |
|
726 apply metis |
|
727 apply(auto) |
|
728 apply(simp only: L_flat_Prf) |
|
729 apply(simp) |
|
730 apply(auto) |
|
731 apply(drule_tac x="projval r1 c v" in spec) |
|
732 apply(drule mp) |
|
733 apply (metis v4_proj2) |
|
734 apply(rotate_tac 1) |
|
735 apply(drule_tac x="sa" in meta_spec) |
|
736 apply(drule_tac x="va" in meta_spec) |
|
737 apply(simp) |
|
738 apply(rotate_tac 2) |
|
739 apply(drule_tac x="sa" in meta_spec) |
|
740 apply(drule_tac x="projval r1 c v" in meta_spec) |
|
741 apply(drule meta_mp) |
|
742 apply(frule_tac c="c" in v3_proj) |
|
743 apply metis |
|
744 |
|
745 apply(frule PMatch1(1)) |
|
746 apply(drule PMatch1(2)) |
|
747 apply(clarify) |
|
748 apply(subst (asm) (2) not_def) |
|
749 apply(drule mp) |
|
750 thm v3_proj |
|
751 apply(rule v3_proj) |
|
752 apply(rule conjI) |
|
753 apply (metis v4_proj2) |
|
754 apply (metis v3_proj) |
|
755 v3_proj |
|
756 |
|
757 |
|
758 |
|
759 |
|
760 |
|
761 section {* Sulzmann's Ordering of values *} |
|
762 |
|
763 inductive ValOrd :: "val \<Rightarrow> rexp \<Rightarrow> val \<Rightarrow> bool" ("_ \<succ>_ _" [100, 100, 100] 100) |
|
764 where |
|
765 "v2 \<succ>r2 v2' \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1 v2')" |
|
766 | "\<lbrakk>v1 \<succ>r1 v1'; v1 \<noteq> v1'\<rbrakk> \<Longrightarrow> (Seq v1 v2) \<succ>(SEQ r1 r2) (Seq v1' v2')" |
|
767 | "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Right v2)" |
|
768 | "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Left v1)" |
|
769 | "v2 \<succ>r2 v2' \<Longrightarrow> (Right v2) \<succ>(ALT r1 r2) (Right v2')" |
|
770 | "v1 \<succ>r1 v1' \<Longrightarrow> (Left v1) \<succ>(ALT r1 r2) (Left v1')" |
|
771 | "Void \<succ>EMPTY Void" |
|
772 | "(Char c) \<succ>(CHAR c) (Char c)" |
|
773 |
|
774 inductive ValOrd2 :: "val \<Rightarrow> val \<Rightarrow> bool" ("_ 2\<succ> _" [100, 100] 100) |
|
775 where |
|
776 "v2 2\<succ> v2' \<Longrightarrow> (Seq v1 v2) 2\<succ> (Seq v1 v2')" |
|
777 | "\<lbrakk>v1 2\<succ> v1'; v1 \<noteq> v1'\<rbrakk> \<Longrightarrow> (Seq v1 v2) 2\<succ> (Seq v1' v2')" |
|
778 | "length (flat v1) \<ge> length (flat v2) \<Longrightarrow> (Left v1) 2\<succ> (Right v2)" |
|
779 | "length (flat v2) > length (flat v1) \<Longrightarrow> (Right v2) 2\<succ> (Left v1)" |
|
780 | "v2 2\<succ> v2' \<Longrightarrow> (Right v2) 2\<succ> (Right v2')" |
|
781 | "v1 2\<succ> v1' \<Longrightarrow> (Left v1) 2\<succ> (Left v1')" |
|
782 | "Void 2\<succ> Void" |
|
783 | "(Char c) 2\<succ> (Char c)" |
|
784 |
|
785 lemma Ord1: |
|
786 "v1 \<succ>r v2 \<Longrightarrow> v1 2\<succ> v2" |
|
787 apply(induct rule: ValOrd.induct) |
|
788 apply(auto intro: ValOrd2.intros) |
|
789 done |
|
790 |
|
791 lemma Ord2: |
|
792 "v1 2\<succ> v2 \<Longrightarrow> \<exists>r. v1 \<succ>r v2" |
|
793 apply(induct v1 v2 rule: ValOrd2.induct) |
|
794 apply(auto intro: ValOrd.intros) |
|
795 done |
|
796 |
|
797 lemma Ord3: |
|
798 "\<lbrakk>v1 2\<succ> v2; \<turnstile> v1 : r\<rbrakk> \<Longrightarrow> v1 \<succ>r v2" |
|
799 apply(induct v1 v2 arbitrary: r rule: ValOrd2.induct) |
|
800 apply(auto intro: ValOrd.intros elim: Prf.cases) |
|
801 done |
|
802 |
|
803 lemma ValOrd_refl: |
|
804 assumes "\<turnstile> v : r" |
|
805 shows "v \<succ>r v" |
|
806 using assms |
|
807 apply(induct) |
|
808 apply(auto intro: ValOrd.intros) |
|
809 done |
|
810 |
|
811 lemma ValOrd_total: |
|
812 shows "\<lbrakk>\<turnstile> v1 : r; \<turnstile> v2 : r\<rbrakk> \<Longrightarrow> v1 \<succ>r v2 \<or> v2 \<succ>r v1" |
|
813 apply(induct r arbitrary: v1 v2) |
|
814 apply(auto) |
|
815 apply(erule Prf.cases) |
|
816 apply(simp_all)[5] |
|
817 apply(erule Prf.cases) |
|
818 apply(simp_all)[5] |
|
819 apply(erule Prf.cases) |
|
820 apply(simp_all)[5] |
|
821 apply (metis ValOrd.intros(7)) |
|
822 apply(erule Prf.cases) |
|
823 apply(simp_all)[5] |
|
824 apply(erule Prf.cases) |
|
825 apply(simp_all)[5] |
|
826 apply (metis ValOrd.intros(8)) |
|
827 apply(erule Prf.cases) |
|
828 apply(simp_all)[5] |
|
829 apply(erule Prf.cases) |
|
830 apply(simp_all)[5] |
|
831 apply(clarify) |
|
832 apply(case_tac "v1a = v1b") |
|
833 apply(simp) |
|
834 apply(rule ValOrd.intros(1)) |
|
835 apply (metis ValOrd.intros(1)) |
|
836 apply(rule ValOrd.intros(2)) |
|
837 apply(auto)[2] |
|
838 apply(erule contrapos_np) |
|
839 apply(rule ValOrd.intros(2)) |
|
840 apply(auto) |
|
841 apply(erule Prf.cases) |
|
842 apply(simp_all)[5] |
|
843 apply(erule Prf.cases) |
|
844 apply(simp_all)[5] |
|
845 apply (metis Ord1 Ord3 Prf.intros(2) ValOrd2.intros(6)) |
|
846 apply(rule ValOrd.intros) |
|
847 apply(erule contrapos_np) |
|
848 apply(rule ValOrd.intros) |
|
849 apply (metis le_eq_less_or_eq neq_iff) |
|
850 apply(erule Prf.cases) |
|
851 apply(simp_all)[5] |
|
852 apply(rule ValOrd.intros) |
|
853 apply(erule contrapos_np) |
|
854 apply(rule ValOrd.intros) |
|
855 apply (metis le_eq_less_or_eq neq_iff) |
|
856 apply(rule ValOrd.intros) |
|
857 apply(erule contrapos_np) |
|
858 apply(rule ValOrd.intros) |
|
859 by metis |
|
860 |
|
861 lemma ValOrd_anti: |
|
862 shows "\<lbrakk>\<turnstile> v1 : r; \<turnstile> v2 : r; v1 \<succ>r v2; v2 \<succ>r v1\<rbrakk> \<Longrightarrow> v1 = v2" |
|
863 apply(induct r arbitrary: v1 v2) |
|
864 apply(erule Prf.cases) |
|
865 apply(simp_all)[5] |
|
866 apply(erule Prf.cases) |
|
867 apply(simp_all)[5] |
|
868 apply(erule Prf.cases) |
|
869 apply(simp_all)[5] |
|
870 apply(erule Prf.cases) |
|
871 apply(simp_all)[5] |
|
872 apply(erule Prf.cases) |
|
873 apply(simp_all)[5] |
|
874 apply(erule Prf.cases) |
|
875 apply(simp_all)[5] |
|
876 apply(erule Prf.cases) |
|
877 apply(simp_all)[5] |
|
878 apply(erule ValOrd.cases) |
|
879 apply(simp_all)[8] |
|
880 apply(erule ValOrd.cases) |
|
881 apply(simp_all)[8] |
|
882 apply(erule ValOrd.cases) |
|
883 apply(simp_all)[8] |
|
884 apply(erule Prf.cases) |
|
885 apply(simp_all)[5] |
|
886 apply(erule Prf.cases) |
|
887 apply(simp_all)[5] |
|
888 apply(erule ValOrd.cases) |
|
889 apply(simp_all)[8] |
|
890 apply(erule ValOrd.cases) |
|
891 apply(simp_all)[8] |
|
892 apply(erule ValOrd.cases) |
|
893 apply(simp_all)[8] |
|
894 apply(erule ValOrd.cases) |
|
895 apply(simp_all)[8] |
|
896 apply(erule Prf.cases) |
|
897 apply(simp_all)[5] |
|
898 apply(erule ValOrd.cases) |
|
899 apply(simp_all)[8] |
|
900 apply(erule ValOrd.cases) |
|
901 apply(simp_all)[8] |
|
902 apply(erule ValOrd.cases) |
|
903 apply(simp_all)[8] |
|
904 apply(erule ValOrd.cases) |
|
905 apply(simp_all)[8] |
|
906 done |
|
907 |
|
908 lemma refl_on_ValOrd: |
|
909 "refl_on (Values r s) {(v1, v2). v1 \<succ>r v2 \<and> v1 \<in> Values r s \<and> v2 \<in> Values r s}" |
|
910 unfolding refl_on_def |
|
911 apply(auto) |
|
912 apply(rule ValOrd_refl) |
|
913 apply(simp add: Values_def) |
|
914 done |
|
915 |
|
916 |
|
917 section {* Posix definition *} |
|
918 |
|
919 definition POSIX :: "val \<Rightarrow> rexp \<Rightarrow> bool" |
|
920 where |
|
921 "POSIX v r \<equiv> (\<turnstile> v : r \<and> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v \<succ>r v'))" |
|
922 |
|
923 definition POSIX2 :: "val \<Rightarrow> rexp \<Rightarrow> bool" |
|
924 where |
|
925 "POSIX2 v r \<equiv> (\<turnstile> v : r \<and> (\<forall>v'. (\<turnstile> v' : r \<and> flat v = flat v') \<longrightarrow> v 2\<succ> v'))" |
|
926 |
|
927 lemma "POSIX v r = POSIX2 v r" |
|
928 unfolding POSIX_def POSIX2_def |
|
929 apply(auto) |
|
930 apply(rule Ord1) |
|
931 apply(auto) |
|
932 apply(rule Ord3) |
|
933 apply(auto) |
|
934 done |
|
935 |
|
936 section {* POSIX for some constructors *} |
|
937 |
|
938 lemma POSIX_SEQ1: |
|
939 assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2" |
|
940 shows "POSIX v1 r1" |
|
941 using assms |
|
942 unfolding POSIX_def |
|
943 apply(auto) |
|
944 apply(drule_tac x="Seq v' v2" in spec) |
|
945 apply(simp) |
|
946 apply(erule impE) |
|
947 apply(rule Prf.intros) |
|
948 apply(simp) |
|
949 apply(simp) |
|
950 apply(erule ValOrd.cases) |
|
951 apply(simp_all) |
|
952 apply(clarify) |
|
953 by (metis ValOrd_refl) |
|
954 |
|
955 lemma POSIX_SEQ2: |
|
956 assumes "POSIX (Seq v1 v2) (SEQ r1 r2)" "\<turnstile> v1 : r1" "\<turnstile> v2 : r2" |
|
957 shows "POSIX v2 r2" |
|
958 using assms |
|
959 unfolding POSIX_def |
|
960 apply(auto) |
|
961 apply(drule_tac x="Seq v1 v'" in spec) |
|
962 apply(simp) |
|
963 apply(erule impE) |
|
964 apply(rule Prf.intros) |
|
965 apply(simp) |
|
966 apply(simp) |
|
967 apply(erule ValOrd.cases) |
|
968 apply(simp_all) |
|
969 done |
|
970 |
|
971 lemma POSIX_ALT2: |
|
972 assumes "POSIX (Left v1) (ALT r1 r2)" |
|
973 shows "POSIX v1 r1" |
|
974 using assms |
|
975 unfolding POSIX_def |
|
976 apply(auto) |
|
977 apply(erule Prf.cases) |
|
978 apply(simp_all)[5] |
|
979 apply(drule_tac x="Left v'" in spec) |
|
980 apply(simp) |
|
981 apply(drule mp) |
|
982 apply(rule Prf.intros) |
|
983 apply(auto) |
|
984 apply(erule ValOrd.cases) |
|
985 apply(simp_all) |
|
986 done |
|
987 |
|
988 lemma POSIX_ALT1a: |
|
989 assumes "POSIX (Right v2) (ALT r1 r2)" |
|
990 shows "POSIX v2 r2" |
|
991 using assms |
|
992 unfolding POSIX_def |
|
993 apply(auto) |
|
994 apply(erule Prf.cases) |
|
995 apply(simp_all)[5] |
|
996 apply(drule_tac x="Right v'" in spec) |
|
997 apply(simp) |
|
998 apply(drule mp) |
|
999 apply(rule Prf.intros) |
|
1000 apply(auto) |
|
1001 apply(erule ValOrd.cases) |
|
1002 apply(simp_all) |
|
1003 done |
|
1004 |
|
1005 lemma POSIX_ALT1b: |
|
1006 assumes "POSIX (Right v2) (ALT r1 r2)" |
|
1007 shows "(\<forall>v'. (\<turnstile> v' : r2 \<and> flat v' = flat v2) \<longrightarrow> v2 \<succ>r2 v')" |
|
1008 using assms |
|
1009 apply(drule_tac POSIX_ALT1a) |
|
1010 unfolding POSIX_def |
|
1011 apply(auto) |
|
1012 done |
|
1013 |
|
1014 lemma POSIX_ALT_I1: |
|
1015 assumes "POSIX v1 r1" |
|
1016 shows "POSIX (Left v1) (ALT r1 r2)" |
|
1017 using assms |
|
1018 unfolding POSIX_def |
|
1019 apply(auto) |
|
1020 apply (metis Prf.intros(2)) |
|
1021 apply(rotate_tac 2) |
|
1022 apply(erule Prf.cases) |
|
1023 apply(simp_all)[5] |
|
1024 apply(auto) |
|
1025 apply(rule ValOrd.intros) |
|
1026 apply(auto) |
|
1027 apply(rule ValOrd.intros) |
|
1028 by simp |
|
1029 |
|
1030 lemma POSIX_ALT_I2: |
|
1031 assumes "POSIX v2 r2" "\<forall>v'. \<turnstile> v' : r1 \<longrightarrow> length (flat v2) > length (flat v')" |
|
1032 shows "POSIX (Right v2) (ALT r1 r2)" |
|
1033 using assms |
|
1034 unfolding POSIX_def |
|
1035 apply(auto) |
|
1036 apply (metis Prf.intros) |
|
1037 apply(rotate_tac 3) |
|
1038 apply(erule Prf.cases) |
|
1039 apply(simp_all)[5] |
|
1040 apply(auto) |
|
1041 apply(rule ValOrd.intros) |
|
1042 apply metis |
|
1043 done |
|
1044 |
|
1045 lemma mkeps_POSIX: |
|
1046 assumes "nullable r" |
|
1047 shows "POSIX (mkeps r) r" |
|
1048 using assms |
|
1049 apply(induct r) |
|
1050 apply(auto)[1] |
|
1051 apply(simp add: POSIX_def) |
|
1052 apply(auto)[1] |
|
1053 apply (metis Prf.intros(4)) |
|
1054 apply(erule Prf.cases) |
|
1055 apply(simp_all)[5] |
|
1056 apply (metis ValOrd.intros) |
|
1057 apply(simp) |
|
1058 apply(auto)[1] |
|
1059 apply(simp add: POSIX_def) |
|
1060 apply(auto)[1] |
|
1061 apply (metis mkeps.simps(2) mkeps_nullable nullable.simps(5)) |
|
1062 apply(rotate_tac 6) |
|
1063 apply(erule Prf.cases) |
|
1064 apply(simp_all)[5] |
|
1065 apply (simp add: mkeps_flat) |
|
1066 apply(case_tac "mkeps r1a = v1") |
|
1067 apply(simp) |
|
1068 apply (metis ValOrd.intros(1)) |
|
1069 apply (rule ValOrd.intros(2)) |
|
1070 apply metis |
|
1071 apply(simp) |
|
1072 (* ALT case *) |
|
1073 thm mkeps.simps |
|
1074 apply(simp) |
|
1075 apply(erule disjE) |
|
1076 apply(simp) |
|
1077 apply (metis POSIX_ALT_I1) |
|
1078 (* *) |
|
1079 apply(auto)[1] |
|
1080 thm POSIX_ALT_I1 |
|
1081 apply (metis POSIX_ALT_I1) |
|
1082 apply(simp (no_asm) add: POSIX_def) |
|
1083 apply(auto)[1] |
|
1084 apply(rule Prf.intros(3)) |
|
1085 apply(simp only: POSIX_def) |
|
1086 apply(rotate_tac 4) |
|
1087 apply(erule Prf.cases) |
|
1088 apply(simp_all)[5] |
|
1089 thm mkeps_flat |
|
1090 apply(simp add: mkeps_flat) |
|
1091 apply(auto)[1] |
|
1092 thm Prf_flat_L nullable_correctness |
|
1093 apply (metis Prf_flat_L nullable_correctness) |
|
1094 apply(rule ValOrd.intros) |
|
1095 apply(subst (asm) POSIX_def) |
|
1096 apply(clarify) |
|
1097 apply(drule_tac x="v2" in spec) |
|
1098 by simp |
|
1099 |
|
1100 |
|
1101 |
|
1102 text {* |
|
1103 Injection value is related to r |
|
1104 *} |
|
1105 |
|
1106 |
|
1107 |
880 text {* |
1108 text {* |
881 The string behind the injection value is an added c |
1109 The string behind the injection value is an added c |
882 *} |
1110 *} |
883 |
1111 |
884 lemma v4: |
|
885 assumes "\<turnstile> v : der c r" |
|
886 shows "flat (injval r c v) = c # (flat v)" |
|
887 using assms |
|
888 apply(induct arbitrary: v rule: der.induct) |
|
889 apply(simp) |
|
890 apply(erule Prf.cases) |
|
891 apply(simp_all)[5] |
|
892 apply(simp) |
|
893 apply(erule Prf.cases) |
|
894 apply(simp_all)[5] |
|
895 apply(simp) |
|
896 apply(case_tac "c = c'") |
|
897 apply(simp) |
|
898 apply(auto)[1] |
|
899 apply(erule Prf.cases) |
|
900 apply(simp_all)[5] |
|
901 apply(simp) |
|
902 apply(erule Prf.cases) |
|
903 apply(simp_all)[5] |
|
904 apply(simp) |
|
905 apply(erule Prf.cases) |
|
906 apply(simp_all)[5] |
|
907 apply(simp) |
|
908 apply(case_tac "nullable r1") |
|
909 apply(simp) |
|
910 apply(erule Prf.cases) |
|
911 apply(simp_all (no_asm_use))[5] |
|
912 apply(auto)[1] |
|
913 apply(erule Prf.cases) |
|
914 apply(simp_all)[5] |
|
915 apply(clarify) |
|
916 apply(simp only: injval.simps flat.simps) |
|
917 apply(auto)[1] |
|
918 apply (metis mkeps_flat) |
|
919 apply(simp) |
|
920 apply(erule Prf.cases) |
|
921 apply(simp_all)[5] |
|
922 done |
|
923 |
|
924 lemma v4_proj: |
|
925 assumes "\<turnstile> v : r" and "\<exists>s. (flat v) = c # s" |
|
926 shows "c # flat (projval r c v) = flat v" |
|
927 using assms |
|
928 apply(induct rule: Prf.induct) |
|
929 prefer 4 |
|
930 apply(simp) |
|
931 prefer 4 |
|
932 apply(simp) |
|
933 prefer 2 |
|
934 apply(simp) |
|
935 prefer 2 |
|
936 apply(simp) |
|
937 apply(auto) |
|
938 by (metis Cons_eq_append_conv) |
|
939 |
|
940 lemma v4_proj2: |
|
941 assumes "\<turnstile> v : r" and "(flat v) = c # s" |
|
942 shows "flat (projval r c v) = s" |
|
943 using assms |
|
944 by (metis list.inject v4_proj) |
|
945 |
1112 |
946 lemma injval_inj: "inj_on (injval r c) {v. \<turnstile> v : der c r}" |
1113 lemma injval_inj: "inj_on (injval r c) {v. \<turnstile> v : der c r}" |
947 apply(induct c r rule: der.induct) |
1114 apply(induct c r rule: der.induct) |
948 unfolding inj_on_def |
1115 unfolding inj_on_def |
949 apply(auto)[1] |
1116 apply(auto)[1] |