diff -r 779e1d9fbf3e -r 86167563a1ed MyhillNerode.thy --- a/MyhillNerode.thy Wed Oct 20 14:11:14 2010 +0000 +++ b/MyhillNerode.thy Thu Oct 21 13:42:08 2010 +0000 @@ -7,7 +7,7 @@ definition lang_seq :: "string set \ string set \ string set" ("_ ; _" [100,100] 100) where - "L1 ; L2 = {s1@s2 | s1 s2. s1 \ L1 \ s2 \ L2}" + "L1 ; L2 = {s1 @ s2 | s1 s2. s1 \ L1 \ s2 \ L2}" lemma lang_seq_empty: shows "{[]} ; L = L" @@ -33,8 +33,10 @@ lemma lang_seq_assoc: shows "(L1 ; L2) ; L3 = L1 ; (L2 ; L3)" -by (simp add: lang_seq_def Collect_def mem_def expand_fun_eq) - (metis append_assoc) +unfolding lang_seq_def +apply(auto) +apply(metis) +by (metis append_assoc) lemma lang_seq_minus: shows "(L1; L2) - {[]} = @@ -245,485 +247,8 @@ | "L_rexp (STAR r) = (L_rexp r)\" end -definition - Ls :: "rexp set \ string set" -where - "Ls R = (\r\R. (L r))" -lemma Ls_union: - "Ls (R1 \ R2) = (Ls R1) \ (Ls R2)" -unfolding Ls_def by auto - -text {* helper function for termination proofs *} -fun - Left :: "rexp \ rexp" -where - "Left (SEQ r1 r2) = r1" - -text {* dagger function *} - -function - dagger :: "rexp \ char \ rexp list" ("_ \ _") -where - c1: "(NULL \ c) = []" -| c2: "(EMPTY) \ c = []" -| c3: "(CHAR c') \ c = (if c = c' then [EMPTY] else [])" -| c4: "(ALT r1 r2) \ c = r1 \ c @ r2 \ c" -| c5: "(SEQ NULL r2) \ c = []" -| c6: "(SEQ EMPTY r2) \ c = r2 \ c" -| c7: "(SEQ (CHAR c') r2) \ c = (if c = c' then [r2] else [])" -| c8: "(SEQ (SEQ r11 r12) r2) \ c = (SEQ r11 (SEQ r12 r2)) \ c" -| c9: "(SEQ (ALT r11 r12) r2) \ c = (SEQ r11 r2) \ c @ (SEQ r12 r2) \ c" -| c10: "(SEQ (STAR r1) r2) \ c = r2 \ c @ [SEQ r' (SEQ (STAR r1) r2). r' \ r1 \ c]" -| c11: "(STAR r) \ c = [SEQ r' (STAR r) . r' \ r \ c]" -by (pat_completeness) (auto) - -termination dagger - by (relation "measures [\(r, c). size r, \(r, c). size (Left r)]") (simp_all) - -lemma dagger_correctness: - "Ls (set r \ c) = {s. c#s \ L r}" -proof (induct rule: dagger.induct) - case (1 c) - show "Ls (set NULL \ c) = {s. c#s \ L NULL}" by (simp add: Ls_def) -next - case (2 c) - show "Ls (set EMPTY \ c) = {s. c#s \ L EMPTY}" by (simp add: Ls_def) -next - case (3 c' c) - show "Ls (set CHAR c' \ c) = {s. c#s \ L (CHAR c')}" by (simp add: Ls_def) -next - case (4 r1 r2 c) - have ih1: "Ls (set r1 \ c) = {s. c#s \ L r1}" by fact - have ih2: "Ls (set r2 \ c) = {s. c#s \ L r2}" by fact - show "Ls (set ALT r1 r2 \ c) = {s. c#s \ L (ALT r1 r2)}" - by (simp add: Ls_union ih1 ih2 Collect_disj_eq) -next - case (5 r2 c) - show "Ls (set SEQ NULL r2 \ c) = {s. c#s \ L (SEQ NULL r2)}" by (simp add: Ls_def lang_seq_null) -next - case (6 r2 c) - have ih: "Ls (set r2 \ c) = {s. c#s \ L r2}" by fact - show "Ls (set SEQ EMPTY r2 \ c) = {s. c#s \ L (SEQ EMPTY r2)}" - by (simp add: ih lang_seq_empty) -next - case (7 c' r2 c) - show "Ls (set SEQ (CHAR c') r2 \ c) = {s. c#s \ L (SEQ (CHAR c') r2)}" - by (simp add: Ls_def lang_seq_def) -next - case (8 r11 r12 r2 c) - have ih: "Ls (set SEQ r11 (SEQ r12 r2) \ c) = {s. c#s \ L (SEQ r11 (SEQ r12 r2))}" by fact - show "Ls (set SEQ (SEQ r11 r12) r2 \ c) = {s. c#s \ L (SEQ (SEQ r11 r12) r2)}" - by (simp add: ih lang_seq_assoc) -next - case (9 r11 r12 r2 c) - have ih1: "Ls (set SEQ r11 r2 \ c) = {s. c#s \ L (SEQ r11 r2)}" by fact - have ih2: "Ls (set SEQ r12 r2 \ c) = {s. c#s \ L (SEQ r12 r2)}" by fact - show "Ls (set SEQ (ALT r11 r12) r2 \ c) = {s. c#s \ L (SEQ (ALT r11 r12) r2)}" - by (simp add: Ls_union ih1 ih2 lang_seq_union Collect_disj_eq) -next - case (10 r1 r2 c) - have ih2: "Ls (set r2 \ c) = {s. c#s \ L r2}" by fact - have ih1: "Ls (set r1 \ c) = {s. c#s \ L r1}" by fact - have "Ls (set SEQ (STAR r1) r2 \ c) = Ls (set r2 \ c) \ (Ls (set r1 \ c); ((L r1)\ ; L r2))" - by (auto simp add: lang_seq_def Ls_def) - also have "\ = {s. c#s \ L r2} \ ({s. c#s \ L r1} ; ((L r1)\ ; L r2))" using ih1 ih2 by simp - also have "\ = {s. c#s \ L r2} \ ({s. c#s \ L r1} ; (L r1)\) ; L r2" by (simp add: lang_seq_assoc) - also have "\ = {s. c#s \ L r2} \ {s. c#s \ (L r1)\} ; L r2" by (simp add: zzz) - also have "\ = {s. c#s \ L r2} \ {s. c#s \ (L r1)\ ; L r2}" - by (auto simp add: lang_seq_def Cons_eq_append_conv) - also have "\ = {s. c#s \ (L r1)\ ; L r2}" - by (force simp add: lang_seq_def) - finally show "Ls (set SEQ (STAR r1) r2 \ c) = {s. c#s \ L (SEQ (STAR r1) r2)}" by simp -next - case (11 r c) - have ih: "Ls (set r \ c) = {s. c#s \ L r}" by fact - have "Ls (set (STAR r) \ c) = Ls (set r \ c) ; (L r)\" - by (auto simp add: lang_seq_def Ls_def) - also have "\ = {s. c#s \ L r} ; (L r)\" using ih by simp - also have "\ = {s. c#s \ (L r)\}" using zzz by simp - finally show "Ls (set (STAR r) \ c) = {s. c#s \ L (STAR r)}" by simp -qed - - -text {* matcher function (based on the "list"-dagger function) *} -fun - first_True :: "bool list \ bool" -where - "first_True [] = False" -| "first_True (x#xs) = (if x then True else first_True xs)" - -lemma not_first_True[simp]: - shows "(\(first_True xs)) = (\x \ set xs. \x)" -by (induct xs) (auto) - -lemma first_True: - shows "(first_True xs) = (\x \ set xs. x)" -by (induct xs) (auto) - -text {* matcher function *} - -function - matcher :: "rexp \ string \ bool" ("_ ! _") -where - "NULL ! s = False" -| "EMPTY ! s = (s =[])" -| "CHAR c ! s = (s = [c])" -| "ALT r1 r2 ! s = (r1 ! s \ r2 ! s)" -| "STAR r ! [] = True" -| "STAR r ! c#s = first_True [SEQ (r') (STAR r) ! s. r' \ r \ c]" -| "SEQ r1 r2 ! [] = (r1 ! [] \ r2 ! [])" -| "SEQ NULL r2 ! (c#s) = False" -| "SEQ EMPTY r2 ! (c#s) = (r2 ! c#s)" -| "SEQ (CHAR c') r2 ! (c#s) = (if c'=c then r2 ! s else False)" -| "SEQ (SEQ r11 r12) r2 ! (c#s) = (SEQ r11 (SEQ r12 r2) ! c#s)" -| "SEQ (ALT r11 r12) r2 ! (c#s) = ((SEQ r11 r2) ! (c#s) \ (SEQ r12 r2) ! (c#s))" -| "SEQ (STAR r1) r2 ! (c#s) = (r2 ! (c#s) \ first_True [SEQ r' (SEQ (STAR r1) r2) ! s. r' \ r1 \ c])" -by (pat_completeness) (auto) - -termination matcher - by(relation "measures [\(r,s). length s, \(r,s). size r, \(r,s). size (Left r)]") (simp_all) - -text {* positive correctness of the matcher *} -lemma matcher1: - shows "r ! s \ s \ L r" -proof (induct rule: matcher.induct) - case (1 s) - have "NULL ! s" by fact - then show "s \ L NULL" by simp -next - case (2 s) - have "EMPTY ! s" by fact - then show "s \ L EMPTY" by simp -next - case (3 c s) - have "CHAR c ! s" by fact - then show "s \ L (CHAR c)" by simp -next - case (4 r1 r2 s) - have ih1: "r1 ! s \ s \ L r1" by fact - have ih2: "r2 ! s \ s \ L r2" by fact - have "ALT r1 r2 ! s" by fact - with ih1 ih2 show "s \ L (ALT r1 r2)" by auto -next - case (5 r) - have "STAR r ! []" by fact - then show "[] \ L (STAR r)" by auto -next - case (6 r c s) - have ih1: "\rx. \rx \ set r \ c; SEQ rx (STAR r) ! s\ \ s \ L (SEQ rx (STAR r))" by fact - have as: "STAR r ! c#s" by fact - then obtain r' where imp1: "r' \ set r \ c" and imp2: "SEQ r' (STAR r) ! s" - by (auto simp add: first_True) - from imp2 imp1 have "s \ L (SEQ r' (STAR r))" using ih1 by simp - then have "s \ L r' ; (L r)\" by simp - then have "s \ Ls (set r \ c) ; (L r)\" using imp1 by (auto simp add: Ls_def lang_seq_def) - then have "s \ {s. c#s \ L r} ; (L r)\" by (auto simp add: dagger_correctness) - then have "s \ {s. c#s \ (L r)\}" by (simp add: zzz) - then have "c#s \ {[c]}; {s. c#s \ (L r)\}" by (auto simp add: lang_seq_def) - then have "c#s \ (L r)\" by (auto simp add: lang_seq_def) - then show "c#s \ L (STAR r)" by simp -next - case (7 r1 r2) - have ih1: "r1 ! [] \ [] \ L r1" by fact - have ih2: "r2 ! [] \ [] \ L r2" by fact - have as: "SEQ r1 r2 ! []" by fact - then have "r1 ! [] \ r2 ! []" by simp - then show "[] \ L (SEQ r1 r2)" using ih1 ih2 by (simp add: lang_seq_def) -next - case (8 r2 c s) - have "SEQ NULL r2 ! c#s" by fact - then show "c#s \ L (SEQ NULL r2)" by simp -next - case (9 r2 c s) - have ih1: "r2 ! c#s \ c#s \ L r2" by fact - have "SEQ EMPTY r2 ! c#s" by fact - then show "c#s \ L (SEQ EMPTY r2)" using ih1 by (simp add: lang_seq_def) -next - case (10 c' r2 c s) - have ih1: "\c' = c; r2 ! s\ \ s \ L r2" by fact - have "SEQ (CHAR c') r2 ! c#s" by fact - then show "c#s \ L (SEQ (CHAR c') r2)" - using ih1 by (auto simp add: lang_seq_def split: if_splits) -next - case (11 r11 r12 r2 c s) - have ih1: "SEQ r11 (SEQ r12 r2) ! c#s \ c#s \ L (SEQ r11 (SEQ r12 r2))" by fact - have "SEQ (SEQ r11 r12) r2 ! c#s" by fact - then have "c#s \ L (SEQ r11 (SEQ r12 r2))" using ih1 by simp - then show "c#s \ L (SEQ (SEQ r11 r12) r2)" by (simp add: lang_seq_assoc) -next - case (12 r11 r12 r2 c s) - have ih1: "SEQ r11 r2 ! c#s \ c#s \ L (SEQ r11 r2)" by fact - have ih2: "SEQ r12 r2 ! c#s \ c#s \ L (SEQ r12 r2)" by fact - have "SEQ (ALT r11 r12) r2 ! c#s" by fact - then show "c#s \ L (SEQ (ALT r11 r12) r2)" - using ih1 ih2 by (auto simp add: lang_seq_union) -next - case (13 r1 r2 c s) - have ih1: "r2 ! c#s \ c#s \ L r2" by fact - have ih2: "\r'. \r' \ set r1 \ c; SEQ r' (SEQ (STAR r1) r2) ! s\ \ - s \ L (SEQ r' (SEQ (STAR r1) r2))" by fact - have "SEQ (STAR r1) r2 ! c#s" by fact - then have "(r2 ! c#s) \ (\r' \ set r1 \ c. SEQ r' (SEQ (STAR r1) r2) ! s)" by (auto simp add: first_True) - moreover - { assume "r2 ! c#s" - with ih1 have "c#s \ L r2" by simp - then have "c # s \ L r1\ ; L r2" - by (auto simp add: lang_seq_def) - then have "c#s \ L (SEQ (STAR r1) r2)" by simp - } - moreover - { assume "\r' \ set r1 \ c. SEQ r' (SEQ (STAR r1) r2) ! s" - then obtain r' where imp1: "r' \ set r1 \ c" and imp2: "SEQ r' (SEQ (STAR r1) r2) ! s" by blast - from imp2 imp1 have "s \ L (SEQ r' (SEQ (STAR r1) r2))" using ih2 by simp - then have "s \ L r' ; ((L r1)\ ; L r2)" by simp - then have "s \ Ls (set r1 \ c) ; ((L r1)\ ; L r2)" using imp1 by (auto simp add: Ls_def lang_seq_def) - then have "s \ {s. c#s \ L r1} ; ((L r1)\ ; L r2)" by (simp add: dagger_correctness) - then have "s \ ({s. c#s \ L r1} ; (L r1)\) ; L r2" by (simp add: lang_seq_assoc) - then have "s \ {s. c#s \ (L r1)\} ; L r2" by (simp add: zzz) - then have "c#s \ {[c]}; ({s. c#s \ (L r1)\}; L r2)" by (auto simp add: lang_seq_def) - then have "c#s \ ({[c]}; {s. c#s \ (L r1)\}) ; L r2" by (simp add: lang_seq_assoc) - then have "c#s \ (L r1)\; L r2" by (auto simp add: lang_seq_def) - then have "c#s \ L (SEQ (STAR r1) r2)" by simp - } - ultimately show "c#s \ L (SEQ (STAR r1) r2)" by blast -qed - -text {* negative correctness of the matcher *} -lemma matcher2: - shows "\ r ! s \ s \ L r" -proof (induct rule: matcher.induct) - case (1 s) - have "\ NULL ! s" by fact - then show "s \ L NULL" by simp -next - case (2 s) - have "\ EMPTY ! s" by fact - then show "s \ L EMPTY" by simp -next - case (3 c s) - have "\ CHAR c ! s" by fact - then show "s \ L (CHAR c)" by simp -next - case (4 r1 r2 s) - have ih2: "\ r1 ! s \ s \ L r1" by fact - have ih4: "\ r2 ! s \ s \ L r2" by fact - have "\ ALT r1 r2 ! s" by fact - then show "s \ L (ALT r1 r2)" by (simp add: ih2 ih4) -next - case (5 r) - have "\ STAR r ! []" by fact - then show "[] \ L (STAR r)" by simp -next - case (6 r c s) - have ih: "\rx. \rx \ set r \ c; \SEQ rx (STAR r) ! s\ \ s \ L (SEQ rx (STAR r))" by fact - have as: "\ STAR r ! c#s" by fact - then have "\r'\ set r \ c. \ (SEQ r' (STAR r) ! s)" by simp - then have "\r'\ set r \ c. s \ L (SEQ r' (STAR r))" using ih by auto - then have "\r'\ set r \ c. s \ L r' ; ((L r)\)" by simp - then have "s \ (Ls (set r \ c)) ; ((L r)\)" by (auto simp add: Ls_def lang_seq_def) - then have "s \ {s. c#s \ L r} ; ((L r)\)" by (simp add: dagger_correctness) - then have "s \ {s. c#s \ (L r)\}" by (simp add: zzz) - then have "c#s \ {[c]} ; {s. c#s \ (L r)\}" by (auto simp add: lang_seq_assoc lang_seq_def) - then have "c#s \ (L r)\" by (simp add: lang_seq_def) - then show "c#s \ L (STAR r)" by simp -next - case (7 r1 r2) - have ih2: "\ r1 ! [] \ [] \ L r1" by fact - have ih4: "\ r2 ! [] \ [] \ L r2" by fact - have "\ SEQ r1 r2 ! []" by fact - then have "\ r1 ! [] \ \ r2 ! []" by simp - then show "[] \ L (SEQ r1 r2)" using ih2 ih4 - by (auto simp add: lang_seq_def) -next - case (8 r2 c s) - have "\ SEQ NULL r2 ! c#s" by fact - then show "c#s \ L (SEQ NULL r2)" by (simp add: lang_seq_null) -next - case (9 r2 c s) - have ih1: "\ r2 ! c#s \ c#s \ L r2" by fact - have "\ SEQ EMPTY r2 ! c#s" by fact - then show "c#s \ L (SEQ EMPTY r2)" - using ih1 by (simp add: lang_seq_def) -next - case (10 c' r2 c s) - have ih2: "\c' = c; \r2 ! s\ \ s \ L r2" by fact - have "\ SEQ (CHAR c') r2 ! c#s" by fact - then show "c#s \ L (SEQ (CHAR c') r2)" - using ih2 by (auto simp add: lang_seq_def) -next - case (11 r11 r12 r2 c s) - have ih2: "\ SEQ r11 (SEQ r12 r2) ! c#s \ c#s \ L (SEQ r11 (SEQ r12 r2))" by fact - have "\ SEQ (SEQ r11 r12) r2 ! c#s" by fact - then show "c#s \ L (SEQ (SEQ r11 r12) r2)" - using ih2 by (auto simp add: lang_seq_def) -next - case (12 r11 r12 r2 c s) - have ih2: "\ SEQ r11 r2 ! c#s \ c#s \ L (SEQ r11 r2)" by fact - have ih4: "\ SEQ r12 r2 ! c#s \ c#s \ L (SEQ r12 r2)" by fact - have "\ SEQ (ALT r11 r12) r2 ! c#s" by fact - then show " c#s \ L (SEQ (ALT r11 r12) r2)" - using ih2 ih4 by (simp add: lang_seq_union) -next - case (13 r1 r2 c s) - have ih1: "\r2 ! c#s \ c#s \ L r2" by fact - have ih2: "\rx. \rx \ set r1 \ c; \ SEQ rx (SEQ (STAR r1) r2) ! s\ - \ s \ L (SEQ rx (SEQ (STAR r1) r2))" by fact - have as: "\ SEQ (STAR r1) r2 ! c#s" by fact - then have as1: "\r2 ! c#s" and as2: "\r1'\set r1 \ c. \ SEQ r1' (SEQ (STAR r1) r2) ! s" by simp_all - from as1 have bs1: "c#s \ L r2" using ih1 by simp - from as2 have "\r1'\set r1 \ c. \ SEQ r1' (SEQ (STAR r1) r2) ! s" by simp - then have "\r1'\set r1 \ c. s \ L (SEQ r1' (SEQ (STAR r1) r2))" using ih2 by simp - then have "\r1'\set r1 \ c. s \ L r1'; ((L r1)\; L r2)" by simp - then have "s \ (Ls (set r1 \ c)) ; ((L r1)\; L r2)" by (auto simp add: Ls_def lang_seq_def) - then have "s \ {s. c#s \ L r1} ; ((L r1)\; L r2)" by (simp add: dagger_correctness) - then have "s \ ({s. c#s \ L r1} ; (L r1)\); L r2" by (simp add: lang_seq_assoc) - then have "s \ {s. c#s \ (L r1)\}; L r2" by (simp add: zzz) - then have "c#s \ {[c]}; ({s. c#s \ (L r1)\}; L r2)" by (auto simp add: lang_seq_def) - then have "c#s \ (L r1)\; L r2" using bs1 by (auto simp add: lang_seq_def Cons_eq_append_conv) - then show "c#s \ L (SEQ (STAR r1) r2)" by simp -qed - -section {* Questions *} - -text {* - - Why was the key lemma k1 omitted; is there an easy, non-induction - way for obtaining this property? - - Why was False included in the definition of the STAR-clause in - the matcher? Has this something to do with executing the code? - -*} - -section {* Code *} - -export_code dagger in SML module_name Dagger file - -export_code matcher in SML module_name Dagger file - - -section {* Examples *} - -text {* since now everything is based on lists, the evaluation is quite fast *} - -value "NULL ! []" -value "(CHAR (CHR ''a'')) ! [CHR ''a'']" -value "((CHAR a) ! [a,a])" -value "(STAR (CHAR a)) ! []" -value "(STAR (CHAR a)) ! [a,a]" -value "(SEQ (CHAR (CHR ''a'')) (SEQ (STAR (CHAR (CHR ''b''))) (CHAR (CHR ''c'')))) ! ''abbbbc''" -value "(SEQ (CHAR (CHR ''a'')) (SEQ (STAR (CHAR (CHR ''b''))) (CHAR (CHR ''c'')))) ! ''abbcbbc''" - -section {* Slind et al's matcher based on derivatives *} - -fun - nullable :: "rexp \ bool" -where - "nullable (NULL) = False" -| "nullable (EMPTY) = True" -| "nullable (CHAR c) = False" -| "nullable (ALT r1 r2) = ((nullable r1) \ (nullable r2))" -| "nullable (SEQ r1 r2) = ((nullable r1) \ (nullable r2))" -| "nullable (STAR r) = True" - -lemma nullable: - shows "([] \ L r) = nullable r" -by (induct r) - (auto simp add: lang_seq_def) - -fun - der :: "char \ rexp \ rexp" -where - "der c (NULL) = NULL" -| "der c (EMPTY) = NULL" -| "der c (CHAR c') = (if c=c' then EMPTY else NULL)" -| "der c (ALT r1 r2) = ALT (der c r1) (der c r2)" -| "der c (SEQ r1 r2) = ALT (SEQ (der c r1) r2) (if nullable r1 then der c r2 else NULL)" -| "der c (STAR r) = SEQ (der c r) (STAR r)" - -lemma k2: - assumes b: "s \ L1\" - and a: "s \ []" - shows "s \ (L1; (L1\))" -using b a -apply(induct rule: Star.induct) -apply(simp) -apply(case_tac "s1=[]") -apply(simp) -apply(simp add: lang_seq_def) -apply(blast) -done - - -lemma der_correctness: - shows "(s \ L (der c r)) = ((c#s) \ L r)" -proof (induct r arbitrary: s) - case (NULL s) - show "(s \ L (der c NULL)) = (c#s \ L NULL)" by simp -next - case (EMPTY s) - show "(s \ L (der c EMPTY)) = (c#s \ L EMPTY)" by simp -next - case (CHAR c' s) - show "(s \ L (der c (CHAR c'))) = (c#s \ L (CHAR c'))" by simp -next - case (SEQ r1 r2 s) - have ih1: "\s. (s \ L (der c r1)) = (c#s \ L r1)" by fact - have ih2: "\s. (s \ L (der c r2)) = (c#s \ L r2)" by fact - show "(s \ L (der c (SEQ r1 r2))) = (c#s \ L (SEQ r1 r2))" - using ih1 ih2 - by (auto simp add: nullable[symmetric] lang_seq_def Cons_eq_append_conv) -next - case (ALT r1 r2 s) - have ih1: "\s. (s \ L (der c r1)) = (c#s \ L r1)" by fact - have ih2: "\s. (s \ L (der c r2)) = (c#s \ L r2)" by fact - show "(s \ L (der c (ALT r1 r2))) = (c#s \ L (ALT r1 r2))" - using ih1 ih2 by (auto simp add: lang_seq_def) -next - case (STAR r s) - have ih1: "\s. (s \ L (der c r)) = (c#s \ L r)" by fact - show "(s \ L (der c (STAR r))) = (c#s \ L (STAR r))" - using ih1 - apply(simp) - apply(auto simp add: lang_seq_def) - apply(drule lang_seq_append) - apply(assumption) - apply(simp) - apply(subst lang_star_cases) - apply(simp) - thm k1 - apply(drule k2) - apply(simp) - apply(simp add: lang_seq_def) - apply(erule exE)+ - apply(erule conjE)+ - apply(auto simp add: lang_seq_def Cons_eq_append_conv) - apply(drule k1) - apply(simp) - apply(simp add: lang_seq_def) - apply(erule exE)+ - apply(erule conjE)+ - apply(auto simp add: lang_seq_def Cons_eq_append_conv) - done -qed - -fun - derivative :: "string \ rexp \ rexp" -where - "derivative [] r = r" -| "derivative (c#s) r = derivative s (der c r)" - -fun - slind_matcher :: "rexp \ string \ bool" -where - "slind_matcher r s = nullable (derivative s r)" - -lemma slind_matcher: - shows "slind_matcher r s = (s \ L r)" -by (induct s arbitrary: r) - (auto simp add: nullable der_correctness) - -export_code slind_matcher in SML module_name Slind file - - - -(* ******************************************** now is the codes writen by chunhan ************************************* *) +(* ************ now is the codes writen by chunhan ************************************* *) section {* Arden's Lemma revised *} @@ -922,7 +447,6 @@ lemma seq_rhs_r_prop1: "L (seq_rhs_r rhs r) = (L rhs);(L r)" -apply (rule set_ext, rule iffI) apply (auto simp:seq_rhs_r_def image_def lang_seq_def) apply (rule_tac x = "s1 @ s1a" in exI, rule_tac x = "s2a" in exI, simp) apply (rule_tac x = a in exI, rule_tac x = b in exI, simp) @@ -936,13 +460,11 @@ lemma del_x_paired_prop1: "\distinct_rhs rhs; (X, r) \ rhs\ \ X ; L r \ L (del_x_paired rhs X) = L rhs" -apply (simp add:del_x_paired_def) -apply (rule set_ext, rule iffI, simp) -apply (erule disjE, rule_tac x = X in exI, rule_tac x = r in exI, simp) -apply (clarify, rule_tac x = Xa in exI, rule_tac x = ra in exI, simp) -apply (clarsimp, drule_tac x = Xa in spec, drule_tac x = ra in spec) -apply (auto simp:distinct_rhs_def) -done + apply (simp add:del_x_paired_def) + apply (simp add: distinct_rhs_def) + apply(auto simp add: lang_seq_def) + apply(metis) + done lemma arden_variate_prop: assumes "(X, rx) \ rhs"