doc-src/TutorialI/Trie/document/Trie.tex
changeset 9674 f789d2490669
parent 9541 d17c0b34d5c8
child 9698 f0740137a65d
--- a/doc-src/TutorialI/Trie/document/Trie.tex	Mon Aug 21 19:17:07 2000 +0200
+++ b/doc-src/TutorialI/Trie/document/Trie.tex	Mon Aug 21 19:29:27 2000 +0200
@@ -7,7 +7,7 @@
 list of (letter,trie) pairs.  Abstracting over the alphabet \isa{'a} and the
 values \isa{'v} we define a trie as follows:%
 \end{isamarkuptext}%
-\isacommand{datatype}\ ('a,'v)trie\ =\ Trie\ \ {"}'v\ option{"}\ \ {"}('a\ *\ ('a,'v)trie)list{"}%
+\isacommand{datatype}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie\ {\isacharequal}\ Trie\ \ {\isachardoublequote}{\isacharprime}v\ option{\isachardoublequote}\ \ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a\ {\isacharasterisk}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie{\isacharparenright}list{\isachardoublequote}%
 \begin{isamarkuptext}%
 \noindent
 The first component is the optional value, the second component the
@@ -15,48 +15,48 @@
 which is fine because products are datatypes as well.
 We define two selector functions:%
 \end{isamarkuptext}%
-\isacommand{consts}\ value\ ::\ {"}('a,'v)trie\ {\isasymRightarrow}\ 'v\ option{"}\isanewline
-\ \ \ \ \ \ \ alist\ ::\ {"}('a,'v)trie\ {\isasymRightarrow}\ ('a\ *\ ('a,'v)trie)list{"}\isanewline
-\isacommand{primrec}\ {"}value(Trie\ ov\ al)\ =\ ov{"}\isanewline
-\isacommand{primrec}\ {"}alist(Trie\ ov\ al)\ =\ al{"}%
+\isacommand{consts}\ value\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie\ {\isasymRightarrow}\ {\isacharprime}v\ option{\isachardoublequote}\isanewline
+\ \ \ \ \ \ \ alist\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a\ {\isacharasterisk}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie{\isacharparenright}list{\isachardoublequote}\isanewline
+\isacommand{primrec}\ {\isachardoublequote}value{\isacharparenleft}Trie\ ov\ al{\isacharparenright}\ {\isacharequal}\ ov{\isachardoublequote}\isanewline
+\isacommand{primrec}\ {\isachardoublequote}alist{\isacharparenleft}Trie\ ov\ al{\isacharparenright}\ {\isacharequal}\ al{\isachardoublequote}%
 \begin{isamarkuptext}%
 \noindent
 Association lists come with a generic lookup function:%
 \end{isamarkuptext}%
-\isacommand{consts}\ \ \ assoc\ ::\ {"}('key\ *\ 'val)list\ {\isasymRightarrow}\ 'key\ {\isasymRightarrow}\ 'val\ option{"}\isanewline
-\isacommand{primrec}\ {"}assoc\ []\ x\ =\ None{"}\isanewline
-\ \ \ \ \ \ \ \ {"}assoc\ (p\#ps)\ x\ =\isanewline
-\ \ \ \ \ \ \ \ \ \ \ (let\ (a,b)\ =\ p\ in\ if\ a=x\ then\ Some\ b\ else\ assoc\ ps\ x){"}%
+\isacommand{consts}\ \ \ assoc\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}key\ {\isacharasterisk}\ {\isacharprime}val{\isacharparenright}list\ {\isasymRightarrow}\ {\isacharprime}key\ {\isasymRightarrow}\ {\isacharprime}val\ option{\isachardoublequote}\isanewline
+\isacommand{primrec}\ {\isachardoublequote}assoc\ {\isacharbrackleft}{\isacharbrackright}\ x\ {\isacharequal}\ None{\isachardoublequote}\isanewline
+\ \ \ \ \ \ \ \ {\isachardoublequote}assoc\ {\isacharparenleft}p{\isacharhash}ps{\isacharparenright}\ x\ {\isacharequal}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}let\ {\isacharparenleft}a{\isacharcomma}b{\isacharparenright}\ {\isacharequal}\ p\ in\ if\ a{\isacharequal}x\ then\ Some\ b\ else\ assoc\ ps\ x{\isacharparenright}{\isachardoublequote}%
 \begin{isamarkuptext}%
 Now we can define the lookup function for tries. It descends into the trie
 examining the letters of the search string one by one. As
 recursion on lists is simpler than on tries, let us express this as primitive
 recursion on the search string argument:%
 \end{isamarkuptext}%
-\isacommand{consts}\ \ \ lookup\ ::\ {"}('a,'v)trie\ {\isasymRightarrow}\ 'a\ list\ {\isasymRightarrow}\ 'v\ option{"}\isanewline
-\isacommand{primrec}\ {"}lookup\ t\ []\ =\ value\ t{"}\isanewline
-\ \ \ \ \ \ \ \ {"}lookup\ t\ (a\#as)\ =\ (case\ assoc\ (alist\ t)\ a\ of\isanewline
+\isacommand{consts}\ \ \ lookup\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}v\ option{\isachardoublequote}\isanewline
+\isacommand{primrec}\ {\isachardoublequote}lookup\ t\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharequal}\ value\ t{\isachardoublequote}\isanewline
+\ \ \ \ \ \ \ \ {\isachardoublequote}lookup\ t\ {\isacharparenleft}a{\isacharhash}as{\isacharparenright}\ {\isacharequal}\ {\isacharparenleft}case\ assoc\ {\isacharparenleft}alist\ t{\isacharparenright}\ a\ of\isanewline
 \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ None\ {\isasymRightarrow}\ None\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ |\ Some\ at\ {\isasymRightarrow}\ lookup\ at\ as){"}%
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharbar}\ Some\ at\ {\isasymRightarrow}\ lookup\ at\ as{\isacharparenright}{\isachardoublequote}%
 \begin{isamarkuptext}%
 As a first simple property we prove that looking up a string in the empty
 trie \isa{Trie~None~[]} always returns \isa{None}. The proof merely
 distinguishes the two cases whether the search string is empty or not:%
 \end{isamarkuptext}%
