doc-src/Exercises/2003/a6/a6.thy
author wenzelm
Sat, 29 May 2004 15:05:25 +0200
changeset 14830 faa4865ba1ce
parent 14526 51dc6c7b1fd7
permissions -rw-r--r--
removed norm_typ; improved output; refer to Pretty.pp;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
14526
mehta
parents:
diff changeset
     1
(*<*)
mehta
parents:
diff changeset
     2
theory a6 = Main:
mehta
parents:
diff changeset
     3
(*>*)
mehta
parents:
diff changeset
     4
mehta
parents:
diff changeset
     5
mehta
parents:
diff changeset
     6
subsection {* Optimising Compiler Verification *}
mehta
parents:
diff changeset
     7
mehta
parents:
diff changeset
     8
text {*
mehta
parents:
diff changeset
     9
Section 3.3 of the Isabelle tutorial describes an expression compiler for a stack machine. In this exercise we will build and verify an optimising expression compiler for a register machine.
mehta
parents:
diff changeset
    10
*}
mehta
parents:
diff changeset
    11
mehta
parents:
diff changeset
    12
text {*\subsubsection*{The Source Language: Expressions}*}
mehta
parents:
diff changeset
    13
mehta
parents:
diff changeset
    14
text {*
mehta
parents:
diff changeset
    15
The arithmetic expressions we will work with consist of variables, constants, and an arbitrary binary operator @{text "oper"}.
mehta
parents:
diff changeset
    16
*}
mehta
parents:
diff changeset
    17
mehta
parents:
diff changeset
    18
consts oper :: "nat \<Rightarrow> nat \<Rightarrow> nat";
mehta
parents:
diff changeset
    19
mehta
parents:
diff changeset
    20
types var = string;
mehta
parents:
diff changeset
    21
mehta
parents:
diff changeset
    22
datatype exp = 
mehta
parents:
diff changeset
    23
    Const nat 
mehta
parents:
diff changeset
    24
  | Var var
mehta
parents:
diff changeset
    25
  | Op exp exp
mehta
parents:
diff changeset
    26
;
mehta
parents:
diff changeset
    27
mehta
parents:
diff changeset
    28
text {*
mehta
parents:
diff changeset
    29
The state in which an expression is evaluated is modelled by an {\em environment} function that maps variables to constants.
mehta
parents:
diff changeset
    30
*}
mehta
parents:
diff changeset
    31
mehta
parents:
diff changeset
    32
types env = "var \<Rightarrow> nat";
mehta
parents:
diff changeset
    33
mehta
parents:
diff changeset
    34
text {*
mehta
parents:
diff changeset
    35
Define a function @{text "value"} that evaluates an expression in a given environment.
mehta
parents:
diff changeset
    36
*}
mehta
parents:
diff changeset
    37
mehta
parents:
diff changeset
    38
consts value :: "exp \<Rightarrow> env \<Rightarrow> nat";
mehta
parents:
diff changeset
    39
mehta
parents:
diff changeset
    40
text {*\subsubsection*{The Register Machine}*}
mehta
parents:
diff changeset
    41
mehta
parents:
diff changeset
    42
text {*
mehta
parents:
diff changeset
    43
As the name suggests, a register machine uses a collection of registers to store intermediate results. There exists a special register, called the accumulator, that serves as an implicit argument to each instruction. The rest of the registers make up the register file, and can be randomly accessed using an index. 
mehta
parents:
diff changeset
    44
*}
mehta
parents:
diff changeset
    45
mehta
parents:
diff changeset
    46
types regIndex = nat;
mehta
parents:
diff changeset
    47
mehta
parents:
diff changeset
    48
datatype cell = 
mehta
parents:
diff changeset
    49
    Acc
mehta
parents:
diff changeset
    50
  | Reg regIndex;
mehta
parents:
diff changeset
    51
mehta
parents:
diff changeset
    52
mehta
parents:
diff changeset
    53
text {*
mehta
parents:
diff changeset
    54
The state of the register machine is denoted by a function that maps storage cells to constants.
mehta
parents:
diff changeset
    55
*}
mehta
parents:
diff changeset
    56
mehta
parents:
diff changeset
    57
types state = "cell \<Rightarrow> nat";
mehta
parents:
diff changeset
    58
mehta
parents:
diff changeset
    59
text {*
mehta
parents:
diff changeset
    60
The instruction set for the register machine is defined as follows:
mehta
parents:
diff changeset
    61
*}
mehta
parents:
diff changeset
    62
