# HG changeset patch # User Christian Urban # Date 1384712180 0 # Node ID 90796ee3c17a4d0831946aa081c4db2db5a9f16d # Parent 6518475020fc669b56f2efea585903b922a6c519 added diff -r 6518475020fc -r 90796ee3c17a hws/hw07.pdf Binary file hws/hw07.pdf has changed diff -r 6518475020fc -r 90796ee3c17a hws/hw07.tex --- a/hws/hw07.tex Fri Nov 15 10:29:04 2013 +0000 +++ b/hws/hw07.tex Sun Nov 17 18:16:20 2013 +0000 @@ -13,11 +13,21 @@ \section*{Homework 7} \begin{enumerate} -\item Suppose the following grammar for positive numbers +\item Suppose the context-sensitive grammar \begin{center} +\begin{tabular}{lcl} +$S$ & $\rightarrow$ & $bSAA\;|\; \epsilon$\\ +$A$ & $\rightarrow$ & $a$\\ +$bA$ & $\rightarrow$ & $Ab$\\ +\end{tabular} +\end{center} -\end{center} +where $S$ is the starting symbol of the grammar. +Give a derivation of the string $"\!aaabaaabb"$. +What can you say about the number of as and bs in the +strings recognised by this grammar. + \item Consider the following grammar @@ -41,14 +51,27 @@ \texttt{The trainer trains the student team} \end{center} -\item {\bf (Optional)} The task is to match strings where the letters are in alphabetical order---for example, -\texttt{abcfjz} would pass, but \texttt{acb} would not. Whitespace should be ignored---for example -\texttt{ab c d} should pass. The point is to try to get the regular expression as short as possible! -See: +\item Transform the grammar \begin{center} -\url{http://callumacrae.github.com/regex-tuesday/challenge11.html} +\begin{tabular}{lcl} +$A$ & $\rightarrow$ & $0A1 \;|\; BB$\\ +$B$ & $\rightarrow$ & $\epsilon \;|\; 2B$ +\end{tabular} \end{center} + +\noindent +into Chomsky normal form. + + +%\item {\bf (Optional)} The task is to match strings where the letters are in alphabetical order---for example, +%\texttt{abcfjz} would pass, but \texttt{acb} would not. Whitespace should be ignored---for example +%\texttt{ab c d} should pass. The point is to try to get the regular expression as short as possible! +%See: + +%\begin{center} +%\url{http://callumacrae.github.com/regex-tuesday/challenge11.html} +%\end{center} \end{enumerate} \end{document} diff -r 6518475020fc -r 90796ee3c17a progs/Matcher2.thy --- a/progs/Matcher2.thy Fri Nov 15 10:29:04 2013 +0000 +++ b/progs/Matcher2.thy Sun Nov 17 18:16:20 2013 +0000 @@ -52,6 +52,7 @@ lemma seq_union: shows "A ;; (B \ C) = A ;; B \ A ;; C" + and "(B \ C) ;; A = B ;; A \ C ;; A" by (auto simp add: Seq_def) lemma seq_Union: @@ -62,6 +63,30 @@ "[] \ A ;; B \ ([] \ A \ [] \ B)" by (simp add: Seq_def) +lemma seq_assoc: + shows "A ;; (B ;; C) = (A ;; B) ;; C" +apply(auto simp add: Seq_def) +apply(metis append_assoc) +apply(metis) +done + + +section {* Power for Sets *} + +fun + pow :: "string set \ nat \ string set" ("_ \ _" [101, 102] 101) +where + "A \ 0 = {[]}" +| "A \ (Suc n) = A ;; (A \ n)" + +lemma pow_empty [simp]: + shows "[] \ A \ n \ (n = 0 \ [] \ A)" +by (induct n) (auto) + +lemma pow_plus: + "A \ (n + m) = A \ n ;; A \ m" +by (induct n) (simp_all add: seq_assoc) + section {* Kleene Star for Sets *} inductive_set @@ -71,33 +96,43 @@ start[intro]: "[] \ A\" | step[intro]: "\s1 \ A; s2 \ A\\ \ s1 @ s2 \ A\" - text {* A Standard Property of Star *} -lemma star_cases: - shows "A\ = {[]} \ A ;; A\" -unfolding Seq_def -by (auto) (metis Star.simps) - lemma star_decomp: assumes a: "c # x \ A\" shows "\a b. x = a @ b \ c # a \ A \ b \ A\" +using a using a by (induct x\"c # x" rule: Star.induct) (auto simp add: append_eq_Cons_conv) -section {* Power for Sets *} +lemma star_cases: + shows "A\ = {[]} \ A ;; A\" +unfolding Seq_def +by (auto) (metis Star.simps) -fun - pow :: "string set \ nat \ string set" ("_ \ _" [101, 102] 101) -where - "A \ 0 = {[]}" -| "A \ (Suc n) = A ;; (A \ n)" +lemma Star_in_Pow: + assumes a: "s \ A\" + shows "\n. s \ A \ n" +using a +apply(induct) +apply(auto) +apply(rule_tac x="Suc n" in exI) +apply(auto simp add: Seq_def) +done + +lemma Pow_in_Star: + assumes a: "s \ A \ n" + shows "s \ A\" +using a +by (induct n arbitrary: s) (auto simp add: Seq_def) -lemma pow_empty [simp]: - shows "[] \ A \ n \ (n = 0 \ [] \ A)" -by (induct n) (auto) +lemma Star_def2: + shows "A\ = (\n. A \ n)" +using Star_in_Pow Pow_in_Star +by (auto) + section {* Semantics of Regular Expressions *} @@ -265,4 +300,39 @@ (simp_all add: nullable_correctness der_correctness Der_def) + + +lemma bb: "A \ n ;; A\ \ A\" +apply(induct n) +apply(simp) +apply(simp) +apply(subst aa[symmetric]) +apply(subst Seq_def) +apply(auto) +done + +lemma cc: "(\i\{..n}. A \ i) ;; A\ \ A\" +apply(induct n) +apply(simp) +apply(simp add: Suc_Union del: pow.simps) +apply(simp only: seq_union) +apply(rule Un_least) +defer +apply(simp) +apply(rule bb) +done + +lemma "A\ ;; A\ = A\" +apply(simp add: Seq_def) +apply(rule subset_antisym) +defer +apply(auto)[1] +apply(auto simp add: Star_def2) + +apply(subst (asm) Star_def2) +apply(subst (asm) Star_def2) +apply(auto) + + + end \ No newline at end of file