doc-src/TutorialI/CodeGen/document/CodeGen.tex
changeset 9541 d17c0b34d5c8
parent 9458 c613cd06d5cf
child 9673 1b2d4f995b13
--- a/doc-src/TutorialI/CodeGen/document/CodeGen.tex	Fri Aug 04 23:02:11 2000 +0200
+++ b/doc-src/TutorialI/CodeGen/document/CodeGen.tex	Sun Aug 06 15:26:53 2000 +0200
@@ -10,10 +10,10 @@
 a fixed set of binary operations: instead the expression contains the
 appropriate function itself.%
 \end{isamarkuptext}%
-\isacommand{types}~'v~binop~=~{"}'v~{\isasymRightarrow}~'v~{\isasymRightarrow}~'v{"}\isanewline
-\isacommand{datatype}~('a,'v)expr~=~Cex~'v\isanewline
-~~~~~~~~~~~~~~~~~~~~~|~Vex~'a\isanewline
-~~~~~~~~~~~~~~~~~~~~~|~Bex~{"}'v~binop{"}~~{"}('a,'v)expr{"}~~{"}('a,'v)expr{"}%
+\isacommand{types}\ 'v\ binop\ =\ {"}'v\ {\isasymRightarrow}\ 'v\ {\isasymRightarrow}\ 'v{"}\isanewline
+\isacommand{datatype}\ ('a,'v)expr\ =\ Cex\ 'v\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ |\ Vex\ 'a\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ |\ Bex\ {"}'v\ binop{"}\ \ {"}('a,'v)expr{"}\ \ {"}('a,'v)expr{"}%
 \begin{isamarkuptext}%
 \noindent
 The three constructors represent constants, variables and the application of
@@ -22,20 +22,20 @@
 The value of an expression w.r.t.\ an environment that maps variables to
 values is easily defined:%
 \end{isamarkuptext}%
-\isacommand{consts}~value~::~{"}('a,'v)expr~{\isasymRightarrow}~('a~{\isasymRightarrow}~'v)~{\isasymRightarrow}~'v{"}\isanewline
+\isacommand{consts}\ value\ ::\ {"}('a,'v)expr\ {\isasymRightarrow}\ ('a\ {\isasymRightarrow}\ 'v)\ {\isasymRightarrow}\ 'v{"}\isanewline
 \isacommand{primrec}\isanewline
-{"}value~(Cex~v)~env~=~v{"}\isanewline
-{"}value~(Vex~a)~env~=~env~a{"}\isanewline
-{"}value~(Bex~f~e1~e2)~env~=~f~(value~e1~env)~(value~e2~env){"}%
+{"}value\ (Cex\ v)\ env\ =\ v{"}\isanewline
+{"}value\ (Vex\ a)\ env\ =\ env\ a{"}\isanewline
+{"}value\ (Bex\ f\ e1\ e2)\ env\ =\ f\ (value\ e1\ env)\ (value\ e2\ env){"}%
 \begin{isamarkuptext}%
 The stack machine has three instructions: load a constant value onto the
 stack, load the contents of a certain address onto the stack, and apply a
 binary operation to the two topmost elements of the stack, replacing them by
 the result. As for \isa{expr}, addresses and values are type parameters:%
 \end{isamarkuptext}%
-\isacommand{datatype}~('a,'v)~instr~=~Const~'v\isanewline
-~~~~~~~~~~~~~~~~~~~~~~~|~Load~'a\isanewline
-~~~~~~~~~~~~~~~~~~~~~~~|~Apply~{"}'v~binop{"}%
+\isacommand{datatype}\ ('a,'v)\ instr\ =\ Const\ 'v\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ |\ Load\ 'a\isanewline
+\ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ \ |\ Apply\ {"}'v\ binop{"}%
 \begin{isamarkuptext}%
 The execution of the stack machine is modelled by a function
 \isa{exec} that takes a list of instructions, a store (modelled as a
@@ -44,13 +44,13 @@
 and returns the stack at the end of the execution---the store remains
 unchanged:%
 \end{isamarkuptext}%
-\isacommand{consts}~exec~::~{"}('a,'v)instr~list~{\isasymRightarrow}~('a{\isasymRightarrow}'v)~{\isasymRightarrow}~'v~list~{\isasymRightarrow}~'v~list{"}\isanewline
+\isacommand{consts}\ exec\ ::\ {"}('a,'v)instr\ list\ {\isasymRightarrow}\ ('a{\isasymRightarrow}'v)\ {\isasymRightarrow}\ 'v\ list\ {\isasymRightarrow}\ 'v\ list{"}\isanewline
 \isacommand{primrec}\isanewline
-{"}exec~[]~s~vs~=~vs{"}\isanewline
-{"}exec~(i\#is)~s~vs~=~(case~i~of\isanewline
-~~~~Const~v~~{\isasymRightarrow}~exec~is~s~(v\#vs)\isanewline
-~~|~Load~a~~~{\isasymRightarrow}~exec~is~s~((s~a)\#vs)\isanewline
-~~|~Apply~f~~{\isasymRightarrow}~exec~is~s~((f~(hd~vs)~(hd(tl~vs)))\#(tl(tl~vs)))){"}%
+{"}exec\ []\ s\ vs\ =\ vs{"}\isanewline
+{"}exec\ (i\#is)\ s\ vs\ =\ (case\ i\ of\isanewline
+\ \ \ \ Const\ v\ \ {\isasymRightarrow}\ exec\ is\ s\ (v\#vs)\isanewline
+\ \ |\ Load\ a\ \ \ {\isasymRightarrow}\ exec\ is\ s\ ((s\ a)\#vs)\isanewline
+\ \ |\ Apply\ f\ \ {\isasymRightarrow}\ exec\ is\ s\ ((f\ (hd\ vs)\ (hd(tl\ vs)))\#(tl(tl\ vs)))){"}%
 \begin{isamarkuptext}%
 \noindent
 Recall that \isa{hd} and \isa{tl}
@@ -64,29 +64,29 @@
 The compiler is a function from expressions to a list of instructions. Its
 definition is pretty much obvious:%
 \end{isamarkuptext}%
-\isacommand{consts}~comp~::~{"}('a,'v)expr~{\isasymRightarrow}~('a,'v)instr~list{"}\isanewline
+\isacommand{consts}\ comp\ ::\ {"}('a,'v)expr\ {\isasymRightarrow}\ ('a,'v)instr\ list{"}\isanewline
 \isacommand{primrec}\isanewline
-{"}comp~(Cex~v)~~~~~~~=~[Const~v]{"}\isanewline
-{"}comp~(Vex~a)~~~~~~~=~[Load~a]{"}\isanewline
-{"}comp~(Bex~f~e1~e2)~=~(comp~e2)~@~(comp~e1)~@~[Apply~f]{"}%
+{"}comp\ (Cex\ v)\ \ \ \ \ \ \ =\ [Const\ v]{"}\isanewline
+{"}comp\ (Vex\ a)\ \ \ \ \ \ \ =\ [Load\ a]{"}\isanewline
+{"}comp\ (Bex\ f\ e1\ e2)\ =\ (comp\ e2)\ @\ (comp\ e1)\ @\ [Apply\ f]{"}%
 \begin{isamarkuptext}%
 Now we have to prove the correctness of the compiler, i.e.\ that the
 execution of a compiled expression results in the value of the expression:%
 \end{isamarkuptext}%
-\isacommand{theorem}~{"}exec~(comp~e)~s~[]~=~[value~e~s]{"}%
+\isacommand{theorem}\ {"}exec\ (comp\ e)\ s\ []\ =\ [value\ e\ s]{"}%
 \begin{isamarkuptext}%
 \noindent
 This theorem needs to be generalized to%
 \end{isamarkuptext}%
-\isacommand{theorem}~{"}{\isasymforall}vs.~exec~(comp~e)~s~vs~=~(value~e~s)~\#~vs{"}%
+\isacommand{theorem}\ {"}{\isasymforall}vs.\ exec\ (comp\ e)\ s\ vs\ =\ (value\ e\ s)\ \#\ vs{"}%
 \begin{isamarkuptxt}%
 \noindent
 which is proved by induction on \isa{e} followed by simplification, once
 we have the following lemma about executing the concatenation of two
 instruction sequences:%
 \end{isamarkuptxt}%
-\isacommand{lemma}~exec\_app[simp]:\isanewline
-~~{"}{\isasymforall}vs.~exec~(xs@ys)~s~vs~=~exec~ys~s~(exec~xs~s~vs){"}%
+\isacommand{lemma}\ exec\_app[simp]:\isanewline
+\ \ {"}{\isasymforall}vs.\ exec\ (xs@ys)\ s\ vs\ =\ exec\ ys\ s\ (exec\ xs\ s\ vs){"}%
 \begin{isamarkuptxt}%
 \noindent
 This requires induction on \isa{xs} and ordinary simplification for the
@@ -94,14 +94,14 @@
 that contains two \isa{case}-expressions over instructions. Thus we add
 automatic case splitting as well, which finishes the proof:%
 \end{isamarkuptxt}%
-\isacommand{by}(induct\_tac~xs,~simp,~simp~split:~instr.split)%
+\isacommand{by}(induct\_tac\ xs,\ simp,\ simp\ split:\ instr.split)%
 \begin{isamarkuptext}%
 \noindent
 Note that because \isaindex{auto} performs simplification, it can
 also be modified in the same way \isa{simp} can. Thus the proof can be
 rewritten as%
 \end{isamarkuptext}%
-\isacommand{by}(induct\_tac~xs,~auto~split:~instr.split)%
+\isacommand{by}(induct\_tac\ xs,\ auto\ split:\ instr.split)%
 \begin{isamarkuptext}%
 \noindent
 Although this is more compact, it is less clear for the reader of the proof.