ProgTutorial/FirstSteps.thy
author Christian Urban <urbanc@in.tum.de>
Sat, 07 Nov 2009 01:03:37 +0100
changeset 375 92f7328dc5cc
parent 374 304426a9aecf
child 376 76b1b679e845
permissions -rw-r--r--
added type work and updated to Isabelle and poly 5.3
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory FirstSteps
25
e2f9f94b26d4 Antiquotation setup is now contained in theory Base.
berghofe
parents: 21
diff changeset
     2
imports Base
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
     5
(*<*)
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
     6
setup{*
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
     7
open_file_with_prelude 
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
     8
  "FirstSteps_Code.thy"
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
     9
  ["theory FirstSteps", "imports Main", "begin"]
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
    10
*}
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
    11
(*>*)
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
    12
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
chapter {* First Steps *}
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
    15
text {*
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
    16
   \begin{flushright}
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
    17
  {\em
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
    18
  ``We will most likely never realize the full importance of painting the Tower,\\ 
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
    19
  that it is the essential element in the conservation of metal works and the\\ 
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
    20
  more meticulous the paint job, the longer the tower shall endure.''} \\[1ex]
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 323
diff changeset
    21
  Gustave Eiffel, In his book {\em The 300-Meter Tower}.\footnote{The Eiffel Tower has been 
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
    22
  re-painted 18 times since its initial construction, an average of once every 
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
    23
  seven years. It takes more than one year for a team of 25 painters to paint the tower 
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 323
diff changeset
    24
  from top to bottom.}
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
    25
  \end{flushright}
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
    26
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 323
diff changeset
    27
  \medskip
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    28
  Isabelle programming is done in ML. Just like lemmas and proofs, ML-code for
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    29
  Isabelle must be part of a theory. If you want to follow the code given in
54
1783211b3494 tuned; added document antiquotation ML_response_fake_both
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
    30
  this chapter, we assume you are working inside the theory starting with
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 249
diff changeset
    32
  \begin{quote}
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    33
  \begin{tabular}{@ {}l}
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 34
diff changeset
    34
  \isacommand{theory} FirstSteps\\
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    35
  \isacommand{imports} Main\\
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    36
  \isacommand{begin}\\
6
007e09485351 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 5
diff changeset
    37
  \ldots
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    38
  \end{tabular}
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 249
diff changeset
    39
  \end{quote}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
    40
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 323
diff changeset
    41
  We also generally assume you are working with the logic HOL. The examples
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 323
diff changeset
    42
  that will be given might need to be adapted if you work in a different logic.
14
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
    43
*}
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    44
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
    45
section {* Including ML-Code\label{sec:include} *}
14
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
    46
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
    47
text {*
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    48
  The easiest and quickest way to include code in a theory is by using the
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    49
  \isacommand{ML}-command. For example:
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    50
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    51
  \begin{isabelle}
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    52
  \begin{graybox}
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    53
  \isacommand{ML}~@{text "\<verbopen>"}\isanewline
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    54
  \hspace{5mm}@{ML "3 + 4"}\isanewline
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    55
  @{text "\<verbclose>"}\isanewline
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    56
  @{text "> 7"}\smallskip
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    57
  \end{graybox}
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    58
  \end{isabelle}
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    59
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    60
  Like normal Isabelle scripts, \isacommand{ML}-commands can be evaluated by
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    61
  using the advance and undo buttons of your Isabelle environment. The code
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    62
  inside the \isacommand{ML}-command can also contain value and function
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    63
  bindings, for example
193
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
    64
*}
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
    65
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
    66
ML %gray {* 
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
    67
  val r = Unsynchronized.ref 0
193
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
    68
  fun f n = n + 1 
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
    69
*}
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
    70
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
    71
text {*
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    72
  and even those can be undone when the proof script is retracted.  As
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    73
  mentioned in the Introduction, we will drop the \isacommand{ML}~@{text
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    74
  "\<verbopen> \<dots> \<verbclose>"} scaffolding whenever we show code. The lines
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    75
  prefixed with @{text [quotes] ">"} are not part of the code, rather they
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    76
  indicate what the response is when the code is evaluated.  There are also
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    77
  the commands \isacommand{ML\_val} and \isacommand{ML\_prf} for including
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    78
  ML-code. The first evaluates the given code, but any effect on the theory,
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    79
  in which the code is embedded, is suppressed. The second needs to be used if
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    80
  ML-code is defined inside a proof. For example
253
3cfd9a8a6de1 added comments about ML_prf and ML_val
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    81
254
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    82
  \begin{quote}
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    83
  \begin{isabelle}
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    84
  \isacommand{lemma}~@{text "test:"}\isanewline
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    85
  \isacommand{shows}~@{text [quotes] "True"}\isanewline
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    86
  \isacommand{ML\_prf}~@{text "\<verbopen>"}~@{ML "writeln \"Trivial!\""}~@{text "\<verbclose>"}\isanewline
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    87
  \isacommand{oops}
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    88
  \end{isabelle}
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    89
  \end{quote}
253
3cfd9a8a6de1 added comments about ML_prf and ML_val
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    90
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    91
  However, both commands will only play minor roles in this tutorial (we will
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
    92
  always arrange that the ML-code is defined outside proofs).
10
df09e49b19bf many changes in the FirstSteps section
Christian Urban <urbanc@in.tum.de>
parents: 6
diff changeset
    93
375
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
    94
  
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
    95
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
    96
102
5e309df58557 general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
Christian Urban <urbanc@in.tum.de>
parents: 101
diff changeset
    97
  Once a portion of code is relatively stable, you usually want to export it
235
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    98
  to a separate ML-file. Such files can then be included somewhere inside a 
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    99
  theory by using the command \isacommand{use}. For example
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
   100
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 249
diff changeset
   101
  \begin{quote}
235
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
   102
  \begin{tabular}{@ {}l}
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
   103
  \isacommand{theory} FirstSteps\\
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
   104
  \isacommand{imports} Main\\
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
   105
  \isacommand{uses}~@{text "(\"file_to_be_included.ML\")"} @{text "\<dots>"}\\
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
   106
  \isacommand{begin}\\
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
   107
  \ldots\\
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
   108
  \isacommand{use}~@{text "\"file_to_be_included.ML\""}\\
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
   109
  \ldots
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
   110
  \end{tabular}
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 249
diff changeset
   111
  \end{quote}
235
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
   112
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
   113
  The \isacommand{uses}-command in the header of the theory is needed in order 
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
   114
  to indicate the dependency of the theory on the ML-file. Alternatively, the 
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
   115
  file can be included by just writing in the header
14
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
   116
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 249
diff changeset
   117
  \begin{quote}
14
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
   118
  \begin{tabular}{@ {}l}
54
1783211b3494 tuned; added document antiquotation ML_response_fake_both
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   119
  \isacommand{theory} FirstSteps\\
14
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
   120
  \isacommand{imports} Main\\
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   121
  \isacommand{uses} @{text "\"file_to_be_included.ML\""} @{text "\<dots>"}\\
14
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
   122
  \isacommand{begin}\\
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
   123
  \ldots
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
   124
  \end{tabular}
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 249
diff changeset
   125
  \end{quote}
235
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
   126
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   127
  Note that no parentheses are given this time. Note also that the included
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   128
  ML-file should not contain any \isacommand{use} itself. Otherwise Isabelle
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   129
  is unable to record all file dependencies, which is a nuisance if you have
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   130
  to track down errors.
14
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
   131
*}
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
   132
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   133
section {* Printing and Debugging\label{sec:printing} *}
14
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
   134
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
   135
text {*
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   136
  During development you might find it necessary to inspect data in your
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   137
  code. This can be done in a ``quick-and-dirty'' fashion using the function
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   138
  @{ML_ind writeln in Output}. For example
10
df09e49b19bf many changes in the FirstSteps section
Christian Urban <urbanc@in.tum.de>
parents: 6
diff changeset
   139
317
d69214e47ef9 added an experimental antiquotation to replace eventually ML_response_fake
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
   140
  @{ML_response_eq [display,gray] "writeln \"any string\"" "\"any string\"" with "(op =)"}
10
df09e49b19bf many changes in the FirstSteps section
Christian Urban <urbanc@in.tum.de>
parents: 6
diff changeset
   141
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   142
  will print out @{text [quotes] "any string"} inside the response buffer of
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   143
  Isabelle.  This function expects a string as argument. If you develop under
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   144
  PolyML, then there is a convenient, though again ``quick-and-dirty'', method
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   145
  for converting values into strings, namely the function 
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   146
  @{ML_ind makestring in PolyML}:
10
df09e49b19bf many changes in the FirstSteps section
Christian Urban <urbanc@in.tum.de>
parents: 6
diff changeset
   147
317
d69214e47ef9 added an experimental antiquotation to replace eventually ML_response_fake
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
   148
  @{ML_response_eq [display,gray] "writeln (PolyML.makestring 1)" "\"1\"" with "(op =)"} 
12
2f1736cb8f26 various changes by Alex and Christian
Christian Urban <urbanc@in.tum.de>
parents: 11
diff changeset
   149
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   150
  However, @{ML makestring in PolyML} only works if the type of what
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   151
  is converted is monomorphic and not a function.
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   152
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   153
  The function @{ML "writeln"} should only be used for testing purposes,
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   154
  because any output this function generates will be overwritten as soon as an
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   155
  error is raised. For printing anything more serious and elaborate, the
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   156
  function @{ML_ind tracing in Output} is more appropriate. This function writes all
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   157
  output into a separate tracing buffer. For example:
14
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
   158
317
d69214e47ef9 added an experimental antiquotation to replace eventually ML_response_fake
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
   159
  @{ML_response_eq [display,gray] "tracing \"foo\"" "\"foo\"" with "(op =)"}
14
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
   160
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   161
  It is also possible to redirect the ``channel'' where the string @{text
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   162
  "foo"} is printed to a separate file, e.g., in order to prevent ProofGeneral from
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   163
  choking on massive amounts of trace output. This redirection can be achieved
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   164
  with the code:
43
02f76f1b6e7b added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   165
*}
14
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
   166
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   167
ML{*val strip_specials =
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   168
let
43
02f76f1b6e7b added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   169
  fun strip ("\^A" :: _ :: cs) = strip cs
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   170
    | strip (c :: cs) = c :: strip cs
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   171
    | strip [] = [];
306
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   172
in 
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   173
  implode o strip o explode 
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   174
end
42
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   175
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   176
fun redirect_tracing stream =
cd612b489504 tuned mostly antiquotation and text
Christian Urban <urbanc@in.tum.de>
parents: 41
diff changeset
   177
 Output.tracing_fn := (fn s =>
14
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
   178
    (TextIO.output (stream, (strip_specials s));
43
02f76f1b6e7b added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   179
     TextIO.output (stream, "\n");
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   180
     TextIO.flushOut stream)) *}
