src/HOL/Imperative_HOL/Overview.thy
author paulson
Thu, 12 Sep 2019 14:51:50 +0100
changeset 70689 67360d50ebb3
parent 69597 ff784d5a5bfb
child 71847 da12452c9be2
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Imperative_HOL/Overview.thy
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
     2
    Author:     Florian Haftmann, TU Muenchen
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
     3
*)
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
     4
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
     5
(*<*)
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
     6
theory Overview
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 63167
diff changeset
     7
imports Imperative_HOL "HOL-Library.LaTeXsugar"
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
     8
begin
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
     9
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    10
(* type constraints with spacing *)
61143
5f898411ce87 eliminated \<Colon> from syntax of constraints;
wenzelm
parents: 58622
diff changeset
    11
no_syntax (output)
5f898411ce87 eliminated \<Colon> from syntax of constraints;
wenzelm
parents: 58622
diff changeset
    12
  "_constrain" :: "logic => type => logic"  ("_::_" [4, 0] 3)
5f898411ce87 eliminated \<Colon> from syntax of constraints;
wenzelm
parents: 58622
diff changeset
    13
  "_constrain" :: "prop' => type => prop'"  ("_::_" [4, 0] 3)
5f898411ce87 eliminated \<Colon> from syntax of constraints;
wenzelm
parents: 58622
diff changeset
    14
5f898411ce87 eliminated \<Colon> from syntax of constraints;
wenzelm
parents: 58622
diff changeset
    15
syntax (output)
5f898411ce87 eliminated \<Colon> from syntax of constraints;
wenzelm
parents: 58622
diff changeset
    16
  "_constrain" :: "logic => type => logic"  ("_ :: _" [4, 0] 3)
5f898411ce87 eliminated \<Colon> from syntax of constraints;
wenzelm
parents: 58622
diff changeset
    17
  "_constrain" :: "prop' => type => prop'"  ("_ :: _" [4, 0] 3)
5f898411ce87 eliminated \<Colon> from syntax of constraints;
wenzelm
parents: 58622
diff changeset
    18
(*>*)
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    19
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    20
text \<open>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    21
  \<open>Imperative HOL\<close> is a leightweight framework for reasoning
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    22
  about imperative data structures in \<open>Isabelle/HOL\<close>
58622
aa99568f56de more antiquotations;
wenzelm
parents: 56240
diff changeset
    23
  @{cite "Nipkow-et-al:2002:tutorial"}.  Its basic ideas are described in
aa99568f56de more antiquotations;
wenzelm
parents: 56240
diff changeset
    24
  @{cite "Bulwahn-et-al:2008:imp_HOL"}.  However their concrete
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    25
  realisation has changed since, due to both extensions and
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    26
  refinements.  Therefore this overview wants to present the framework
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    27
  \qt{as it is} by now.  It focusses on the user-view, less on matters