mehta
parents:
diff changeset
    63
datatype instr = 
mehta
parents:
diff changeset
    64
  LI nat        
mehta
parents:
diff changeset
    65
  -- "Load Immediate: loads a constant into the accumulator." 
mehta
parents:
diff changeset
    66
| LOAD regIndex 
mehta
parents:
diff changeset
    67
  -- "Loads the contents of a register into the accumulator."
mehta
parents:
diff changeset
    68
| STORE regIndex 
mehta
parents:
diff changeset
    69
  -- "Saves the contents of the accumulator in a register." 
mehta
parents:
diff changeset
    70
| OPER regIndex 
mehta
parents:
diff changeset
    71
  -- {* Performs the binary operation @{text "oper"}.*}
mehta
parents:
diff changeset
    72
    -- "The first argument is taken from a register."
mehta
parents:
diff changeset
    73
    -- "The second argument is taken from the accumulator." 
mehta
parents:
diff changeset
    74
    -- "The result of the computation is stored in the accumulator."
mehta
parents:
diff changeset
    75
;
mehta
parents:
diff changeset
    76
mehta
parents:
diff changeset
    77
text {*
mehta
parents:
diff changeset
    78
A program is a list of such instructions. The result of running a program is a change of state of the register machine. Define a function @{text "exec"} that models this.
mehta
parents:
diff changeset
    79
*}
mehta
parents:
diff changeset
    80
mehta
parents:
diff changeset
    81
consts exec :: "state \<Rightarrow> instr list \<Rightarrow> state"
mehta
parents:
diff changeset
    82
mehta
parents:
diff changeset
    83
text {*\subsubsection*{Compilation}*}
mehta
parents:
diff changeset
    84
mehta
parents:
diff changeset
    85
text {*
mehta
parents:
diff changeset
    86
The task now is to translate an expression into a sequence of instructions that computes it. At the end of execution, the result should be stored in the accumulator.
mehta
parents:
diff changeset
    87
mehta
parents:
diff changeset
    88
Before execution, the values of each variable need to be stored somewhere in the register file. A {\it mapping} function maps variables to positions in the register file.
mehta
parents:
diff changeset
    89
*}
mehta
parents:
diff changeset
    90
mehta
parents:
diff changeset
    91
types map = "var \<Rightarrow> regIndex";
mehta
parents:
diff changeset
    92
mehta
parents:
diff changeset
    93
text {*
mehta
parents:
diff changeset
    94
Define a function @{text "cmp"} that compiles an expression into a sequence of instructions. The evaluation should proceed in a bottom-up depth-first manner.
mehta
parents:
diff changeset
    95
mehta
parents:
diff changeset
    96
State and prove a theorem expressing the correctness of @{text "cmp"}.
mehta
parents:
diff changeset
    97
mehta
parents:
diff changeset
    98
Hints:
mehta
parents:
diff changeset
    99
\begin{itemize}
mehta
parents:
diff changeset
   100
  \item The compilation function is dependent on the mapping function.
mehta
parents:
diff changeset
   101
  \item The compilation function needs some way of storing intermediate results. It should be clever enough to reuse registers it no longer needs.
mehta
parents:
diff changeset
   102
  \item It may be helpful to assume that at each recursive call, compilation is only allowed to use registers with indices greater than a given value to store intermediate results.
mehta
parents:
diff changeset
   103
\end{itemize}
mehta
parents:
diff changeset
   104
*}
mehta
parents:
diff changeset
   105
mehta
parents:
diff changeset
   106
text {*\subsubsection*{Compiler Optimisation: Common Subexpressions}*}
mehta
parents:
diff changeset
   107
mehta
parents:
diff changeset
   108
text {*
mehta
parents:
diff changeset
   109
In the previous section, the compiler @{text "cmp"} was allowed to evaluate a subexpression every time it occurred. In situations where arithmetic operations are costly, one may want to compute commonly occurring subexpressions only once.
mehta
parents:
diff changeset
   110
mehta
parents:
diff changeset
   111
For example, to compute @{text "(a op b) op (a op b)"}, @{text "cmp"} was allowed three calls to @{text "oper"}, when only two were needed.
mehta
parents:
diff changeset
   112
mehta
parents:
diff changeset
   113
Develop an optimised compiler @{text "optCmp"}, that evaluates every commonly occurring subexpression only once. Prove its correctness.
mehta
parents:
diff changeset
   114
*}
mehta
parents:
diff changeset
   115
mehta
parents:
diff changeset
   116
(*<*) end (*>*)