-\isacommand{lemma}\ [simp]:\ {"}lookup\ (Trie\ None\ [])\ as\ =\ None{"}\isanewline
-\isacommand{by}(case\_tac\ as,\ auto)%
+\isacommand{lemma}\ {\isacharbrackleft}simp{\isacharbrackright}{\isacharcolon}\ {\isachardoublequote}lookup\ {\isacharparenleft}Trie\ None\ {\isacharbrackleft}{\isacharbrackright}{\isacharparenright}\ as\ {\isacharequal}\ None{\isachardoublequote}\isanewline
+\isacommand{by}{\isacharparenleft}case{\isacharunderscore}tac\ as{\isacharcomma}\ auto{\isacharparenright}%
 \begin{isamarkuptext}%
 Things begin to get interesting with the definition of an update function
 that adds a new (string,value) pair to a trie, overwriting the old value
 associated with that string:%
 \end{isamarkuptext}%
-\isacommand{consts}\ update\ ::\ {"}('a,'v)trie\ {\isasymRightarrow}\ 'a\ list\ {\isasymRightarrow}\ 'v\ {\isasymRightarrow}\ ('a,'v)trie{"}\isanewline
+\isacommand{consts}\ update\ {\isacharcolon}{\isacharcolon}\ {\isachardoublequote}{\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie\ {\isasymRightarrow}\ {\isacharprime}a\ list\ {\isasymRightarrow}\ {\isacharprime}v\ {\isasymRightarrow}\ {\isacharparenleft}{\isacharprime}a{\isacharcomma}{\isacharprime}v{\isacharparenright}trie{\isachardoublequote}\isanewline
 \isacommand{primrec}\isanewline
-\ \ {"}update\ t\ []\ \ \ \ \ v\ =\ Trie\ (Some\ v)\ (alist\ t){"}\isanewline
-\ \ {"}update\ t\ (a\#as)\ v\ =\isanewline
-\ \ \ \ \ (let\ tt\ =\ (case\ assoc\ (alist\ t)\ a\ of\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ None\ {\isasymRightarrow}\ Trie\ None\ []\ |\ Some\ at\ {\isasymRightarrow}\ at)\isanewline
-\ \ \ \ \ \ in\ Trie\ (value\ t)\ ((a,update\ tt\ as\ v)\#alist\ t)){"}%
+\ \ {\isachardoublequote}update\ t\ {\isacharbrackleft}{\isacharbrackright}\ \ \ \ \ v\ {\isacharequal}\ Trie\ {\isacharparenleft}Some\ v{\isacharparenright}\ {\isacharparenleft}alist\ t{\isacharparenright}{\isachardoublequote}\isanewline
+\ \ {\isachardoublequote}update\ t\ {\isacharparenleft}a{\isacharhash}as{\isacharparenright}\ v\ {\isacharequal}\isanewline
+\ \ \ \ \ {\isacharparenleft}let\ tt\ {\isacharequal}\ {\isacharparenleft}case\ assoc\ {\isacharparenleft}alist\ t{\isacharparenright}\ a\ of\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ None\ {\isasymRightarrow}\ Trie\ None\ {\isacharbrackleft}{\isacharbrackright}\ {\isacharbar}\ Some\ at\ {\isasymRightarrow}\ at{\isacharparenright}\isanewline
+\ \ \ \ \ \ in\ Trie\ {\isacharparenleft}value\ t{\isacharparenright}\ {\isacharparenleft}{\isacharparenleft}a{\isacharcomma}update\ tt\ as\ v{\isacharparenright}{\isacharhash}alist\ t{\isacharparenright}{\isacharparenright}{\isachardoublequote}%
 \begin{isamarkuptext}%
 \noindent
 The base case is obvious. In the recursive case the subtrie
@@ -70,8 +70,8 @@
 expand all \isa{let}s and to split all \isa{case}-constructs over
 options:%
 \end{isamarkuptext}%
-\isacommand{lemmas}\ [simp]\ =\ Let\_def\isanewline
-\isacommand{lemmas}\ [split]\ =\ option.split%
+\isacommand{lemmas}\ {\isacharbrackleft}simp{\isacharbrackright}\ {\isacharequal}\ Let{\isacharunderscore}def\isanewline
+\isacommand{lemmas}\ {\isacharbrackleft}split{\isacharbrackright}\ {\isacharequal}\ option{\isachardot}split%
 \begin{isamarkuptext}%
 \noindent
 The reason becomes clear when looking (probably after a failed proof
@@ -81,8 +81,8 @@
 Our main goal is to prove the correct interaction of \isa{update} and
 \isa{lookup}:%
 \end{isamarkuptext}%
-\isacommand{theorem}\ {"}{\isasymforall}t\ v\ bs.\ lookup\ (update\ t\ as\ v)\ bs\ =\isanewline
-\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ (if\ as=bs\ then\ Some\ v\ else\ lookup\ t\ bs){"}%
+\isacommand{theorem}\ {\isachardoublequote}{\isasymforall}t\ v\ bs{\isachardot}\ lookup\ {\isacharparenleft}update\ t\ as\ v{\isacharparenright}\ bs\ {\isacharequal}\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ {\isacharparenleft}if\ as{\isacharequal}bs\ then\ Some\ v\ else\ lookup\ t\ bs{\isacharparenright}{\isachardoublequote}%
 \begin{isamarkuptxt}%
 \noindent
 Our plan is to induct on \isa{as}; hence the remaining variables are
@@ -93,7 +93,7 @@
 \isa{as} is instantiated.
 The start of the proof is completely conventional:%
 \end{isamarkuptxt}%
-\isacommand{apply}(induct\_tac\ as,\ auto)%
+\isacommand{apply}{\isacharparenleft}induct{\isacharunderscore}tac\ as{\isacharcomma}\ auto{\isacharparenright}%
 \begin{isamarkuptxt}%
 \noindent
 Unfortunately, this time we are left with three intimidating looking subgoals:
@@ -106,7 +106,7 @@
 well now. It turns out that instead of induction, case distinction
 suffices:%
 \end{isamarkuptxt}%
-\isacommand{by}(case\_tac[!]\ bs,\ auto)%
+\isacommand{by}{\isacharparenleft}case{\isacharunderscore}tac{\isacharbrackleft}{\isacharbang}{\isacharbrackright}\ bs{\isacharcomma}\ auto{\isacharparenright}%
 \begin{isamarkuptext}%
 \noindent
 All methods ending in \isa{tac} take an optional first argument that