src/HOL/Isar_Examples/Expr_Compiler.thy
author wenzelm
Fri, 25 Jun 2010 11:48:37 +0200
changeset 37536 c62aa9281101
parent 35416 d8d7d1b785af
child 37671 fa53d267dab3
permissions -rw-r--r--
explicit treatment of UTF8 sequences as Isabelle symbols;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33026
8f35633c4922 modernized session Isar_Examples;
wenzelm
parents: 31758
diff changeset
     1
(*  Title:      HOL/Isar_Examples/Expr_Compiler.thy
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
     2
    Author:     Markus Wenzel, TU Muenchen
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
     3
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
     4
Correctness of a simple expression/stack-machine compiler.
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
     5
*)
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
     6
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
     7
header {* Correctness of a simple expression compiler *}
7748
5b9c45b21782 improved presentation;
wenzelm
parents: 7740
diff changeset
     8
31758
3edd5f813f01 observe standard theory naming conventions;
wenzelm
parents: 23373
diff changeset
     9
theory Expr_Compiler
3edd5f813f01 observe standard theory naming conventions;
wenzelm
parents: 23373
diff changeset
    10
imports Main
3edd5f813f01 observe standard theory naming conventions;
wenzelm
parents: 23373
diff changeset
    11
begin
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    12
7869
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    13
text {*
7874
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    14
 This is a (rather trivial) example of program verification.  We model
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    15
 a compiler for translating expressions to stack machine instructions,
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    16
 and prove its correctness wrt.\ some evaluation semantics.
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    17
*}
7869
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    18
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    19
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    20
subsection {* Binary operations *}
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    21
6744
9727e83f0578 changed {| |} verbatim syntax to {* *};
wenzelm
parents: 6517
diff changeset
    22
text {*
7869
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    23
 Binary operations are just functions over some type of values.  This
7982
d534b897ce39 improved presentation;
wenzelm
parents: 7874
diff changeset
    24
 is both for abstract syntax and semantics, i.e.\ we use a ``shallow
7874
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    25
 embedding'' here.
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    26
*}
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    27
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    28
types
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    29
  'val binop = "'val => 'val => 'val"
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    30
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    31
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    32
subsection {* Expressions *}
7869
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    33
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    34
text {*
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    35
 The language of expressions is defined as an inductive type,
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    36
 consisting of variables, constants, and binary operations on
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    37
 expressions.
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    38
*}
7869
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    39
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    40
datatype ('adr, 'val) expr =
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    41
  Variable 'adr |
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    42
  Constant 'val |
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    43
  Binop "'val binop" "('adr, 'val) expr" "('adr, 'val) expr"
7869
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    44
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    45
text {*
7874
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    46
 Evaluation (wrt.\ some environment of variable assignments) is
180364256231 improved presentation;
wenzelm
parents: 7869
diff changeset
    47
 defined by primitive recursion over the structure of expressions.
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    48
*}
7869
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    49
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    50
consts
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    51
  eval :: "('adr, 'val) expr => ('adr => 'val) => 'val"
7869
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    52
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    53
primrec
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    54
  "eval (Variable x) env = env x"
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    55
  "eval (Constant c) env = c"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    56
  "eval (Binop f e1 e2) env = f (eval e1 env) (eval e2 env)"
7869
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    57
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    58
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    59
subsection {* Machine *}
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    60
6744
9727e83f0578 changed {| |} verbatim syntax to {* *};
wenzelm
parents: 6517
diff changeset
    61
text {*
6792
c68035d06cce tuned comments;
wenzelm
parents: 6746
diff changeset
    62
 Next we model a simple stack machine, with three instructions.
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    63
*}
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    64
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    65
datatype ('adr, 'val) instr =
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    66
  Const 'val |
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    67
  Load 'adr |
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    68
  Apply "'val binop"
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    69
6744
9727e83f0578 changed {| |} verbatim syntax to {* *};
wenzelm
parents: 6517
diff changeset
    70
