src/Doc/Codegen/Refinement.thy
author wenzelm
Sat, 05 Jan 2019 17:24:33 +0100
changeset 69597 ff784d5a5bfb
parent 69505 cc2d676d5395
child 69660 2bc2a8599369
permissions -rw-r--r--
isabelle update -u control_cartouches;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
38405
7935b334893e sketch of new outline
haftmann
parents:
diff changeset
     1
theory Refinement
69422
472af2d7835d clarified session dependencies: faster build_doc/build_release;
wenzelm
parents: 68484
diff changeset
     2
imports Setup
38405
7935b334893e sketch of new outline
haftmann
parents:
diff changeset
     3
begin
7935b334893e sketch of new outline
haftmann
parents:
diff changeset
     4
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
     5
section \<open>Program and datatype refinement \label{sec:refinement}\<close>
38405
7935b334893e sketch of new outline
haftmann
parents:
diff changeset
     6
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
     7
text \<open>
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
     8
  Code generation by shallow embedding (cf.~\secref{sec:principle})
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
     9
  allows to choose code equations and datatype constructors freely,
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    10
  given that some very basic syntactic properties are met; this
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    11
  flexibility opens up mechanisms for refinement which allow to extend
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    12
  the scope and quality of generated code dramatically.
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    13
\<close>
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    14
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    15
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    16
subsection \<open>Program refinement\<close>
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    17
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    18
text \<open>
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    19
  Program refinement works by choosing appropriate code equations
40352
8fd36f8a5cb7 more precise text
haftmann
parents: 39745
diff changeset
    20
  explicitly (cf.~\secref{sec:equations}); as example, we use Fibonacci
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    21
  numbers:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    22
\<close>
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    23
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    24
fun %quote fib :: "nat \<Rightarrow> nat" where
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    25
    "fib 0 = 0"
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    26
  | "fib (Suc 0) = Suc 0"
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    27
  | "fib (Suc (Suc n)) = fib n + fib (Suc n)"
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    28
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    29
text \<open>
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    30
  \noindent The runtime of the corresponding code grows exponential due
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    31
  to two recursive calls:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    32
\<close>
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    33
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    34
text %quotetypewriter \<open>
39683
f75a01ee6c41 prefer typewrite tag over raw latex environment
haftmann
parents: 39664
diff changeset
    35
  @{code_stmts fib (consts) fib (Haskell)}
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    36
\<close>
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    37
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    38
text \<open>
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    39
  \noindent A more efficient implementation would use dynamic
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    40
  programming, e.g.~sharing of common intermediate results between
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    41
  recursive calls.  This idea is expressed by an auxiliary operation
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    42
  which computes a Fibonacci number and its successor simultaneously:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    43
\<close>
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    44
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    45
definition %quote fib_step :: "nat \<Rightarrow> nat \<times> nat" where
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    46
  "fib_step n = (fib (Suc n), fib n)"
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    47
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    48
text \<open>
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    49
  \noindent This operation can be implemented by recursion using
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    50
  dynamic programming:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    51
\<close>
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    52
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    53
lemma %quote [code]:
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    54
  "fib_step 0 = (Suc 0, 0)"
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    55
  "fib_step (Suc n) = (let (m, q) = fib_step n in (m + q, m))"
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    56
  by (simp_all add: fib_step_def)
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    57
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    58
text \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    59
  \noindent What remains is to implement \<^const>\<open>fib\<close> by \<^const>\<open>fib_step\<close> as follows:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    60
\<close>
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    61
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    62
lemma %quote [code]:
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    63
  "fib 0 = 0"
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    64
  "fib (Suc n) = fst (fib_step n)"
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    65
  by (simp_all add: fib_step_def)
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    66
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    67
text \<open>
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    68
  \noindent The resulting code shows only linear growth of runtime:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    69