40358
b6ed3364753d added note on countable types
haftmann
parents: 39717
diff changeset
    28
  of construction.  For details study of the theory sources is
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    29
  encouraged.
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    30
\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    31
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    32
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    33
section \<open>A polymorphic heap inside a monad\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    34
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    35
text \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    36
  Heaps (\<^type>\<open>heap\<close>) can be populated by values of class \<^class>\<open>heap\<close>; HOL's default types are already instantiated to class \<^class>\<open>heap\<close>.  Class \<^class>\<open>heap\<close> is a subclass of \<^class>\<open>countable\<close>;  see
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    37
  theory \<open>Countable\<close> for ways to instantiate types as \<^class>\<open>countable\<close>.
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    38
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    39
  The heap is wrapped up in a monad \<^typ>\<open>'a Heap\<close> by means of the
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    40
  following specification:
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    41
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    42
  \begin{quote}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    43
    \<^datatype>\<open>Heap\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    44
  \end{quote}
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    45
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    46
  Unwrapping of this monad type happens through
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    47
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    48
  \begin{quote}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    49
    \<^term_type>\<open>execute\<close> \\
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    50
    @{thm execute.simps [no_vars]}
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    51
  \end{quote}
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    52
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    53
  This allows for equational reasoning about monadic expressions; the
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    54
  fact collection \<open>execute_simps\<close> contains appropriate rewrites
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    55
  for all fundamental operations.
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    56
39610
23bdf736d84f tuned text
haftmann
parents: 39307
diff changeset
    57
  Primitive fine-granular control over heaps is available through rule
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    58
  \<open>Heap_cases\<close>:
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    59
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    60
  \begin{quote}
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    61
    @{thm [break] Heap_cases [no_vars]}
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    62
  \end{quote}
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    63
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    64
  Monadic expression involve the usual combinators:
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    65
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    66
  \begin{quote}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    67
    \<^term_type>\<open>return\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    68
    \<^term_type>\<open>bind\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    69
    \<^term_type>\<open>raise\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    70
  \end{quote}
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    71
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    72
  This is also associated with nice monad do-syntax.  The \<^typ>\<open>string\<close> argument to \<^const>\<open>raise\<close> is just a codified comment.
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    73
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    74
  Among a couple of generic combinators the following is helpful for
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    75
  establishing invariants:
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    76
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    77
  \begin{quote}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    78
    \<^term_type>\<open>assert\<close> \\
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    79
    @{thm assert_def [no_vars]}
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    80
  \end{quote}
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    81
\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    82
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    83
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    84
section \<open>Relational reasoning about \<^type>\<open>Heap\<close> expressions\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    85
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    86
text \<open>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    87
  To establish correctness of imperative programs, predicate
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    88
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    89
  \begin{quote}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
    90
    \<^term_type>\<open>effect\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    91
  \end{quote}
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    92
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    93
  provides a simple relational calculus.  Primitive rules are \<open>effectI\<close> and \<open>effectE\<close>, rules appropriate for reasoning about
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    94
  imperative operations are available in the \<open>effect_intros\<close> and
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
    95
  \<open>effect_elims\<close> fact collections.
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    96
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    97
  Often non-failure of imperative computations does not depend
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    98
  on the heap at all;  reasoning then can be easier using predicate
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
    99
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   100
  \begin{quote}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   101
    \<^term_type>\<open>success\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   102
  \end{quote}
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   103
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   104
  Introduction rules for \<^const>\<open>success\<close> are available in the
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   105
  \<open>success_intro\<close> fact collection.
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   106
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   107
  \<^const>\<open>execute\<close>, \<^const>\<open>effect\<close>, \<^const>\<open>success\<close> and \<^const>\<open>bind\<close>
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   108
  are related by rules \<open>execute_bind_success\<close>, \<open>success_bind_executeI\<close>, \<open>success_bind_effectI\<close>, \<open>effect_bindI\<close>, \<open>effect_bindE\<close> and \<open>execute_bind_eq_SomeI\<close>.
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   109
\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   110
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   111
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   112
section \<open>Monadic data structures\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   113
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   114
text \<open>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   115
  The operations for monadic data structures (arrays and references)
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   116
  come in two flavours:
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   117
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   118
  \begin{itemize}
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   119
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   120
     \item Operations on the bare heap; their number is kept minimal
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   121
       to facilitate proving.
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   122
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   123
     \item Operations on the heap wrapped up in a monad; these are designed
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   124
       for executing.
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   125
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   126
  \end{itemize}
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   127
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   128
  Provided proof rules are such that they reduce monad operations to
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   129
  operations on bare heaps.
39717
e9bec0b43449 added hint on reference equality
haftmann
parents: 39610
diff changeset
   130
e9bec0b43449 added hint on reference equality
haftmann
parents: 39610
diff changeset
   131
  Note that HOL equality coincides with reference equality and may be