text {*
7869
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    71
 Execution of a list of stack machine instructions is easily defined
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    72
 as follows.
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    73
*}
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    74
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    75
consts
7761
7fab9592384f improved presentation;
wenzelm
parents: 7748
diff changeset
    76
  exec :: "(('adr, 'val) instr) list
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    77
    => 'val list => ('adr => 'val) => 'val list"
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    78
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    79
primrec
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    80
  "exec [] stack env = stack"
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    81
  "exec (instr # instrs) stack env =
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    82
    (case instr of
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    83
      Const c => exec instrs (c # stack) env
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    84
    | Load x => exec instrs (env x # stack) env
7761
7fab9592384f improved presentation;
wenzelm
parents: 7748
diff changeset
    85
    | Apply f => exec instrs (f (hd stack) (hd (tl stack))
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    86
                   # (tl (tl stack))) env)"
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    87
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 33026
diff changeset
    88
definition execute :: "(('adr, 'val) instr) list => ('adr => 'val) => 'val" where
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    89
  "execute instrs env == hd (exec instrs [] env)"
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    90
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    91
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    92
subsection {* Compiler *}
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    93
6744
9727e83f0578 changed {| |} verbatim syntax to {* *};
wenzelm
parents: 6517
diff changeset
    94
text {*
7869
c007f801cd59 improved presentation;
wenzelm
parents: 7761
diff changeset
    95
 We are ready to define the compilation function of expressions to
6792
c68035d06cce tuned comments;
wenzelm
parents: 6746
diff changeset
    96
 lists of stack machine instructions.
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
    97
*}
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    98
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
    99
consts
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   100
  compile :: "('adr, 'val) expr => (('adr, 'val) instr) list"
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   101
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   102
primrec
8031
826c6222cca2 renamed comp to compile (avoids clash with Relation.comp);
wenzelm
parents: 7982
diff changeset
   103
  "compile (Variable x) = [Load x]"
826c6222cca2 renamed comp to compile (avoids clash with Relation.comp);
wenzelm
parents: 7982
diff changeset
   104
  "compile (Constant c) = [Const c]"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   105
  "compile (Binop f e1 e2) = compile e2 @ compile e1 @ [Apply f]"
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   106
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   107
6744
9727e83f0578 changed {| |} verbatim syntax to {* *};
wenzelm
parents: 6517
diff changeset
   108
text {*
8051
wenzelm
parents: 8031
diff changeset
   109
 The main result of this development is the correctness theorem for
wenzelm
parents: 8031
diff changeset
   110
 $\idt{compile}$.  We first establish a lemma about $\idt{exec}$ and
wenzelm
parents: 8031
diff changeset
   111
 list append.
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   112
*}
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   113
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   114
lemma exec_append:
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   115
  "exec (xs @ ys) stack env =
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   116
    exec ys (exec xs stack env) env"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 18193
diff changeset
   117
proof (induct xs arbitrary: stack)
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   118
  case Nil
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   119
  show ?case by simp
11809
c9ffdd63dd93 tuned induction proofs;
wenzelm
parents: 10007
diff changeset
   120
next
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   121
  case (Cons x xs)
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   122
  show ?case
11809
c9ffdd63dd93 tuned induction proofs;
wenzelm
parents: 10007
diff changeset
   123
  proof (induct x)
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 20523
diff changeset
   124
    case Const
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 20523
diff changeset
   125
    from Cons show ?case by simp
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   126
  next
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 20523
diff changeset
   127
    case Load
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 20523
diff changeset
   128
    from Cons show ?case by simp
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   129
  next
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 20523
diff changeset
   130
    case Apply
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 20523
diff changeset
   131
    from Cons show ?case by simp
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   132
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   133
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   134
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   135
theorem correctness: "execute (compile e) env = eval e env"
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   136
proof -
18193
54419506df9e tuned document;
wenzelm
parents: 18153
diff changeset
   137
  have "\<And>stack. exec (compile e) stack env = eval e env # stack"
11809
c9ffdd63dd93 tuned induction proofs;
wenzelm
parents: 10007
diff changeset
   138
  proof (induct e)
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   139
    case Variable show ?case by simp
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   140
  next
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   141
    case Constant show ?case by simp
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   142
  next
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   143
    case Binop then show ?case by (simp add: exec_append)
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   144
  qed
23373
ead82c82da9e tuned proofs: avoid implicit prems;
wenzelm
parents: 20523
diff changeset
   145
  then show ?thesis by (simp add: execute_def)
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   146
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   147
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   148
8051
wenzelm
parents: 8031
diff changeset
   149
text {*
wenzelm
parents: 8031
diff changeset
   150
 \bigskip In the proofs above, the \name{simp} method does quite a lot
wenzelm
parents: 8031
diff changeset
   151
 of work behind the scenes (mostly ``functional program execution'').
wenzelm
parents: 8031
diff changeset
   152
 Subsequently, the same reasoning is elaborated in detail --- at most
wenzelm
parents: 8031
diff changeset
   153
 one recursive function definition is used at a time.  Thus we get a
wenzelm
parents: 8031
diff changeset
   154
 better idea of what is actually going on.
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   155
*}
8051
wenzelm
parents: 8031
diff changeset
   156
13524
604d0f3622d6 *** empty log message ***
wenzelm
parents: 11809
diff changeset
   157
lemma exec_append':
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   158
  "exec (xs @ ys) stack env = exec ys (exec xs stack env) env"
20503
503ac4c5ef91 induct method: renamed 'fixing' to 'arbitrary';
wenzelm
parents: 18193
diff changeset
   159
proof (induct xs arbitrary: stack)
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   160
  case (Nil s)
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   161
  have "exec ([] @ ys) s env = exec ys s env" by simp
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   162
  also have "... = exec ys (exec [] s env) env" by simp
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   163
  finally show ?case .
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   164
next
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   165
  case (Cons x xs s)
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   166
  show ?case
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   167
  proof (induct x)
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   168
    case (Const val)
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   169
    have "exec ((Const val # xs) @ ys) s env = exec (Const val # xs @ ys) s env"
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   170
      by simp
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   171
    also have "... = exec (xs @ ys) (val # s) env" by simp
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   172
    also from Cons have "... = exec ys (exec xs (val # s) env) env" .
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   173
    also have "... = exec ys (exec (Const val # xs) s env) env" by simp
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   174
    finally show ?case .
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   175
  next
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   176
    case (Load adr)
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   177
    from Cons show ?case by simp -- {* same as above *}
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   178
  next
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20503
diff changeset
   179
    case (Apply fn)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20503
diff changeset
   180
    have "exec ((Apply fn # xs) @ ys) s env =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20503
diff changeset
   181
        exec (Apply fn # xs @ ys) s env" by simp
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   182
    also have "... =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20503
diff changeset
   183
        exec (xs @ ys) (fn (hd s) (hd (tl s)) # (tl (tl s))) env" by simp
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   184
    also from Cons have "... =
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20503
diff changeset
   185
        exec ys (exec xs (fn (hd s) (hd (tl s)) # tl (tl s)) env) env" .
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20503
diff changeset
   186
    also have "... = exec ys (exec (Apply fn # xs) s env) env" by simp
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   187
    finally show ?case .
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   188
  qed
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   189
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   190
13537
f506eb568121 avoid duplicate fact bindings;
wenzelm
parents: 13524
diff changeset
   191
theorem correctness': "execute (compile e) env = eval e env"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   192
proof -
18193
54419506df9e tuned document;
wenzelm
parents: 18153
diff changeset
   193
  have exec_compile: "\<And>stack. exec (compile e) stack env = eval e env # stack"
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   194
  proof (induct e)
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   195
    case (Variable adr s)
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   196
    have "exec (compile (Variable adr)) s env = exec [Load adr] s env"
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   197
      by simp
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   198
    also have "... = env adr # s" by simp
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   199
    also have "env adr = eval (Variable adr) env" by simp
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   200
    finally show ?case .
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   201
  next
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   202
    case (Constant val s)
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   203
    show ?case by simp -- {* same as above *}
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   204
  next
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20503
diff changeset
   205
    case (Binop fn e1 e2 s)
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20503
diff changeset
   206
    have "exec (compile (Binop fn e1 e2)) s env =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20503
diff changeset
   207
        exec (compile e2 @ compile e1 @ [Apply fn]) s env" by simp
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20503
diff changeset
   208
    also have "... = exec [Apply fn]
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   209
        (exec (compile e1) (exec (compile e2) s env) env) env"
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   210
      by (simp only: exec_append)
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   211
    also have "exec (compile e2) s env = eval e2 env # s" by fact
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   212
    also have "exec (compile e1) ... env = eval e1 env # ..." by fact
20523
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20503
diff changeset
   213
    also have "exec [Apply fn] ... env =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20503
diff changeset
   214
        fn (hd ...) (hd (tl ...)) # (tl (tl ...))" by simp
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20503
diff changeset
   215
    also have "... = fn (eval e1 env) (eval e2 env) # s" by simp
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20503
diff changeset
   216
    also have "fn (eval e1 env) (eval e2 env) =
36a59e5d0039 Major update to function package, including new syntax and the (only theoretical)
krauss
parents: 20503
diff changeset
   217
        eval (Binop fn e1 e2) env"
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   218
      by simp
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   219
    finally show ?case .
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   220
  qed
8051
wenzelm
parents: 8031
diff changeset
   221
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   222
  have "execute (compile e) env = hd (exec (compile e) [] env)"
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   223
    by (simp add: execute_def)
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   224
  also from exec_compile
18153
a084aa91f701 tuned proofs;
wenzelm
parents: 16417
diff changeset
   225
    have "exec (compile e) [] env = [eval e env]" .
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   226
  also have "hd ... = eval e env" by simp
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   227
  finally show ?thesis .
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   228
qed
6444
2ebe9e630cab Miscellaneous Isabelle/Isar examples for Higher-Order Logic.
wenzelm
parents:
diff changeset
   229
10007
64bf7da1994a isar-strip-terminators;
wenzelm
parents: 9659
diff changeset
   230
end