\<close>
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    70
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    71
text %quotetypewriter \<open>
39683
f75a01ee6c41 prefer typewrite tag over raw latex environment
haftmann
parents: 39664
diff changeset
    72
  @{code_stmts fib (consts) fib fib_step (Haskell)}
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    73
\<close>
38451
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    74
4c065e97ecee added section on program refinement
haftmann
parents: 38437
diff changeset
    75
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    76
subsection \<open>Datatype refinement\<close>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    77
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    78
text \<open>
38459
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
    79
  Selecting specific code equations \emph{and} datatype constructors
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
    80
  leads to datatype refinement.  As an example, we will develop an
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
    81
  alternative representation of the queue example given in
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
    82
  \secref{sec:queue_example}.  The amortised representation is
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
    83
  convenient for generating code but exposes its \qt{implementation}
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
    84
  details, which may be cumbersome when proving theorems about it.
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
    85
  Therefore, here is a simple, straightforward representation of
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
    86
  queues:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
    87
\<close>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    88
58310
91ea607a34d8 updated news
blanchet
parents: 58305
diff changeset
    89
datatype %quote 'a queue = Queue "'a list"
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    90
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    91
definition %quote empty :: "'a queue" where
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    92
  "empty = Queue []"
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    93
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    94
primrec %quote enqueue :: "'a \<Rightarrow> 'a queue \<Rightarrow> 'a queue" where
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    95
  "enqueue x (Queue xs) = Queue (xs @ [x])"
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    96
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    97
fun %quote dequeue :: "'a queue \<Rightarrow> 'a option \<times> 'a queue" where
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    98
    "dequeue (Queue []) = (None, Queue [])"
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
    99
  | "dequeue (Queue (x # xs)) = (Some x, Queue xs)"
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   100
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   101
text \<open>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   102
  \noindent This we can use directly for proving;  for executing,
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   103
  we provide an alternative characterisation:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   104
\<close>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   105
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   106
definition %quote AQueue :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a queue" where
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   107
  "AQueue xs ys = Queue (ys @ rev xs)"
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   108
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   109
code_datatype %quote AQueue
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   110
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   111
text \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   112
  \noindent Here we define a \qt{constructor} \<^const>\<open>AQueue\<close> which
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
   113
  is defined in terms of \<open>Queue\<close> and interprets its arguments
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   114
  according to what the \emph{content} of an amortised queue is supposed
38459
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
   115
  to be.
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
   116
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
   117
  The prerequisite for datatype constructors is only syntactical: a
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
   118
  constructor must be of type \<open>\<tau> = \<dots> \<Rightarrow> \<kappa> \<alpha>\<^sub>1 \<dots> \<alpha>\<^sub>n\<close> where \<open>{\<alpha>\<^sub>1, \<dots>, \<alpha>\<^sub>n}\<close> is exactly the set of \emph{all} type variables in
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 69422
diff changeset
   119
  \<open>\<tau>\<close>; then \<open>\<kappa>\<close> is its corresponding datatype.  The
38459
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
   120
  HOL datatype package by default registers any new datatype with its
38511
abf95b39d65c use command_def more consciously
haftmann
parents: 38502
diff changeset
   121
  constructors, but this may be changed using @{command_def
38459
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
   122
  code_datatype}; the currently chosen constructors can be inspected
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
   123
  using the @{command print_codesetup} command.
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
   124
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
   125
  Equipped with this, we are able to prove the following equations
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   126
  for our primitive queue operations which \qt{implement} the simple
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   127
  queues in an amortised fashion:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   128
\<close>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   129
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   130
lemma %quote empty_AQueue [code]:
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   131
  "empty = AQueue [] []"
40754
e3d4f2522a5f added equation for Queue;
haftmann
parents: 40352
diff changeset
   132
  by (simp add: AQueue_def empty_def)
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   133
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   134
lemma %quote enqueue_AQueue [code]:
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   135
  "enqueue x (AQueue xs ys) = AQueue (x # xs) ys"
40754
e3d4f2522a5f added equation for Queue;
haftmann
parents: 40352
diff changeset
   136
  by (simp add: AQueue_def)
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   137
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   138
lemma %quote dequeue_AQueue [code]:
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   139
  "dequeue (AQueue xs []) =
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   140
    (if xs = [] then (None, AQueue [] [])
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   141
    else dequeue (AQueue [] (rev xs)))"
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   142
  "dequeue (AQueue xs (y # ys)) = (Some y, AQueue xs ys)"
40754
e3d4f2522a5f added equation for Queue;
haftmann
parents: 40352
diff changeset
   143
  by (simp_all add: AQueue_def)
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   144
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   145
text \<open>
40754
e3d4f2522a5f added equation for Queue;
haftmann
parents: 40352
diff changeset
   146
  \noindent It is good style, although no absolute requirement, to