e9bec0b43449 added hint on reference equality
haftmann
parents: 39610
diff changeset
   132
  used as primitive executable operation.
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   133
\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   134
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   135
subsection \<open>Arrays\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   136
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   137
text \<open>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   138
  Heap operations:
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   139
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   140
  \begin{quote}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   141
    \<^term_type>\<open>Array.alloc\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   142
    \<^term_type>\<open>Array.present\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   143
    \<^term_type>\<open>Array.get\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   144
    \<^term_type>\<open>Array.set\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   145
    \<^term_type>\<open>Array.length\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   146
    \<^term_type>\<open>Array.update\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   147
    \<^term_type>\<open>Array.noteq\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   148
  \end{quote}
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   149
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   150
  Monad operations:
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   151
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   152
  \begin{quote}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   153
    \<^term_type>\<open>Array.new\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   154
    \<^term_type>\<open>Array.of_list\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   155
    \<^term_type>\<open>Array.make\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   156
    \<^term_type>\<open>Array.len\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   157
    \<^term_type>\<open>Array.nth\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   158
    \<^term_type>\<open>Array.upd\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   159
    \<^term_type>\<open>Array.map_entry\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   160
    \<^term_type>\<open>Array.swap\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   161
    \<^term_type>\<open>Array.freeze\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   162
  \end{quote}
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   163
\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   164
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   165
subsection \<open>References\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   166
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   167
text \<open>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   168
  Heap operations:
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   169
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   170
  \begin{quote}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   171
    \<^term_type>\<open>Ref.alloc\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   172
    \<^term_type>\<open>Ref.present\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   173
    \<^term_type>\<open>Ref.get\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   174
    \<^term_type>\<open>Ref.set\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   175
    \<^term_type>\<open>Ref.noteq\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   176
  \end{quote}
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   177
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   178
  Monad operations:
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   179
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   180
  \begin{quote}
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   181
    \<^term_type>\<open>Ref.ref\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   182
    \<^term_type>\<open>Ref.lookup\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   183
    \<^term_type>\<open>Ref.update\<close> \\
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 66453
diff changeset
   184
    \<^term_type>\<open>Ref.change\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   185
  \end{quote}
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   186
\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   187
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   188
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   189
section \<open>Code generation\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   190
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   191
text \<open>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   192
  Imperative HOL sets up the code generator in a way that imperative
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   193
  operations are mapped to suitable counterparts in the target
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   194
  language.  For \<open>Haskell\<close>, a suitable \<open>ST\<close> monad is used;
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   195
  for \<open>SML\<close>, \<open>Ocaml\<close> and \<open>Scala\<close> unit values ensure
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   196
  that the evaluation order is the same as you would expect from the
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   197
  original monadic expressions.  These units may look cumbersome; the
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   198
  target language variants \<open>SML_imp\<close>, \<open>Ocaml_imp\<close> and
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   199
  \<open>Scala_imp\<close> make some effort to optimize some of them away.
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   200
\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   201
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   202
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   203
section \<open>Some hints for using the framework\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   204
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   205
text \<open>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   206
  Of course a framework itself does not by itself indicate how to make
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   207
  best use of it.  Here some hints drawn from prior experiences with
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   208
  Imperative HOL:
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   209
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   210
  \begin{itemize}
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   211
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   212
    \item Proofs on bare heaps should be strictly separated from those
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   213
      for monadic expressions.  The first capture the essence, while the
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   214
      latter just describe a certain wrapping-up.
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   215
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   216
    \item A good methodology is to gradually improve an imperative
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   217
      program from a functional one.  In the extreme case this means
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   218
      that an original functional program is decomposed into suitable
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   219
      operations with exactly one corresponding imperative operation.
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   220
      Having shown suitable correspondence lemmas between those, the
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   221
      correctness prove of the whole imperative program simply
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   222
      consists of composing those.
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   223
      
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   224
    \item Whether one should prefer equational reasoning (fact
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   225
      collection \<open>execute_simps\<close> or relational reasoning (fact
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   226
      collections \<open>effect_intros\<close> and \<open>effect_elims\<close>) depends
39610
23bdf736d84f tuned text
haftmann
parents: 39307
diff changeset
   227
      on the problems to solve.  For complex expressions or
23bdf736d84f tuned text
haftmann
parents: 39307
diff changeset
   228
      expressions involving binders, the relation style usually is
23bdf736d84f tuned text
haftmann
parents: 39307
diff changeset
   229
      superior but requires more proof text.
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   230
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   231
    \item Note that you can extend the fact collections of Imperative
39610
23bdf736d84f tuned text
haftmann
parents: 39307
diff changeset
   232
      HOL yourself whenever appropriate.
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   233
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   234
  \end{itemize}
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 62390
diff changeset
   235
\<close>
39307
8d42d668b5b0 added Imperative HOL overview
haftmann
parents:
diff changeset
   236
62390
842917225d56 more canonical names
nipkow
parents: 61143
diff changeset
   237
end