14
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
   181
43
02f76f1b6e7b added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   182
text {*
307
f4fa6540e280 suggestions by Peter Homeier
Christian Urban <urbanc@in.tum.de>
parents: 306
diff changeset
   183
  Calling now
f4fa6540e280 suggestions by Peter Homeier
Christian Urban <urbanc@in.tum.de>
parents: 306
diff changeset
   184
f4fa6540e280 suggestions by Peter Homeier
Christian Urban <urbanc@in.tum.de>
parents: 306
diff changeset
   185
  @{ML [display,gray] "redirect_tracing (TextIO.openOut \"foo.bar\")"}
f4fa6540e280 suggestions by Peter Homeier
Christian Urban <urbanc@in.tum.de>
parents: 306
diff changeset
   186
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   187
  will cause that all tracing information is printed into the file @{text "foo.bar"}.
75
Christian Urban <urbanc@in.tum.de>
parents: 73
diff changeset
   188
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   189
  You can print out error messages with the function @{ML_ind error in Library}; for 
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   190
  example:
75
Christian Urban <urbanc@in.tum.de>
parents: 73
diff changeset
   191
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
   192
  @{ML_response_fake [display,gray] 
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
   193
  "if 0=1 then true else (error \"foo\")" 
122
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   194
"Exception- ERROR \"foo\" raised
79696161ae16 polished
Christian Urban <urbanc@in.tum.de>
parents: 120
diff changeset
   195
At command \"ML\"."}
75
Christian Urban <urbanc@in.tum.de>
parents: 73
diff changeset
   196
307
f4fa6540e280 suggestions by Peter Homeier
Christian Urban <urbanc@in.tum.de>
parents: 306
diff changeset
   197
  This function raises the exception @{text ERROR}, which will then 
304
14173c0e8688 polished comment for error function
Christian Urban <urbanc@in.tum.de>
parents: 301
diff changeset
   198
  be displayed by the infrastructure.
14173c0e8688 polished comment for error function
Christian Urban <urbanc@in.tum.de>
parents: 301
diff changeset
   199
14173c0e8688 polished comment for error function
Christian Urban <urbanc@in.tum.de>
parents: 301
diff changeset
   200
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   201
  \footnote{\bf FIXME Mention how to work with @{ML_ind debug in Toplevel} and 
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   202
  @{ML_ind profiling in Toplevel}.}
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   203
*}
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   204
322
2b7c08d90e2e included some tests
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
   205