e3d4f2522a5f added equation for Queue;
haftmann
parents: 40352
diff changeset
   147
  provide code equations for the original artefacts of the implemented
e3d4f2522a5f added equation for Queue;
haftmann
parents: 40352
diff changeset
   148
  type, if possible; in our case, these are the datatype constructor
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   149
  \<^const>\<open>Queue\<close> and the case combinator \<^const>\<open>case_queue\<close>:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   150
\<close>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   151
40754
e3d4f2522a5f added equation for Queue;
haftmann
parents: 40352
diff changeset
   152
lemma %quote Queue_AQueue [code]:
e3d4f2522a5f added equation for Queue;
haftmann
parents: 40352
diff changeset
   153
  "Queue = AQueue []"
e3d4f2522a5f added equation for Queue;
haftmann
parents: 40352
diff changeset
   154
  by (simp add: AQueue_def fun_eq_iff)
e3d4f2522a5f added equation for Queue;
haftmann
parents: 40352
diff changeset
   155
55422
6445a05a1234 compile
blanchet
parents: 53125
diff changeset
   156
lemma %quote case_queue_AQueue [code]:
6445a05a1234 compile
blanchet
parents: 53125
diff changeset
   157
  "case_queue f (AQueue xs ys) = f (ys @ rev xs)"
40754
e3d4f2522a5f added equation for Queue;
haftmann
parents: 40352
diff changeset
   158
  by (simp add: AQueue_def)
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   159
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   160
text \<open>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   161
  \noindent The resulting code looks as expected:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   162
\<close>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   163
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   164
text %quotetypewriter \<open>
55422
6445a05a1234 compile
blanchet
parents: 53125
diff changeset
   165
  @{code_stmts empty enqueue dequeue Queue case_queue (SML)}
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   166
\<close>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   167
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   168
text \<open>
38459
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
   169
  The same techniques can also be applied to types which are not
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   170
  specified as datatypes, e.g.~type \<^typ>\<open>int\<close> is originally specified
38511
abf95b39d65c use command_def more consciously
haftmann
parents: 38502
diff changeset
   171
  as quotient type by means of @{command_def typedef}, but for code
38459
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
   172
  generation constants allowing construction of binary numeral values
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   173
  are used as constructors for \<^typ>\<open>int\<close>.
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   174
38459
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
   175
  This approach however fails if the representation of a type demands
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
   176
  invariants; this issue is discussed in the next section.
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   177
\<close>
38459
cfe74b0eecb1 reworked section on simple datatype refinement
haftmann
parents: 38451
diff changeset
   178
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   179
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   180
subsection \<open>Datatype refinement involving invariants \label{sec:invariant}\<close>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   181
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   182
text \<open>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   183
  Datatype representation involving invariants require a dedicated
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   184
  setup for the type and its primitive operations.  As a running
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   185
  example, we implement a type \<^typ>\<open>'a dlist\<close> of lists consisting
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   186
  of distinct elements.
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   187
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   188
  The specification of \<^typ>\<open>'a dlist\<close> itself can be found in theory
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   189
  \<^theory>\<open>HOL-Library.Dlist\<close>.
