Myhill_1.thy
changeset 170 b1258b7d2789
parent 166 7743d2ad71d1
child 179 edacc141060f
equal deleted inserted replaced
169:b794db0b79db 170:b1258b7d2789
     1 theory Myhill_1
     1 theory Myhill_1
     2 imports Regular
     2 imports More_Regular_Set
     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 
    10 by simp
    10 by simp
    11 
    11 
    12 text {* Myhill-Nerode relation *}
    12 text {* Myhill-Nerode relation *}
    13 
    13 
    14 definition
    14 definition
    15   str_eq_rel :: "lang \<Rightarrow> (string \<times> string) set" ("\<approx>_" [100] 100)
    15   str_eq_rel :: "'a lang \<Rightarrow> ('a list \<times> 'a list) set" ("\<approx>_" [100] 100)
    16 where
    16 where
    17   "\<approx>A \<equiv> {(x, y).  (\<forall>z. x @ z \<in> A \<longleftrightarrow> y @ z \<in> A)}"
    17   "\<approx>A \<equiv> {(x, y).  (\<forall>z. x @ z \<in> A \<longleftrightarrow> y @ z \<in> A)}"
    18 
    18 
    19 definition 
    19 definition 
    20   finals :: "lang \<Rightarrow> lang set"
    20   finals :: "'a lang \<Rightarrow> 'a lang set"
    21 where
    21 where
    22   "finals A \<equiv> {\<approx>A `` {s} | s . s \<in> A}"
    22   "finals A \<equiv> {\<approx>A `` {s} | s . s \<in> A}"
    23 
    23 
    24 lemma lang_is_union_of_finals: 
    24 lemma lang_is_union_of_finals: 
    25   shows "A = \<Union> finals A"
    25   shows "A = \<Union> finals A"
    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 trm = 
    40 datatype 'a trm = 
    41    Lam "rexp"            (* Lambda-marker *)
    41    Lam "'a rexp"            (* Lambda-marker *)
    42  | Trn "lang" "rexp"     (* Transition *)
    42  | Trn "'a lang" "'a rexp"     (* Transition *)
    43 
    43 
    44 fun 
    44 fun 
    45   L_trm::"trm \<Rightarrow> lang"
    45   lang_trm::"'a trm \<Rightarrow> 'a lang"
    46 where
    46 where
    47   "L_trm (Lam r) = L_rexp r" 
    47   "lang_trm (Lam r) = lang r" 
    48 | "L_trm (Trn X r) = X \<cdot> L_rexp r"
    48 | "lang_trm (Trn X r) = X \<cdot> lang r"
    49 
    49 
    50 fun 
    50 fun 
    51   L_rhs::"trm set \<Rightarrow> lang"
    51   lang_rhs::"('a trm) set \<Rightarrow> 'a lang"
    52 where 
    52 where 
    53   "L_rhs rhs = \<Union> (L_trm ` rhs)"
    53   "lang_rhs rhs = \<Union> (lang_trm ` rhs)"
    54 
    54 
    55 lemma L_rhs_set:
    55 lemma lang_rhs_set:
    56   shows "L_rhs {Trn X r | r. P r} = \<Union>{L_trm (Trn X r) | r. P r}"
    56   shows "lang_rhs {Trn X r | r. P r} = \<Union>{lang_trm (Trn X r) | r. P r}"
    57 by (auto)
    57 by (auto)
    58 
    58 
    59 lemma L_rhs_union_distrib:
    59 lemma lang_rhs_union_distrib:
    60   fixes A B::"trm set"
    60   fixes A B::"('a trm) set"
    61   shows "L_rhs A \<union> L_rhs B = L_rhs (A \<union> B)"
    61   shows "lang_rhs A \<union> lang_rhs B = lang_rhs (A \<union> B)"
    62 by simp
    62 by simp
    63 
    63 
    64 
    64 
    65 text {* Transitions between equivalence classes *}
    65 text {* Transitions between equivalence classes *}
    66 
    66 
    67 definition 
    67 definition 
    68   transition :: "lang \<Rightarrow> char \<Rightarrow> lang \<Rightarrow> bool" ("_ \<Turnstile>_\<Rightarrow>_" [100,100,100] 100)
    68   transition :: "'a lang \<Rightarrow> 'a \<Rightarrow> 'a lang \<Rightarrow> bool" ("_ \<Turnstile>_\<Rightarrow>_" [100,100,100] 100)
    69 where
    69 where
    70   "Y \<Turnstile>c\<Rightarrow> X \<equiv> Y \<cdot> {[c]} \<subseteq> X"
    70   "Y \<Turnstile>c\<Rightarrow> X \<equiv> Y \<cdot> {[c]} \<subseteq> X"
    71 
    71 
    72 text {* Initial equational system *}
    72 text {* Initial equational system *}
    73 
    73 
    74 definition
    74 definition
    75   "Init_rhs CS X \<equiv>  
    75   "Init_rhs CS X \<equiv>  
    76       if ([] \<in> X) then 
    76       if ([] \<in> X) then 
    77           {Lam EMPTY} \<union> {Trn Y (CHAR c) | Y c. Y \<in> CS \<and> Y \<Turnstile>c\<Rightarrow> X}
    77           {Lam One} \<union> {Trn Y (Atom c) | Y c. Y \<in> CS \<and> Y \<Turnstile>c\<Rightarrow> X}
    78       else 
    78       else 
    79           {Trn Y (CHAR c)| Y c. Y \<in> CS \<and> Y \<Turnstile>c\<Rightarrow> X}"
    79           {Trn Y (Atom c)| Y c. Y \<in> CS \<and> Y \<Turnstile>c\<Rightarrow> X}"
    80 
    80 
    81 definition 
    81 definition 
    82   "Init CS \<equiv> {(X, Init_rhs CS X) | X.  X \<in> CS}"
    82   "Init CS \<equiv> {(X, Init_rhs CS X) | X.  X \<in> CS}"
    83 
    83 
    84 
    84 
    85 section {* Arden Operation on equations *}
    85 section {* Arden Operation on equations *}
    86 
    86 
    87 fun 
    87 fun 
    88   Append_rexp :: "rexp \<Rightarrow> trm \<Rightarrow> trm"
    88   Append_rexp :: "'a rexp \<Rightarrow> 'a trm \<Rightarrow> 'a trm"
    89 where
    89 where
    90   "Append_rexp r (Lam rexp)   = Lam (SEQ rexp r)"
    90   "Append_rexp r (Lam rexp)   = Lam (Times rexp r)"
    91 | "Append_rexp r (Trn X rexp) = Trn X (SEQ rexp r)"
    91 | "Append_rexp r (Trn X rexp) = Trn X (Times rexp r)"
    92 
    92 
    93 
    93 
    94 definition
    94 definition
    95   "Append_rexp_rhs rhs rexp \<equiv> (Append_rexp rexp) ` rhs"
    95   "Append_rexp_rhs rhs rexp \<equiv> (Append_rexp rexp) ` rhs"
    96 
    96 
    97 definition 
    97 definition 
    98   "Arden X rhs \<equiv> 
    98   "Arden X rhs \<equiv> 
    99      Append_rexp_rhs (rhs - {Trn X r | r. Trn X r \<in> rhs}) (STAR (\<Uplus> {r. Trn X r \<in> rhs}))"
    99      Append_rexp_rhs (rhs - {Trn X r | r. Trn X r \<in> rhs}) (Star (\<Uplus> {r. Trn X r \<in> rhs}))"
   100 
   100 
   101 
   101 
   102 section {* Substitution Operation on equations *}
   102 section {* Substitution Operation on equations *}
   103 
   103 
   104 definition 
   104 definition 
   105   "Subst rhs X xrhs \<equiv> 
   105   "Subst rhs X xrhs \<equiv> 
   106         (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}))"
   107 
   107 
   108 definition
   108 definition
   109   Subst_all :: "(lang \<times> trm set) set \<Rightarrow> lang \<Rightarrow> trm set \<Rightarrow> (lang \<times> trm set) set"
   109   Subst_all :: "('a lang \<times> ('a trm) set) set \<Rightarrow> 'a lang \<Rightarrow> ('a trm) set \<Rightarrow> ('a lang \<times> ('a trm) set) set"
   110 where
   110 where
   111   "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}"
   112 
   112 
   113 definition
   113 definition
   114   "Remove ES X xrhs \<equiv> 
   114   "Remove ES X xrhs \<equiv> 
   141 definition 
   141 definition 
   142   "distinctness ES \<equiv> 
   142   "distinctness ES \<equiv> 
   143      \<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'"
   144 
   144 
   145 definition 
   145 definition 
   146   "soundness ES \<equiv> \<forall>(X, rhs) \<in> ES. X = L_rhs rhs"
   146   "soundness ES \<equiv> \<forall>(X, rhs) \<in> ES. X = lang_rhs rhs"
   147 
   147 
   148 definition 
   148 definition 
   149   "ardenable rhs \<equiv> (\<forall> Y r. Trn Y r \<in> rhs \<longrightarrow> [] \<notin> L_rexp r)"
   149   "ardenable rhs \<equiv> (\<forall> Y r. Trn Y r \<in> rhs \<longrightarrow> [] \<notin> lang r)"
   150 
   150 
   151 definition 
   151 definition 
   152   "ardenable_all ES \<equiv> \<forall>(X, rhs) \<in> ES. ardenable rhs"
   152   "ardenable_all ES \<equiv> \<forall>(X, rhs) \<in> ES. ardenable rhs"
   153 
   153 
   154 definition
   154 definition
   221     done
   221     done
   222 qed
   222 qed
   223 
   223 
   224 lemma trm_soundness:
   224 lemma trm_soundness:
   225   assumes finite:"finite rhs"
   225   assumes finite:"finite 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}))"
   226   shows "lang_rhs ({Trn X r| r. Trn X r \<in> rhs}) = X \<cdot> (lang (\<Uplus>{r. Trn X r \<in> rhs}))"
   227 proof -
   227 proof -
   228   have "finite {r. Trn X r \<in> rhs}" 
   228   have "finite {r. Trn X r \<in> rhs}" 
   229     by (rule finite_Trn[OF finite]) 
   229     by (rule finite_Trn[OF finite]) 
   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}))"
   230   then show "lang_rhs ({Trn X r| r. Trn X r \<in> rhs}) = X \<cdot> (lang (\<Uplus>{r. Trn X r \<in> rhs}))"
   231     by (simp only: L_rhs_set L_trm.simps) (auto simp add: Seq_def)
   231     by (simp only: lang_rhs_set lang_trm.simps) (auto simp add: conc_def)
   232 qed
   232 qed
   233 
   233 
   234 lemma lang_of_append_rexp:
   234 lemma lang_of_append_rexp:
   235   "L_trm (Append_rexp r trm) = L_trm trm \<cdot> L_rexp r"
   235   "lang_trm (Append_rexp r trm) = lang_trm trm \<cdot> lang r"
   236 by (induct rule: Append_rexp.induct)
   236 by (induct rule: Append_rexp.induct)
   237    (auto simp add: seq_assoc)
   237    (auto simp add: conc_assoc)
   238 
   238 
   239 lemma lang_of_append_rexp_rhs:
   239 lemma lang_of_append_rexp_rhs:
   240   "L_rhs (Append_rexp_rhs rhs r) = L_rhs rhs \<cdot> L_rexp r"
   240   "lang_rhs (Append_rexp_rhs rhs r) = lang_rhs rhs \<cdot> lang r"
   241 unfolding Append_rexp_rhs_def
   241 unfolding Append_rexp_rhs_def
   242 by (auto simp add: Seq_def lang_of_append_rexp)
   242 by (auto simp add: conc_def lang_of_append_rexp)
   243 
   243 
   244 
   244 
   245 subsubsection {* Intial Equational System *}
   245 subsubsection {* Intial Equational System *}
   246 
   246 
   247 lemma defined_by_str:
   247 lemma defined_by_str:
   261     unfolding Y_def quotient_def by auto
   261     unfolding Y_def quotient_def by auto
   262   moreover
   262   moreover
   263   have "X = \<approx>A `` {s @ [c]}" 
   263   have "X = \<approx>A `` {s @ [c]}" 
   264     using has_str in_CS defined_by_str by blast
   264     using has_str in_CS defined_by_str by blast
   265   then have "Y \<cdot> {[c]} \<subseteq> X" 
   265   then have "Y \<cdot> {[c]} \<subseteq> X" 
   266     unfolding Y_def Image_def Seq_def
   266     unfolding Y_def Image_def conc_def
   267     unfolding str_eq_rel_def
   267     unfolding str_eq_rel_def
   268     by clarsimp
   268     by clarsimp
   269   moreover
   269   moreover
   270   have "s \<in> Y" unfolding Y_def 
   270   have "s \<in> Y" unfolding Y_def 
   271     unfolding Image_def str_eq_rel_def by simp
   271     unfolding Image_def str_eq_rel_def by simp
   272   ultimately show thesis using that by blast
   272   ultimately show thesis using that by blast
   273 qed
   273 qed
   274 
   274 
   275 lemma l_eq_r_in_eqs:
   275 lemma l_eq_r_in_eqs:
   276   assumes X_in_eqs: "(X, rhs) \<in> Init (UNIV // \<approx>A)"
   276   assumes X_in_eqs: "(X, rhs) \<in> Init (UNIV // \<approx>A)"
   277   shows "X = L_rhs rhs"
   277   shows "X = lang_rhs rhs"
   278 proof 
   278 proof 
   279   show "X \<subseteq> L_rhs rhs"
   279   show "X \<subseteq> lang_rhs rhs"
   280   proof
   280   proof
   281     fix x
   281     fix x
   282     assume in_X: "x \<in> X"
   282     assume in_X: "x \<in> X"
   283     { assume empty: "x = []"
   283     { assume empty: "x = []"
   284       then have "x \<in> L_rhs rhs" using X_in_eqs in_X
   284       then have "x \<in> lang_rhs rhs" using X_in_eqs in_X
   285 	unfolding Init_def Init_rhs_def
   285 	unfolding Init_def Init_rhs_def
   286         by auto
   286         by auto
   287     }
   287     }
   288     moreover
   288     moreover
   289     { assume not_empty: "x \<noteq> []"
   289     { assume not_empty: "x \<noteq> []"
   290       then obtain s c where decom: "x = s @ [c]"
   290       then obtain s c where decom: "x = s @ [c]"
   291 	using rev_cases by blast
   291 	using rev_cases by blast
   292       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
   293       then obtain Y where "Y \<in> UNIV // \<approx>A" "Y \<cdot> {[c]} \<subseteq> X" "s \<in> Y"
   293       then obtain Y where "Y \<in> UNIV // \<approx>A" "Y \<cdot> {[c]} \<subseteq> X" "s \<in> Y"
   294         using decom in_X every_eqclass_has_transition by blast
   294         using decom in_X every_eqclass_has_transition by metis
   295       then have "x \<in> L_rhs {Trn Y (CHAR c)| Y c. Y \<in> UNIV // \<approx>A \<and> Y \<Turnstile>c\<Rightarrow> X}"
   295       then have "x \<in> lang_rhs {Trn Y (Atom c)| Y c. Y \<in> UNIV // \<approx>A \<and> Y \<Turnstile>c\<Rightarrow> X}"
   296         unfolding transition_def
   296         unfolding transition_def
   297 	using decom by (force simp add: Seq_def)
   297 	using decom by (force simp add: conc_def)
   298       then have "x \<in> L_rhs rhs" using X_in_eqs in_X
   298       then have "x \<in> lang_rhs rhs" using X_in_eqs in_X
   299 	unfolding Init_def Init_rhs_def by simp
   299 	unfolding Init_def Init_rhs_def by simp
   300     }
   300     }
   301     ultimately show "x \<in> L_rhs rhs" by blast
   301     ultimately show "x \<in> lang_rhs rhs" by blast
   302   qed
   302   qed
   303 next
   303 next
   304   show "L_rhs rhs \<subseteq> X" using X_in_eqs
   304   show "lang_rhs rhs \<subseteq> X" using X_in_eqs
   305     unfolding Init_def Init_rhs_def transition_def
   305     unfolding Init_def Init_rhs_def transition_def
   306     by auto 
   306     by auto 
   307 qed
   307 qed
   308 
   308 
   309 lemma test:
       
   310   assumes X_in_eqs: "(X, rhs) \<in> Init (UNIV // \<approx>A)"
       
   311   shows "X = \<Union> (L_trm `  rhs)"
       
   312 using assms l_eq_r_in_eqs by (simp)
       
   313 
   309 
   314 lemma finite_Init_rhs: 
   310 lemma finite_Init_rhs: 
       
   311   fixes CS::"(('a::finite) lang) set"
   315   assumes finite: "finite CS"
   312   assumes finite: "finite CS"
   316   shows "finite (Init_rhs CS X)"
   313   shows "finite (Init_rhs CS X)"
   317 proof-
   314 proof-
   318   def S \<equiv> "{(Y, c)| Y c. Y \<in> CS \<and> Y \<cdot> {[c]} \<subseteq> X}" 
   315   def S \<equiv> "{(Y, c)| Y c::'a. Y \<in> CS \<and> Y \<cdot> {[c]} \<subseteq> X}" 
   319   def h \<equiv> "\<lambda> (Y, c). Trn Y (CHAR c)"
   316   def h \<equiv> "\<lambda> (Y, c::'a). Trn Y (Atom c)"
   320   have "finite (CS \<times> (UNIV::char set))" using finite by auto
   317   have "finite (CS \<times> (UNIV::('a::finite) set))" using finite by auto
   321   then have "finite S" using S_def 
   318   then have "finite S" using S_def 
   322     by (rule_tac B = "CS \<times> UNIV" in finite_subset) (auto)
   319     by (rule_tac B = "CS \<times> UNIV" in finite_subset) (auto)
   323   moreover have "{Trn Y (CHAR c) |Y c. Y \<in> CS \<and> Y \<cdot> {[c]} \<subseteq> X} = h ` S"
   320   moreover have "{Trn Y (Atom c) |Y c::'a. Y \<in> CS \<and> Y \<cdot> {[c]} \<subseteq> X} = h ` S"
   324     unfolding S_def h_def image_def by auto
   321     unfolding S_def h_def image_def by auto
   325   ultimately
   322   ultimately
   326   have "finite {Trn Y (CHAR c) |Y c. Y \<in> CS \<and> Y \<cdot> {[c]} \<subseteq> X}" by auto
   323   have "finite {Trn Y (Atom c) |Y c. Y \<in> CS \<and> Y \<cdot> {[c]} \<subseteq> X}" by auto
   327   then show "finite (Init_rhs CS X)" unfolding Init_rhs_def transition_def by simp
   324   then show "finite (Init_rhs CS X)" unfolding Init_rhs_def transition_def by simp
   328 qed
   325 qed
   329 
   326 
       
   327 
   330 lemma Init_ES_satisfies_invariant:
   328 lemma Init_ES_satisfies_invariant:
       
   329   fixes A::"(('a::finite) lang)"
   331   assumes finite_CS: "finite (UNIV // \<approx>A)"
   330   assumes finite_CS: "finite (UNIV // \<approx>A)"
   332   shows "invariant (Init (UNIV // \<approx>A))"
   331   shows "invariant (Init (UNIV // \<approx>A))"
   333 proof (rule invariantI)
   332 proof (rule invariantI)
   334   show "soundness (Init (UNIV // \<approx>A))"
   333   show "soundness (Init (UNIV // \<approx>A))"
   335     unfolding soundness_def 
   334     unfolding soundness_def 
   350 qed
   349 qed
   351 
   350 
   352 subsubsection {* Interation step *}
   351 subsubsection {* Interation step *}
   353 
   352 
   354 lemma Arden_keeps_eq:
   353 lemma Arden_keeps_eq:
   355   assumes l_eq_r: "X = L_rhs rhs"
   354   assumes l_eq_r: "X = lang_rhs rhs"
   356   and not_empty: "ardenable rhs"
   355   and not_empty: "ardenable rhs"
   357   and finite: "finite rhs"
   356   and finite: "finite rhs"
   358   shows "X = L_rhs (Arden X rhs)"
   357   shows "X = lang_rhs (Arden X rhs)"
   359 proof -
   358 proof -
   360   def A \<equiv> "L_rexp (\<Uplus>{r. Trn X r \<in> rhs})"
   359   def A \<equiv> "lang (\<Uplus>{r. Trn X r \<in> rhs})"
   361   def b \<equiv> "{Trn X r | r. Trn X r \<in> rhs}"
   360   def b \<equiv> "{Trn X r | r. Trn X r \<in> rhs}"
   362   def B \<equiv> "L_rhs (rhs - b)"
   361   def B \<equiv> "lang_rhs (rhs - b)"
   363   have not_empty2: "[] \<notin> A" 
   362   have not_empty2: "[] \<notin> A" 
   364     using finite_Trn[OF finite] not_empty
   363     using finite_Trn[OF finite] not_empty
   365     unfolding A_def ardenable_def by simp
   364     unfolding A_def ardenable_def by simp
   366   have "X = L_rhs rhs" using l_eq_r by simp
   365   have "X = lang_rhs rhs" using l_eq_r by simp
   367   also have "\<dots> = L_rhs (b \<union> (rhs - b))" unfolding b_def by auto
   366   also have "\<dots> = lang_rhs (b \<union> (rhs - b))" unfolding b_def by auto
   368   also have "\<dots> = L_rhs b \<union> B" unfolding B_def by (simp only: L_rhs_union_distrib)
   367   also have "\<dots> = lang_rhs b \<union> B" unfolding B_def by (simp only: lang_rhs_union_distrib)
   369   also have "\<dots> = X \<cdot> A \<union> B"
   368   also have "\<dots> = X \<cdot> A \<union> B"
   370     unfolding b_def
   369     unfolding b_def
   371     unfolding trm_soundness[OF finite]
   370     unfolding trm_soundness[OF finite]
   372     unfolding A_def
   371     unfolding A_def
   373     by blast
   372     by blast
   374   finally have "X = X \<cdot> A \<union> B" . 
   373   finally have "X = X \<cdot> A \<union> B" . 
   375   then have "X = B \<cdot> A\<star>"
   374   then have "X = B \<cdot> A\<star>"
   376     by (simp add: arden[OF not_empty2])
   375     by (simp add: arden[OF not_empty2])
   377   also have "\<dots> = L_rhs (Arden X rhs)"
   376   also have "\<dots> = lang_rhs (Arden X rhs)"
   378     unfolding Arden_def A_def B_def b_def
   377     unfolding Arden_def A_def B_def b_def
   379     by (simp only: lang_of_append_rexp_rhs L_rexp.simps)
   378     by (simp only: lang_of_append_rexp_rhs lang.simps)
   380   finally show "X = L_rhs (Arden X rhs)" by simp
   379   finally show "X = lang_rhs (Arden X rhs)" by simp
   381 qed 
   380 qed 
   382 
   381 
   383 lemma Append_keeps_finite:
   382 lemma Append_keeps_finite:
   384   "finite rhs \<Longrightarrow> finite (Append_rexp_rhs rhs r)"
   383   "finite rhs \<Longrightarrow> finite (Append_rexp_rhs rhs r)"
   385 by (auto simp:Append_rexp_rhs_def)
   384 by (auto simp: Append_rexp_rhs_def)
   386 
   385 
   387 lemma Arden_keeps_finite:
   386 lemma Arden_keeps_finite:
   388   "finite rhs \<Longrightarrow> finite (Arden X rhs)"
   387   "finite rhs \<Longrightarrow> finite (Arden X rhs)"
   389 by (auto simp:Arden_def Append_keeps_finite)
   388 by (auto simp: Arden_def Append_keeps_finite)
   390 
   389 
   391 lemma Append_keeps_nonempty:
   390 lemma Append_keeps_nonempty:
   392   "ardenable rhs \<Longrightarrow> ardenable (Append_rexp_rhs rhs r)"
   391   "ardenable rhs \<Longrightarrow> ardenable (Append_rexp_rhs rhs r)"
   393 apply (auto simp:ardenable_def Append_rexp_rhs_def)
   392 apply (auto simp: ardenable_def Append_rexp_rhs_def)
   394 by (case_tac x, auto simp:Seq_def)
   393 by (case_tac x, auto simp: conc_def)
   395 
   394 
   396 lemma nonempty_set_sub:
   395 lemma nonempty_set_sub:
   397   "ardenable rhs \<Longrightarrow> ardenable (rhs - A)"
   396   "ardenable rhs \<Longrightarrow> ardenable (rhs - A)"
   398 by (auto simp:ardenable_def)
   397 by (auto simp:ardenable_def)
   399 
   398 
   409 lemma Subst_keeps_nonempty:
   408 lemma Subst_keeps_nonempty:
   410   "\<lbrakk>ardenable rhs; ardenable xrhs\<rbrakk> \<Longrightarrow> ardenable (Subst rhs X xrhs)"
   409   "\<lbrakk>ardenable rhs; ardenable xrhs\<rbrakk> \<Longrightarrow> ardenable (Subst rhs X xrhs)"
   411 by (simp only: Subst_def Append_keeps_nonempty nonempty_set_union nonempty_set_sub)
   410 by (simp only: Subst_def Append_keeps_nonempty nonempty_set_union nonempty_set_sub)
   412 
   411 
   413 lemma Subst_keeps_eq:
   412 lemma Subst_keeps_eq:
   414   assumes substor: "X = L_rhs xrhs"
   413   assumes substor: "X = lang_rhs xrhs"
   415   and finite: "finite rhs"
   414   and finite: "finite rhs"
   416   shows "L_rhs (Subst rhs X xrhs) = L_rhs rhs" (is "?Left = ?Right")
   415   shows "lang_rhs (Subst rhs X xrhs) = lang_rhs rhs" (is "?Left = ?Right")
   417 proof-
   416 proof-
   418   def A \<equiv> "L_rhs (rhs - {Trn X r | r. Trn X r \<in> rhs})"
   417   def A \<equiv> "lang_rhs (rhs - {Trn X r | r. Trn X r \<in> rhs})"
   419   have "?Left = A \<union> L_rhs (Append_rexp_rhs xrhs (\<Uplus>{r. Trn X r \<in> rhs}))"
   418   have "?Left = A \<union> lang_rhs (Append_rexp_rhs xrhs (\<Uplus>{r. Trn X r \<in> rhs}))"
   420     unfolding Subst_def
   419     unfolding Subst_def
   421     unfolding L_rhs_union_distrib[symmetric]
   420     unfolding lang_rhs_union_distrib[symmetric]
   422     by (simp add: A_def)
   421     by (simp add: A_def)
   423   moreover have "?Right = A \<union> L_rhs {Trn X r | r. Trn X r \<in> rhs}"
   422   moreover have "?Right = A \<union> lang_rhs {Trn X r | r. Trn X r \<in> rhs}"
   424   proof-
   423   proof-
   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
   424     have "rhs = (rhs - {Trn X r | r. Trn X r \<in> rhs}) \<union> ({Trn X r | r. Trn X r \<in> rhs})" by auto
   426     thus ?thesis 
   425     thus ?thesis 
   427       unfolding A_def
   426       unfolding A_def
   428       unfolding L_rhs_union_distrib
   427       unfolding lang_rhs_union_distrib
   429       by simp
   428       by simp
   430   qed
   429   qed
   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}" 
   430   moreover 
       
   431   have "lang_rhs (Append_rexp_rhs xrhs (\<Uplus>{r. Trn X r \<in> rhs})) = lang_rhs {Trn X r | r. Trn X r \<in> rhs}" 
   432     using finite substor by (simp only: lang_of_append_rexp_rhs trm_soundness)
   432     using finite substor by (simp only: lang_of_append_rexp_rhs trm_soundness)
   433   ultimately show ?thesis by simp
   433   ultimately show ?thesis by simp
   434 qed
   434 qed
   435 
   435 
   436 lemma Subst_keeps_finite_rhs:
   436 lemma Subst_keeps_finite_rhs:
   439 
   439 
   440 lemma Subst_all_keeps_finite:
   440 lemma Subst_all_keeps_finite:
   441   assumes finite: "finite ES"
   441   assumes finite: "finite ES"
   442   shows "finite (Subst_all ES Y yrhs)"
   442   shows "finite (Subst_all ES Y yrhs)"
   443 proof -
   443 proof -
   444   def eqns \<equiv> "{(X::lang, rhs) |X rhs. (X, rhs) \<in> ES}"
   444   def eqns \<equiv> "{(X::'a lang, rhs) |X rhs. (X, rhs) \<in> ES}"
   445   def h \<equiv> "\<lambda>(X::lang, rhs). (X, Subst rhs Y yrhs)"
   445   def h \<equiv> "\<lambda>(X::'a lang, rhs). (X, Subst rhs Y yrhs)"
   446   have "finite (h ` eqns)" using finite h_def eqns_def by auto
   446   have "finite (h ` eqns)" using finite h_def eqns_def by auto
   447   moreover 
   447   moreover 
   448   have "Subst_all ES Y yrhs = h ` eqns" unfolding h_def eqns_def Subst_all_def by auto
   448   have "Subst_all ES Y yrhs = h ` eqns" unfolding h_def eqns_def Subst_all_def by auto
   449   ultimately
   449   ultimately
   450   show "finite (Subst_all ES Y yrhs)" by simp
   450   show "finite (Subst_all ES Y yrhs)" by simp
   454   "\<lbrakk>finite_rhs ES; finite yrhs\<rbrakk> \<Longrightarrow> finite_rhs (Subst_all ES Y yrhs)"
   454   "\<lbrakk>finite_rhs ES; finite yrhs\<rbrakk> \<Longrightarrow> finite_rhs (Subst_all ES Y yrhs)"
   455 by (auto intro:Subst_keeps_finite_rhs simp add:Subst_all_def finite_rhs_def)
   455 by (auto intro:Subst_keeps_finite_rhs simp add:Subst_all_def finite_rhs_def)
   456 
   456 
   457 lemma append_rhs_keeps_cls:
   457 lemma append_rhs_keeps_cls:
   458   "rhss (Append_rexp_rhs rhs r) = rhss rhs"
   458   "rhss (Append_rexp_rhs rhs r) = rhss rhs"
   459 apply (auto simp:rhss_def Append_rexp_rhs_def)
   459 apply (auto simp: rhss_def Append_rexp_rhs_def)
   460 apply (case_tac xa, auto simp:image_def)
   460 apply (case_tac xa, auto simp: image_def)
   461 by (rule_tac x = "SEQ ra r" in exI, rule_tac x = "Trn x ra" in bexI, simp+)
   461 by (rule_tac x = "Times ra r" in exI, rule_tac x = "Trn x ra" in bexI, simp+)
   462 
   462 
   463 lemma Arden_removes_cl:
   463 lemma Arden_removes_cl:
   464   "rhss (Arden Y yrhs) = rhss yrhs - {Y}"
   464   "rhss (Arden Y yrhs) = rhss yrhs - {Y}"
   465 apply (simp add:Arden_def append_rhs_keeps_cls)
   465 apply (simp add:Arden_def append_rhs_keeps_cls)
   466 by (auto simp:rhss_def)
   466 by (auto simp: rhss_def)
   467 
   467 
   468 lemma lhss_keeps_cls:
   468 lemma lhss_keeps_cls:
   469   "lhss (Subst_all ES Y yrhs) = lhss ES"
   469   "lhss (Subst_all ES Y yrhs) = lhss ES"
   470 by (auto simp:lhss_def Subst_all_def)
   470 by (auto simp: lhss_def Subst_all_def)
   471 
   471 
   472 lemma Subst_updates_cls:
   472 lemma Subst_updates_cls:
   473   "X \<notin> rhss xrhs \<Longrightarrow> 
   473   "X \<notin> rhss xrhs \<Longrightarrow> 
   474       rhss (Subst rhs X xrhs) = rhss rhs \<union> rhss xrhs - {X}"
   474       rhss (Subst rhs X xrhs) = rhss rhs \<union> rhss xrhs - {X}"
   475 apply (simp only:Subst_def append_rhs_keeps_cls rhss_union_distrib)
   475 apply (simp only:Subst_def append_rhs_keeps_cls rhss_union_distrib)
   476 by (auto simp:rhss_def)
   476 by (auto simp: rhss_def)
   477 
   477 
   478 lemma Subst_all_keeps_validity:
   478 lemma Subst_all_keeps_validity:
   479   assumes sc: "validity (ES \<union> {(Y, yrhs)})"        (is "validity ?A")
   479   assumes sc: "validity (ES \<union> {(Y, yrhs)})"        (is "validity ?A")
   480   shows "validity (Subst_all ES Y (Arden Y yrhs))"  (is "validity ?B")
   480   shows "validity (Subst_all ES Y (Arden Y yrhs))"  (is "validity ?B")
   481 proof -
   481 proof -
   488     proof-
   488     proof-
   489       have "lhss ?B = lhss ES" by (auto simp add:lhss_def Subst_all_def)
   489       have "lhss ?B = lhss ES" by (auto simp add:lhss_def Subst_all_def)
   490       moreover have "rhss xrhs' \<subseteq> lhss ES"
   490       moreover have "rhss xrhs' \<subseteq> lhss ES"
   491       proof-
   491       proof-
   492         have "rhss xrhs' \<subseteq>  rhss xrhs \<union> rhss (Arden Y yrhs) - {Y}"
   492         have "rhss xrhs' \<subseteq>  rhss xrhs \<union> rhss (Arden Y yrhs) - {Y}"
   493         proof-
   493         proof -
   494           have "Y \<notin> rhss (Arden Y yrhs)" 
   494           have "Y \<notin> rhss (Arden Y yrhs)" 
   495             using Arden_removes_cl by simp
   495             using Arden_removes_cl by auto
   496           thus ?thesis using xrhs_xrhs' by (auto simp:Subst_updates_cls)
   496           thus ?thesis using xrhs_xrhs' by (auto simp: Subst_updates_cls)
   497         qed
   497         qed
   498         moreover have "rhss xrhs \<subseteq> lhss ES \<union> {Y}" using X_in sc
   498         moreover have "rhss xrhs \<subseteq> lhss ES \<union> {Y}" using X_in sc
   499           apply (simp only:validity_def lhss_union_distrib)
   499           apply (simp only:validity_def lhss_union_distrib)
   500           by (drule_tac x = "(X, xrhs)" in bspec, auto simp:lhss_def)
   500           by (drule_tac x = "(X, xrhs)" in bspec, auto simp:lhss_def)
   501         moreover have "rhss (Arden Y yrhs) \<subseteq> lhss ES \<union> {Y}" 
   501         moreover have "rhss (Arden Y yrhs) \<subseteq> lhss ES \<union> {Y}" 
   502           using sc 
   502           using sc 
   503           by (auto simp add:Arden_removes_cl validity_def lhss_def)
   503           by (auto simp add: Arden_removes_cl validity_def lhss_def)
   504         ultimately show ?thesis by auto
   504         ultimately show ?thesis by auto
   505       qed
   505       qed
   506       ultimately show ?thesis by simp
   506       ultimately show ?thesis by simp
   507     qed
   507     qed
   508   } thus ?thesis by (auto simp only:Subst_all_def validity_def)
   508   } thus ?thesis by (auto simp only:Subst_all_def validity_def)
   510 
   510 
   511 lemma Subst_all_satisfies_invariant:
   511 lemma Subst_all_satisfies_invariant:
   512   assumes invariant_ES: "invariant (ES \<union> {(Y, yrhs)})"
   512   assumes invariant_ES: "invariant (ES \<union> {(Y, yrhs)})"
   513   shows "invariant (Subst_all ES Y (Arden Y yrhs))"
   513   shows "invariant (Subst_all ES Y (Arden Y yrhs))"
   514 proof (rule invariantI)
   514 proof (rule invariantI)
   515   have Y_eq_yrhs: "Y = L_rhs yrhs" 
   515   have Y_eq_yrhs: "Y = lang_rhs yrhs" 
   516     using invariant_ES by (simp only:invariant_def soundness_def, blast)
   516     using invariant_ES by (simp only:invariant_def soundness_def, blast)
   517    have finite_yrhs: "finite yrhs" 
   517    have finite_yrhs: "finite yrhs" 
   518     using invariant_ES by (auto simp:invariant_def finite_rhs_def)
   518     using invariant_ES by (auto simp:invariant_def finite_rhs_def)
   519   have nonempty_yrhs: "ardenable yrhs" 
   519   have nonempty_yrhs: "ardenable yrhs" 
   520     using invariant_ES by (auto simp:invariant_def ardenable_all_def)
   520     using invariant_ES by (auto simp:invariant_def ardenable_all_def)
   521   show "soundness (Subst_all ES Y (Arden Y yrhs))"
   521   show "soundness (Subst_all ES Y (Arden Y yrhs))"
   522   proof -
   522   proof -
   523     have "Y = L_rhs (Arden Y yrhs)" 
   523     have "Y = lang_rhs (Arden Y yrhs)" 
   524       using Y_eq_yrhs invariant_ES finite_yrhs
   524       using Y_eq_yrhs invariant_ES finite_yrhs
   525       using finite_Trn[OF finite_yrhs]
   525       using finite_Trn[OF finite_yrhs]
   526       apply(rule_tac Arden_keeps_eq)
   526       apply(rule_tac Arden_keeps_eq)
   527       apply(simp_all)
   527       apply(simp_all)
   528       unfolding invariant_def ardenable_all_def ardenable_def
   528       unfolding invariant_def ardenable_all_def ardenable_def
   529       apply(auto)
   529       apply(auto)
   530       done
   530       done
   531     thus ?thesis using invariant_ES
   531     thus ?thesis using invariant_ES
   532       unfolding invariant_def finite_rhs_def2 soundness_def Subst_all_def
   532       unfolding invariant_def finite_rhs_def2 soundness_def Subst_all_def
   533       by (auto simp add: Subst_keeps_eq simp del: L_rhs.simps)
   533       by (auto simp add: Subst_keeps_eq simp del: lang_rhs.simps)
   534   qed
   534   qed
   535   show "finite (Subst_all ES Y (Arden Y yrhs))" 
   535   show "finite (Subst_all ES Y (Arden Y yrhs))" 
   536     using invariant_ES by (simp add:invariant_def Subst_all_keeps_finite)
   536     using invariant_ES by (simp add:invariant_def Subst_all_keeps_finite)
   537   show "distinctness (Subst_all ES Y (Arden Y yrhs))" 
   537   show "distinctness (Subst_all ES Y (Arden Y yrhs))" 
   538     using invariant_ES 
   538     using invariant_ES 
   555       by (simp add:invariant_def finite_rhs_def)
   555       by (simp add:invariant_def finite_rhs_def)
   556     moreover have "finite (Arden Y yrhs)"
   556     moreover have "finite (Arden Y yrhs)"
   557     proof -
   557     proof -
   558       have "finite yrhs" using invariant_ES 
   558       have "finite yrhs" using invariant_ES 
   559         by (auto simp:invariant_def finite_rhs_def)
   559         by (auto simp:invariant_def finite_rhs_def)
   560       thus ?thesis using Arden_keeps_finite by simp
   560       thus ?thesis using Arden_keeps_finite by auto
   561     qed
   561     qed
   562     ultimately show ?thesis 
   562     ultimately show ?thesis 
   563       by (simp add:Subst_all_keeps_finite_rhs)
   563       by (simp add:Subst_all_keeps_finite_rhs)
   564   qed
   564   qed
   565   show "validity (Subst_all ES Y (Arden Y yrhs))"
   565   show "validity (Subst_all ES Y (Arden Y yrhs))"
   566     using invariant_ES Subst_all_keeps_validity by (simp add:invariant_def)
   566     using invariant_ES Subst_all_keeps_validity by (auto simp add: invariant_def)
   567 qed
   567 qed
   568 
   568 
   569 lemma Remove_in_card_measure:
   569 lemma Remove_in_card_measure:
   570   assumes finite: "finite ES"
   570   assumes finite: "finite ES"
   571   and     in_ES: "(X, rhs) \<in> ES"
   571   and     in_ES: "(X, rhs) \<in> ES"
   572   shows "(Remove ES X rhs, ES) \<in> measure card"
   572   shows "(Remove ES X rhs, ES) \<in> measure card"
   573 proof -
   573 proof -
   574   def f \<equiv> "\<lambda> x. ((fst x)::lang, Subst (snd x) X (Arden X rhs))"
   574   def f \<equiv> "\<lambda> x. ((fst x)::'a lang, Subst (snd x) X (Arden X rhs))"
   575   def ES' \<equiv> "ES - {(X, rhs)}"
   575   def ES' \<equiv> "ES - {(X, rhs)}"
   576   have "Subst_all ES' X (Arden X rhs) = f ` ES'" 
   576   have "Subst_all ES' X (Arden X rhs) = f ` ES'" 
   577     apply (auto simp: Subst_all_def f_def image_def)
   577     apply (auto simp: Subst_all_def f_def image_def)
   578     by (rule_tac x = "(Y, yrhs)" in bexI, simp+)
   578     by (rule_tac x = "(Y, yrhs)" in bexI, simp+)
   579   then have "card (Subst_all ES' X (Arden X rhs)) \<le> card ES'"
   579   then have "card (Subst_all ES' X (Arden X rhs)) \<le> card ES'"
   672 
   672 
   673 
   673 
   674 subsubsection {* Conclusion of the proof *}
   674 subsubsection {* Conclusion of the proof *}
   675 
   675 
   676 lemma Solve:
   676 lemma Solve:
       
   677   fixes A::"('a::finite) lang"
   677   assumes fin: "finite (UNIV // \<approx>A)"
   678   assumes fin: "finite (UNIV // \<approx>A)"
   678   and     X_in: "X \<in> (UNIV // \<approx>A)"
   679   and     X_in: "X \<in> (UNIV // \<approx>A)"
   679   shows "\<exists>rhs. Solve X (Init (UNIV // \<approx>A)) = {(X, rhs)} \<and> invariant {(X, rhs)}"
   680   shows "\<exists>rhs. Solve X (Init (UNIV // \<approx>A)) = {(X, rhs)} \<and> invariant {(X, rhs)}"
   680 proof -
   681 proof -
   681   def Inv \<equiv> "\<lambda>ES. invariant ES \<and> (\<exists>rhs. (X, rhs) \<in> ES)"
   682   def Inv \<equiv> "\<lambda>ES. invariant ES \<and> (\<exists>rhs. (X, rhs) \<in> ES)"
   712   show "\<exists>rhs. Solve X (Init (UNIV // \<approx>A)) = {(X, rhs)} \<and> invariant {(X, rhs)}" 
   713   show "\<exists>rhs. Solve X (Init (UNIV // \<approx>A)) = {(X, rhs)} \<and> invariant {(X, rhs)}" 
   713     unfolding Solve_def by (rule while_rule)
   714     unfolding Solve_def by (rule while_rule)
   714 qed
   715 qed
   715 
   716 
   716 lemma every_eqcl_has_reg:
   717 lemma every_eqcl_has_reg:
       
   718   fixes A::"('a::finite) lang"
   717   assumes finite_CS: "finite (UNIV // \<approx>A)"
   719   assumes finite_CS: "finite (UNIV // \<approx>A)"
   718   and X_in_CS: "X \<in> (UNIV // \<approx>A)"
   720   and X_in_CS: "X \<in> (UNIV // \<approx>A)"
   719   shows "\<exists>r. X = L_rexp r" 
   721   shows "\<exists>r. X = lang r" 
   720 proof -
   722 proof -
   721   from finite_CS X_in_CS 
   723   from finite_CS X_in_CS 
   722   obtain xrhs where Inv_ES: "invariant {(X, xrhs)}"
   724   obtain xrhs where Inv_ES: "invariant {(X, xrhs)}"
   723     using Solve by metis
   725     using Solve by metis
   724 
   726 
   733   
   735   
   734   have "finite A" using Inv_ES unfolding A_def invariant_def finite_rhs_def
   736   have "finite A" using Inv_ES unfolding A_def invariant_def finite_rhs_def
   735     using Arden_keeps_finite by auto
   737     using Arden_keeps_finite by auto
   736   then have fin: "finite {r. Lam r \<in> A}" by (rule finite_Lam)
   738   then have fin: "finite {r. Lam r \<in> A}" by (rule finite_Lam)
   737 
   739 
   738   have "X = L_rhs xrhs" using Inv_ES unfolding invariant_def soundness_def
   740   have "X = lang_rhs xrhs" using Inv_ES unfolding invariant_def soundness_def
   739     by simp
   741     by simp
   740   then have "X = L_rhs A" using Inv_ES 
   742   then have "X = lang_rhs A" using Inv_ES 
   741     unfolding A_def invariant_def ardenable_all_def finite_rhs_def 
   743     unfolding A_def invariant_def ardenable_all_def finite_rhs_def 
   742     by (rule_tac Arden_keeps_eq) (simp_all add: finite_Trn)
   744     by (rule_tac Arden_keeps_eq) (simp_all add: finite_Trn)
   743   then have "X = L_rhs {Lam r | r. Lam r \<in> A}" using eq by simp
   745   then have "X = lang_rhs {Lam r | r. Lam r \<in> A}" using eq by simp
   744   then have "X = L_rexp (\<Uplus>{r. Lam r \<in> A})" using fin by auto
   746   then have "X = lang (\<Uplus>{r. Lam r \<in> A})" using fin by auto
   745   then show "\<exists>r. X = L_rexp r" by blast
   747   then show "\<exists>r. X = lang r" by blast
   746 qed
   748 qed
   747 
   749 
   748 lemma bchoice_finite_set:
   750 lemma bchoice_finite_set:
   749   assumes a: "\<forall>x \<in> S. \<exists>y. x = f y" 
   751   assumes a: "\<forall>x \<in> S. \<exists>y. x = f y" 
   750   and     b: "finite S"
   752   and     b: "finite S"
   754 apply(rule_tac x="fa ` S" in exI)
   756 apply(rule_tac x="fa ` S" in exI)
   755 apply(auto)
   757 apply(auto)
   756 done
   758 done
   757 
   759 
   758 theorem Myhill_Nerode1:
   760 theorem Myhill_Nerode1:
       
   761   fixes A::"('a::finite) lang"
   759   assumes finite_CS: "finite (UNIV // \<approx>A)"
   762   assumes finite_CS: "finite (UNIV // \<approx>A)"
   760   shows   "\<exists>r. A = L_rexp r"
   763   shows   "\<exists>r. A = lang r"
   761 proof -
   764 proof -
   762   have fin: "finite (finals A)" 
   765   have fin: "finite (finals A)" 
   763     using finals_in_partitions finite_CS by (rule finite_subset)
   766     using finals_in_partitions finite_CS by (rule finite_subset)
   764   have "\<forall>X \<in> (UNIV // \<approx>A). \<exists>r. X = L_rexp r" 
   767   have "\<forall>X \<in> (UNIV // \<approx>A). \<exists>r. X = lang r" 
   765     using finite_CS every_eqcl_has_reg by blast
   768     using finite_CS every_eqcl_has_reg by blast
   766   then have a: "\<forall>X \<in> finals A. \<exists>r. X = L_rexp r"
   769   then have a: "\<forall>X \<in> finals A. \<exists>r. X = lang r"
   767     using finals_in_partitions by auto
   770     using finals_in_partitions by auto
   768   then obtain rs::"rexp set" where "\<Union> (finals A) = \<Union>(L_rexp ` rs)" "finite rs"
   771   then obtain rs::"('a rexp) set" where "\<Union> (finals A) = \<Union>(lang ` rs)" "finite rs"
   769     using fin by (auto dest: bchoice_finite_set)
   772     using fin by (auto dest: bchoice_finite_set)
   770   then have "A = L_rexp (\<Uplus>rs)" 
   773   then have "A = lang (\<Uplus>rs)" 
   771     unfolding lang_is_union_of_finals[symmetric] by simp
   774     unfolding lang_is_union_of_finals[symmetric] by simp
   772   then show "\<exists>r. A = L_rexp r" by blast
   775   then show "\<exists>r. A = lang r" by blast
   773 qed 
   776 qed 
   774 
   777 
   775 
   778 
   776 end
   779 end