(* FIXME
207
d3cd633e8240 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   206
ML {* reset Toplevel.debug *}
d3cd633e8240 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   207
d3cd633e8240 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   208
ML {* fun dodgy_fun () = (raise TYPE ("",[],[]); 1) *}
193
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   209
207
d3cd633e8240 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   210
ML {* fun innocent () = dodgy_fun () *}
d3cd633e8240 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   211
ML {* exception_trace (fn () => cterm_of @{theory} (Bound 0)) *}
d3cd633e8240 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   212
ML {* exception_trace (fn () => innocent ()) *}
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   213
207
d3cd633e8240 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   214
ML {* Toplevel.program (fn () => cterm_of @{theory} (Bound 0)) *}
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   215
207
d3cd633e8240 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 204
diff changeset
   216
ML {* Toplevel.program (fn () => innocent ()) *}
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   217
*)
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   218
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   219
text {*
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
   220
  Most often you want to inspect data of Isabelle's basic data
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   221
  structures, namely @{ML_type term}, @{ML_type cterm} and @{ML_type
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   222
  thm}. Isabelle contains elaborate pretty-printing functions for printing
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   223
  them (see Section \ref{sec:pretty}), but for quick-and-dirty solutions they
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   224
  are a bit unwieldy. One way to transform a term into a string is to use the
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   225
  function @{ML_ind  string_of_term in Syntax} from the structure @{ML_struct
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   226
  Syntax}. For more convenience, we bind this function to the toplevel.
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   227
*}
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   228
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   229
ML{*val string_of_term = Syntax.string_of_term*}
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   230
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   231
text {*
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   232
  It can now be used as follows
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   233
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   234
  @{ML_response_fake [display,gray]
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   235
  "string_of_term @{context} @{term \"1::nat\"}"
317
d69214e47ef9 added an experimental antiquotation to replace eventually ML_response_fake
Christian Urban <urbanc@in.tum.de>
parents: 316
diff changeset
   236
  "\"\\^E\\^Fterm\\^E\\^E\\^Fconst\\^Fname=HOL.one_class.one\\^E1\\^E\\^F\\^E\\^E\\^F\\^E\""} 
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   237
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   238
  We obtain a string corrsponding to the term @{term [show_types] "1::nat"} with some
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   239
  additional information encoded in it. The string can be properly printed by
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   240
  using either the function @{ML  writeln} or @{ML  tracing}:
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   241
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   242
  @{ML_response_fake [display,gray]
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   243
  "writeln (string_of_term @{context} @{term \"1::nat\"})"
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   244
  "\"1\""}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   245
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   246
  or
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   247
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   248
  @{ML_response_fake [display,gray]
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   249
  "tracing (string_of_term @{context} @{term \"1::nat\"})"
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   250
  "\"1\""}
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   251
343
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
   252
  If there are more than one term to be printed, you can use the 
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   253
  function @{ML_ind commas in Library} to separate them.
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   254
*} 
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   255
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   256
ML{*fun string_of_terms ctxt ts =  
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   257
  commas (map (string_of_term ctxt) ts)*}
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   258
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   259
text {*
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   260
  You can also print out terms together with typing information.
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   261
  For this you need to set the reference @{ML_ind show_types in Syntax} 
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   262
  to @{ML true}.
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   263
*}
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   264
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   265
ML{*show_types := true*}
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   266
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   267
text {*
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   268
  Now @{ML string_of_term} prints out
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   269
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   270
  @{ML_response_fake [display, gray]
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   271
  "tracing (string_of_term @{context} @{term \"(1::nat, x)\"})"
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   272
  "(1::nat, x::'a)"}
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   273
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   274
  where @{text 1} and @{text x} are displayed with their inferred type.
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   275
  Even more type information can be printed by setting 
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   276
  the reference @{ML_ind show_all_types in Syntax} to @{ML true}. 
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   277
  We obtain then
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   278
*}
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   279
(*<*)ML %linenos {*show_all_types := true*}
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   280
(*>*)
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   281
text {*
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   282
  @{ML_response_fake [display, gray]
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   283
  "tracing (string_of_term @{context} @{term \"(1::nat, x)\"})"
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   284
  "(Pair::nat \<Rightarrow> 'a \<Rightarrow> nat \<times> 'a) (1::nat) (x::'a)"}
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   285
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   286
  where @{term Pair} is the term-constructor for products. 
371
e6f583366779 solved problem with mixfix.
Christian Urban <urbanc@in.tum.de>
parents: 370
diff changeset
   287
  Other references that influence printing are @{ML_ind show_brackets in Syntax} 
e6f583366779 solved problem with mixfix.
Christian Urban <urbanc@in.tum.de>
parents: 370
diff changeset
   288
  and @{ML_ind show_sorts in Syntax}. 
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   289
*}
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   290
(*<*)ML %linenos {*show_types := false; show_all_types := false*}
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   291
(*>*)
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   292
text {*
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   293
  A @{ML_type cterm} can be transformed into a string by the following function.
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   294
*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   295
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   296
ML{*fun string_of_cterm ctxt ct =  
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   297
  string_of_term ctxt (term_of ct)*}
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   298
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   299
text {*
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   300
  In this example the function @{ML_ind term_of in Thm} extracts the @{ML_type
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   301
  term} from a @{ML_type cterm}.  More than one @{ML_type cterm}s can again be
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   302
  printed with @{ML commas}.
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   303
*} 
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   304
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   305
ML{*fun string_of_cterms ctxt cts =  
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   306
  commas (map (string_of_cterm ctxt) cts)*}
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   307
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   308
text {*
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   309
  The easiest way to get the string of a theorem is to transform it
369
74ba778b09c9 tuned index
Christian Urban <urbanc@in.tum.de>
parents: 361
diff changeset
   310
  into a @{ML_type term} using the function @{ML_ind prop_of in Thm}. 
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   311
*}
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   312
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 249
diff changeset
   313
ML{*fun string_of_thm ctxt thm =
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   314
  string_of_term ctxt (prop_of thm)*}
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   315
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   316
text {*
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   317
  Theorems include schematic variables, such as @{text "?P"}, 
343
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
   318
  @{text "?Q"} and so on. They are needed in Isabelle in order to able to
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   319
  instantiate theorems when they are applied. For example the theorem 
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   320
  @{thm [source] conjI} shown below can be used for any (typable) 
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
   321
  instantiation of @{text "?P"} and @{text "?Q"}. 
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   322
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   323
  @{ML_response_fake [display, gray]
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   324
  "tracing (string_of_thm @{context} @{thm conjI})"
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   325
  "\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q"}
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   326
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   327
  However, in order to improve the readability when printing theorems, we
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   328
  convert these schematic variables into free variables using the function
316
74f0a06f751f further polishing of index generation
Christian Urban <urbanc@in.tum.de>
parents: 315
diff changeset
   329
  @{ML_ind  import in Variable}. This is similar to statements like @{text
343
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
   330
  "conjI[no_vars]"} on Isabelle's user-level.
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   331
*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   332
158
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
   333
ML{*fun no_vars ctxt thm =
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
   334
let 
263
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   335
  val ((_, [thm']), _) = Variable.import true [thm] ctxt
158
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
   336
in
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
   337
  thm'
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
   338
end
d7944bdf7b3f removed infix_conv and moved function no_vars into the FirstSteps chapter
Christian Urban <urbanc@in.tum.de>
parents: 157
diff changeset
   339
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 249
diff changeset
   340
fun string_of_thm_no_vars ctxt thm =
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   341
  string_of_term ctxt (prop_of (no_vars ctxt thm))*}
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   342
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   343
text {* 
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   344
  With this function, theorem @{thm [source] conjI} is now printed as follows:
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   345
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   346
  @{ML_response_fake [display, gray]
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   347
  "tracing (string_of_thm_no_vars @{context} @{thm conjI})"
196
840b49bfb1cf fixed `str_of_thms' output in example + small changes
griff
parents: 192
diff changeset
   348
  "\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> P \<and> Q"}
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   349
  
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   350
  Again the function @{ML commas} helps with printing more than one theorem. 
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   351
*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   352
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 249
diff changeset
   353
ML{*fun string_of_thms ctxt thms =  
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 249
diff changeset
   354
  commas (map (string_of_thm ctxt) thms)
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   355
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 249
diff changeset
   356
fun string_of_thms_no_vars ctxt thms =  
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 249
diff changeset
   357
  commas (map (string_of_thm_no_vars ctxt) thms) *}
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   358
305
2ac9dc1a95b4 added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
parents: 304
diff changeset
   359
text {*
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   360
  \begin{readmore}
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   361
  The simple conversion functions from Isabelle's main datatypes to 
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   362
  @{ML_type string}s are implemented in @{ML_file "Pure/Syntax/syntax.ML"}. 
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   363
  The references that change the printing information are declared in 
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   364
  @{ML_file "Pure/Syntax/printer.ML"}.
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   365
  \end{readmore}
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   366
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   367
  Note that for printing out several ``parcels'' of information that
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   368
  semantically belong together, like a warning message consisting of 
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   369
  a term and its type, you should try to keep this information together in a
343
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
   370
  single string. Therefore do \emph{not} print out information as
306
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   371
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   372
@{ML_response_fake [display,gray]
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   373
"tracing \"First half,\"; 
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   374
tracing \"and second half.\""
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   375
"First half,
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   376
and second half."}
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   377
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   378
  but as a single string with appropriate formatting. For example
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   379
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   380
@{ML_response_fake [display,gray]
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   381
"tracing (\"First half,\" ^ \"\\n\" ^ \"and second half.\")"
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   382
"First half,
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   383
and second half."}
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   384
  
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   385
  To ease this kind of string manipulations, there are a number
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   386
  of library functions in Isabelle. For example,  the function 
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   387
  @{ML_ind cat_lines in Library} concatenates a list of strings 
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   388
  and inserts newlines in between each element. 
305
2ac9dc1a95b4 added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
parents: 304
diff changeset
   389
2ac9dc1a95b4 added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
parents: 304
diff changeset
   390
  @{ML_response [display, gray]
2ac9dc1a95b4 added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
parents: 304
diff changeset
   391
  "cat_lines [\"foo\", \"bar\"]"
2ac9dc1a95b4 added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
parents: 304
diff changeset
   392
  "\"foo\\nbar\""}
306
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   393
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   394
  Section \ref{sec:pretty} will also explain the infrastructure that Isabelle 
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   395
  provides for more elaborate pretty printing. 
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   396
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   397
  \begin{readmore}
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   398
  Most of the basic string functions of Isabelle are defined in 
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   399
  @{ML_file "Pure/library.ML"}.
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   400
  \end{readmore}
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   401
305
2ac9dc1a95b4 added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
parents: 304
diff changeset
   402
*}
2ac9dc1a95b4 added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
parents: 304
diff changeset
   403
2ac9dc1a95b4 added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
parents: 304
diff changeset
   404
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   405
section {* Combinators\label{sec:combinators} *}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   406
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   407
text {*
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
   408
  For beginners perhaps the most puzzling parts in the existing code of Isabelle are
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   409
  the combinators. At first they seem to greatly obstruct the
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   410
  comprehension of the code, but after getting familiar with them, they
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   411
  actually ease the understanding and also the programming.
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   412
373
28a49fe024c9 added structure index
Christian Urban <urbanc@in.tum.de>
parents: 372
diff changeset
   413
  The simplest combinator is @{ML_ind I in Library}, which is just the 
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   414
  identity function defined as
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   415
*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   416
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   417
ML{*fun I x = x*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   418
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   419
text {* 
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   420
  Another simple combinator is @{ML_ind K in Library}, defined as 
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   421
*}
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   422
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   423
ML{*fun K x = fn _ => x*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   424
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   425
text {*
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   426
  @{ML K} ``wraps'' a function around @{text "x"} that ignores its argument. As a 
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   427
  result, @{ML K} defines a constant function always returning @{text x}.
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   428
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   429
  The next combinator is reverse application, @{ML_ind  "|>" in Basics}, defined as: 
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   430
*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   431
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   432
ML{*fun x |> f = f x*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   433
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   434
text {* While just syntactic sugar for the usual function application,
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   435
  the purpose of this combinator is to implement functions in a  
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   436
  ``waterfall fashion''. Consider for example the function *}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   437
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   438
ML %linenosgray{*fun inc_by_five x =
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   439
  x |> (fn x => x + 1)
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   440
    |> (fn x => (x, x))
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   441
    |> fst
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   442
    |> (fn x => x + 4)*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   443
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   444
text {*
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   445
  which increments its argument @{text x} by 5. It does this by first incrementing 
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   446
  the argument by 1 (Line 2); then storing the result in a pair (Line 3); taking 
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   447
  the first component of the pair (Line 4) and finally incrementing the first 
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   448
  component by 4 (Line 5). This kind of cascading manipulations of values is quite
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   449
  common when dealing with theories (for example by adding a definition, followed by
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   450
  lemmas and so on). The reverse application allows you to read what happens in 
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   451
  a top-down manner. This kind of coding should also be familiar, 
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 249
diff changeset
   452
  if you have been exposed to Haskell's {\it do}-notation. Writing the function 
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 249
diff changeset
   453
  @{ML inc_by_five} using the reverse application is much clearer than writing
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   454
*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   455
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   456
ML{*fun inc_by_five x = fst ((fn x => (x, x)) (x + 1)) + 4*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   457
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   458
text {* or *}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   459
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   460
ML{*fun inc_by_five x = 
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 323
diff changeset
   461
  ((fn x => x + 4) o fst o (fn x => (x, x)) o (fn x => x + 1)) x*}
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   462
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   463
text {* and typographically more economical than *}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   464
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   465
ML{*fun inc_by_five x =
257
ce0f60d0351e corrected index references and minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   466
let val y1 = x + 1
ce0f60d0351e corrected index references and minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   467
    val y2 = (y1, y1)
ce0f60d0351e corrected index references and minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   468
    val y3 = fst y2
ce0f60d0351e corrected index references and minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   469
    val y4 = y3 + 4
ce0f60d0351e corrected index references and minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   470
in y4 end*}
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   471
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   472
text {* 
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   473
  Another reason why the let-bindings in the code above are better to be
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   474
  avoided: it is more than easy to get the intermediate values wrong, not to 
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   475
  mention the nightmares the maintenance of this code causes!
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   476
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   477
  In Isabelle a ``real world'' example for a function written in 
178
fb8f22dd8ad0 adapted to latest Attrib.setup changes and more work on the simple induct chapter
Christian Urban <urbanc@in.tum.de>
parents: 177
diff changeset
   478
  the waterfall fashion might be the following code:
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 171
diff changeset
   479
*}
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   480
193
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   481
ML %linenosgray{*fun apply_fresh_args f ctxt =
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   482
    f |> fastype_of
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   483
      |> binder_types 
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   484
      |> map (pair "z")
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   485
      |> Variable.variant_frees ctxt [f]
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   486
      |> map Free  
257
ce0f60d0351e corrected index references and minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   487
      |> curry list_comb f *}
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   488
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 171
diff changeset
   489
text {*
266
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   490
  This function takes a term and a context as argument. If the term is of function
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   491
  type, then @{ML "apply_fresh_args"} returns the term with distinct variables 
343
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
   492
  applied to it. For example below three variables are applied to the term 
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   493
  @{term [show_types] "P::nat \<Rightarrow> int \<Rightarrow> unit \<Rightarrow> bool"}:
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   494
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   495
  @{ML_response_fake [display,gray]
266
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   496
"let
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   497
  val t = @{term \"P::nat \<Rightarrow> int \<Rightarrow> unit \<Rightarrow> bool\"}
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   498
  val ctxt = @{context}
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   499
in 
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   500
  apply_fresh_args t ctxt
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   501
   |> string_of_term ctxt
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   502
   |> tracing
266
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   503
end" 
183
8bb4eaa2ec92 a simplification suggested by Stefan and some polishing
Christian Urban <urbanc@in.tum.de>
parents: 178
diff changeset
   504
  "P z za zb"}
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 171
diff changeset
   505
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   506
  You can read off this behaviour from how @{ML apply_fresh_args} is coded: in
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   507
  Line 2, the function @{ML_ind fastype_of in Term} calculates the type of the
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   508
  term; @{ML_ind binder_types in Term} in the next line produces the list of
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   509
  argument types (in the case above the list @{text "[nat, int, unit]"}); Line
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   510
  4 pairs up each type with the string @{text "z"}; the function @{ML_ind
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   511
  variant_frees in Variable} generates for each @{text "z"} a unique name
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   512
  avoiding the given @{text f}; the list of name-type pairs is turned into a
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   513
  list of variable terms in Line 6, which in the last line is applied by the
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   514
  function @{ML_ind list_comb in Term} to the term. In this last step we have
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   515
  to use the function @{ML_ind curry in Library}, because @{ML list_comb}
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   516
  expects the function and the variables list as a pair. 
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   517
  
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   518
  Functions like @{ML apply_fresh_args} are often needed when constructing 
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   519
  terms with fresh variables. The
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   520
  infrastructure helps tremendously to avoid any name clashes. Consider for
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   521
  example:
252
f380b13b25a7 added an example
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   522
f380b13b25a7 added an example
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   523
   @{ML_response_fake [display,gray]
266
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   524
"let
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   525
  val t = @{term \"za::'a \<Rightarrow> 'b \<Rightarrow> 'c\"}
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   526
  val ctxt = @{context}
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   527
in
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   528
  apply_fresh_args t ctxt 
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   529
   |> string_of_term ctxt
301
2728e8daebc0 replaced "writeln" with "tracing"
Christian Urban <urbanc@in.tum.de>
parents: 299
diff changeset
   530
   |> tracing
266
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   531
end"
252
f380b13b25a7 added an example
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   532
  "za z zb"}
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 171
diff changeset
   533
  
266
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   534
  where the @{text "za"} is correctly avoided.
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   535
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   536
  The combinator @{ML_ind "#>" in Basics} is the reverse function composition. 
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   537
  It can be used to define the following function
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   538
*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   539
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   540
ML{*val inc_by_six = 
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   541
  (fn x => x + 1) #> 
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   542
  (fn x => x + 2) #> 
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   543
  (fn x => x + 3)*}
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   544
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   545
text {* 
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   546
  which is the function composed of first the increment-by-one function and then
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   547
  increment-by-two, followed by increment-by-three. Again, the reverse function 
375
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   548
  composition allows you to read the code top-down. This combinator is often used
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   549
  for setup function inside the \isacommand{setup}-command. These function have to be
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   550
  of type @{ML_type "theory -> theory"} in order to install, for example, some new 
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   551
  data inside the current theory. More than one such setup function can be composed 
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   552
  with @{ML "#>"}. For example
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   553
*}
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   554
375
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   555
setup %gray {* let
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   556
  val (ival1, setup_ival1) = Attrib.config_int "ival1" 1
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   557
  val (ival2, setup_ival2) = Attrib.config_int "ival2" 2
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   558
in
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   559
  setup_ival1 #>
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   560
  setup_ival2
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   561
end *}
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   562
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   563
text {*
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   564
  after this the configuration values @{text ival1} and @{text ival2} are known
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   565
  in the current theory and can be manipulated by the user (for more information 
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   566
  about configuration values see Section~\ref{sec:storing}, for more about setup 
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   567
  functions see Section~\ref{sec:theories}). 
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   568
  
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   569
  The remaining combinators we describe in this section add convenience for the
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   570
  ``waterfall method'' of writing functions. The combinator @{ML_ind tap in
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   571
  Basics} allows you to get hold of an intermediate result (to do some
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   572
  side-calculations for instance). The function
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   573
 *}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   574
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   575
ML %linenosgray{*fun inc_by_three x =
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   576
     x |> (fn x => x + 1)
240
d111f5988e49 replaced explode by Symbol.explode
Christian Urban <urbanc@in.tum.de>
parents: 239
diff changeset
   577
       |> tap (fn x => tracing (PolyML.makestring x))
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   578
       |> (fn x => x + 2)*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   579
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   580
text {* 
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   581
  increments the argument first by @{text "1"} and then by @{text "2"}. In the
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   582
  middle (Line 3), however, it uses @{ML tap} for printing the ``plus-one''
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   583
  intermediate result. The function @{ML tap} can only be used for
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   584
  side-calculations, because any value that is computed cannot be merged back
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   585
  into the ``main waterfall''. To do this, you can use the next combinator.
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   586
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   587
  The combinator @{ML_ind "`" in Basics} (a backtick) is similar to @{ML tap},
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   588
  but applies a function to the value and returns the result together with the
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   589
  value (as a pair). It is defined as
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   590
*}
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   591
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   592
ML{*fun `f = fn x => (f x, x)*}
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   593
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   594
text {*
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   595
  An example for this combinator is the function
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   596
*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   597
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   598
ML{*fun inc_as_pair x =
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   599
     x |> `(fn x => x + 1)
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   600
       |> (fn (x, y) => (x, y + 1))*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   601
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   602
text {*
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   603
  which takes @{text x} as argument, and then increments @{text x}, but also keeps
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   604
  @{text x}. The intermediate result is therefore the pair @{ML "(x + 1, x)"
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   605
  for x}. After that, the function increments the right-hand component of the
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   606
  pair. So finally the result will be @{ML "(x + 1, x + 1)" for x}.
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   607
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   608
  The combinators @{ML_ind "|>>" in Basics} and @{ML_ind "||>" in Basics} are
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   609
  defined for functions manipulating pairs. The first applies the function to
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   610
  the first component of the pair, defined as
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   611
*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   612
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   613
ML{*fun (x, y) |>> f = (f x, y)*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   614
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   615
text {*
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   616
  and the second combinator to the second component, defined as
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   617
*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   618
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   619
ML{*fun (x, y) ||> f = (x, f y)*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   620
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   621
text {*
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   622
  These two functions can, for example, be used to avoid explicit @{text "lets"} for
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   623
  intermediate values in functions that return pairs. As an example, suppose you
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   624
  want to separate a list of integers into two lists according to a
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   625
  treshold. If the treshold is @{ML "5"}, the list @{ML "[1,6,2,5,3,4]"}
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   626
  should be separated to @{ML "([1,2,3,4], [6,5])"}.  This function can be
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   627
  implemented as
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   628
*}
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   629
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   630
ML{*fun separate i [] = ([], [])
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   631
  | separate i (x::xs) =
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   632
      let 
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   633
        val (los, grs) = separate i xs
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   634
      in
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   635
        if i <= x then (los, x::grs) else (x::los, grs)
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   636
      end*}
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   637
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   638
text {*
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   639
  where the return value of the recursive call is bound explicitly to
309
Christian Urban <urbanc@in.tum.de>
parents: 308
diff changeset
   640
  the pair @{ML "(los, grs)" for los grs}. You can implement this function
Christian Urban <urbanc@in.tum.de>
parents: 308
diff changeset
   641
  more concisely as
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   642
*}
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   643
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   644
ML{*fun separate i [] = ([], [])
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   645
  | separate i (x::xs) =
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   646
      if i <= x 
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   647
      then separate i xs ||> cons x
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   648
      else separate i xs |>> cons x*}
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   649
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   650
text {*
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   651
  avoiding the explicit @{text "let"}. While in this example the gain in
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   652
  conciseness is only small, in more complicated situations the benefit of
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   653
  avoiding @{text "lets"} can be substantial.
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   654
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   655
  With the combinator @{ML_ind "|->" in Basics} you can re-combine the 
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   656
  elements from a pair. This combinator is defined as
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   657
*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   658
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   659
ML{*fun (x, y) |-> f = f x y*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   660
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   661
text {* 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   662
  and can be used to write the following roundabout version 
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   663
  of the @{text double} function: 
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   664
*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   665
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   666
ML{*fun double x =
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   667
      x |>  (fn x => (x, x))
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   668
        |-> (fn x => fn y => x + y)*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   669
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   670
text {* 
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   671
  The combinator @{ML_ind ||>> in Basics} plays a central rôle whenever your
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   672
  task is to update a theory and the update also produces a side-result (for
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   673
  example a theorem). Functions for such tasks return a pair whose second
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   674
  component is the theory and the fist component is the side-result. Using
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   675
  @{ML ||>>}, you can do conveniently the update and also
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   676
  accumulate the side-results. Consider the following simple function.
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   677
*}
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   678
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   679
ML %linenosgray{*fun acc_incs x =
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   680
    x |> (fn x => ("", x)) 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   681
      ||>> (fn x => (x, x + 1))
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   682
      ||>> (fn x => (x, x + 1))
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   683
      ||>> (fn x => (x, x + 1))*}
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   684
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   685
text {*
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   686
  The purpose of Line 2 is to just pair up the argument with a dummy value (since
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   687
  @{ML ||>>} operates on pairs). Each of the next three lines just increment 
280
a63ca3a9b44a spell checked
griff
parents: 279
diff changeset
   688
  the value by one, but also nest the intermediate results to the left. For example 
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   689
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   690
  @{ML_response [display,gray]
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   691
  "acc_incs 1"
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   692
  "((((\"\", 1), 2), 3), 4)"}
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   693
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   694
  You can continue this chain with:
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   695
  
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   696
  @{ML_response [display,gray]
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   697
  "acc_incs 1 ||>> (fn x => (x, x + 2))"
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   698
  "(((((\"\", 1), 2), 3), 4), 6)"}
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   699
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   700
  \footnote{\bf FIXME: maybe give a ``real world'' example for this combinator.}
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   701
*}
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   702
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   703
text {*
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   704
  Recall that @{ML "|>"} is the reverse function application. Recall also that
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   705
  the related reverse function composition is @{ML "#>"}. In fact all the
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   706
  combinators @{ML "|->"}, @{ML "|>>"} , @{ML "||>"} and @{ML "||>>"}
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   707
  described above have related combinators for function composition, namely
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   708
  @{ML_ind "#->" in Basics}, @{ML_ind "#>>" in Basics}, @{ML_ind "##>" in
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   709
  Basics} and @{ML_ind "##>>" in Basics}. Using @{ML "#->"}, for example, the
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   710
  function @{text double} can also be written as:
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   711
*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   712
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   713
ML{*val double =
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   714
            (fn x => (x, x))
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   715
        #-> (fn x => fn y => x + y)*}
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   716
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   717
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   718
text {* 
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   719
  When using combinators for writing functions in waterfall fashion, it is
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   720
  sometimes necessary to do some ``plumbing'' in order to fit functions
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   721
  together. We have already seen such plumbing in the function @{ML
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   722
  apply_fresh_args}, where @{ML curry} is needed for making the function @{ML
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   723
  list_comb}, which works over pairs to fit with the combinator @{ML "|>"}. Such 
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   724
  plumbing is also needed in situations where a function operate over lists, 
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   725
  but one calculates only with a single element. An example is the function 
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   726
  @{ML_ind check_terms in Syntax}, whose purpose is to simultaneously type-check 
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   727
  a list of terms. Consider the code:
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   728
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   729
  @{ML_response_fake [display, gray]
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   730
  "let
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   731
  val ctxt = @{context}
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   732
in
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 323
diff changeset
   733
  map (Syntax.parse_term ctxt) [\"m + n\", \"m * n\", \"m - (n::nat)\"] 
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   734
  |> Syntax.check_terms ctxt
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   735
  |> string_of_terms ctxt
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   736
  |> tracing
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   737
end"
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 323
diff changeset
   738
  "m + n, m * n, m - n"}
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   739
*}
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   740
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   741
text {*
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   742
  In this example we obtain three terms (using the function 
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   743
  @{ML_ind parse_term in Syntax}) whose variables @{text m} and @{text n} 
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   744
  are of type @{typ "nat"}. If you have only a single term, then @{ML
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 323
diff changeset
   745
  check_terms in Syntax} needs plumbing. This can be done with the function
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   746
  @{ML_ind singleton in Library}.\footnote{There is already a function @{ML check_term in
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   747
  Syntax} in the file @{ML_file "Pure/Syntax/syntax.ML"} that is implemented 
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   748
  in terms of @{ML singleton} and @{ML check_terms in Syntax}.} For example
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   749
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   750
  @{ML_response_fake [display, gray, linenos]
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   751
  "let 
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   752
  val ctxt = @{context}
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   753
in
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   754
  Syntax.parse_term ctxt \"m - (n::nat)\" 
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   755
  |> singleton (Syntax.check_terms ctxt)
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   756
  |> string_of_term ctxt
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   757
  |> tracing
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   758
end"
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   759
  "m - n"}
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   760
   
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   761
  where in Line 5, the function operating over lists fits with the
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   762
  single term generated in Line 4.
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   763
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   764
  \begin{readmore}
196
840b49bfb1cf fixed `str_of_thms' output in example + small changes
griff
parents: 192
diff changeset
   765
  The most frequently used combinators are defined in the files @{ML_file
840b49bfb1cf fixed `str_of_thms' output in example + small changes
griff
parents: 192
diff changeset
   766
  "Pure/library.ML"}
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   767
  and @{ML_file "Pure/General/basics.ML"}. Also \isccite{sec:ML-linear-trans} 
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   768
  contains further information about combinators.
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   769
  \end{readmore}
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   770
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   771
  \footnote{\bf FIXME: find a good exercise for combinators}
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   772
  \footnote{\bf FIXME: say something about calling conventions} 
15
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents: 14
diff changeset
   773
*}
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents: 14
diff changeset
   774
10
df09e49b19bf many changes in the FirstSteps section
Christian Urban <urbanc@in.tum.de>
parents: 6
diff changeset
   775
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
   776
section {* ML-Antiquotations *}
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   777
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   778
text {*
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   779
  Recall from Section \ref{sec:include} that code in Isabelle is always
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   780
  embedded in a theory.  The main advantage of this is that the code can
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   781
  contain references to entities defined on the logical level of Isabelle. By
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   782
  this we mean references to definitions, theorems, terms and so on. These
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   783
  reference are realised in Isabelle with ML-antiquotations, often just called
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   784
  antiquotations.\footnote{Note that there are two kinds of antiquotations in
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   785
  Isabelle, which have very different purposes and infrastructures. The first
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   786
  kind, described in this section, are \emph{\index*{ML-antiquotation}}. They
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   787
  are used to refer to entities (like terms, types etc) from Isabelle's logic
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   788
  layer inside ML-code. The other kind of antiquotations are
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   789
  \emph{document}\index{document antiquotation} antiquotations. They are used
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   790
  only in the text parts of Isabelle and their purpose is to print logical
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   791
  entities inside \LaTeX-documents. Document antiquotations are part of the
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   792
  user level and therefore we are not interested in them in this Tutorial,
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   793
  except in Appendix \ref{rec:docantiquotations} where we show how to
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   794
  implement your own document antiquotations.}  Syntactically antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   795
  are indicated by the @{ML_text @}-sign followed by text wrapped in @{text
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   796
  "{\<dots>}"}.  For example, one can print out the name of the current theory with
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   797
  the code
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 34
diff changeset
   798
  
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 71
diff changeset
   799
  @{ML_response [display,gray] "Context.theory_name @{theory}" "\"FirstSteps\""}
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 34
diff changeset
   800
 
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
   801
  where @{text "@{theory}"} is an antiquotation that is substituted with the
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   802
  current theory (remember that we assumed we are inside the theory 
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 87
diff changeset
   803
  @{text FirstSteps}). The name of this theory can be extracted using
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   804
  the function @{ML_ind theory_name in Context}. 
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
   805
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 87
diff changeset
   806
  Note, however, that antiquotations are statically linked, that is their value is
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   807
  determined at ``compile-time'', not at ``run-time''. For example the function
43
02f76f1b6e7b added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   808
*}
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
   809
69
19106a9975c1 highligted the background of ML-code
Christian Urban <urbanc@in.tum.de>
parents: 68
diff changeset
   810
ML{*fun not_current_thyname () = Context.theory_name @{theory} *}
43
02f76f1b6e7b added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   811
02f76f1b6e7b added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   812
text {*
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   813
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 87
diff changeset
   814
  does \emph{not} return the name of the current theory, if it is run in a 
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
   815
  different theory. Instead, the code above defines the constant function 
58
f3794c231898 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 54
diff changeset
   816
  that always returns the string @{text [quotes] "FirstSteps"}, no matter where the
43
02f76f1b6e7b added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   817
  function is called. Operationally speaking,  the antiquotation @{text "@{theory}"} is 
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
   818
  \emph{not} replaced with code that will look up the current theory in 
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
   819
  some data structure and return it. Instead, it is literally
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
   820
  replaced with the value representing the theory name.
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   821
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   822
  In a similar way you can use antiquotations to refer to proved theorems: 
133
3e94ccc0f31e polishing and start of the section about attributes
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   823
  @{text "@{thm \<dots>}"} for a single theorem
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 34
diff changeset
   824
72
7b8c4fe235aa added an antiquotation option [gray] for gray boxes around displays
Christian Urban <urbanc@in.tum.de>
parents: 71
diff changeset
   825
  @{ML_response_fake [display,gray] "@{thm allI}" "(\<And>x. ?P x) \<Longrightarrow> \<forall>x. ?P x"}
75
Christian Urban <urbanc@in.tum.de>
parents: 73
diff changeset
   826
133
3e94ccc0f31e polishing and start of the section about attributes
Christian Urban <urbanc@in.tum.de>
parents: 132
diff changeset
   827
  and @{text "@{thms \<dots>}"} for more than one
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   828
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   829
@{ML_response_fake [display,gray] "@{thms conj_ac}" 
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   830
"(?P \<and> ?Q) = (?Q \<and> ?P)
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   831
(?P \<and> ?Q \<and> ?R) = (?Q \<and> ?P \<and> ?R)
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   832
((?P \<and> ?Q) \<and> ?R) = (?P \<and> ?Q \<and> ?R)"}  
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   833
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   834
  The point of these antiquotations is that referring to theorems in this way
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   835
  makes your code independent from what theorems the user might have stored
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   836
  under this name (this becomes especially important when you deal with
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   837
  theorem lists; see Section \ref{sec:storing}).
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   838
375
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   839
  It is also possible to prove lemmas with the antiquotation @{text "@{lemma \<dots> by \<dots>}"}
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   840
  whose first argument is a statement (possibly many of them separated by @{text "and"},
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   841
  and the second is a proof. For example
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   842
*}
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   843
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   844
ML{*val foo_thm = @{lemma "True" and "True" by simp simp} *}
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   845
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   846
text {*
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   847
  which can be printed out as follows
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   848
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   849
  @{ML_response_fake [gray,display]
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   850
"foo_thm |> string_of_thms @{context}
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   851
         |> tracing"
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   852
  "True, True"}
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   853
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   854
  You can also refer to the current simpset via an antiquotation. To illustrate 
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   855
  this we implement the function that extracts the theorem names stored in a 
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   856
  simpset.
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
   857
*}
75
Christian Urban <urbanc@in.tum.de>
parents: 73
diff changeset
   858
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 138
diff changeset
   859
ML{*fun get_thm_names_from_ss simpset =
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
   860
let
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
   861
  val {simps,...} = MetaSimplifier.dest_ss simpset
70
bbb2d5f1d58d deleted the fixme about simpsets
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   862
in
163
2319cff107f0 removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
Christian Urban <urbanc@in.tum.de>
parents: 162
diff changeset
   863
  map #1 simps
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
   864
end*}
54
1783211b3494 tuned; added document antiquotation ML_response_fake_both
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   865
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
   866
text {*
339
c588e8422737 used a better implementation of \index in Latex; added more to the theorem section
Christian Urban <urbanc@in.tum.de>
parents: 330
diff changeset
   867
  The function @{ML_ind dest_ss in MetaSimplifier} returns a record containing all
193
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   868
  information stored in the simpset, but we are only interested in the names of the
250
ab9e09076462 some polishing; added together with Jasmin more examples to the pretty printing section
Christian Urban <urbanc@in.tum.de>
parents: 249
diff changeset
   869
  simp-rules. Now you can feed in the current simpset into this function. 
193
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   870
  The current simpset can be referred to using the antiquotation @{text "@{simpset}"}.
81
8fda6b452f28 polished
Christian Urban <urbanc@in.tum.de>
parents: 78
diff changeset
   871
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
   872
  @{ML_response_fake [display,gray] 
149
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 138
diff changeset
   873
  "get_thm_names_from_ss @{simpset}" 
253ea99c1441 polished
Christian Urban <urbanc@in.tum.de>
parents: 138
diff changeset
   874
  "[\"Nat.of_nat_eq_id\", \"Int.of_int_eq_id\", \"Nat.One_nat_def\", \<dots>]"}
10
df09e49b19bf many changes in the FirstSteps section
Christian Urban <urbanc@in.tum.de>
parents: 6
diff changeset
   875
196
840b49bfb1cf fixed `str_of_thms' output in example + small changes
griff
parents: 192
diff changeset
   876
  Again, this way of referencing simpsets makes you independent from additions
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   877
  of lemmas to the simpset by the user, which can potentially cause loops in your
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   878
  code.
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   879
251
cccb25ee1309 a few additions
Christian Urban <urbanc@in.tum.de>
parents: 250
diff changeset
   880
  On the ML-level of Isabelle, you often have to work with qualified names.
315
de49d5780f57 simplified a bit the index generation
Christian Urban <urbanc@in.tum.de>
parents: 314
diff changeset
   881
  These are strings with some additional information, such as positional
de49d5780f57 simplified a bit the index generation
Christian Urban <urbanc@in.tum.de>
parents: 314
diff changeset
   882
  information and qualifiers. Such qualified names can be generated with the
de49d5780f57 simplified a bit the index generation
Christian Urban <urbanc@in.tum.de>
parents: 314
diff changeset
   883
  antiquotation @{text "@{binding \<dots>}"}. For example
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   884
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   885
  @{ML_response [display,gray]
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   886
  "@{binding \"name\"}"
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   887
  "name"}
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   888
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   889
  An example where a qualified name is needed is the function 
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   890
  @{ML_ind define in LocalTheory}.  This function is used below to define 
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   891
  the constant @{term "TrueConj"} as the conjunction @{term "True \<and> True"}.
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   892
*}
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   893
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   894
local_setup %gray {* 
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   895
  LocalTheory.define Thm.internalK 
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   896
    ((@{binding "TrueConj"}, NoSyn), 
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   897
      (Attrib.empty_binding, @{term "True \<and> True"})) #> snd *}
192
2fff636e1fa0 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 190
diff changeset
   898
193
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   899
text {* 
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   900
  Now querying the definition you obtain:
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   901
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   902
  \begin{isabelle}
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   903
  \isacommand{thm}~@{text "TrueConj_def"}\\
225
7859fc59249a section for further material about simple inductive
Christian Urban <urbanc@in.tum.de>
parents: 215
diff changeset
   904
  @{text "> "}~@{thm TrueConj_def}
193
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   905
  \end{isabelle}
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   906
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   907
  \begin{readmore}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   908
  The basic operations on bindings are implemented in 
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   909
  @{ML_file "Pure/General/binding.ML"}.
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   910
  \end{readmore}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   911
351
f118240ab44a updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 350
diff changeset
   912
  \footnote{\bf FIXME give a better example why bindings are important}
f118240ab44a updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 350
diff changeset
   913
  \footnote{\bf FIXME give a pointer to \isacommand{local\_setup}; if not, then explain
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   914
  why @{ML snd} is needed.}
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   915
  \footnote{\bf FIXME: There should probably a separate section on binding, long-names
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
   916
  and sign.}
193
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   917
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   918
  It is also possible to define your own antiquotations. But you should
315
de49d5780f57 simplified a bit the index generation
Christian Urban <urbanc@in.tum.de>
parents: 314
diff changeset
   919
  exercise care when introducing new ones, as they can also make your code
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   920
  also difficult to read. In the next chapter we describe how to construct
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   921
  terms with the (build in) antiquotation @{text "@{term \<dots>}"}.  A restriction
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   922
  of this antiquotation is that it does not allow you to use schematic
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   923
  variables in terms. If you want to have an antiquotation that does not have
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
   924
  this restriction, you can implement your own using the function @{ML_ind
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   925
  inline in ML_Antiquote} from the structure @{ML_struct ML_Antiquote}. The code
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   926
  for the antiquotation @{text "term_pat"} is as follows.
43
02f76f1b6e7b added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   927
*}
02f76f1b6e7b added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
Christian Urban <urbanc@in.tum.de>
parents: 42
diff changeset
   928
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   929
ML %linenosgray{*let
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   930
  val parser = Args.context -- Scan.lift Args.name_source
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   931
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   932
  fun term_pat (ctxt, str) =
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   933
     str |> ProofContext.read_term_pattern ctxt
264
311830b43f8f minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   934
         |> ML_Syntax.print_term
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   935
         |> ML_Syntax.atomic
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   936
in
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   937
  ML_Antiquote.inline "term_pat" (parser >> term_pat)
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   938
end*}
263
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   939
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   940
text {*
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   941
  The parser in Line 2 provides us with a context and a string; this string is
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 323
diff changeset
   942
  transformed into a term using the function @{ML_ind read_term_pattern in
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   943
  ProofContext} (Line 5); the next two lines transform the term into a string
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   944
  so that the ML-system can understand it. (All these functions will be explained
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   945
  in more detail in later sections.) An example for this antiquotation is:
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   946
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   947
  @{ML_response_fake [display,gray]
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   948
  "@{term_pat \"Suc (?x::nat)\"}"
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   949
  "Const (\"Suc\", \"nat \<Rightarrow> nat\") $ Var ((\"x\", 0), \"nat\")"}
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   950
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
   951
  which shows the internal representation of the term @{text "Suc ?x"}.
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   952
263
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   953
  \begin{readmore}
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   954
  The file @{ML_file "Pure/ML/ml_antiquote.ML"} contains the the definitions
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
   955
  for most antiquotations. Most of the basic operations on ML-syntax are implemented 
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
   956
  in @{ML_file "Pure/ML/ml_syntax.ML"}.
263
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   957
  \end{readmore}
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
   958
*}
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
   959
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
   960
section {* Storing Data in Isabelle\label{sec:storing} *}
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   961
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
   962
text {*
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   963
  Isabelle provides mechanisms for storing (and retrieving) arbitrary
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   964
  data. Before we delve into the details, let us digress a bit. Conventional
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   965
  wisdom has it that the type-system of ML ensures that  an
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   966
  @{ML_type "'a list"}, say, can only hold elements of the same type, namely
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   967
  @{ML_type "'a"}. Despite this wisdom, however, it is possible to implement a
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   968
  universal type in ML, although by some arguably accidental features of ML.
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   969
  This universal type can be used to store data of different type into a single list.
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   970
  In fact, it allows one to inject and to project data of \emph{arbitrary} type. This is
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   971
  in contrast to datatypes, which only allow injection and projection of data for
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   972
  some \emph{fixed} collection of types. In light of the conventional wisdom cited
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   973
  above it is important to keep in mind that the universal type does not
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   974
  destroy type-safety of ML: storing and accessing the data can only be done
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   975
  in a type-safe manner.
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
   976
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   977
  \begin{readmore}
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   978
  In Isabelle the universal type is implemented as the type @{ML_type
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   979
  Universal.universal} in the file
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   980
  @{ML_file "Pure/ML-Systems/universal.ML"}.
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   981
  \end{readmore}
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   982
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   983
  We will show the usage of the universal type by storing an integer and
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   984
  a boolean into a single list. Let us first define injection and projection 
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   985
  functions for booleans and integers into and from the type @{ML_type Universal.universal}.
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
   986
*}
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
   987
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   988
ML{*local
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   989
  val fn_int  = Universal.tag () : int  Universal.tag
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   990
  val fn_bool = Universal.tag () : bool Universal.tag
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   991
in
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   992
  val inject_int   = Universal.tagInject fn_int;
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   993
  val inject_bool  = Universal.tagInject fn_bool;
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   994
  val project_int  = Universal.tagProject fn_int;
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   995
  val project_bool = Universal.tagProject fn_bool
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   996
end*}
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   997
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   998
text {*
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   999
  Using the injection functions, we can inject the integer @{ML_text "13"} 
330
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1000
  and the boolean value @{ML_text "true"} into @{ML_type Universal.universal}, and
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1001
  then store them in a @{ML_type "Universal.universal list"} as follows:
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
  1002
*}
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
  1003
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1004
ML{*val foo_list = 
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1005
let
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1006
  val thirteen  = inject_int 13
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1007
  val truth_val = inject_bool true
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1008
in
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1009
  [thirteen, truth_val]
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1010
end*}
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1011
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1012
text {*
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1013
  The data can be retrieved with the projection functions defined above.
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1014
  
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1015
  @{ML_response_fake [display, gray]
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1016
"project_int (nth foo_list 0); 
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1017
project_bool (nth foo_list 1)" 
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1018
"13
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1019
true"}
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1020
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1021
  Notice that we access the integer as an integer and the boolean as
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1022
  a boolean. If we attempt to access the integer as a boolean, then we get 
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1023
  a runtime error. 
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1024
  
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1025
  @{ML_response_fake [display, gray]
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1026
"project_bool (nth foo_list 0)"  
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1027
"*** Exception- Match raised"}
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1028
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1029
  This runtime error is the reason why ML is still type-sound despite
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1030
  containing a universal type.
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1031
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1032
  Now, Isabelle heavily uses this mechanism for storing all sorts of
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1033
  data: theorem lists, simpsets, facts etc.  Roughly speaking, there are two
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1034
  places where data can be stored in Isabelle: in \emph{theories} and in \emph{proof
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1035
  contexts}. Data such as simpsets are ``global'' and therefore need to be stored
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1036
  in a theory (simpsets need to be maintained across proofs and even across
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1037
  theories).  On the other hand, data such as facts change inside a proof and
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1038
  are only relevant to the proof at hand. Therefore such data needs to be 
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1039
  maintained inside a proof context, which represents ``local'' data.
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1040
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1041
  For theories and proof contexts there are, respectively, the functors 
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1042
  @{ML_funct_ind TheoryDataFun} and @{ML_funct_ind ProofDataFun} that help
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1043
  with the data storage. Below we show how to implement a table in which you
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1044
  can store theorems and look them up according to a string key. The
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1045
  intention in this example is to be able to look up introduction rules for logical
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1046
  connectives. Such a table might be useful in an automatic proof procedure
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1047
  and therefore it makes sense to store this data inside a theory.  
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1048
  Consequently we use the functor @{ML_funct TheoryDataFun}.
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1049
  The code for the table is:
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
  1050
*}
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
  1051
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1052
ML %linenosgray{*structure Data = TheoryDataFun
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1053
  (type T = thm Symtab.table
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1054
   val empty = Symtab.empty
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1055
   val copy = I
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1056
   val extend = I
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1057
   fun merge _ = Symtab.merge (K true))*}
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1058
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1059
text {*
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1060
  In order to store data in a theory, we have to specify the type of the data
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1061
  (Line 2). In this case we specify the type @{ML_type "thm Symtab.table"},
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1062
  which stands for a table in which @{ML_type string}s can be looked up
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1063
  producing an associated @{ML_type thm}. We also have to specify four
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1064
  functions to use this functor: namely how to initialise the data storage
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1065
  (Line 3), how to copy it (Line 4), how to extend it (Line 5) and how two
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1066
  tables should be merged (Line 6). These functions correspond roughly to the
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1067
  operations performed on theories and we just give some sensible 
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1068
  defaults.\footnote{\bf FIXME: Say more about the
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1069
  assumptions of these operations.} The result structure @{ML_text Data}
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1070
  contains functions for accessing the table (@{ML Data.get}) and for updating
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1071
  it (@{ML Data.map}). There are also two more functions (@{ML Data.init} and
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1072
  @{ML Data.put}), which however are not relevant here. Below we define two
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1073
  auxiliary functions, which help us with accessing the table.
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1074
*}
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1075
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1076
ML{*val lookup = Symtab.lookup o Data.get
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1077
fun update k v = Data.map (Symtab.update (k, v))*}
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1078
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1079
text {*
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1080
  Since we want to store introduction rules associated with their 
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1081
  logical connective, we can fill the table as follows.
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1082
*}
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1083
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1084
setup %gray {*
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1085
  update "op &"   @{thm conjI} #>
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1086
  update "op -->" @{thm impI}  #>
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1087
  update "All"    @{thm allI}
326
f76135c6c527 more work on the tutorial
Christian Urban <urbanc@in.tum.de>
parents: 325
diff changeset
  1088
*}
f76135c6c527 more work on the tutorial
Christian Urban <urbanc@in.tum.de>
parents: 325
diff changeset
  1089
f76135c6c527 more work on the tutorial
Christian Urban <urbanc@in.tum.de>
parents: 325
diff changeset
  1090
text {*
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1091
  The use of the command \isacommand{setup} makes sure the table in the 
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1092
  \emph{current} theory is updated (this is explained further in 
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1093
  section~\ref{sec:theories}). The lookup can now be performed as follows.
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1094
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1095
  @{ML_response_fake [display, gray]
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1096
"lookup @{theory} \"op &\""
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1097
"SOME \"\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q\""}
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1098
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1099
  An important point to note is that these tables (and data in general)
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1100
  need to be treated in a purely functional fashion. Although
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1101
  we can update the table as follows
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1102
*}
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1103
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1104
setup %gray {* update "op &" @{thm TrueI} *}
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1105
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1106
text {*
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1107
  and accordingly, @{ML lookup} now produces the introduction rule for @{term "True"}
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1108
  
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1109
@{ML_response_fake [display, gray]
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1110
"lookup @{theory} \"op &\""
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1111
"SOME \"True\""}
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1112
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1113
  there are no references involved. This is one of the most fundamental
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1114
  coding conventions for programming in Isabelle. References  
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1115
  interfere with the multithreaded execution model of Isabelle and also
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1116
  defeat its undo-mechanism. To see the latter, consider the 
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1117
  following data container where we maintain a reference to a list of 
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1118
  integers.
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1119
*}
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1120
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1121
ML{*structure WrongRefData = TheoryDataFun
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1122
  (type T = (int list) Unsynchronized.ref
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1123
   val empty = Unsynchronized.ref []
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1124
   val copy = I
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1125
   val extend = I
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1126
   fun merge _ = fst)*}
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1127
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1128
text {*
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1129
  We initialise the reference with the empty list. Consequently a first 
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1130
  lookup produces @{ML "ref []" in Unsynchronized}.
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1131
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1132
  @{ML_response_fake [display,gray]
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1133
  "WrongRefData.get @{theory}"
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1134
  "ref []"}
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1135
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1136
  For updating the reference we use the following function
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1137
*}
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1138
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1139
ML{*fun ref_update n = WrongRefData.map 
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1140
      (fn r => let val _ = r := n::(!r) in r end)*}
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1141
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1142
text {*
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1143
  which takes an integer and adds it to the content of the reference.
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1144
  As before, we update the reference with the command 
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1145
  \isacommand{setup}. 
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1146
*}
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1147
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1148
setup %gray {* ref_update 1 *}
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1149
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1150
text {*
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1151
  A lookup in the current theory gives then the expected list 
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1152
  @{ML "ref [1]" in Unsynchronized}.
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1153
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1154
  @{ML_response_fake [display,gray]
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1155
  "WrongRefData.get @{theory}"
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1156
  "ref [1]"}
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1157
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1158
  So far everything is as expected. But, the trouble starts if we attempt to
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1159
  backtrack to the ``point'' before the \isacommand{setup}-command. There, we
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1160
  would expect that the list is empty again. But since it is stored in a
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1161
  reference, Isabelle has no control over it. So it is not empty, but still
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1162
  @{ML "ref [1]" in Unsynchronized}. Adding to the trouble, if we execute the
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1163
  \isacommand{setup}-command again, we do not obtain @{ML "ref [1]" in
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1164
  Unsynchronized}, but
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1165
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1166
  @{ML_response_fake [display,gray]
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1167
  "WrongRefData.get @{theory}"
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1168
  "ref [1, 1]"}
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1169
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1170
  Now imagine how often you go backwards and forwards in your proof scripts. 
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1171
  By using references in Isabelle code, you are bound to cause all
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1172
  hell to break loose. Therefore observe the coding convention: 
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1173
  Do not use references for storing data!
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1174
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1175
  \begin{readmore}
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1176
  The functors for data storage are defined in @{ML_file "Pure/context.ML"}.
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1177
  Isabelle contains implementations of several container data structures,
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1178
  including association lists in @{ML_file "Pure/General/alist.ML"},
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1179
  directed graphs in @{ML_file "Pure/General/graph.ML"}, and 
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1180
  tables and symtables in @{ML_file "Pure/General/table.ML"}.
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1181
  \end{readmore}
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1182
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1183
  Storing data in a proof context is done in a similar fashion. As mentioned
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1184
  before, the corresponding functor is @{ML_funct_ind ProofDataFun}. With the
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1185
  following code we can store a list of terms in a proof context.
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1186
*}
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1187
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1188
ML{*structure Data = ProofDataFun
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1189
  (type T = term list
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1190
   fun init _ = [])*}
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1191
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1192
text {*
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1193
  The function we have to specify has to produce a list once a context 
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1194
  is initialised (possibly taking the theory into account from which the 
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1195
  context is derived). We choose here to just return the empty list. Next 
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1196
  we define two auxiliary functions for updating the list with a given
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1197
  term and printing the list. 
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1198
*}
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1199
330
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1200
ML{*fun update trm = Data.map (fn trms => trm::trms)
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1201
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1202
fun print ctxt =
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1203
  case (Data.get ctxt) of
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1204
    [] => tracing "Empty!"
330
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1205
  | trms => tracing (string_of_terms ctxt trms)*}
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1206
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1207
text {*
330
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1208
  Next we start with the context generated by the antiquotation 
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1209
  @{text "@{context}"} and update it in various ways. 
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1210
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1211
  @{ML_response_fake [display,gray]
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1212
"let
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1213
  val ctxt0 = @{context}
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1214
  val ctxt1 = ctxt0 |> update @{term \"False\"}
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1215
                    |> update @{term \"True \<and> True\"} 
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1216
  val ctxt2 = ctxt0 |> update @{term \"1::nat\"}
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1217
  val ctxt3 = ctxt2 |> update @{term \"2::nat\"}
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1218
in
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1219
  print ctxt0; 
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1220
  print ctxt1;
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1221
  print ctxt2;
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1222
  print ctxt3
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1223
end"
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1224
"Empty!
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1225
True \<and> True, False
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1226
1
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1227
2, 1"}
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1228
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1229
  Many functions in Isabelle manage and update data in a similar
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1230
  fashion. Consequently, such calculation with contexts occur frequently in
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1231
  Isabelle code, although the ``context flow'' is usually only linear.
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1232
  Note also that the calculation above has no effect on the underlying
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1233
  theory. Once we throw away the contexts, we have no access to their
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1234
  associated data. This is different to theories, where the command 
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1235
  \isacommand{setup} registers the data with the current and future 
330
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1236
  theories, and therefore one can access the data potentially 
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1237
  indefinitely.
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1238
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1239
  For convenience there is an abstract layer, namely the type @{ML_type Context.generic}, 
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1240
  for treating theories and proof contexts more uniformly. This type is defined as follows
330
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1241
*}
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1242
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1243
ML_val{*datatype generic = 
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1244
  Theory of theory 
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1245
| Proof of proof*}
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1246
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1247
text {*
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1248
  \footnote{\bf FIXME: say more about generic contexts.}
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1249
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1250
  There are two special instances of the data storage mechanism described 
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1251
  above. The first instance implements named theorem lists using the functor
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1252
  @{ML_funct_ind Named_Thms}. This is because storing theorems in a list 
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1253
  is such a common task.  To obtain a named theorem list, you just declare
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1254
*}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1255
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1256
ML{*structure FooRules = Named_Thms
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1257
  (val name = "foo" 
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1258
   val description = "Theorems for foo") *}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1259
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1260
text {*
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1261
  and set up the @{ML_struct FooRules} with the command
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1262
*}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1263
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1264
setup %gray {* FooRules.setup *}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1265
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1266
text {*
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1267
  This code declares a data container where the theorems are stored,
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1268
  an attribute @{text foo} (with the @{text add} and @{text del} options
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1269
  for adding and deleting theorems) and an internal ML-interface to retrieve and 
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1270
  modify the theorems.
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1271
  Furthermore, the theorems are made available on the user-level under the name 
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1272
  @{text foo}. For example you can declare three lemmas to be a member of the 
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1273
  theorem list @{text foo} by:
326
f76135c6c527 more work on the tutorial
Christian Urban <urbanc@in.tum.de>
parents: 325
diff changeset
  1274
*}
f76135c6c527 more work on the tutorial
Christian Urban <urbanc@in.tum.de>
parents: 325
diff changeset
  1275
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1276
lemma rule1[foo]: "A" sorry
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1277
lemma rule2[foo]: "B" sorry
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1278
lemma rule3[foo]: "C" sorry
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1279
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1280
text {* and undeclare the first one by: *}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1281
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1282
declare rule1[foo del]
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1283
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1284
text {* You can query the remaining ones with:
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1285
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1286
  \begin{isabelle}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1287
  \isacommand{thm}~@{text "foo"}\\
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1288
  @{text "> ?C"}\\
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1289
  @{text "> ?B"}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1290
  \end{isabelle}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1291
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1292
  On the ML-level, we can add theorems to the list with @{ML FooRules.add_thm}:
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1293
*} 
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1294
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1295
setup %gray {* Context.theory_map (FooRules.add_thm @{thm TrueI}) *}
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1296
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1297
text {*
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1298
  The rules in the list can be retrieved using the function 
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1299
  @{ML FooRules.get}:
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1300
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1301
  @{ML_response_fake [display,gray] 
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1302
  "FooRules.get @{context}" 
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1303
  "[\"True\", \"?C\",\"?B\"]"}
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1304
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1305
  Note that this function takes a proof context as argument. This might be 
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1306
  confusing, since the theorem list is stored as theory data. It becomes clear by knowing
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1307
  that the proof context contains the information about the current theory and so the function
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1308
  can access the theorem list in the theory via the context. 
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1309
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1310
  \begin{readmore}
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1311
  For more information about named theorem lists see 
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1312
  @{ML_file "Pure/Tools/named_thms.ML"}.
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1313
  \end{readmore}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1314
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1315
  The second special instance of the data storage mechanism are configuration
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1316
  values. They are used to enable users to configure tools without having to
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1317
  resort to the ML-level (and also to avoid references). Assume you want the
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1318
  user to control three values, say @{text bval} containing a boolean, @{text
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1319
  ival} containing an integer and @{text sval} containing a string. These
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1320
  values can be declared by
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1321
*}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1322
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1323
ML{*val (bval, setup_bval) = Attrib.config_bool "bval" false
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1324
val (ival, setup_ival) = Attrib.config_int "ival" 0
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1325
val (sval, setup_sval) = Attrib.config_string "sval" "some string" *}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1326
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1327
text {* 
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1328
  where each value needs to be given a default. To enable these values on the 
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1329
  user-level, they need to be set up with
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1330
*}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1331
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1332
setup %gray {* 
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1333
  setup_bval #> 
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1334
  setup_ival #>
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1335
  setup_sval
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1336
*}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1337
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1338
text {*
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1339
  The user can now manipulate the values from the user-level of Isabelle 
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1340
  with the command
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1341
*}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1342
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1343
declare [[bval = true, ival = 3]]
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1344
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1345
text {*
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1346
  On the ML-level these values can be retrieved using the 
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1347
  function @{ML_ind get in Config} from a proof context
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1348
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1349
  @{ML_response [display,gray] 
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1350
  "Config.get @{context} bval" 
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1351
  "true"}
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1352
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1353
  or directly from a theory using the function @{ML_ind get_thy in Config}
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1354
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1355
  @{ML_response [display,gray] 
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1356
  "Config.get_thy @{theory} bval" 
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1357
  "true"}
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1358
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1359
  It is also possible to manipulate the configuration values
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1360
  from the ML-level with the functions @{ML_ind put in Config}
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1361
  and @{ML_ind put_thy in Config}. For example
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1362
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1363
  @{ML_response [display,gray]
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1364
  "let
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1365
  val ctxt = @{context}
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1366
  val ctxt' = Config.put sval \"foo\" ctxt
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1367
  val ctxt'' = Config.put sval \"bar\" ctxt'
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1368
in
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1369
  (Config.get ctxt sval, 
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1370
   Config.get ctxt' sval, 
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1371
   Config.get ctxt'' sval)
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1372
end"
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1373
  "(\"some string\", \"foo\", \"bar\")"}
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1374
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1375
  \begin{readmore}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1376
  For more information about configuration values see 
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1377
  the files @{ML_file "Pure/Isar/attrib.ML"} and 
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1378
  @{ML_file "Pure/config.ML"}.
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1379
  \end{readmore}
343
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1380
*}
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1381
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1382
section {* Summary *}
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1383
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1384
text {*
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1385
  This chapter describes the combinators that are used in Isabelle, as well
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1386
  as a simple printing infrastructure for @{ML_type term}, @{ML_type cterm}
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1387
  and @{ML_type thm}. The section on ML-antiquotations shows how to refer 
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1388
  statically to entities from the logic level of Isabelle. Isabelle also
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1389
  contains mechanisms for storing arbitrary data in theory and proof 
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1390
  contexts.
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1391
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1392
  \begin{conventions}
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1393
  \begin{itemize}
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1394
  \item Print messages that belong together in a single string.
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1395
  \item Do not use references in Isabelle code.
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1396
  \end{itemize}
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1397
  \end{conventions}
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1398
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1399
*}
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1400
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1401
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
  1402
(**************************************************)
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
  1403
(* bak                                            *)
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
  1404
(**************************************************)
263
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
  1405
322
2b7c08d90e2e included some tests
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1406
(*
2b7c08d90e2e included some tests
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1407
section {* Do Not Try This At Home! *}
2b7c08d90e2e included some tests
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1408
2b7c08d90e2e included some tests
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1409
ML {* val my_thms = ref ([]: thm list) *}
2b7c08d90e2e included some tests
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1410
2b7c08d90e2e included some tests
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1411
attribute_setup my_thm_bad =
2b7c08d90e2e included some tests
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1412
  {* Scan.succeed (Thm.declaration_attribute (fn th => fn ctxt =>
2b7c08d90e2e included some tests
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1413
      (my_thms := th :: ! my_thms; ctxt))) *} "bad attribute"
2b7c08d90e2e included some tests
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1414
2b7c08d90e2e included some tests
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1415
declare sym [my_thm_bad]
2b7c08d90e2e included some tests
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1416
declare refl [my_thm_bad]
2b7c08d90e2e included some tests
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1417
2b7c08d90e2e included some tests
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1418
ML "!my_thms"
2b7c08d90e2e included some tests
Christian Urban <urbanc@in.tum.de>
parents: 318
diff changeset
  1419
*)
196
840b49bfb1cf fixed `str_of_thms' output in example + small changes
griff
parents: 192
diff changeset
  1420
end