66405
82e2291cabff be more explicit on type dlist
haftmann
parents: 59377
diff changeset
   190
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   191
  The first step is to decide on which representation the abstract
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   192
  type (in our example \<^typ>\<open>'a dlist\<close>) should be implemented.
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   193
  Here we choose \<^typ>\<open>'a list\<close>.  Then a conversion from the concrete
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   194
  type to the abstract type must be specified, here:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   195
\<close>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   196
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   197
text %quote \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   198
  \<^term_type>\<open>Dlist\<close>
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   199
\<close>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   200
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   201
text \<open>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   202
  \noindent Next follows the specification of a suitable \emph{projection},
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   203
  i.e.~a conversion from abstract to concrete type:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   204
\<close>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   205
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   206
text %quote \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   207
  \<^term_type>\<open>list_of_dlist\<close>
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   208
\<close>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   209
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   210
text \<open>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   211
  \noindent This projection must be specified such that the following
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   212
  \emph{abstract datatype certificate} can be proven:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   213
\<close>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   214
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   215
lemma %quote [code abstype]:
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   216
  "Dlist (list_of_dlist dxs) = dxs"
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   217
  by (fact Dlist_list_of_dlist)
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   218
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   219
text \<open>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   220
  \noindent Note that so far the invariant on representations
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   221
  (\<^term_type>\<open>distinct\<close>) has never been mentioned explicitly:
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   222
  the invariant is only referred to implicitly: all values in
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   223
  set \<^term>\<open>{xs. list_of_dlist (Dlist xs) = xs}\<close> are invariant,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   224
  and in our example this is exactly \<^term>\<open>{xs. distinct xs}\<close>.
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   225
  
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   226
  The primitive operations on \<^typ>\<open>'a dlist\<close> are specified
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   227
  indirectly using the projection \<^const>\<open>list_of_dlist\<close>.  For
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   228
  the empty \<open>dlist\<close>, \<^const>\<open>Dlist.empty\<close>, we finally want
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   229
  the code equation
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   230
\<close>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   231
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   232
text %quote \<open>
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   233
  \<^term>\<open>Dlist.empty = Dlist []\<close>
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   234
\<close>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   235
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   236
text \<open>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   237
  \noindent This we have to prove indirectly as follows:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   238
\<close>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   239
52637
1501ebe39711 attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents: 48985
diff changeset
   240
lemma %quote [code]:
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   241
  "list_of_dlist Dlist.empty = []"
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   242
  by (fact list_of_dlist_empty)
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   243
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   244
text \<open>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   245
  \noindent This equation logically encodes both the desired code
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   246
  equation and that the expression \<^const>\<open>Dlist\<close> is applied to obeys
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   247
  the implicit invariant.  Equations for insertion and removal are
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   248
  similar:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   249
\<close>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   250
52637
1501ebe39711 attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents: 48985
diff changeset
   251
lemma %quote [code]:
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   252
  "list_of_dlist (Dlist.insert x dxs) = List.insert x (list_of_dlist dxs)"
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   253
  by (fact list_of_dlist_insert)
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   254
52637
1501ebe39711 attribute "code" declares concrete and abstract code equations uniformly; added explicit "code equation" instead
haftmann
parents: 48985
diff changeset
   255
lemma %quote [code]:
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   256
  "list_of_dlist (Dlist.remove x dxs) = remove1 x (list_of_dlist dxs)"
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   257
  by (fact list_of_dlist_remove)
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   258
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   259
text \<open>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   260
  \noindent Then the corresponding code is as follows:
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   261
\<close>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   262
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   263
text %quotetypewriter \<open>
39683
f75a01ee6c41 prefer typewrite tag over raw latex environment
haftmann
parents: 39664
diff changeset
   264
  @{code_stmts Dlist.empty Dlist.insert Dlist.remove list_of_dlist (Haskell)}
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   265
\<close>
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   266
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   267
text \<open>
58620
7435b6a3f72e more antiquotations;
wenzelm
parents: 58310
diff changeset
   268
  See further @{cite "Haftmann-Kraus-Kuncar-Nipkow:2013:data_refinement"}
53125
f4c6f8f6515b reference to datatype refinment paper
haftmann
parents: 53015
diff changeset
   269
  for the meta theory of datatype refinement involving invariants.
f4c6f8f6515b reference to datatype refinment paper
haftmann
parents: 53015
diff changeset
   270
38502
c4b7ae8ea82e added quick and dirty section on invariants
haftmann
parents: 38459
diff changeset
   271
  Typical data structures implemented by representations involving
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   272
  invariants are available in the library, theory \<^theory>\<open>HOL-Library.Mapping\<close>
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   273
  specifies key-value-mappings (type \<^typ>\<open>('a, 'b) mapping\<close>);
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
   274
  these can be implemented by red-black-trees (theory \<^theory>\<open>HOL-Library.RBT\<close>).
59377
056945909f60 modernized cartouches
haftmann
parents: 58620
diff changeset
   275
\<close>
38437
ffb1c5bf0425 adaptation to new outline
haftmann
parents: 38405
diff changeset
   276
38405
7935b334893e sketch of new outline
haftmann
parents:
diff changeset
   277
end