diff -r 05e5d68c9627 -r f1be8028a4a9 Nominal/activities/cas09/Lec3.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Nominal/activities/cas09/Lec3.thy Wed Mar 30 17:27:34 2016 +0100 @@ -0,0 +1,437 @@ +(***************************************************************** + + Isabelle Tutorial + ----------------- + + 2st June 2009, Beijing + +*) + +theory Lec3 + imports "Main" +begin + + +definition + lang_seq :: "string set \ string set \ string set" ("_ ; _") +where + "L1 ; L2 = {s1@s2 | s1 s2. s1 \ L1 \ s2 \ L2}" + +fun + lang_pow :: "string set \ nat \ string set" ("_ \ _") +where + "L \ 0 = {[]}" +| "L \ (Suc i) = L ; (L \ i)" + +definition + lang_star :: "string set \ string set" ("_\") +where + "L\ \ \i. (L \ i)" + +lemma lang_seq_cases: + shows "(s \ L1 ; L2) = (\s1 s2. s = s1@s2 \ s1\L1 \ s2\L2)" +by (simp add: lang_seq_def) + +lemma lang_seq_union: + shows "(L1 \ L2);L3 = (L1;L3) \ (L2;L3)" + and "L1;(L2 \ L3) = (L1;L2) \ (L1;L3)" +unfolding lang_seq_def by auto + +lemma lang_seq_empty: + shows "{[]} ; L = L" + and "L ; {[]} = L" +unfolding lang_seq_def by auto + +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) + +lemma silly: + shows "[] \ L \ 0" +by simp + +lemma lang_star_empty: + shows "{[]} \ (L\) = L\" +unfolding lang_star_def +by (auto intro: silly) + +lemma lang_star_in_empty: + shows "[] \ L\" +unfolding lang_star_def +by (auto intro: silly) + +lemma lang_seq_subseteq: + shows "L \ (L'\) ; L" + and "L \ L ; (L'\)" +proof - + have "L = {[]} ; L" using lang_seq_empty by simp + also have "\ \ ({[]} ; L) \ ((L'\) ; L)" by auto + also have "\ = ({[]} \ (L'\)) ; L" by (simp add: lang_seq_union[symmetric]) + also have "\ = (L'\); L" using lang_star_empty by simp + finally show "L \ (L'\); L" by simp +next + show "L \ L ; (L'\)" sorry +qed + +lemma lang_star_subseteq: + shows "L ; (L\) \ (L\)" +unfolding lang_star_def lang_seq_def +apply(auto) +apply(rule_tac x="Suc xa" in exI) +apply(auto simp add: lang_seq_def) +done + +(* regular expressions *) + +datatype rexp = + EMPTY +| CHAR char +| SEQ rexp rexp +| ALT rexp rexp +| STAR rexp + +fun + L :: "rexp \ string set" +where + "L(EMPTY) = {[]}" +| "L(CHAR c) = {[c]}" +| "L(SEQ r1 r2) = (L r1) ; (L r2)" +| "L(ALT r1 r2) = (L r1) \ (L r2)" +| "L(STAR r) = (L r)\" + +definition + Ls :: "rexp set \ string set" +where + "Ls R = (\r\R. (L r))" + +lemma + shows "Ls {} = {}" +unfolding Ls_def by simp + +lemma Ls_union: + "Ls (R1 \ R2) = (Ls R1) \ (Ls R2)" +unfolding Ls_def by auto + +function + dagger :: "rexp \ char \ rexp set" ("_ \ _") +where + r1: "(EMPTY) \ c = {}" +| r2: "(CHAR c') \ c = (if c = c' then {EMPTY} else {})" +| r3: "(ALT r1 r2) \ c = r1 \ c \ r2 \ c" +| r4: "(SEQ EMPTY r2) \ c = r2 \ c" +| r5: "(SEQ (CHAR c') r2) \ c = (if c= c' then {r2} else {})" +| r6: "(SEQ (SEQ r11 r12) r2) \ c = (SEQ r11 (SEQ r12 r2)) \ c" +| r7: "(SEQ (ALT r11 r12) r2) \ c = (SEQ r11 r2) \ c \ (SEQ r12 r2) \ c" +| r8: "(SEQ (STAR r1) r2) \ c = + r2 \ c \ {SEQ (SEQ r' (STAR r1)) r2 | r'. r' \ r1 \ c}" +| r9: "(STAR r) \ c = {SEQ r' (STAR r) | r'. r' \ r \ c}" +by (pat_completeness) (auto) + +termination + dagger sorry + +definition + OR :: "bool set \ bool" +where + "OR S \ (\b\S. b)" + +function + matcher :: "rexp \ string \ bool" ("_ ! _") +where + s01: "EMPTY ! s = (s =[])" +| s02: "CHAR c ! s = (s = [c])" +| s03: "ALT r1 r2 ! s = (r1 ! s \ r2 ! s)" +| s04: "STAR r ! [] = True" +| s05: "STAR r ! c#s = (False \ OR {SEQ (r') (STAR r)!s | r'. r' \ r \ c})" +| s06: "SEQ r1 r2 ! [] = (r1 ! [] \ r2 ! [])" +| s07: "SEQ EMPTY r2 ! (c#s) = (r2 ! c#s)" +| s08: "SEQ (CHAR c') r2 ! (c#s) = (if c'=c then r2 ! s else False)" +| s09: "SEQ (SEQ r11 r12) r2 ! (c#s) = (SEQ r11 (SEQ r12 r2) ! c#s)" +| s10: "SEQ (ALT r11 r12) r2 ! (c#s) = ((SEQ r11 r2) ! (c#s) \ (SEQ r12 r2) ! (c#s))" +| s11: "SEQ (STAR r1) r2 ! (c#s) = + (r2 ! (c#s) \ OR {SEQ r' (SEQ (STAR r1) r2) ! s | r'. r' \ r1 \ c})" +by (pat_completeness) (auto) + +termination + matcher sorry + +lemma "(CHAR a) ! [a]" by auto +lemma "\(CHAR a) ! [a,a]" by auto +lemma "(STAR (CHAR a)) ! []" by auto +lemma "(STAR (CHAR a)) ! [a,a]" by (auto simp add: OR_def) +lemma "(SEQ (CHAR a) (SEQ (STAR (CHAR b)) (CHAR c))) ! [a,b,b,b,c]" + by (auto simp add: OR_def) + +lemma holes: + assumes a: "Ls (r \ c) = {s. c#s \ L r}" + shows "Ls (r \ c) ; L (STAR r) = {s''. c#s'' \ L (STAR r)}" +proof - + have "Ls (r \ c) ; L (STAR r) = {s. c#s \ L r} ; L (STAR r)" by (simp add: a) + also have "\ = {s'. c#s' \ (L r ; L (STAR r))}" sorry + also have "\ = {s''. c#s'' \ L (STAR r)}" sorry + finally show "Ls (r \ c) ; L (STAR r) = {s''. c#s'' \ L (STAR r)}" by simp +qed + +lemma eq: + shows "Ls (STAR r) \ c = (Ls (r \ c) ; L (STAR r))" +proof + show "Ls STAR r \ c \ Ls r \ c ; L (STAR r)" + by (auto simp add: lang_star_def lang_seq_def Ls_def) (blast) +next + show "Ls r \ c ; L (STAR r) \ Ls STAR r \ c" + apply(auto simp add: lang_star_def lang_seq_def Ls_def) + apply(rule_tac x="SEQ xa (STAR r)" in exI) + apply(simp add: lang_star_def lang_seq_def) + apply(blast) + done +qed + +(* correctness of the matcher *) +lemma dagger_holes: + "Ls (r \ c) = {s. c#s \ L r}" +proof (induct rule: dagger.induct) + case (1 c) + show "Ls (EMPTY \ c) = {s. c#s \ L EMPTY}" + by (simp add: Ls_def) +next + case (2 c' c) + show "Ls (CHAR c') \ c = {s. c#s \ L (CHAR c')}" + proof (cases "c=c'") + assume "c=c'" + then show "Ls (CHAR c') \ c = {s. c#s \ L (CHAR c')}" + by (simp add: Ls_def) + next + assume "c\c'" + then show "Ls (CHAR c') \ c = {s. c#s \ L (CHAR c')}" + by (simp add: Ls_def) + qed +next + case (3 r1 r2 c) + have ih1: "Ls r1 \ c = {s. c#s \ L r1}" by fact + have ih2: "Ls r2 \ c = {s. c#s \ L r2}" by fact + show "Ls (ALT r1 r2) \ c = {s. c#s \ L (ALT r1 r2)}" + by (auto simp add: Ls_union ih1 ih2) +next + case (4 r2 c) + have ih: "Ls r2 \ c = {s. c#s \ L r2}" by fact + show "Ls (SEQ EMPTY r2) \ c = {s. c#s \ L (SEQ EMPTY r2)}" + by (simp add: ih lang_seq_empty) +next + case (5 c' r2 c) + show "Ls (SEQ (CHAR c') r2) \ c = {s. c#s \ L (SEQ (CHAR c') r2)}" + proof (cases "c=c'") + assume "c=c'" + then show "Ls (SEQ (CHAR c') r2) \ c = {s. c#s \ L (SEQ (CHAR c') r2)}" + by (simp add: Ls_def lang_seq_def) + next + assume "c\c'" + then show "Ls (SEQ (CHAR c') r2) \ c = {s. c#s \ L (SEQ (CHAR c') r2)}" + by (simp add: Ls_def lang_seq_def) + qed +next + case (6 r11 r12 r2 c) + have ih: "Ls (SEQ r11 (SEQ r12 r2)) \ c = {s. c#s \ L (SEQ r11 (SEQ r12 r2))}" by fact + show "Ls (SEQ (SEQ r11 r12) r2) \ c = {s. c # s \ L (SEQ (SEQ r11 r12) r2)}" + by (simp add: ih lang_seq_assoc) +next + case (7 r11 r12 r2 c) + have ih1: "Ls (SEQ r11 r2) \ c = {s. c#s \ L (SEQ r11 r2)}" by fact + have ih2: "Ls (SEQ r12 r2) \ c = {s. c#s \ L (SEQ r12 r2)}" by fact + show "Ls (SEQ (ALT r11 r12) r2) \ c = {s. c#s \ L (SEQ (ALT r11 r12) r2)}" + by (auto simp add: Ls_union ih1 ih2 lang_seq_union) +next + case (8 r1 r2 c) + have ih1: "Ls r2 \ c = {s. c#s \ L r2}" by fact + have ih2: "Ls r1 \ c = {s. c#s \ L r1}" by fact + show "Ls (SEQ (STAR r1) r2) \ c = {s. c#s \ L (SEQ (STAR r1) r2)}" + sorry +next + case (9 r c) + have ih: "Ls r \ c = {s. c#s \ L r}" by fact + show "Ls (STAR r) \ c = {s. c#s \ L (STAR r)}" + by (simp only: eq holes[OF ih] del: r9) +qed + +(* correctness of the matcher *) +lemma macher_holes: + shows "r ! s \ s \ L r" + and "\ r ! s \ s \ L r" +proof (induct rule: matcher.induct) + case (1 s) + { case 1 + have "EMPTY ! s" by fact + then show "s \ L EMPTY" by simp + next + case 2 + have "\ EMPTY ! s" by fact + then show "s \ L EMPTY" by simp + } +next + case (2 c s) + { case 1 + have "CHAR c ! s" by fact + then show "s \ L (CHAR c)" by simp + next + case 2 + have "\ CHAR c ! s" by fact + then show "s \ L (CHAR c)" by simp + } +next + case (3 r1 r2 s) + have ih1: "r1 ! s \ s \ L r1" by fact + have ih2: "\ r1 ! s \ s \ L r1" by fact + have ih3: "r2 ! s \ s \ L r2" by fact + have ih4: "\ r2 ! s \ s \ L r2" by fact + { case 1 + have "ALT r1 r2 ! s" by fact + then show "s \ L (ALT r1 r2)" by (auto simp add: ih1 ih3) + next + case 2 + have "\ ALT r1 r2 ! s" by fact + then show "s \ L (ALT r1 r2)" by (simp add: ih2 ih4) + } +next + case (4 r) + { case 1 + have "STAR r ! []" by fact + then show "[] \ L (STAR r)" by (simp add: lang_star_in_empty) + next + case 2 + have "\ STAR r ! []" by fact + then show "[] \ L (STAR r)" by (simp) + } +next + case (5 r c s) + have ih1: "\rx. SEQ rx (STAR r) ! s \ s \ L (SEQ rx (STAR r))" by fact + have ih2: "\rx. \SEQ rx (STAR r) ! s \ s \ L (SEQ rx (STAR r))" by fact + { case 1 + have as: "STAR r ! c#s" by fact + then have "\r' \ r \ c. SEQ r' (STAR r) ! s" by (auto simp add: OR_def) + then obtain r' where imp1: "r' \ r \ c" and imp2: "SEQ r' (STAR r) ! s" by blast + from imp2 have "s \ L (SEQ r' (STAR r))" using ih1 by simp + then have "s \ L r' ; L (STAR r)" by simp + then have "c#s \ {[c]} ; (L r' ; L (STAR r))" by (simp add: lang_seq_def) + also have "\ \ L r ; L (STAR r)" using imp1 + apply(auto simp add: lang_seq_def) sorry + also have "\ \ L (STAR r)" by (simp add: lang_star_subseteq) + finally show "c#s \ L (STAR r)" by simp + next + case 2 + have as: "\ STAR r ! c#s" by fact + then have "\r'\ r \ c. \ (SEQ r' (STAR r) ! s)" + by (auto simp add: OR_def) + then have "\r'\ r \ c. s \ L (SEQ r' (STAR r))" using ih2 by auto + then obtain r' where "r'\ r \ c \ s \ L (SEQ r' (STAR r))" by auto + + show "c#s \ L (STAR r)" sorry + } +next + case (6 r1 r2) + have ih1: "r1 ! [] \ [] \ L r1" by fact + have ih2: "\ r1 ! [] \ [] \ L r1" by fact + have ih3: "r2 ! [] \ [] \ L r2" by fact + have ih4: "\ r2 ! [] \ [] \ L r2" by fact + { case 1 + have as: "SEQ r1 r2 ! []" by fact + then have "r1 ! [] \ r2 ! []" by (simp) + then show "[] \ L (SEQ r1 r2)" using ih1 ih3 by (simp add: lang_seq_def) + next + case 2 + 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 (7 r2 c s) + have ih1: "r2 ! c#s \ c#s \ L r2" by fact + have ih2: "\ r2 ! c#s \ c#s \ L r2" by fact + { case 1 + 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 2 + have "\ SEQ EMPTY r2 ! c#s" by fact + then show "c#s \ L (SEQ EMPTY r2)" + using ih2 by (simp add: lang_seq_def) + } +next + case (8 c' r2 c s) + have ih1: "\c' = c; r2 ! s\ \ s \ L r2" by fact + have ih2: "\c' = c; \r2 ! s\ \ s \ L r2" by fact + { case 1 + 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 2 + 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 (9 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 ih2: "\ SEQ r11 (SEQ r12 r2) ! c#s \ c#s \ L (SEQ r11 (SEQ r12 r2))" by fact + { case 1 + have "SEQ (SEQ r11 r12) r2 ! c#s" by fact + then show "c#s \ L (SEQ (SEQ r11 r12) r2)" + using ih1 + apply(auto simp add: lang_seq_def) + apply(rule_tac x="s1@s1a" in exI) + apply(rule_tac x="s2a" in exI) + apply(simp) + apply(blast) + done + next + case 2 + 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 (10 r11 r12 r2 c s) + have ih1: "SEQ r11 r2 ! c#s \ c#s \ L (SEQ r11 r2)" by fact + have ih2: "\ SEQ r11 r2 ! c#s \ c#s \ L (SEQ r11 r2)" by fact + have ih3: "SEQ r12 r2 ! c#s \ c#s \ L (SEQ r12 r2)" by fact + have ih4: "\ SEQ r12 r2 ! c#s \ c#s \ L (SEQ r12 r2)" by fact + { case 1 + have "SEQ (ALT r11 r12) r2 ! c#s" by fact + then show "c#s \ L (SEQ (ALT r11 r12) r2)" + using ih1 ih3 by (auto simp add: lang_seq_union) + next + case 2 + 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 (11 r1 r2 c s) + have ih1: "r2 ! c#s \ c#s \ L r2" by fact + have ih2: "\r2 ! c#s \ c#s \ L r2" by fact + { case 1 + have "SEQ (STAR r1) r2 ! c#s" by fact + then show "c#s \ L (SEQ (STAR r1) r2)" + using ih1 sorry + next + case 2 + have "\ SEQ (STAR r1) r2 ! c#s" by fact + then show "c#s \ L (SEQ (STAR r1) r2)" + using ih2 sorry + } +qed + + +end + + + + + + + + + +