Slides/Slides1.thy
changeset 212 3629680a20a2
parent 211 a9e4acbf7b00
child 213 dda2e90de8a2
equal deleted inserted replaced
211:a9e4acbf7b00 212:3629680a20a2
   312   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   312   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   313   \mode<presentation>{
   313   \mode<presentation>{
   314   \begin{frame}[c]
   314   \begin{frame}[c]
   315   \frametitle{\LARGE The Myhill-Nerode Theorem}
   315   \frametitle{\LARGE The Myhill-Nerode Theorem}
   316 
   316 
   317   \mbox{}\\[5cm]
   317   \begin{center}
   318 
   318   \only<1>{%
       
   319   \begin{tikzpicture}[scale=3]
       
   320   \draw[very thick] (0.5,0.5) circle (.6cm);
       
   321   \end{tikzpicture}}%
       
   322   \only<2->{%
       
   323   \begin{tikzpicture}[scale=3]
       
   324   \draw[very thick] (0.5,0.5) circle (.6cm);
       
   325   \clip[draw] (0.5,0.5) circle (.6cm);
       
   326   \draw[step=2mm, very thick] (-1.4,-1.4) grid (1.4,1.4);
       
   327   \end{tikzpicture}}
       
   328   \end{center}
       
   329   
   319   \begin{itemize}
   330   \begin{itemize}
   320   \item \smath{\text{finite}\, (U\!N\!IV /\!/ \approx_A) \;\Leftrightarrow\; A\; \text{is regular}}
   331   \item \smath{\text{finite}\, (U\!N\!IV /\!/ \approx_A) \;\Leftrightarrow\; A\; \text{is regular}}
   321   \end{itemize}
   332   \end{itemize}
   322 
   333 
   323   \only<2->{
   334   \begin{textblock}{5}(2.1,5.3)
       
   335   \begin{tikzpicture}
       
   336   \node at (0,0) [single arrow, fill=red,text=white, minimum height=2cm]
       
   337   {$U\!N\!IV$};
       
   338   \draw (-0.3,-1.1) node {\begin{tabular}{l}set of all\\[-1mm] strings\end{tabular}};
       
   339   \end{tikzpicture}
       
   340   \end{textblock}
       
   341 
       
   342   \only<2->{%
       
   343   \begin{textblock}{5}(9.4,7.2)
       
   344   \begin{tikzpicture}
       
   345   \node at (0,0) [shape border rotate=180,single arrow, fill=red,text=white, minimum height=2cm]
       
   346   {@{text "\<lbrakk>x\<rbrakk>"}$_{\approx_{A}}$};
       
   347   \draw (0.9,-1.1) node {\begin{tabular}{l}equivalence class\end{tabular}};
       
   348   \end{tikzpicture}
       
   349   \end{textblock}}
       
   350 
       
   351   \only<3->{
   324   \begin{textblock}{11.9}(1.7,3)
   352   \begin{textblock}{11.9}(1.7,3)
   325   \begin{block}{}
   353   \begin{block}{}
   326   \begin{minipage}{11.4cm}\raggedright
   354   \begin{minipage}{11.4cm}\raggedright
   327   Two directions:\medskip\\
   355   Two directions:\medskip\\
   328 
       
   329   \begin{tabular}{@ {}ll}
   356   \begin{tabular}{@ {}ll}
   330   1.)\;finite $\Rightarrow$ regular\\
   357   1.)\;finite $\Rightarrow$ regular\\
   331   \;\;\;\smath{\text{finite}\,(U\!N\!IV /\!/ \approx_A) \Rightarrow \exists r.\;A = {\cal L}(r)}\\[3mm]
   358   \;\;\;\smath{\text{finite}\,(U\!N\!IV /\!/ \approx_A) \Rightarrow \exists r.\;A = {\cal L}(r)}\\[3mm]
   332   2.)\;regular $\Rightarrow$ finite\\
   359   2.)\;regular $\Rightarrow$ finite\\
   333   \;\;\;\smath{\text{finite}\, (U\!N\!IV /\!/ \approx_{{\cal L}(r)})}
   360   \;\;\;\smath{\text{finite}\, (U\!N\!IV /\!/ \approx_{{\cal L}(r)})}
   345 
   372 
   346 text_raw {*
   373 text_raw {*
   347   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   374   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   348   \mode<presentation>{
   375   \mode<presentation>{
   349   \begin{frame}[c]
   376   \begin{frame}[c]
   350   \frametitle{\LARGE Final States}
   377   \frametitle{\LARGE Initial and Final {\sout{\textcolor{gray}{States}}}}
   351 
   378 
   352   \mbox{}\\[3cm]
   379   \begin{textblock}{8}(10, 2)
   353 
   380   \textcolor{black}{Equivalence Classes}
   354   \begin{itemize}
   381   \end{textblock}
   355   \item \smath{\text{final}_A\,X \dn}\\
   382 
   356   \smath{\hspace{6mm}X \in (U\!N\!IV /\!/\approx_A) \;\wedge\; \forall s \in X.\; s \in A}
   383 
       
   384   \begin{center}
       
   385   \begin{tikzpicture}[scale=3]
       
   386   \draw[very thick] (0.5,0.5) circle (.6cm);
       
   387   \clip[draw] (0.5,0.5) circle (.6cm);
       
   388   \draw[step=2mm, very thick] (-1.4,-1.4) grid (1.4,1.4);
       
   389   \only<2->{\draw[blue, fill] (0.0, 0.6) rectangle (0.2, 0.8);}
       
   390   \only<3->{\draw[red, fill] (0.2, 0.2) rectangle (0.4, 0.4);
       
   391   \draw[red, fill] (0.4, 0.8) rectangle (0.6, 1.0);
       
   392   \draw[red, fill] (0.6, 0.0) rectangle (0.8, 0.2);
       
   393   \draw[red, fill] (0.8, 0.4) rectangle (1.0, 0.6);}
       
   394   \end{tikzpicture}
       
   395   \end{center}
       
   396 
       
   397   \begin{itemize}
       
   398   \item \smath{\text{final}_A\,X \dn \{[\!|x|\!]_{\approx_{A}}\;|\;x \in A\}}
   357   \smallskip
   399   \smallskip
   358   \item we can prove: \smath{A = \bigcup \{X.\;\text{final}_A\,X\}}
   400   \item we can prove: \smath{A = \bigcup \{X.\;\text{final}_A\,X\}}
   359 
   401   \end{itemize}
   360   \end{itemize}
   402 
   361 
   403   \only<2->{%
   362   \end{frame}}
   404   \begin{textblock}{5}(2.1,4.6)
   363   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
   405   \begin{tikzpicture}
   364 *}
   406   \node at (0,0) [single arrow, fill=blue,text=white, minimum height=2cm]
   365 
   407   {$[] \in X$};
   366 
   408   \end{tikzpicture}
   367 text_raw {*
   409   \end{textblock}}
   368   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
   410 
   369   \mode<presentation>{
   411   \only<3->{%
   370   \begin{frame}[c]
   412   \begin{textblock}{5}(10,7.4)
   371   \frametitle{\LARGE Transitions between Eq-Classes}
   413   \begin{tikzpicture}
   372 
   414   \node at (0,0) [shape border rotate=180,single arrow, fill=red,text=white, minimum height=2cm]
   373   \smath{L = \{[c]\}}
   415   {a final};
   374 
   416   \end{tikzpicture}
   375   \begin{tabular}{@ {\hspace{-7mm}}cc}
   417   \end{textblock}}
       
   418 
       
   419   \end{frame}}
       
   420   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
       
   421 *}
       
   422 
       
   423 
       
   424 text_raw {*
       
   425   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%
       
   426   \mode<presentation>{
       
   427   \begin{frame}<-1>[c]
       
   428   \frametitle{\begin{tabular}{@ {}l}\LARGE% 
       
   429   Transitions between Eq-Classes\end{tabular}}
       
   430 
       
   431   \begin{center}
       
   432   \begin{tikzpicture}[scale=3]
       
   433   \draw[very thick] (0.5,0.5) circle (.6cm);
       
   434   \clip[draw] (0.5,0.5) circle (.6cm);
       
   435   \draw[step=2mm, very thick] (-1.4,-1.4) grid (1.4,1.4);
       
   436   \draw[blue, fill] (0.0, 0.6) rectangle (0.2, 0.8);
       
   437   \draw[blue, fill] (0.8, 0.4) rectangle (1.0, 0.6);
       
   438   \draw[white] (0.1,0.7) node {$X$};
       
   439   \draw[white] (0.9,0.5) node {$Y$};
       
   440   \end{tikzpicture}
       
   441   \end{center}
       
   442 
       
   443   \begin{center}
       
   444   \smath{X \stackrel{c}{\longrightarrow} Y \;\dn\; X ; c \subseteq Y}
       
   445   \end{center}
       
   446 
       
   447   \onslide<8>{
   376   \begin{tabular}{c}
   448   \begin{tabular}{c}
   377   \begin{tikzpicture}[shorten >=1pt,node distance=2cm,auto, ultra thick]
   449   \begin{tikzpicture}[shorten >=1pt,node distance=2cm,auto, ultra thick]
   378   \tikzstyle{state}=[circle,thick,draw=blue!75,fill=blue!20,minimum size=0mm]
   450   \tikzstyle{state}=[circle,thick,draw=blue!75,fill=blue!20,minimum size=0mm]
   379 
   451   \node[state,initial] (q_0) {$R_1$};
   380   %\draw[help lines] (0,0) grid (3,2);
   452   \end{tikzpicture}
   381 
   453   \end{tabular}}
   382   \node[state,initial]   (q_0)                        {$R_1$};
       
   383   \node[state,accepting] (q_1) [above right of=q_0]   {$R_2$};
       
   384   \node[state]           (q_2) [below right of=q_0]   {$R_3$};
       
   385 
       
   386   \path[->] (q_0) edge                node        {c} (q_1)
       
   387                   edge                node [swap] {$\Sigma-{c}$} (q_2)
       
   388             (q_2) edge [loop below]   node        {$\Sigma$} ()
       
   389             (q_1) edge                node        {$\Sigma$} (q_2);
       
   390   \end{tikzpicture}
       
   391   \end{tabular}
       
   392   &
       
   393   \begin{tabular}[t]{ll}
       
   394   \\[-20mm]
       
   395   \multicolumn{2}{l}{\smath{U\!N\!IV /\!/\approx_L} produces}\\[4mm]
       
   396 
       
   397   \smath{R_1}: & \smath{\{[]\}}\\
       
   398   \smath{R_2}: & \smath{\{[c]\}}\\
       
   399   \smath{R_3}: & \smath{U\!N\!IV - \{[], [c]\}}\\[6mm]
       
   400   \multicolumn{2}{l}{\onslide<2->{\smath{X \stackrel{c}{\longrightarrow} Y \dn X ; [c] \subseteq Y}}}
       
   401   \end{tabular}
       
   402 
       
   403   \end{tabular}
       
   404 
   454 
   405   \end{frame}}
   455   \end{frame}}
   406   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
   456   %%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%%     
   407 *}
   457 *}
   408 
   458