392 shows "finals A \<subseteq> (UNIV // \<approx>A)" |
346 shows "finals A \<subseteq> (UNIV // \<approx>A)" |
393 unfolding finals_def |
347 unfolding finals_def |
394 unfolding quotient_def |
348 unfolding quotient_def |
395 by auto |
349 by auto |
396 |
350 |
397 section {* Direction @{text "finite partition \<Rightarrow> regular language"}*} |
351 |
398 |
352 section {* Equational systems *} |
399 text {* |
|
400 The relationship between equivalent classes can be described by an |
|
401 equational system. For example, in equational system \eqref{example_eqns}, |
|
402 $X_0, X_1$ are equivalent classes. The first equation says every string in |
|
403 $X_0$ is obtained either by appending one $b$ to a string in $X_0$ or by |
|
404 appending one $a$ to a string in $X_1$ or just be an empty string |
|
405 (represented by the regular expression $\lambda$). Similary, the second |
|
406 equation tells how the strings inside $X_1$ are composed. |
|
407 |
|
408 \begin{equation}\label{example_eqns} |
|
409 \begin{aligned} |
|
410 X_0 & = X_0 b + X_1 a + \lambda \\ |
|
411 X_1 & = X_0 a + X_1 b |
|
412 \end{aligned} |
|
413 \end{equation} |
|
414 |
|
415 \noindent |
|
416 The summands on the right hand side is represented by the following data |
|
417 type @{text "rhs_item"}, mnemonic for 'right hand side item'. Generally, |
|
418 there are two kinds of right hand side items, one kind corresponds to pure |
|
419 regular expressions, like the $\lambda$ in \eqref{example_eqns}, the other |
|
420 kind corresponds to transitions from one one equivalent class to another, |
|
421 like the $X_0 b, X_1 a$ etc. |
|
422 |
|
423 *} |
|
424 |
353 |
425 datatype rhs_item = |
354 datatype rhs_item = |
426 Lam "rexp" (* Lambda *) |
355 Lam "rexp" (* Lambda-marker *) |
427 | Trn "lang" "rexp" (* Transition *) |
356 | Trn "lang" "rexp" (* Transition *) |
428 |
357 |
429 |
358 |
430 text {* |
359 overloading L_rhs_item \<equiv> "L:: rhs_item \<Rightarrow> lang" |
431 In this formalization, pure regular expressions like $\lambda$ is |
|
432 repsented by @{text "Lam(EMPTY)"}, while transitions like $X_0 a$ is |
|
433 represented by @{term "Trn X\<^isub>0 (CHAR a)"}. |
|
434 *} |
|
435 |
|
436 text {* |
|
437 Every right-hand side item @{text "itm"} defines a language given |
|
438 by @{text "L(itm)"}, defined as: |
|
439 *} |
|
440 |
|
441 overloading L_rhs_e \<equiv> "L:: rhs_item \<Rightarrow> lang" |
|
442 begin |
360 begin |
443 fun L_rhs_e:: "rhs_item \<Rightarrow> lang" |
361 fun L_rhs_item:: "rhs_item \<Rightarrow> lang" |
444 where |
362 where |
445 "L_rhs_e (Lam r) = L r" |
363 "L_rhs_item (Lam r) = L r" |
446 | "L_rhs_e (Trn X r) = X ;; L r" |
364 | "L_rhs_item (Trn X r) = X ;; L r" |
447 end |
365 end |
448 |
|
449 text {* |
|
450 The right hand side of every equation is represented by a set of |
|
451 items. The string set defined by such a set @{text "itms"} is given |
|
452 by @{text "L(itms)"}, defined as: |
|
453 *} |
|
454 |
366 |
455 overloading L_rhs \<equiv> "L:: rhs_item set \<Rightarrow> lang" |
367 overloading L_rhs \<equiv> "L:: rhs_item set \<Rightarrow> lang" |
456 begin |
368 begin |
457 fun L_rhs:: "rhs_item set \<Rightarrow> lang" |
369 fun L_rhs:: "rhs_item set \<Rightarrow> lang" |
458 where |
370 where |
459 "L_rhs rhs = \<Union> (L ` rhs)" |
371 "L_rhs rhs = \<Union> (L ` rhs)" |
460 end |
372 end |
461 |
373 |
462 text {* |
374 definition |
463 Given a set of equivalence classes @{text "CS"} and one equivalence class @{text "X"} among |
375 "trns_of rhs X \<equiv> {Trn X r | r. Trn X r \<in> rhs}" |
464 @{text "CS"}, the term @{text "init_rhs CS X"} is used to extract the right hand side of |
376 |
465 the equation describing the formation of @{text "X"}. The definition of @{text "init_rhs"} |
377 text {* Transitions between equivalence classes *} |
466 is: |
|
467 *} |
|
468 |
378 |
469 definition |
379 definition |
470 transition :: "lang \<Rightarrow> rexp \<Rightarrow> lang \<Rightarrow> bool" ("_ \<Turnstile>_\<Rightarrow>_" [100,100,100] 100) |
380 transition :: "lang \<Rightarrow> rexp \<Rightarrow> lang \<Rightarrow> bool" ("_ \<Turnstile>_\<Rightarrow>_" [100,100,100] 100) |
471 where |
381 where |
472 "Y \<Turnstile>r\<Rightarrow> X \<equiv> Y ;; (L r) \<subseteq> X" |
382 "Y \<Turnstile>r\<Rightarrow> X \<equiv> Y ;; (L r) \<subseteq> X" |
|
383 |
|
384 text {* Initial equational system *} |
473 |
385 |
474 definition |
386 definition |
475 "init_rhs CS X \<equiv> |
387 "init_rhs CS X \<equiv> |
476 if ([] \<in> X) then |
388 if ([] \<in> X) then |
477 {Lam EMPTY} \<union> {Trn Y (CHAR c) | Y c. Y \<in> CS \<and> Y \<Turnstile>(CHAR c)\<Rightarrow> X} |
389 {Lam EMPTY} \<union> {Trn Y (CHAR c) | Y c. Y \<in> CS \<and> Y \<Turnstile>(CHAR c)\<Rightarrow> X} |
478 else |
390 else |
479 {Trn Y (CHAR c)| Y c. Y \<in> CS \<and> Y \<Turnstile>(CHAR c)\<Rightarrow> X}" |
391 {Trn Y (CHAR c)| Y c. Y \<in> CS \<and> Y \<Turnstile>(CHAR c)\<Rightarrow> X}" |
480 |
392 |
481 text {* |
393 definition |
482 In the definition of @{text "init_rhs"}, the term |
394 "eqs CS \<equiv> {(X, init_rhs CS X) | X. X \<in> CS}" |
483 @{text "{Trn Y (CHAR c)| Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X}"} appearing on both branches |
395 |
484 describes the formation of strings in @{text "X"} out of transitions, while |
396 |
485 the term @{text "{Lam(EMPTY)}"} describes the empty string which is intrinsically contained in |
397 |
486 @{text "X"} rather than by transition. This @{text "{Lam(EMPTY)}"} corresponds to |
398 section {* Arden Operation on equations *} |
487 the $\lambda$ in \eqref{example_eqns}. |
399 |
488 |
400 text {* |
489 With the help of @{text "init_rhs"}, the equitional system descrbing the formation of every |
401 The function @{text "attach_rexp r item"} SEQ-composes @{text r} to the |
490 equivalent class inside @{text "CS"} is given by the following @{text "eqs(CS)"}. |
402 right of every rhs-item. |
491 *} |
|
492 |
|
493 |
|
494 definition "eqs CS \<equiv> {(X, init_rhs CS X) | X. X \<in> CS}" |
|
495 |
|
496 |
|
497 |
|
498 (************ arden's lemma variation ********************) |
|
499 |
|
500 text {* |
|
501 The following @{text "trns_of rhs X"} returns all @{text "X"}-items in @{text "rhs"}. |
|
502 *} |
|
503 |
|
504 definition |
|
505 "trns_of rhs X \<equiv> {Trn X r | r. Trn X r \<in> rhs}" |
|
506 |
|
507 text {* |
|
508 The following @{text "attach_rexp rexp' itm"} attach |
|
509 the regular expression @{text "rexp'"} to |
|
510 the right of right hand side item @{text "itm"}. |
|
511 *} |
403 *} |
512 |
404 |
513 fun |
405 fun |
514 attach_rexp :: "rexp \<Rightarrow> rhs_item \<Rightarrow> rhs_item" |
406 attach_rexp :: "rexp \<Rightarrow> rhs_item \<Rightarrow> rhs_item" |
515 where |
407 where |
516 "attach_rexp rexp' (Lam rexp) = Lam (SEQ rexp rexp')" |
408 "attach_rexp r (Lam rexp) = Lam (SEQ rexp r)" |
517 | "attach_rexp rexp' (Trn X rexp) = Trn X (SEQ rexp rexp')" |
409 | "attach_rexp r (Trn X rexp) = Trn X (SEQ rexp r)" |
518 |
410 |
519 text {* |
|
520 The following @{text "append_rhs_rexp rhs rexp"} attaches |
|
521 @{text "rexp"} to every item in @{text "rhs"}. |
|
522 *} |
|
523 |
411 |
524 definition |
412 definition |
525 "append_rhs_rexp rhs rexp \<equiv> (attach_rexp rexp) ` rhs" |
413 "append_rhs_rexp rhs rexp \<equiv> (attach_rexp rexp) ` rhs" |
526 |
414 |
527 text {* |
415 definition |
528 With the help of the two functions immediately above, Ardens' |
416 "arden_op X rhs \<equiv> |
529 transformation on right hand side @{text "rhs"} is implemented |
417 append_rhs_rexp (rhs - trns_of rhs X) (STAR (\<Uplus> {r. Trn X r \<in> rhs}))" |
530 by the following function @{text "arden_variate X rhs"}. |
418 |
531 After this transformation, the recursive occurence of @{text "X"} |
419 |
532 in @{text "rhs"} will be eliminated, while the string set defined |
420 section {* Substitution Operation on equations *} |
533 by @{text "rhs"} is kept unchanged. |
|
534 *} |
|
535 |
|
536 definition |
|
537 "arden_variate X rhs \<equiv> |
|
538 append_rhs_rexp (rhs - trns_of rhs X) (STAR (\<Uplus> {r. Trn X r \<in> rhs}))" |
|
539 |
|
540 |
|
541 (*********** substitution of ES *************) |
|
542 |
421 |
543 text {* |
422 text {* |
544 Suppose the equation defining @{text "X"} is $X = xrhs$, |
423 Suppose and equation @{text "X = xrhs"}, @{text "subst_op"} substitutes |
545 the purpose of @{text "rhs_subst"} is to substitute all occurences of @{text "X"} in |
424 all occurences of @{text "X"} in @{text "rhs"} by @{text "xrhs"}. |
546 @{text "rhs"} by @{text "xrhs"}. |
425 *} |
547 A litte thought may reveal that the final result |
426 |
548 should be: first append $(a_1 | a_2 | \ldots | a_n)$ to every item of @{text "xrhs"} and then |
427 definition |
549 union the result with all non-@{text "X"}-items of @{text "rhs"}. |
428 "subst_op rhs X xrhs \<equiv> |
550 *} |
|
551 |
|
552 definition |
|
553 "rhs_subst rhs X xrhs \<equiv> |
|
554 (rhs - (trns_of rhs X)) \<union> (append_rhs_rexp xrhs (\<Uplus> {r. Trn X r \<in> rhs}))" |
429 (rhs - (trns_of rhs X)) \<union> (append_rhs_rexp xrhs (\<Uplus> {r. Trn X r \<in> rhs}))" |
555 |
430 |
556 text {* |
431 text {* |
557 Suppose the equation defining @{text "X"} is $X = xrhs$, the follwing |
432 @{text "eqs_subst ES X xrhs"} substitutes @{text xrhs} into every |
558 @{text "eqs_subst ES X xrhs"} substitute @{text "xrhs"} into every equation |
433 equation of the equational system @{text ES}. |
559 of the equational system @{text "ES"}. |
434 *} |
560 *} |
|
561 |
435 |
562 definition |
436 definition |
563 "eqs_subst ES X xrhs \<equiv> {(Y, rhs_subst yrhs X xrhs) | Y yrhs. (Y, yrhs) \<in> ES}" |
437 "subst_op_all ES X xrhs \<equiv> {(Y, subst_op yrhs X xrhs) | Y yrhs. (Y, yrhs) \<in> ES}" |
|
438 |
|
439 |
|
440 section {* Well-founded iteration *} |
564 |
441 |
565 text {* |
442 text {* |
566 The computation of regular expressions for equivalence classes is accomplished |
443 The computation of regular expressions for equivalence classes is accomplished |
567 using a iteration principle given by the following lemma. |
444 using a iteration principle given by the following lemma. |
568 *} |
445 *} |
599 The particular invariant used to solve our problem is defined by function @{text "Inv(ES)"}, |
476 The particular invariant used to solve our problem is defined by function @{text "Inv(ES)"}, |
600 an invariant over equal system @{text "ES"}. |
477 an invariant over equal system @{text "ES"}. |
601 Every definition starting next till @{text "Inv"} stipulates a property to be satisfied by @{text "ES"}. |
478 Every definition starting next till @{text "Inv"} stipulates a property to be satisfied by @{text "ES"}. |
602 *} |
479 *} |
603 |
480 |
|
481 |
|
482 section {* Invariants *} |
|
483 |
|
484 text {* Every variable is defined at most onece in @{text ES}. *} |
|
485 |
|
486 definition |
|
487 "distinct_equas ES \<equiv> |
|
488 \<forall> X rhs rhs'. (X, rhs) \<in> ES \<and> (X, rhs') \<in> ES \<longrightarrow> rhs = rhs'" |
|
489 |
604 text {* |
490 text {* |
605 Every variable is defined at most onece in @{text "ES"}. |
491 Every equation in @{text ES} (represented by @{text "(X, rhs)"}) |
606 *} |
492 is valid, i.e. @{text "(X = L rhs)"}. |
607 |
493 *} |
608 definition |
494 |
609 "distinct_equas ES \<equiv> |
495 definition |
610 \<forall> X rhs rhs'. (X, rhs) \<in> ES \<and> (X, rhs') \<in> ES \<longrightarrow> rhs = rhs'" |
496 "valid_eqns ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> (X = L rhs)" |
|
497 |
|
498 text {* |
|
499 @{text "rhs_nonempty rhs"} requires regular expressions occuring in |
|
500 transitional items of @{text "rhs"} do not contain empty string. This is |
|
501 necessary for the application of Arden's transformation to @{text "rhs"}. |
|
502 *} |
|
503 |
|
504 definition |
|
505 "rhs_nonempty rhs \<equiv> (\<forall> Y r. Trn Y r \<in> rhs \<longrightarrow> [] \<notin> L r)" |
|
506 |
|
507 text {* |
|
508 The following @{text "ardenable ES"} requires that Arden's transformation |
|
509 is applicable to every equation of equational system @{text "ES"}. |
|
510 *} |
|
511 |
|
512 definition |
|
513 "ardenable ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> rhs_nonempty rhs" |
|
514 |
|
515 (* The following non_empty seems useless. *) |
|
516 definition |
|
517 "non_empty ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> X \<noteq> {}" |
611 |
518 |
612 text {* |
519 text {* |
613 Every equation in @{text "ES"} (represented by @{text "(X, rhs)"}) is valid, i.e. @{text "(X = L rhs)"}. |
520 @{text "finite_rhs ES"} requires every equation in @{text "rhs"} |
614 *} |
521 be finite. |
615 definition |
522 *} |
616 "valid_eqns ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> (X = L rhs)" |
|
617 |
|
618 text {* |
|
619 The following @{text "rhs_nonempty rhs"} requires regular expressions occuring in transitional |
|
620 items of @{text "rhs"} does not contain empty string. This is necessary for |
|
621 the application of Arden's transformation to @{text "rhs"}. |
|
622 *} |
|
623 |
|
624 definition |
|
625 "rhs_nonempty rhs \<equiv> (\<forall> Y r. Trn Y r \<in> rhs \<longrightarrow> [] \<notin> L r)" |
|
626 |
|
627 text {* |
|
628 The following @{text "ardenable ES"} requires that Arden's transformation is applicable |
|
629 to every equation of equational system @{text "ES"}. |
|
630 *} |
|
631 |
|
632 definition |
|
633 "ardenable ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> rhs_nonempty rhs" |
|
634 |
|
635 (* The following non_empty seems useless. *) |
|
636 definition |
|
637 "non_empty ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> X \<noteq> {}" |
|
638 |
|
639 text {* |
|
640 The following @{text "finite_rhs ES"} requires every equation in @{text "rhs"} be finite. |
|
641 *} |
|
642 definition |
523 definition |
643 "finite_rhs ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> finite rhs" |
524 "finite_rhs ES \<equiv> \<forall> X rhs. (X, rhs) \<in> ES \<longrightarrow> finite rhs" |
644 |
525 |
645 text {* |
526 text {* |
646 The following @{text "classes_of rhs"} returns all variables (or equivalent classes) |
527 @{text "classes_of rhs"} returns all variables (or equivalent classes) |
647 occuring in @{text "rhs"}. |
528 occuring in @{text "rhs"}. |
648 *} |
529 *} |
|
530 |
649 definition |
531 definition |
650 "classes_of rhs \<equiv> {X. \<exists> r. Trn X r \<in> rhs}" |
532 "classes_of rhs \<equiv> {X. \<exists> r. Trn X r \<in> rhs}" |
651 |
533 |
652 text {* |
534 text {* |
653 The following @{text "lefts_of ES"} returns all variables |
535 @{text "lefts_of ES"} returns all variables defined by an |
654 defined by equational system @{text "ES"}. |
536 equational system @{text "ES"}. |
655 *} |
537 *} |
656 definition |
538 definition |
657 "lefts_of ES \<equiv> {Y | Y yrhs. (Y, yrhs) \<in> ES}" |
539 "lefts_of ES \<equiv> {Y | Y yrhs. (Y, yrhs) \<in> ES}" |
658 |
540 |
659 text {* |
541 text {* |
660 The following @{text "self_contained ES"} requires that every |
542 The following @{text "self_contained ES"} requires that every variable occuring |
661 variable occuring on the right hand side of equations is already defined by some |
543 on the right hand side of equations is already defined by some equation in @{text "ES"}. |
662 equation in @{text "ES"}. |
544 *} |
|
545 definition |
|
546 "self_contained ES \<equiv> \<forall> (X, xrhs) \<in> ES. classes_of xrhs \<subseteq> lefts_of ES" |
|
547 |
|
548 |
|
549 text {* |
|
550 The invariant @{text "invariant(ES)"} is a conjunction of all the previously defined constaints. |
663 *} |
551 *} |
664 definition |
552 definition |
665 "self_contained ES \<equiv> \<forall> (X, xrhs) \<in> ES. classes_of xrhs \<subseteq> lefts_of ES" |
553 "invariant ES \<equiv> valid_eqns ES \<and> finite ES \<and> distinct_equas ES \<and> ardenable ES \<and> |
666 |
|
667 |
|
668 text {* |
|
669 The invariant @{text "Inv(ES)"} is a conjunction of all the previously defined constaints. |
|
670 *} |
|
671 definition |
|
672 "Inv ES \<equiv> valid_eqns ES \<and> finite ES \<and> distinct_equas ES \<and> ardenable ES \<and> |
|
673 non_empty ES \<and> finite_rhs ES \<and> self_contained ES" |
554 non_empty ES \<and> finite_rhs ES \<and> self_contained ES" |
674 |
555 |
675 subsection {* The proof of this direction *} |
556 subsection {* The proof of this direction *} |
676 |
557 |
677 subsubsection {* Basic properties *} |
558 subsubsection {* Basic properties *} |
968 moreover have "L (append_rhs_rexp xrhs (\<Uplus>{r. Trn X r \<in> rhs})) = L ({Trn X r | r. Trn X r \<in> rhs})" |
849 moreover have "L (append_rhs_rexp xrhs (\<Uplus>{r. Trn X r \<in> rhs})) = L ({Trn X r | r. Trn X r \<in> rhs})" |
969 using finite substor by (simp only:lang_of_append_rhs lang_of_rexp_of) |
850 using finite substor by (simp only:lang_of_append_rhs lang_of_rexp_of) |
970 ultimately show ?thesis by simp |
851 ultimately show ?thesis by simp |
971 qed |
852 qed |
972 |
853 |
973 lemma rhs_subst_keeps_finite_rhs: |
854 lemma subst_op_keeps_finite_rhs: |
974 "\<lbrakk>finite rhs; finite yrhs\<rbrakk> \<Longrightarrow> finite (rhs_subst rhs Y yrhs)" |
855 "\<lbrakk>finite rhs; finite yrhs\<rbrakk> \<Longrightarrow> finite (subst_op rhs Y yrhs)" |
975 by (auto simp:rhs_subst_def append_keeps_finite) |
856 by (auto simp:subst_op_def append_keeps_finite) |
976 |
857 |
977 lemma eqs_subst_keeps_finite: |
858 lemma subst_op_all_keeps_finite: |
978 assumes finite:"finite (ES:: (string set \<times> rhs_item set) set)" |
859 assumes finite:"finite (ES:: (string set \<times> rhs_item set) set)" |
979 shows "finite (eqs_subst ES Y yrhs)" |
860 shows "finite (subst_op_all ES Y yrhs)" |
980 proof - |
861 proof - |
981 have "finite {(Ya, rhs_subst yrhsa Y yrhs) |Ya yrhsa. (Ya, yrhsa) \<in> ES}" |
862 have "finite {(Ya, subst_op yrhsa Y yrhs) |Ya yrhsa. (Ya, yrhsa) \<in> ES}" |
982 (is "finite ?A") |
863 (is "finite ?A") |
983 proof- |
864 proof- |
984 def eqns' \<equiv> "{((Ya::string set), yrhsa)| Ya yrhsa. (Ya, yrhsa) \<in> ES}" |
865 def eqns' \<equiv> "{((Ya::string set), yrhsa)| Ya yrhsa. (Ya, yrhsa) \<in> ES}" |
985 def h \<equiv> "\<lambda> ((Ya::string set), yrhsa). (Ya, rhs_subst yrhsa Y yrhs)" |
866 def h \<equiv> "\<lambda> ((Ya::string set), yrhsa). (Ya, subst_op yrhsa Y yrhs)" |
986 have "finite (h ` eqns')" using finite h_def eqns'_def by auto |
867 have "finite (h ` eqns')" using finite h_def eqns'_def by auto |
987 moreover have "?A = h ` eqns'" by (auto simp:h_def eqns'_def) |
868 moreover have "?A = h ` eqns'" by (auto simp:h_def eqns'_def) |
988 ultimately show ?thesis by auto |
869 ultimately show ?thesis by auto |
989 qed |
870 qed |
990 thus ?thesis by (simp add:eqs_subst_def) |
871 thus ?thesis by (simp add:subst_op_all_def) |
991 qed |
872 qed |
992 |
873 |
993 lemma eqs_subst_keeps_finite_rhs: |
874 lemma subst_op_all_keeps_finite_rhs: |
994 "\<lbrakk>finite_rhs ES; finite yrhs\<rbrakk> \<Longrightarrow> finite_rhs (eqs_subst ES Y yrhs)" |
875 "\<lbrakk>finite_rhs ES; finite yrhs\<rbrakk> \<Longrightarrow> finite_rhs (subst_op_all ES Y yrhs)" |
995 by (auto intro:rhs_subst_keeps_finite_rhs simp add:eqs_subst_def finite_rhs_def) |
876 by (auto intro:subst_op_keeps_finite_rhs simp add:subst_op_all_def finite_rhs_def) |
996 |
877 |
997 lemma append_rhs_keeps_cls: |
878 lemma append_rhs_keeps_cls: |
998 "classes_of (append_rhs_rexp rhs r) = classes_of rhs" |
879 "classes_of (append_rhs_rexp rhs r) = classes_of rhs" |
999 apply (auto simp:classes_of_def append_rhs_rexp_def) |
880 apply (auto simp:classes_of_def append_rhs_rexp_def) |
1000 apply (case_tac xa, auto simp:image_def) |
881 apply (case_tac xa, auto simp:image_def) |
1001 by (rule_tac x = "SEQ ra r" in exI, rule_tac x = "Trn x ra" in bexI, simp+) |
882 by (rule_tac x = "SEQ ra r" in exI, rule_tac x = "Trn x ra" in bexI, simp+) |
1002 |
883 |
1003 lemma arden_variate_removes_cl: |
884 lemma arden_op_removes_cl: |
1004 "classes_of (arden_variate Y yrhs) = classes_of yrhs - {Y}" |
885 "classes_of (arden_op Y yrhs) = classes_of yrhs - {Y}" |
1005 apply (simp add:arden_variate_def append_rhs_keeps_cls trns_of_def) |
886 apply (simp add:arden_op_def append_rhs_keeps_cls trns_of_def) |
1006 by (auto simp:classes_of_def) |
887 by (auto simp:classes_of_def) |
1007 |
888 |
1008 lemma lefts_of_keeps_cls: |
889 lemma lefts_of_keeps_cls: |
1009 "lefts_of (eqs_subst ES Y yrhs) = lefts_of ES" |
890 "lefts_of (subst_op_all ES Y yrhs) = lefts_of ES" |
1010 by (auto simp:lefts_of_def eqs_subst_def) |
891 by (auto simp:lefts_of_def subst_op_all_def) |
1011 |
892 |
1012 lemma rhs_subst_updates_cls: |
893 lemma subst_op_updates_cls: |
1013 "X \<notin> classes_of xrhs \<Longrightarrow> |
894 "X \<notin> classes_of xrhs \<Longrightarrow> |
1014 classes_of (rhs_subst rhs X xrhs) = classes_of rhs \<union> classes_of xrhs - {X}" |
895 classes_of (subst_op rhs X xrhs) = classes_of rhs \<union> classes_of xrhs - {X}" |
1015 apply (simp only:rhs_subst_def append_rhs_keeps_cls |
896 apply (simp only:subst_op_def append_rhs_keeps_cls |
1016 classes_of_union_distrib[THEN sym]) |
897 classes_of_union_distrib[THEN sym]) |
1017 by (auto simp:classes_of_def trns_of_def) |
898 by (auto simp:classes_of_def trns_of_def) |
1018 |
899 |
1019 lemma eqs_subst_keeps_self_contained: |
900 lemma subst_op_all_keeps_self_contained: |
1020 fixes Y |
901 fixes Y |
1021 assumes sc: "self_contained (ES \<union> {(Y, yrhs)})" (is "self_contained ?A") |
902 assumes sc: "self_contained (ES \<union> {(Y, yrhs)})" (is "self_contained ?A") |
1022 shows "self_contained (eqs_subst ES Y (arden_variate Y yrhs))" |
903 shows "self_contained (subst_op_all ES Y (arden_op Y yrhs))" |
1023 (is "self_contained ?B") |
904 (is "self_contained ?B") |
1024 proof- |
905 proof- |
1025 { fix X xrhs' |
906 { fix X xrhs' |
1026 assume "(X, xrhs') \<in> ?B" |
907 assume "(X, xrhs') \<in> ?B" |
1027 then obtain xrhs |
908 then obtain xrhs |
1028 where xrhs_xrhs': "xrhs' = rhs_subst xrhs Y (arden_variate Y yrhs)" |
909 where xrhs_xrhs': "xrhs' = subst_op xrhs Y (arden_op Y yrhs)" |
1029 and X_in: "(X, xrhs) \<in> ES" by (simp add:eqs_subst_def, blast) |
910 and X_in: "(X, xrhs) \<in> ES" by (simp add:subst_op_all_def, blast) |
1030 have "classes_of xrhs' \<subseteq> lefts_of ?B" |
911 have "classes_of xrhs' \<subseteq> lefts_of ?B" |
1031 proof- |
912 proof- |
1032 have "lefts_of ?B = lefts_of ES" by (auto simp add:lefts_of_def eqs_subst_def) |
913 have "lefts_of ?B = lefts_of ES" by (auto simp add:lefts_of_def subst_op_all_def) |
1033 moreover have "classes_of xrhs' \<subseteq> lefts_of ES" |
914 moreover have "classes_of xrhs' \<subseteq> lefts_of ES" |
1034 proof- |
915 proof- |
1035 have "classes_of xrhs' \<subseteq> |
916 have "classes_of xrhs' \<subseteq> |
1036 classes_of xrhs \<union> classes_of (arden_variate Y yrhs) - {Y}" |
917 classes_of xrhs \<union> classes_of (arden_op Y yrhs) - {Y}" |
1037 proof- |
918 proof- |
1038 have "Y \<notin> classes_of (arden_variate Y yrhs)" |
919 have "Y \<notin> classes_of (arden_op Y yrhs)" |
1039 using arden_variate_removes_cl by simp |
920 using arden_op_removes_cl by simp |
1040 thus ?thesis using xrhs_xrhs' by (auto simp:rhs_subst_updates_cls) |
921 thus ?thesis using xrhs_xrhs' by (auto simp:subst_op_updates_cls) |
1041 qed |
922 qed |
1042 moreover have "classes_of xrhs \<subseteq> lefts_of ES \<union> {Y}" using X_in sc |
923 moreover have "classes_of xrhs \<subseteq> lefts_of ES \<union> {Y}" using X_in sc |
1043 apply (simp only:self_contained_def lefts_of_union_distrib[THEN sym]) |
924 apply (simp only:self_contained_def lefts_of_union_distrib[THEN sym]) |
1044 by (drule_tac x = "(X, xrhs)" in bspec, auto simp:lefts_of_def) |
925 by (drule_tac x = "(X, xrhs)" in bspec, auto simp:lefts_of_def) |
1045 moreover have "classes_of (arden_variate Y yrhs) \<subseteq> lefts_of ES \<union> {Y}" |
926 moreover have "classes_of (arden_op Y yrhs) \<subseteq> lefts_of ES \<union> {Y}" |
1046 using sc |
927 using sc |
1047 by (auto simp add:arden_variate_removes_cl self_contained_def lefts_of_def) |
928 by (auto simp add:arden_op_removes_cl self_contained_def lefts_of_def) |
1048 ultimately show ?thesis by auto |
929 ultimately show ?thesis by auto |
1049 qed |
930 qed |
1050 ultimately show ?thesis by simp |
931 ultimately show ?thesis by simp |
1051 qed |
932 qed |
1052 } thus ?thesis by (auto simp only:eqs_subst_def self_contained_def) |
933 } thus ?thesis by (auto simp only:subst_op_all_def self_contained_def) |
1053 qed |
934 qed |
1054 |
935 |
1055 lemma eqs_subst_satisfy_Inv: |
936 lemma subst_op_all_satisfy_invariant: |
1056 assumes Inv_ES: "Inv (ES \<union> {(Y, yrhs)})" |
937 assumes invariant_ES: "invariant (ES \<union> {(Y, yrhs)})" |
1057 shows "Inv (eqs_subst ES Y (arden_variate Y yrhs))" |
938 shows "invariant (subst_op_all ES Y (arden_op Y yrhs))" |
1058 proof - |
939 proof - |
1059 have finite_yrhs: "finite yrhs" |
940 have finite_yrhs: "finite yrhs" |
1060 using Inv_ES by (auto simp:Inv_def finite_rhs_def) |
941 using invariant_ES by (auto simp:invariant_def finite_rhs_def) |
1061 have nonempty_yrhs: "rhs_nonempty yrhs" |
942 have nonempty_yrhs: "rhs_nonempty yrhs" |
1062 using Inv_ES by (auto simp:Inv_def ardenable_def) |
943 using invariant_ES by (auto simp:invariant_def ardenable_def) |
1063 have Y_eq_yrhs: "Y = L yrhs" |
944 have Y_eq_yrhs: "Y = L yrhs" |
1064 using Inv_ES by (simp only:Inv_def valid_eqns_def, blast) |
945 using invariant_ES by (simp only:invariant_def valid_eqns_def, blast) |
1065 have "distinct_equas (eqs_subst ES Y (arden_variate Y yrhs))" |
946 have "distinct_equas (subst_op_all ES Y (arden_op Y yrhs))" |
1066 using Inv_ES |
947 using invariant_ES |
1067 by (auto simp:distinct_equas_def eqs_subst_def Inv_def) |
948 by (auto simp:distinct_equas_def subst_op_all_def invariant_def) |
1068 moreover have "finite (eqs_subst ES Y (arden_variate Y yrhs))" |
949 moreover have "finite (subst_op_all ES Y (arden_op Y yrhs))" |
1069 using Inv_ES by (simp add:Inv_def eqs_subst_keeps_finite) |
950 using invariant_ES by (simp add:invariant_def subst_op_all_keeps_finite) |
1070 moreover have "finite_rhs (eqs_subst ES Y (arden_variate Y yrhs))" |
951 moreover have "finite_rhs (subst_op_all ES Y (arden_op Y yrhs))" |
1071 proof- |
952 proof- |
1072 have "finite_rhs ES" using Inv_ES |
953 have "finite_rhs ES" using invariant_ES |
1073 by (simp add:Inv_def finite_rhs_def) |
954 by (simp add:invariant_def finite_rhs_def) |
1074 moreover have "finite (arden_variate Y yrhs)" |
955 moreover have "finite (arden_op Y yrhs)" |
1075 proof - |
956 proof - |
1076 have "finite yrhs" using Inv_ES |
957 have "finite yrhs" using invariant_ES |
1077 by (auto simp:Inv_def finite_rhs_def) |
958 by (auto simp:invariant_def finite_rhs_def) |
1078 thus ?thesis using arden_variate_keeps_finite by simp |
959 thus ?thesis using arden_op_keeps_finite by simp |
1079 qed |
960 qed |
1080 ultimately show ?thesis |
961 ultimately show ?thesis |
1081 by (simp add:eqs_subst_keeps_finite_rhs) |
962 by (simp add:subst_op_all_keeps_finite_rhs) |
1082 qed |
963 qed |
1083 moreover have "ardenable (eqs_subst ES Y (arden_variate Y yrhs))" |
964 moreover have "ardenable (subst_op_all ES Y (arden_op Y yrhs))" |
1084 proof - |
965 proof - |
1085 { fix X rhs |
966 { fix X rhs |
1086 assume "(X, rhs) \<in> ES" |
967 assume "(X, rhs) \<in> ES" |
1087 hence "rhs_nonempty rhs" using prems Inv_ES |
968 hence "rhs_nonempty rhs" using prems invariant_ES |
1088 by (simp add:Inv_def ardenable_def) |
969 by (simp add:invariant_def ardenable_def) |
1089 with nonempty_yrhs |
970 with nonempty_yrhs |
1090 have "rhs_nonempty (rhs_subst rhs Y (arden_variate Y yrhs))" |
971 have "rhs_nonempty (subst_op rhs Y (arden_op Y yrhs))" |
1091 by (simp add:nonempty_yrhs |
972 by (simp add:nonempty_yrhs |
1092 rhs_subst_keeps_nonempty arden_variate_keeps_nonempty) |
973 subst_op_keeps_nonempty arden_op_keeps_nonempty) |
1093 } thus ?thesis by (auto simp add:ardenable_def eqs_subst_def) |
974 } thus ?thesis by (auto simp add:ardenable_def subst_op_all_def) |
1094 qed |
975 qed |
1095 moreover have "valid_eqns (eqs_subst ES Y (arden_variate Y yrhs))" |
976 moreover have "valid_eqns (subst_op_all ES Y (arden_op Y yrhs))" |
1096 proof- |
977 proof- |
1097 have "Y = L (arden_variate Y yrhs)" |
978 have "Y = L (arden_op Y yrhs)" |
1098 using Y_eq_yrhs Inv_ES finite_yrhs nonempty_yrhs |
979 using Y_eq_yrhs invariant_ES finite_yrhs nonempty_yrhs |
1099 by (rule_tac arden_variate_keeps_eq, (simp add:rexp_of_empty)+) |
980 by (rule_tac arden_op_keeps_eq, (simp add:rexp_of_empty)+) |
1100 thus ?thesis using Inv_ES |
981 thus ?thesis using invariant_ES |
1101 by (clarsimp simp add:valid_eqns_def |
982 by (clarsimp simp add:valid_eqns_def |
1102 eqs_subst_def rhs_subst_keeps_eq Inv_def finite_rhs_def |
983 subst_op_all_def subst_op_keeps_eq invariant_def finite_rhs_def |
1103 simp del:L_rhs.simps) |
984 simp del:L_rhs.simps) |
1104 qed |
985 qed |
1105 moreover have |
986 moreover have |
1106 non_empty_subst: "non_empty (eqs_subst ES Y (arden_variate Y yrhs))" |
987 non_empty_subst: "non_empty (subst_op_all ES Y (arden_op Y yrhs))" |
1107 using Inv_ES by (auto simp:Inv_def non_empty_def eqs_subst_def) |
988 using invariant_ES by (auto simp:invariant_def non_empty_def subst_op_all_def) |
1108 moreover |
989 moreover |
1109 have self_subst: "self_contained (eqs_subst ES Y (arden_variate Y yrhs))" |
990 have self_subst: "self_contained (subst_op_all ES Y (arden_op Y yrhs))" |
1110 using Inv_ES eqs_subst_keeps_self_contained by (simp add:Inv_def) |
991 using invariant_ES subst_op_all_keeps_self_contained by (simp add:invariant_def) |
1111 ultimately show ?thesis using Inv_ES by (simp add:Inv_def) |
992 ultimately show ?thesis using invariant_ES by (simp add:invariant_def) |
1112 qed |
993 qed |
1113 |
994 |
1114 lemma eqs_subst_card_le: |
995 lemma subst_op_all_card_le: |
1115 assumes finite: "finite (ES::(string set \<times> rhs_item set) set)" |
996 assumes finite: "finite (ES::(string set \<times> rhs_item set) set)" |
1116 shows "card (eqs_subst ES Y yrhs) <= card ES" |
997 shows "card (subst_op_all ES Y yrhs) <= card ES" |
1117 proof- |
998 proof- |
1118 def f \<equiv> "\<lambda> x. ((fst x)::string set, rhs_subst (snd x) Y yrhs)" |
999 def f \<equiv> "\<lambda> x. ((fst x)::string set, subst_op (snd x) Y yrhs)" |
1119 have "eqs_subst ES Y yrhs = f ` ES" |
1000 have "subst_op_all ES Y yrhs = f ` ES" |
1120 apply (auto simp:eqs_subst_def f_def image_def) |
1001 apply (auto simp:subst_op_all_def f_def image_def) |
1121 by (rule_tac x = "(Ya, yrhsa)" in bexI, simp+) |
1002 by (rule_tac x = "(Ya, yrhsa)" in bexI, simp+) |
1122 thus ?thesis using finite by (auto intro:card_image_le) |
1003 thus ?thesis using finite by (auto intro:card_image_le) |
1123 qed |
1004 qed |
1124 |
1005 |
1125 lemma eqs_subst_cls_remains: |
1006 lemma subst_op_all_cls_remains: |
1126 "(X, xrhs) \<in> ES \<Longrightarrow> \<exists> xrhs'. (X, xrhs') \<in> (eqs_subst ES Y yrhs)" |
1007 "(X, xrhs) \<in> ES \<Longrightarrow> \<exists> xrhs'. (X, xrhs') \<in> (subst_op_all ES Y yrhs)" |
1127 by (auto simp:eqs_subst_def) |
1008 by (auto simp:subst_op_all_def) |
1128 |
1009 |
1129 lemma card_noteq_1_has_more: |
1010 lemma card_noteq_1_has_more: |
1130 assumes card:"card S \<noteq> 1" |
1011 assumes card:"card S \<noteq> 1" |
1131 and e_in: "e \<in> S" |
1012 and e_in: "e \<in> S" |
1132 and finite: "finite S" |
1013 and finite: "finite S" |
1141 hence "S - {e} \<noteq> {}" using finite by (rule_tac notI, simp) |
1022 hence "S - {e} \<noteq> {}" using finite by (rule_tac notI, simp) |
1142 thus "(\<And>e'. e' \<in> S \<and> e \<noteq> e' \<Longrightarrow> thesis) \<Longrightarrow> thesis" by auto |
1023 thus "(\<And>e'. e' \<in> S \<and> e \<noteq> e' \<Longrightarrow> thesis) \<Longrightarrow> thesis" by auto |
1143 qed |
1024 qed |
1144 |
1025 |
1145 lemma iteration_step: |
1026 lemma iteration_step: |
1146 assumes Inv_ES: "Inv ES" |
1027 assumes invariant_ES: "invariant ES" |
1147 and X_in_ES: "(X, xrhs) \<in> ES" |
1028 and X_in_ES: "(X, xrhs) \<in> ES" |
1148 and not_T: "card ES \<noteq> 1" |
1029 and not_T: "card ES \<noteq> 1" |
1149 shows "\<exists> ES'. (Inv ES' \<and> (\<exists> xrhs'.(X, xrhs') \<in> ES')) \<and> |
1030 shows "\<exists> ES'. (invariant ES' \<and> (\<exists> xrhs'.(X, xrhs') \<in> ES')) \<and> |
1150 (card ES', card ES) \<in> less_than" (is "\<exists> ES'. ?P ES'") |
1031 (card ES', card ES) \<in> less_than" (is "\<exists> ES'. ?P ES'") |
1151 proof - |
1032 proof - |
1152 have finite_ES: "finite ES" using Inv_ES by (simp add:Inv_def) |
1033 have finite_ES: "finite ES" using invariant_ES by (simp add:invariant_def) |
1153 then obtain Y yrhs |
1034 then obtain Y yrhs |
1154 where Y_in_ES: "(Y, yrhs) \<in> ES" and not_eq: "(X, xrhs) \<noteq> (Y, yrhs)" |
1035 where Y_in_ES: "(Y, yrhs) \<in> ES" and not_eq: "(X, xrhs) \<noteq> (Y, yrhs)" |
1155 using not_T X_in_ES by (drule_tac card_noteq_1_has_more, auto) |
1036 using not_T X_in_ES by (drule_tac card_noteq_1_has_more, auto) |
1156 def ES' == "ES - {(Y, yrhs)}" |
1037 def ES' == "ES - {(Y, yrhs)}" |
1157 let ?ES'' = "eqs_subst ES' Y (arden_variate Y yrhs)" |
1038 let ?ES'' = "subst_op_all ES' Y (arden_op Y yrhs)" |
1158 have "?P ?ES''" |
1039 have "?P ?ES''" |
1159 proof - |
1040 proof - |
1160 have "Inv ?ES''" using Y_in_ES Inv_ES |
1041 have "invariant ?ES''" using Y_in_ES invariant_ES |
1161 by (rule_tac eqs_subst_satisfy_Inv, simp add:ES'_def insert_absorb) |
1042 by (rule_tac subst_op_all_satisfy_invariant, simp add:ES'_def insert_absorb) |
1162 moreover have "\<exists>xrhs'. (X, xrhs') \<in> ?ES''" using not_eq X_in_ES |
1043 moreover have "\<exists>xrhs'. (X, xrhs') \<in> ?ES''" using not_eq X_in_ES |
1163 by (rule_tac ES = ES' in eqs_subst_cls_remains, auto simp add:ES'_def) |
1044 by (rule_tac ES = ES' in subst_op_all_cls_remains, auto simp add:ES'_def) |
1164 moreover have "(card ?ES'', card ES) \<in> less_than" |
1045 moreover have "(card ?ES'', card ES) \<in> less_than" |
1165 proof - |
1046 proof - |
1166 have "finite ES'" using finite_ES ES'_def by auto |
1047 have "finite ES'" using finite_ES ES'_def by auto |
1167 moreover have "card ES' < card ES" using finite_ES Y_in_ES |
1048 moreover have "card ES' < card ES" using finite_ES Y_in_ES |
1168 by (auto simp:ES'_def card_gt_0_iff intro:diff_Suc_less) |
1049 by (auto simp:ES'_def card_gt_0_iff intro:diff_Suc_less) |
1169 ultimately show ?thesis |
1050 ultimately show ?thesis |
1170 by (auto dest:eqs_subst_card_le elim:le_less_trans) |
1051 by (auto dest:subst_op_all_card_le elim:le_less_trans) |
1171 qed |
1052 qed |
1172 ultimately show ?thesis by simp |
1053 ultimately show ?thesis by simp |
1173 qed |
1054 qed |
1174 thus ?thesis by blast |
1055 thus ?thesis by blast |
1175 qed |
1056 qed |
1199 by (rule_tac f = card in wf_iter, auto) |
1080 by (rule_tac f = card in wf_iter, auto) |
1200 qed |
1081 qed |
1201 |
1082 |
1202 lemma last_cl_exists_rexp: |
1083 lemma last_cl_exists_rexp: |
1203 assumes ES_single: "ES = {(X, xrhs)}" |
1084 assumes ES_single: "ES = {(X, xrhs)}" |
1204 and Inv_ES: "Inv ES" |
1085 and invariant_ES: "invariant ES" |
1205 shows "\<exists> (r::rexp). L r = X" (is "\<exists> r. ?P r") |
1086 shows "\<exists> (r::rexp). L r = X" (is "\<exists> r. ?P r") |
1206 proof- |
1087 proof- |
1207 def A \<equiv> "arden_variate X xrhs" |
1088 def A \<equiv> "arden_op X xrhs" |
1208 have "?P (\<Uplus>{r. Lam r \<in> A})" |
1089 have "?P (\<Uplus>{r. Lam r \<in> A})" |
1209 proof - |
1090 proof - |
1210 have "L (\<Uplus>{r. Lam r \<in> A}) = L ({Lam r | r. Lam r \<in> A})" |
1091 have "L (\<Uplus>{r. Lam r \<in> A}) = L ({Lam r | r. Lam r \<in> A})" |
1211 proof(rule rexp_of_lam_eq_lam_set) |
1092 proof(rule rexp_of_lam_eq_lam_set) |
1212 show "finite A" |
1093 show "finite A" |
1213 unfolding A_def |
1094 unfolding A_def |
1214 using Inv_ES ES_single |
1095 using invariant_ES ES_single |
1215 by (rule_tac arden_variate_keeps_finite) |
1096 by (rule_tac arden_op_keeps_finite) |
1216 (auto simp add: Inv_def finite_rhs_def) |
1097 (auto simp add: invariant_def finite_rhs_def) |
1217 qed |
1098 qed |
1218 also have "\<dots> = L A" |
1099 also have "\<dots> = L A" |
1219 proof- |
1100 proof- |
1220 have "{Lam r | r. Lam r \<in> A} = A" |
1101 have "{Lam r | r. Lam r \<in> A} = A" |
1221 proof- |
1102 proof- |
1222 have "classes_of A = {}" using Inv_ES ES_single |
1103 have "classes_of A = {}" using invariant_ES ES_single |
1223 unfolding A_def |
1104 unfolding A_def |
1224 by (simp add:arden_variate_removes_cl |
1105 by (simp add:arden_op_removes_cl |
1225 self_contained_def Inv_def lefts_of_def) |
1106 self_contained_def invariant_def lefts_of_def) |
1226 thus ?thesis |
1107 thus ?thesis |
1227 unfolding A_def |
1108 unfolding A_def |
1228 by (auto simp only: classes_of_def, case_tac x, auto) |
1109 by (auto simp only: classes_of_def, case_tac x, auto) |
1229 qed |
1110 qed |
1230 thus ?thesis by simp |
1111 thus ?thesis by simp |
1231 qed |
1112 qed |
1232 also have "\<dots> = X" |
1113 also have "\<dots> = X" |
1233 unfolding A_def |
1114 unfolding A_def |
1234 proof(rule arden_variate_keeps_eq [THEN sym]) |
1115 proof(rule arden_op_keeps_eq [THEN sym]) |
1235 show "X = L xrhs" using Inv_ES ES_single |
1116 show "X = L xrhs" using invariant_ES ES_single |
1236 by (auto simp only:Inv_def valid_eqns_def) |
1117 by (auto simp only:invariant_def valid_eqns_def) |
1237 next |
1118 next |
1238 from Inv_ES ES_single show "[] \<notin> L (\<Uplus>{r. Trn X r \<in> xrhs})" |
1119 from invariant_ES ES_single show "[] \<notin> L (\<Uplus>{r. Trn X r \<in> xrhs})" |
1239 by(simp add:Inv_def ardenable_def rexp_of_empty finite_rhs_def) |
1120 by(simp add:invariant_def ardenable_def rexp_of_empty finite_rhs_def) |
1240 next |
1121 next |
1241 from Inv_ES ES_single show "finite xrhs" |
1122 from invariant_ES ES_single show "finite xrhs" |
1242 by (simp add:Inv_def finite_rhs_def) |
1123 by (simp add:invariant_def finite_rhs_def) |
1243 qed |
1124 qed |
1244 finally show ?thesis by simp |
1125 finally show ?thesis by simp |
1245 qed |
1126 qed |
1246 thus ?thesis by auto |
1127 thus ?thesis by auto |
1247 qed |
1128 qed |