Myhill_1.thy
changeset 166 7743d2ad71d1
parent 162 e93760534354
child 170 b1258b7d2789
equal deleted inserted replaced
165:b04cc5e4e84c 166:7743d2ad71d1
     1 theory Myhill_1
     1 theory Myhill_1
     2 imports Main Folds Regular
     2 imports Regular
     3         "~~/src/HOL/Library/While_Combinator" 
     3         "~~/src/HOL/Library/While_Combinator" 
     4 begin
     4 begin
     5 
     5 
     6 section {* Direction @{text "finite partition \<Rightarrow> regular language"} *}
     6 section {* Direction @{text "finite partition \<Rightarrow> regular language"} *}
     7 
     7 
    35 
    35 
    36 section {* Equational systems *}
    36 section {* Equational systems *}
    37 
    37 
    38 text {* The two kinds of terms in the rhs of equations. *}
    38 text {* The two kinds of terms in the rhs of equations. *}
    39 
    39 
    40 datatype rhs_trm = 
    40 datatype trm = 
    41    Lam "rexp"            (* Lambda-marker *)
    41    Lam "rexp"            (* Lambda-marker *)
    42  | Trn "lang" "rexp"     (* Transition *)
    42  | Trn "lang" "rexp"     (* Transition *)
    43 
    43 
    44 
    44 fun 
    45 overloading L_rhs_trm \<equiv> "L:: rhs_trm \<Rightarrow> lang"
    45   L_trm::"trm \<Rightarrow> lang"
    46 begin
    46 where
    47   fun L_rhs_trm:: "rhs_trm \<Rightarrow> lang"
    47   "L_trm (Lam r) = L_rexp r" 
    48   where
    48 | "L_trm (Trn X r) = X \<cdot> L_rexp r"
    49     "L_rhs_trm (Lam r) = L r" 
    49 
    50   | "L_rhs_trm (Trn X r) = X ;; L r"
    50 fun 
    51 end
    51   L_rhs::"trm set \<Rightarrow> lang"
    52 
    52 where 
    53 overloading L_rhs \<equiv> "L:: rhs_trm set \<Rightarrow> lang"
    53   "L_rhs rhs = \<Union> (L_trm ` rhs)"
    54 begin
       
    55    fun L_rhs:: "rhs_trm set \<Rightarrow> lang"
       
    56    where 
       
    57      "L_rhs rhs = \<Union> (L ` rhs)"
       
    58 end
       
    59 
    54 
    60 lemma L_rhs_set:
    55 lemma L_rhs_set:
    61   shows "L {Trn X r | r. P r} = \<Union>{L (Trn X r) | r. P r}"
    56   shows "L_rhs {Trn X r | r. P r} = \<Union>{L_trm (Trn X r) | r. P r}"
    62 by (auto simp del: L_rhs_trm.simps)
    57 by (auto)
    63 
    58 
    64 lemma L_rhs_union_distrib:
    59 lemma L_rhs_union_distrib:
    65   fixes A B::"rhs_trm set"
    60   fixes A B::"trm set"
    66   shows "L A \<union> L B = L (A \<union> B)"
    61   shows "L_rhs A \<union> L_rhs B = L_rhs (A \<union> B)"
    67 by simp
    62 by simp
    68 
       
    69 
    63 
    70 
    64 
    71 text {* Transitions between equivalence classes *}
    65 text {* Transitions between equivalence classes *}
    72 
    66 
    73 definition 
    67 definition 
    74   transition :: "lang \<Rightarrow> char \<Rightarrow> lang \<Rightarrow> bool" ("_ \<Turnstile>_\<Rightarrow>_" [100,100,100] 100)
    68   transition :: "lang \<Rightarrow> char \<Rightarrow> lang \<Rightarrow> bool" ("_ \<Turnstile>_\<Rightarrow>_" [100,100,100] 100)
    75 where
    69 where
    76   "Y \<Turnstile>c\<Rightarrow> X \<equiv> Y ;; {[c]} \<subseteq> X"
    70   "Y \<Turnstile>c\<Rightarrow> X \<equiv> Y \<cdot> {[c]} \<subseteq> X"
    77 
    71 
    78 text {* Initial equational system *}
    72 text {* Initial equational system *}
    79 
    73 
    80 definition
    74 definition
    81   "Init_rhs CS X \<equiv>  
    75   "Init_rhs CS X \<equiv>  
    89 
    83 
    90 
    84 
    91 section {* Arden Operation on equations *}
    85 section {* Arden Operation on equations *}
    92 
    86 
    93 fun 
    87 fun 
    94   Append_rexp :: "rexp \<Rightarrow> rhs_trm \<Rightarrow> rhs_trm"
    88   Append_rexp :: "rexp \<Rightarrow> trm \<Rightarrow> trm"
    95 where
    89 where
    96   "Append_rexp r (Lam rexp)   = Lam (SEQ rexp r)"
    90   "Append_rexp r (Lam rexp)   = Lam (SEQ rexp r)"
    97 | "Append_rexp r (Trn X rexp) = Trn X (SEQ rexp r)"
    91 | "Append_rexp r (Trn X rexp) = Trn X (SEQ rexp r)"
    98 
    92 
    99 
    93 
   110 definition 
   104 definition 
   111   "Subst rhs X xrhs \<equiv> 
   105   "Subst rhs X xrhs \<equiv> 
   112         (rhs - {Trn X r | r. Trn X r \<in> rhs}) \<union> (Append_rexp_rhs xrhs (\<Uplus> {r. Trn X r \<in> rhs}))"
   106         (rhs - {Trn X r | r. Trn X r \<in> rhs}) \<union> (Append_rexp_rhs xrhs (\<Uplus> {r. Trn X r \<in> rhs}))"
   113 
   107 
   114 definition
   108 definition
   115   Subst_all :: "(lang \<times> rhs_trm set) set \<Rightarrow> lang \<Rightarrow> rhs_trm set \<Rightarrow> (lang \<times> rhs_trm set) set"
   109   Subst_all :: "(lang \<times> trm set) set \<Rightarrow> lang \<Rightarrow> trm set \<Rightarrow> (lang \<times> trm set) set"
   116 where
   110 where
   117   "Subst_all ES X xrhs \<equiv> {(Y, Subst yrhs X xrhs) | Y yrhs. (Y, yrhs) \<in> ES}"
   111   "Subst_all ES X xrhs \<equiv> {(Y, Subst yrhs X xrhs) | Y yrhs. (Y, yrhs) \<in> ES}"
   118 
   112 
   119 definition
   113 definition
   120   "Remove ES X xrhs \<equiv> 
   114   "Remove ES X xrhs \<equiv> 
   147 definition 
   141 definition 
   148   "distinctness ES \<equiv> 
   142   "distinctness ES \<equiv> 
   149      \<forall> X rhs rhs'. (X, rhs) \<in> ES \<and> (X, rhs') \<in> ES \<longrightarrow> rhs = rhs'"
   143      \<forall> X rhs rhs'. (X, rhs) \<in> ES \<and> (X, rhs') \<in> ES \<longrightarrow> rhs = rhs'"
   150 
   144 
   151 definition 
   145 definition 
   152   "soundness ES \<equiv> \<forall>(X, rhs) \<in> ES. X = L rhs"
   146   "soundness ES \<equiv> \<forall>(X, rhs) \<in> ES. X = L_rhs rhs"
   153 
   147 
   154 definition 
   148 definition 
   155   "ardenable rhs \<equiv> (\<forall> Y r. Trn Y r \<in> rhs \<longrightarrow> [] \<notin> L r)"
   149   "ardenable rhs \<equiv> (\<forall> Y r. Trn Y r \<in> rhs \<longrightarrow> [] \<notin> L_rexp r)"
   156 
   150 
   157 definition 
   151 definition 
   158   "ardenable_all ES \<equiv> \<forall>(X, rhs) \<in> ES. ardenable rhs"
   152   "ardenable_all ES \<equiv> \<forall>(X, rhs) \<in> ES. ardenable rhs"
   159 
   153 
   160 definition
   154 definition
   225     apply(erule finite_imageD)
   219     apply(erule finite_imageD)
   226     apply(auto simp add: inj_on_def)
   220     apply(auto simp add: inj_on_def)
   227     done
   221     done
   228 qed
   222 qed
   229 
   223 
   230 lemma rhs_trm_soundness:
   224 lemma trm_soundness:
   231   assumes finite:"finite rhs"
   225   assumes finite:"finite rhs"
   232   shows "L ({Trn X r| r. Trn X r \<in> rhs}) = X ;; (L (\<Uplus>{r. Trn X r \<in> rhs}))"
   226   shows "L_rhs ({Trn X r| r. Trn X r \<in> rhs}) = X \<cdot> (L_rexp (\<Uplus>{r. Trn X r \<in> rhs}))"
   233 proof -
   227 proof -
   234   have "finite {r. Trn X r \<in> rhs}" 
   228   have "finite {r. Trn X r \<in> rhs}" 
   235     by (rule finite_Trn[OF finite]) 
   229     by (rule finite_Trn[OF finite]) 
   236   then show "L ({Trn X r| r. Trn X r \<in> rhs}) = X ;; (L (\<Uplus>{r. Trn X r \<in> rhs}))"
   230   then show "L_rhs ({Trn X r| r. Trn X r \<in> rhs}) = X \<cdot> (L_rexp (\<Uplus>{r. Trn X r \<in> rhs}))"
   237     by (simp only: L_rhs_set L_rhs_trm.simps) (auto simp add: Seq_def)
   231     by (simp only: L_rhs_set L_trm.simps) (auto simp add: Seq_def)
   238 qed
   232 qed
   239 
   233 
   240 lemma lang_of_append_rexp:
   234 lemma lang_of_append_rexp:
   241   "L (Append_rexp r rhs_trm) = L rhs_trm ;; L r"
   235   "L_trm (Append_rexp r trm) = L_trm trm \<cdot> L_rexp r"
   242 by (induct rule: Append_rexp.induct)
   236 by (induct rule: Append_rexp.induct)
   243    (auto simp add: seq_assoc)
   237    (auto simp add: seq_assoc)
   244 
   238 
   245 lemma lang_of_append_rexp_rhs:
   239 lemma lang_of_append_rexp_rhs:
   246   "L (Append_rexp_rhs rhs r) = L rhs ;; L r"
   240   "L_rhs (Append_rexp_rhs rhs r) = L_rhs rhs \<cdot> L_rexp r"
   247 unfolding Append_rexp_rhs_def
   241 unfolding Append_rexp_rhs_def
   248 by (auto simp add: Seq_def lang_of_append_rexp)
   242 by (auto simp add: Seq_def lang_of_append_rexp)
   249 
   243 
   250 
   244 
   251 
   245 subsubsection {* Intial Equational System *}
   252 subsubsection {* Intialization *}
       
   253 
   246 
   254 lemma defined_by_str:
   247 lemma defined_by_str:
   255   assumes "s \<in> X" "X \<in> UNIV // \<approx>A" 
   248   assumes "s \<in> X" "X \<in> UNIV // \<approx>A" 
   256   shows "X = \<approx>A `` {s}"
   249   shows "X = \<approx>A `` {s}"
   257 using assms
   250 using assms
   259 by auto
   252 by auto
   260 
   253 
   261 lemma every_eqclass_has_transition:
   254 lemma every_eqclass_has_transition:
   262   assumes has_str: "s @ [c] \<in> X"
   255   assumes has_str: "s @ [c] \<in> X"
   263   and     in_CS:   "X \<in> UNIV // \<approx>A"
   256   and     in_CS:   "X \<in> UNIV // \<approx>A"
   264   obtains Y where "Y \<in> UNIV // \<approx>A" and "Y ;; {[c]} \<subseteq> X" and "s \<in> Y"
   257   obtains Y where "Y \<in> UNIV // \<approx>A" and "Y \<cdot> {[c]} \<subseteq> X" and "s \<in> Y"
   265 proof -
   258 proof -
   266   def Y \<equiv> "\<approx>A `` {s}"
   259   def Y \<equiv> "\<approx>A `` {s}"
   267   have "Y \<in> UNIV // \<approx>A" 
   260   have "Y \<in> UNIV // \<approx>A" 
   268     unfolding Y_def quotient_def by auto
   261     unfolding Y_def quotient_def by auto
   269   moreover
   262   moreover
   270   have "X = \<approx>A `` {s @ [c]}" 
   263   have "X = \<approx>A `` {s @ [c]}" 
   271     using has_str in_CS defined_by_str by blast
   264     using has_str in_CS defined_by_str by blast
   272   then have "Y ;; {[c]} \<subseteq> X" 
   265   then have "Y \<cdot> {[c]} \<subseteq> X" 
   273     unfolding Y_def Image_def Seq_def
   266     unfolding Y_def Image_def Seq_def
   274     unfolding str_eq_rel_def
   267     unfolding str_eq_rel_def
   275     by clarsimp
   268     by clarsimp
   276   moreover
   269   moreover
   277   have "s \<in> Y" unfolding Y_def 
   270   have "s \<in> Y" unfolding Y_def 
   279   ultimately show thesis using that by blast
   272   ultimately show thesis using that by blast
   280 qed
   273 qed
   281 
   274 
   282 lemma l_eq_r_in_eqs:
   275 lemma l_eq_r_in_eqs:
   283   assumes X_in_eqs: "(X, rhs) \<in> Init (UNIV // \<approx>A)"
   276   assumes X_in_eqs: "(X, rhs) \<in> Init (UNIV // \<approx>A)"
   284   shows "X = L rhs"
   277   shows "X = L_rhs rhs"
   285 proof 
   278 proof 
   286   show "X \<subseteq> L rhs"
   279   show "X \<subseteq> L_rhs rhs"
   287   proof
   280   proof
   288     fix x
   281     fix x
   289     assume in_X: "x \<in> X"
   282     assume in_X: "x \<in> X"
   290     { assume empty: "x = []"
   283     { assume empty: "x = []"
   291       then have "x \<in> L rhs" using X_in_eqs in_X
   284       then have "x \<in> L_rhs rhs" using X_in_eqs in_X
   292 	unfolding Init_def Init_rhs_def
   285 	unfolding Init_def Init_rhs_def
   293         by auto
   286         by auto
   294     }
   287     }
   295     moreover
   288     moreover
   296     { assume not_empty: "x \<noteq> []"
   289     { assume not_empty: "x \<noteq> []"
   297       then obtain s c where decom: "x = s @ [c]"
   290       then obtain s c where decom: "x = s @ [c]"
   298 	using rev_cases by blast
   291 	using rev_cases by blast
   299       have "X \<in> UNIV // \<approx>A" using X_in_eqs unfolding Init_def by auto
   292       have "X \<in> UNIV // \<approx>A" using X_in_eqs unfolding Init_def by auto
   300       then obtain Y where "Y \<in> UNIV // \<approx>A" "Y ;; {[c]} \<subseteq> X" "s \<in> Y"
   293       then obtain Y where "Y \<in> UNIV // \<approx>A" "Y \<cdot> {[c]} \<subseteq> X" "s \<in> Y"
   301         using decom in_X every_eqclass_has_transition by blast
   294         using decom in_X every_eqclass_has_transition by blast
   302       then have "x \<in> L {Trn Y (CHAR c)| Y c. Y \<in> UNIV // \<approx>A \<and> Y \<Turnstile>c\<Rightarrow> X}"
   295       then have "x \<in> L_rhs {Trn Y (CHAR c)| Y c. Y \<in> UNIV // \<approx>A \<and> Y \<Turnstile>c\<Rightarrow> X}"
   303         unfolding transition_def
   296         unfolding transition_def
   304 	using decom by (force simp add: Seq_def)
   297 	using decom by (force simp add: Seq_def)
   305       then have "x \<in> L rhs" using X_in_eqs in_X
   298       then have "x \<in> L_rhs rhs" using X_in_eqs in_X
   306 	unfolding Init_def Init_rhs_def by simp
   299 	unfolding Init_def Init_rhs_def by simp
   307     }
   300     }
   308     ultimately show "x \<in> L rhs" by blast
   301     ultimately show "x \<in> L_rhs rhs" by blast
   309   qed
   302   qed
   310 next
   303 next
   311   show "L rhs \<subseteq> X" using X_in_eqs
   304   show "L_rhs rhs \<subseteq> X" using X_in_eqs
   312     unfolding Init_def Init_rhs_def transition_def
   305     unfolding Init_def Init_rhs_def transition_def
   313     by auto 
   306     by auto 
   314 qed
   307 qed
   315 
   308 
   316 lemma test:
   309 lemma test:
   317   assumes X_in_eqs: "(X, rhs) \<in> Init (UNIV // \<approx>A)"
   310   assumes X_in_eqs: "(X, rhs) \<in> Init (UNIV // \<approx>A)"
   318   shows "X = \<Union> (L `  rhs)"
   311   shows "X = \<Union> (L_trm `  rhs)"
   319 using assms l_eq_r_in_eqs by (simp)
   312 using assms l_eq_r_in_eqs by (simp)
   320 
   313 
   321 lemma finite_Init_rhs: 
   314 lemma finite_Init_rhs: 
   322   assumes finite: "finite CS"
   315   assumes finite: "finite CS"
   323   shows "finite (Init_rhs CS X)"
   316   shows "finite (Init_rhs CS X)"
   324 proof-
   317 proof-
   325   def S \<equiv> "{(Y, c)| Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X}" 
   318   def S \<equiv> "{(Y, c)| Y c. Y \<in> CS \<and> Y \<cdot> {[c]} \<subseteq> X}" 
   326   def h \<equiv> "\<lambda> (Y, c). Trn Y (CHAR c)"
   319   def h \<equiv> "\<lambda> (Y, c). Trn Y (CHAR c)"
   327   have "finite (CS \<times> (UNIV::char set))" using finite by auto
   320   have "finite (CS \<times> (UNIV::char set))" using finite by auto
   328   then have "finite S" using S_def 
   321   then have "finite S" using S_def 
   329     by (rule_tac B = "CS \<times> UNIV" in finite_subset) (auto)
   322     by (rule_tac B = "CS \<times> UNIV" in finite_subset) (auto)
   330   moreover have "{Trn Y (CHAR c) |Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X} = h ` S"
   323   moreover have "{Trn Y (CHAR c) |Y c. Y \<in> CS \<and> Y \<cdot> {[c]} \<subseteq> X} = h ` S"
   331     unfolding S_def h_def image_def by auto
   324     unfolding S_def h_def image_def by auto
   332   ultimately
   325   ultimately
   333   have "finite {Trn Y (CHAR c) |Y c. Y \<in> CS \<and> Y ;; {[c]} \<subseteq> X}" by auto
   326   have "finite {Trn Y (CHAR c) |Y c. Y \<in> CS \<and> Y \<cdot> {[c]} \<subseteq> X}" by auto
   334   then show "finite (Init_rhs CS X)" unfolding Init_rhs_def transition_def by simp
   327   then show "finite (Init_rhs CS X)" unfolding Init_rhs_def transition_def by simp
   335 qed
   328 qed
   336 
   329 
   337 lemma Init_ES_satisfies_invariant:
   330 lemma Init_ES_satisfies_invariant:
   338   assumes finite_CS: "finite (UNIV // \<approx>A)"
   331   assumes finite_CS: "finite (UNIV // \<approx>A)"
   357 qed
   350 qed
   358 
   351 
   359 subsubsection {* Interation step *}
   352 subsubsection {* Interation step *}
   360 
   353 
   361 lemma Arden_keeps_eq:
   354 lemma Arden_keeps_eq:
   362   assumes l_eq_r: "X = L rhs"
   355   assumes l_eq_r: "X = L_rhs rhs"
   363   and not_empty: "ardenable rhs"
   356   and not_empty: "ardenable rhs"
   364   and finite: "finite rhs"
   357   and finite: "finite rhs"
   365   shows "X = L (Arden X rhs)"
   358   shows "X = L_rhs (Arden X rhs)"
   366 proof -
   359 proof -
   367   def A \<equiv> "L (\<Uplus>{r. Trn X r \<in> rhs})"
   360   def A \<equiv> "L_rexp (\<Uplus>{r. Trn X r \<in> rhs})"
   368   def b \<equiv> "{Trn X r | r. Trn X r \<in> rhs}"
   361   def b \<equiv> "{Trn X r | r. Trn X r \<in> rhs}"
   369   def B \<equiv> "L (rhs - b)"
   362   def B \<equiv> "L_rhs (rhs - b)"
   370   have not_empty2: "[] \<notin> A" 
   363   have not_empty2: "[] \<notin> A" 
   371     using finite_Trn[OF finite] not_empty
   364     using finite_Trn[OF finite] not_empty
   372     unfolding A_def ardenable_def by simp
   365     unfolding A_def ardenable_def by simp
   373   have "X = L rhs" using l_eq_r by simp
   366   have "X = L_rhs rhs" using l_eq_r by simp
   374   also have "\<dots> = L (b \<union> (rhs - b))" unfolding b_def by auto
   367   also have "\<dots> = L_rhs (b \<union> (rhs - b))" unfolding b_def by auto
   375   also have "\<dots> = L b \<union> B" unfolding B_def by (simp only: L_rhs_union_distrib)
   368   also have "\<dots> = L_rhs b \<union> B" unfolding B_def by (simp only: L_rhs_union_distrib)
   376   also have "\<dots> = X ;; A \<union> B"
   369   also have "\<dots> = X \<cdot> A \<union> B"
   377     unfolding b_def
   370     unfolding b_def
   378     unfolding rhs_trm_soundness[OF finite]
   371     unfolding trm_soundness[OF finite]
   379     unfolding A_def
   372     unfolding A_def
   380     by blast
   373     by blast
   381   finally have "X = X ;; A \<union> B" . 
   374   finally have "X = X \<cdot> A \<union> B" . 
   382   then have "X = B ;; A\<star>"
   375   then have "X = B \<cdot> A\<star>"
   383     by (simp add: arden[OF not_empty2])
   376     by (simp add: arden[OF not_empty2])
   384   also have "\<dots> = L (Arden X rhs)"
   377   also have "\<dots> = L_rhs (Arden X rhs)"
   385     unfolding Arden_def A_def B_def b_def
   378     unfolding Arden_def A_def B_def b_def
   386     by (simp only: lang_of_append_rexp_rhs L_rexp.simps)
   379     by (simp only: lang_of_append_rexp_rhs L_rexp.simps)
   387   finally show "X = L (Arden X rhs)" by simp
   380   finally show "X = L_rhs (Arden X rhs)" by simp
   388 qed 
   381 qed 
   389 
   382 
   390 lemma Append_keeps_finite:
   383 lemma Append_keeps_finite:
   391   "finite rhs \<Longrightarrow> finite (Append_rexp_rhs rhs r)"
   384   "finite rhs \<Longrightarrow> finite (Append_rexp_rhs rhs r)"
   392 by (auto simp:Append_rexp_rhs_def)
   385 by (auto simp:Append_rexp_rhs_def)
   416 lemma Subst_keeps_nonempty:
   409 lemma Subst_keeps_nonempty:
   417   "\<lbrakk>ardenable rhs; ardenable xrhs\<rbrakk> \<Longrightarrow> ardenable (Subst rhs X xrhs)"
   410   "\<lbrakk>ardenable rhs; ardenable xrhs\<rbrakk> \<Longrightarrow> ardenable (Subst rhs X xrhs)"
   418 by (simp only: Subst_def Append_keeps_nonempty nonempty_set_union nonempty_set_sub)
   411 by (simp only: Subst_def Append_keeps_nonempty nonempty_set_union nonempty_set_sub)
   419 
   412 
   420 lemma Subst_keeps_eq:
   413 lemma Subst_keeps_eq:
   421   assumes substor: "X = L xrhs"
   414   assumes substor: "X = L_rhs xrhs"
   422   and finite: "finite rhs"
   415   and finite: "finite rhs"
   423   shows "L (Subst rhs X xrhs) = L rhs" (is "?Left = ?Right")
   416   shows "L_rhs (Subst rhs X xrhs) = L_rhs rhs" (is "?Left = ?Right")
   424 proof-
   417 proof-
   425   def A \<equiv> "L (rhs - {Trn X r | r. Trn X r \<in> rhs})"
   418   def A \<equiv> "L_rhs (rhs - {Trn X r | r. Trn X r \<in> rhs})"
   426   have "?Left = A \<union> L (Append_rexp_rhs xrhs (\<Uplus>{r. Trn X r \<in> rhs}))"
   419   have "?Left = A \<union> L_rhs (Append_rexp_rhs xrhs (\<Uplus>{r. Trn X r \<in> rhs}))"
   427     unfolding Subst_def
   420     unfolding Subst_def
   428     unfolding L_rhs_union_distrib[symmetric]
   421     unfolding L_rhs_union_distrib[symmetric]
   429     by (simp add: A_def)
   422     by (simp add: A_def)
   430   moreover have "?Right = A \<union> L ({Trn X r | r. Trn X r \<in> rhs})"
   423   moreover have "?Right = A \<union> L_rhs {Trn X r | r. Trn X r \<in> rhs}"
   431   proof-
   424   proof-
   432     have "rhs = (rhs - {Trn X r | r. Trn X r \<in> rhs}) \<union> ({Trn X r | r. Trn X r \<in> rhs})" by auto
   425     have "rhs = (rhs - {Trn X r | r. Trn X r \<in> rhs}) \<union> ({Trn X r | r. Trn X r \<in> rhs})" by auto
   433     thus ?thesis 
   426     thus ?thesis 
   434       unfolding A_def
   427       unfolding A_def
   435       unfolding L_rhs_union_distrib
   428       unfolding L_rhs_union_distrib
   436       by simp
   429       by simp
   437   qed
   430   qed
   438   moreover have "L (Append_rexp_rhs xrhs (\<Uplus>{r. Trn X r \<in> rhs})) = L ({Trn X r | r. Trn X r \<in> rhs})" 
   431   moreover have "L_rhs (Append_rexp_rhs xrhs (\<Uplus>{r. Trn X r \<in> rhs})) = L_rhs {Trn X r | r. Trn X r \<in> rhs}" 
   439     using finite substor by (simp only: lang_of_append_rexp_rhs rhs_trm_soundness)
   432     using finite substor by (simp only: lang_of_append_rexp_rhs trm_soundness)
   440   ultimately show ?thesis by simp
   433   ultimately show ?thesis by simp
   441 qed
   434 qed
   442 
   435 
   443 lemma Subst_keeps_finite_rhs:
   436 lemma Subst_keeps_finite_rhs:
   444   "\<lbrakk>finite rhs; finite yrhs\<rbrakk> \<Longrightarrow> finite (Subst rhs Y yrhs)"
   437   "\<lbrakk>finite rhs; finite yrhs\<rbrakk> \<Longrightarrow> finite (Subst rhs Y yrhs)"
   517 
   510 
   518 lemma Subst_all_satisfies_invariant:
   511 lemma Subst_all_satisfies_invariant:
   519   assumes invariant_ES: "invariant (ES \<union> {(Y, yrhs)})"
   512   assumes invariant_ES: "invariant (ES \<union> {(Y, yrhs)})"
   520   shows "invariant (Subst_all ES Y (Arden Y yrhs))"
   513   shows "invariant (Subst_all ES Y (Arden Y yrhs))"
   521 proof (rule invariantI)
   514 proof (rule invariantI)
   522   have Y_eq_yrhs: "Y = L yrhs" 
   515   have Y_eq_yrhs: "Y = L_rhs yrhs" 
   523     using invariant_ES by (simp only:invariant_def soundness_def, blast)
   516     using invariant_ES by (simp only:invariant_def soundness_def, blast)
   524    have finite_yrhs: "finite yrhs" 
   517    have finite_yrhs: "finite yrhs" 
   525     using invariant_ES by (auto simp:invariant_def finite_rhs_def)
   518     using invariant_ES by (auto simp:invariant_def finite_rhs_def)
   526   have nonempty_yrhs: "ardenable yrhs" 
   519   have nonempty_yrhs: "ardenable yrhs" 
   527     using invariant_ES by (auto simp:invariant_def ardenable_all_def)
   520     using invariant_ES by (auto simp:invariant_def ardenable_all_def)
   528   show "soundness (Subst_all ES Y (Arden Y yrhs))"
   521   show "soundness (Subst_all ES Y (Arden Y yrhs))"
   529   proof -
   522   proof -
   530     have "Y = L (Arden Y yrhs)" 
   523     have "Y = L_rhs (Arden Y yrhs)" 
   531       using Y_eq_yrhs invariant_ES finite_yrhs
   524       using Y_eq_yrhs invariant_ES finite_yrhs
   532       using finite_Trn[OF finite_yrhs]
   525       using finite_Trn[OF finite_yrhs]
   533       apply(rule_tac Arden_keeps_eq)
   526       apply(rule_tac Arden_keeps_eq)
   534       apply(simp_all)
   527       apply(simp_all)
   535       unfolding invariant_def ardenable_all_def ardenable_def
   528       unfolding invariant_def ardenable_all_def ardenable_def
   721 qed
   714 qed
   722 
   715 
   723 lemma every_eqcl_has_reg:
   716 lemma every_eqcl_has_reg:
   724   assumes finite_CS: "finite (UNIV // \<approx>A)"
   717   assumes finite_CS: "finite (UNIV // \<approx>A)"
   725   and X_in_CS: "X \<in> (UNIV // \<approx>A)"
   718   and X_in_CS: "X \<in> (UNIV // \<approx>A)"
   726   shows "\<exists>r::rexp. X = L r" 
   719   shows "\<exists>r. X = L_rexp r" 
   727 proof -
   720 proof -
   728   from finite_CS X_in_CS 
   721   from finite_CS X_in_CS 
   729   obtain xrhs where Inv_ES: "invariant {(X, xrhs)}"
   722   obtain xrhs where Inv_ES: "invariant {(X, xrhs)}"
   730     using Solve by metis
   723     using Solve by metis
   731 
   724 
   740   
   733   
   741   have "finite A" using Inv_ES unfolding A_def invariant_def finite_rhs_def
   734   have "finite A" using Inv_ES unfolding A_def invariant_def finite_rhs_def
   742     using Arden_keeps_finite by auto
   735     using Arden_keeps_finite by auto
   743   then have fin: "finite {r. Lam r \<in> A}" by (rule finite_Lam)
   736   then have fin: "finite {r. Lam r \<in> A}" by (rule finite_Lam)
   744 
   737 
   745   have "X = L xrhs" using Inv_ES unfolding invariant_def soundness_def
   738   have "X = L_rhs xrhs" using Inv_ES unfolding invariant_def soundness_def
   746     by simp
   739     by simp
   747   then have "X = L A" using Inv_ES 
   740   then have "X = L_rhs A" using Inv_ES 
   748     unfolding A_def invariant_def ardenable_all_def finite_rhs_def 
   741     unfolding A_def invariant_def ardenable_all_def finite_rhs_def 
   749     by (rule_tac Arden_keeps_eq) (simp_all add: finite_Trn)
   742     by (rule_tac Arden_keeps_eq) (simp_all add: finite_Trn)
   750   then have "X = L {Lam r | r. Lam r \<in> A}" using eq by simp
   743   then have "X = L_rhs {Lam r | r. Lam r \<in> A}" using eq by simp
   751   then have "X = L (\<Uplus>{r. Lam r \<in> A})" using fin by auto
   744   then have "X = L_rexp (\<Uplus>{r. Lam r \<in> A})" using fin by auto
   752   then show "\<exists>r::rexp. X = L r" by blast
   745   then show "\<exists>r. X = L_rexp r" by blast
   753 qed
   746 qed
   754 
   747 
   755 lemma bchoice_finite_set:
   748 lemma bchoice_finite_set:
   756   assumes a: "\<forall>x \<in> S. \<exists>y. x = f y" 
   749   assumes a: "\<forall>x \<in> S. \<exists>y. x = f y" 
   757   and     b: "finite S"
   750   and     b: "finite S"
   762 apply(auto)
   755 apply(auto)
   763 done
   756 done
   764 
   757 
   765 theorem Myhill_Nerode1:
   758 theorem Myhill_Nerode1:
   766   assumes finite_CS: "finite (UNIV // \<approx>A)"
   759   assumes finite_CS: "finite (UNIV // \<approx>A)"
   767   shows   "\<exists>r::rexp. A = L r"
   760   shows   "\<exists>r. A = L_rexp r"
   768 proof -
   761 proof -
   769   have fin: "finite (finals A)" 
   762   have fin: "finite (finals A)" 
   770     using finals_in_partitions finite_CS by (rule finite_subset)
   763     using finals_in_partitions finite_CS by (rule finite_subset)
   771   have "\<forall>X \<in> (UNIV // \<approx>A). \<exists>r::rexp. X = L r" 
   764   have "\<forall>X \<in> (UNIV // \<approx>A). \<exists>r. X = L_rexp r" 
   772     using finite_CS every_eqcl_has_reg by blast
   765     using finite_CS every_eqcl_has_reg by blast
   773   then have a: "\<forall>X \<in> finals A. \<exists>r::rexp. X = L r"
   766   then have a: "\<forall>X \<in> finals A. \<exists>r. X = L_rexp r"
   774     using finals_in_partitions by auto
   767     using finals_in_partitions by auto
   775   then obtain rs::"rexp set" where "\<Union> (finals A) = \<Union>(L ` rs)" "finite rs"
   768   then obtain rs::"rexp set" where "\<Union> (finals A) = \<Union>(L_rexp ` rs)" "finite rs"
   776     using fin by (auto dest: bchoice_finite_set)
   769     using fin by (auto dest: bchoice_finite_set)
   777   then have "A = L (\<Uplus>rs)" 
   770   then have "A = L_rexp (\<Uplus>rs)" 
   778     unfolding lang_is_union_of_finals[symmetric] by simp
   771     unfolding lang_is_union_of_finals[symmetric] by simp
   779   then show "\<exists>r::rexp. A = L r" by blast
   772   then show "\<exists>r. A = L_rexp r" by blast
   780 qed 
   773 qed 
   781 
   774 
   782 
   775 
   783 end
   776 end