ProgTutorial/First_Steps.thy
author Norbert Schirmer <norbert.schirmer@web.de>
Tue, 21 May 2019 14:37:39 +0200
changeset 572 438703674711
parent 569 f875a25aa72d
child 575 c3dbc04471a9
permissions -rw-r--r--
prefer more result checking in ML antiquotations
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
441
Christian Urban <urbanc@in.tum.de>
parents: 440
diff changeset
     1
theory First_Steps
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
557
77ea2de0ca62 updated for Isabelle 2014
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 554
diff changeset
     4
                                                  
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
     5
chapter \<open>First Steps\label{chp:firststeps}\<close>
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
     7
text \<open>
489
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
     8
  \begin{flushright}
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
     9
  {\em ``The most effective debugging tool is still careful thought,\\ 
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
    10
  coupled with judiciously placed print statements.''} \\[1ex]
Christian Urban <urbanc@in.tum.de>
parents: 484
diff changeset
    11
  Brian Kernighan, in {\em Unix for Beginners}, 1979
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
    12
  \end{flushright}
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
    13
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 323
diff changeset
    14
  \medskip
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
    15
  Isabelle programming is done in ML @{cite "isa-imp"}. Just like lemmas and proofs, ML-code for
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    16
  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
    17
  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
    18
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
    19
  \begin{quote}
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    20
  \begin{tabular}{@ {}l}
441
Christian Urban <urbanc@in.tum.de>
parents: 440
diff changeset
    21
  \isacommand{theory} First\_Steps\\
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    22
  \isacommand{imports} Main\\
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    23
  \isacommand{begin}\\
6
007e09485351 some tuning
Christian Urban <urbanc@in.tum.de>
parents: 5
diff changeset
    24
  \ldots
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    25
  \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
    26
  \end{quote}
157
76cdc8f562fc added more to the simplifier section
Christian Urban <urbanc@in.tum.de>
parents: 156
diff changeset
    27
324
4172c0743cf2 updated foobar_proof example
Christian Urban <urbanc@in.tum.de>
parents: 323
diff changeset
    28
  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
    29
  that will be given might need to be adapted if you work in a different logic.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
    30
\<close>
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
    31
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
    32
section \<open>Including ML-Code\label{sec:include}\<close>
14
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
    33
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
    34
text \<open>
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    35
  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
    36
  \isacommand{ML}-command. For example:
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    37
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    38
  \begin{isabelle}
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    39
  \begin{graybox}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
    40
  \isacommand{ML}~\<open>\<verbopen>\<close>\isanewline
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
    41
  \hspace{5mm}@{ML \<open>3 + 4\<close>}\isanewline
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
    42
  \<open>\<verbclose>\<close>\isanewline
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
    43
  \<open>> 7\<close>\smallskip
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    44
  \end{graybox}
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    45
  \end{isabelle}
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    46
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
    47
  If you work with the
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
    48
  Isabelle/jEdit, then you just have to hover the cursor over the code 
474
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
    49
  and you see the evaluated result in the ``Output'' window.
193
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
    50
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
    51
  As mentioned in the Introduction, we will drop the \isacommand{ML}~\<open>\<verbopen> \<dots> \<verbclose>\<close> scaffolding whenever we show code. The lines
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    52
  prefixed with @{text [quotes] ">"} are not part of the code, rather they
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    53
  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
    54
  the commands \isacommand{ML\_val} and \isacommand{ML\_prf} for including
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    55
  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
    56
  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
    57
  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
    58
254
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    59
  \begin{quote}
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    60
  \begin{isabelle}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
    61
  \isacommand{lemma}~\<open>test:\<close>\isanewline
254
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    62
  \isacommand{shows}~@{text [quotes] "True"}\isanewline
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
    63
  \isacommand{ML\_prf}~\<open>\<verbopen>\<close>~@{ML \<open>writeln "Trivial!"\<close>}~\<open>\<verbclose>\<close>\isanewline
254
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    64
  \isacommand{oops}
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    65
  \end{isabelle}
cb86bf5658e4 some slight polishing
Christian Urban <urbanc@in.tum.de>
parents: 253
diff changeset
    66
  \end{quote}
253
3cfd9a8a6de1 added comments about ML_prf and ML_val
Christian Urban <urbanc@in.tum.de>
parents: 252
diff changeset
    67
502
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
    68
  However, both commands will only play minor roles in this tutorial (we most of
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
    69
  the time make sure 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
    70
538
e9fd5eff62c1 removed "use" for "ML_file"
Christian Urban <urbanc@in.tum.de>
parents: 517
diff changeset
    71
 
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
    72
  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
    73
  to a separate ML-file. Such files can then be included somewhere inside a 
538
e9fd5eff62c1 removed "use" for "ML_file"
Christian Urban <urbanc@in.tum.de>
parents: 517
diff changeset
    74
  theory by using the command \isacommand{ML\_file}. For example
235
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    75
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
    76
  \begin{quote}
235
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    77
  \begin{tabular}{@ {}l}
441
Christian Urban <urbanc@in.tum.de>
parents: 440
diff changeset
    78
  \isacommand{theory} First\_Steps\\
235
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    79
  \isacommand{imports} Main\\
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    80
  \isacommand{begin}\\
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    81
  \ldots\\
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
    82
  \isacommand{ML\_file}~\<open>"file_to_be_included.ML"\<close>\\
235
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    83
  \ldots
dc955603d813 explained uses and use commands
Christian Urban <urbanc@in.tum.de>
parents: 234
diff changeset
    84
  \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
    85
  \end{quote}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
    86
\<close>
14
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
    87
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
    88
section \<open>Printing and Debugging\label{sec:printing}\<close>
14
1c17e99f6f66 added a paragraph about "uses" and started a paragraph about tracing
Christian Urban <urbanc@in.tum.de>
parents: 13
diff changeset
    89
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
    90
text \<open>
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
    91
  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
    92
  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
    93
  @{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
    94
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
    95
  @{ML_matchresult_fake [display,gray] \<open>writeln "any string"\<close> \<open>"any string"\<close>}
10
df09e49b19bf many changes in the FirstSteps section
Christian Urban <urbanc@in.tum.de>
parents: 6
diff changeset
    96
557
77ea2de0ca62 updated for Isabelle 2014
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 554
diff changeset
    97
  will print out @{text [quotes] "any string"}.  
474
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
    98
  This function expects a string as argument. If you develop under
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
    99
  PolyML, then there is a convenient, though again ``quick-and-dirty'', method
421
620a24bf954a added a section to the introduction; described @{make_string}
Christian Urban <urbanc@in.tum.de>
parents: 420
diff changeset
   100
  for converting values into strings, namely the antiquotation 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   101
  \<open>@{make_string}\<close>:
10
df09e49b19bf many changes in the FirstSteps section
Christian Urban <urbanc@in.tum.de>
parents: 6
diff changeset
   102
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
   103
  @{ML_matchresult_fake [display,gray] \<open>writeln (@{make_string} 1)\<close> \<open>"1"\<close>} 
12
2f1736cb8f26 various changes by Alex and Christian
Christian Urban <urbanc@in.tum.de>
parents: 11
diff changeset
   104
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   105
  However, \<open>@{make_string}\<close> only works if the type of what
466
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   106
  is converted is monomorphic and not a function. 
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   107
466
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   108
  You can print out error messages with the function @{ML_ind error in Library}; 
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   109
  for example:
75
Christian Urban <urbanc@in.tum.de>
parents: 73
diff changeset
   110
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   111
  @{ML_response [display,gray] 
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   112
  \<open>if 0 = 1 then true else (error "foo")\<close> 
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   113
\<open>foo\<close>}
75
Christian Urban <urbanc@in.tum.de>
parents: 73
diff changeset
   114
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   115
  This function raises the exception \<open>ERROR\<close>, which will then 
557
77ea2de0ca62 updated for Isabelle 2014
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 554
diff changeset
   116
  be displayed by the infrastructure indicating that it is an error by
77ea2de0ca62 updated for Isabelle 2014
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 554
diff changeset
   117
  painting the output red. Note that this exception is meant 
466
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   118
  for ``user-level'' error messages seen by the ``end-user''. 
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   119
  For messages where you want to indicate a genuine program error
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   120
  use the exception \<open>Fail\<close>. 
304
14173c0e8688 polished comment for error function
Christian Urban <urbanc@in.tum.de>
parents: 301
diff changeset
   121
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   122
  Most often you want to inspect contents of Isabelle's basic data structures,
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   123
  namely @{ML_type term}, @{ML_type typ}, @{ML_type cterm}, @{ML_type ctyp}
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   124
  and @{ML_type thm}. Isabelle provides elaborate pretty-printing functions,
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   125
  which we will explain in more detail in Section \ref{sec:pretty}. For now
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   126
  we just use the functions @{ML_ind writeln in Pretty}  from the structure
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
   127
  @{ML_structure Pretty} and @{ML_ind pretty_term in Syntax} from the structure
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
   128
  @{ML_structure Syntax}. For more convenience, we bind them to the toplevel.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   129
\<close>
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   130
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   131
ML %grayML\<open>val pretty_term = Syntax.pretty_term
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   132
val pwriteln = Pretty.writeln\<close>
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   133
(* We redfine pwriteln to return a value not just a side effect on the output in order to
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   134
use some checking of output with ML_response antiquotation. *)
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   135
ML %invisible\<open>val pretty_term = Syntax.pretty_term
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   136
val pwriteln = YXML.content_of o Pretty.string_of\<close>
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   137
text \<open>
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   138
  They can be used as follows
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   139
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   140
  @{ML_response [display,gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   141
  \<open>pwriteln (pretty_term @{context} @{term "1::nat"})\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   142
  \<open>"1"\<close>}
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   143
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   144
  If there is more than one term to be printed, you can use the 
446
4c32349b9875 added an example to be used for conversions later on
Christian Urban <urbanc@in.tum.de>
parents: 441
diff changeset
   145
  function @{ML_ind commas in Pretty} and @{ML_ind block in Pretty} 
4c32349b9875 added an example to be used for conversions later on
Christian Urban <urbanc@in.tum.de>
parents: 441
diff changeset
   146
  to separate them.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   147
\<close> 
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   148
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   149
ML %grayML\<open>fun pretty_terms ctxt trms =  
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   150
  Pretty.block (Pretty.commas (map (pretty_term ctxt) trms))\<close>
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   151
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   152
text \<open>
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   153
  To print out terms together with their typing information,
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   154
  set the configuration value
451
fc074e669f9f disabled foobar_prove; updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 450
diff changeset
   155
  @{ML_ind show_types in Syntax} to @{ML true}.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   156
\<close>
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   157
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   158
ML %grayML\<open>val show_types_ctxt = Config.put show_types true @{context}\<close>
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   159
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   160
text \<open>
451
fc074e669f9f disabled foobar_prove; updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 450
diff changeset
   161
  Now by using this context @{ML pretty_term} prints out
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   162
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   163
  @{ML_response [display, gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   164
  \<open>pwriteln (pretty_term show_types_ctxt @{term "(1::nat, x)"})\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   165
  \<open>(1::nat, x::'a)\<close>}
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   166
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   167
  where \<open>1\<close> and \<open>x\<close> are displayed with their inferred types.
546
d84867127c5d updated to Isabelle changes
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 544
diff changeset
   168
  Other configuration values that influence
502
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   169
  printing of terms include 
451
fc074e669f9f disabled foobar_prove; updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 450
diff changeset
   170
502
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   171
  \begin{itemize}
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   172
  \item @{ML_ind show_brackets in Syntax} 
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   173
  \item @{ML_ind show_sorts in Syntax}
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   174
  \item @{ML_ind eta_contract in Syntax}
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   175
  \end{itemize}
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   176
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   177
  A @{ML_type cterm} can be printed with the following function.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   178
\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   179
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   180
ML %grayML %grayML\<open>fun pretty_cterm ctxt ctrm =  
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   181
  pretty_term ctxt (Thm.term_of ctrm)\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   182
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   183
text \<open>
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   184
  Here the function @{ML_ind term_of in Thm} extracts the @{ML_type
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   185
  term} from a @{ML_type cterm}.  More than one @{ML_type cterm}s can be
446
4c32349b9875 added an example to be used for conversions later on
Christian Urban <urbanc@in.tum.de>
parents: 441
diff changeset
   186
  printed again with @{ML commas in Pretty}.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   187
\<close> 
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   188
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   189
ML %grayML\<open>fun pretty_cterms ctxt ctrms =  
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   190
  Pretty.block (Pretty.commas (map (pretty_cterm ctxt) ctrms))\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   191
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   192
text \<open>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   193
  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
   194
  into a @{ML_type term} using the function @{ML_ind prop_of in Thm}. 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   195
\<close>
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   196
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   197
ML %grayML\<open>fun pretty_thm ctxt thm =
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   198
  pretty_term ctxt (Thm.prop_of thm)\<close>
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   199
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   200
text \<open>
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   201
  Theorems include schematic variables, such as \<open>?P\<close>, 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   202
  \<open>?Q\<close> and so on. They are instantiated by Isabelle when a theorem is applied.
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   203
  For example, the theorem 
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   204
  @{thm [source] conjI} shown below can be used for any (typable) 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   205
  instantiation of \<open>?P\<close> and \<open>?Q\<close>. 
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   206
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   207
  @{ML_response [display, gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   208
  \<open>pwriteln (pretty_thm @{context} @{thm conjI})\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   209
  \<open>\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q\<close>}
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   210
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   211
  However, to improve the readability when printing theorems, we
466
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   212
  can switch off the question marks as follows:
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   213
\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   214
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   215
ML %grayML\<open>fun pretty_thm_no_vars ctxt thm =
466
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   216
let
26d2f91608ed a little polishing
Christian Urban <urbanc@in.tum.de>
parents: 462
diff changeset
   217
  val ctxt' = Config.put show_question_marks false 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
   218
in
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   219
  pretty_term ctxt' (Thm.prop_of thm)
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   220
end\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   221
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   222
text \<open>
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   223
  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
   224
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   225
  @{ML_response [display, gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   226
  \<open>pwriteln (pretty_thm_no_vars @{context} @{thm conjI})\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   227
  \<open>\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> P \<and> Q\<close>}
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   228
  
467
e11fe0de19a5 soem polishing
Christian Urban <urbanc@in.tum.de>
parents: 466
diff changeset
   229
  Again the functions @{ML commas} and @{ML block in Pretty} help 
e11fe0de19a5 soem polishing
Christian Urban <urbanc@in.tum.de>
parents: 466
diff changeset
   230
  with printing more than one theorem. 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   231
\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   232
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   233
ML %grayML\<open>fun pretty_thms ctxt thms =  
446
4c32349b9875 added an example to be used for conversions later on
Christian Urban <urbanc@in.tum.de>
parents: 441
diff changeset
   234
  Pretty.block (Pretty.commas (map (pretty_thm ctxt) thms))
190
ca0ac2e75f6d more one the simple-inductive chapter
Christian Urban <urbanc@in.tum.de>
parents: 189
diff changeset
   235
440
a0b280dd4bc7 partially moved from string_of_term to pretty_term
Christian Urban <urbanc@in.tum.de>
parents: 423
diff changeset
   236
fun pretty_thms_no_vars ctxt thms =  
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   237
  Pretty.block (Pretty.commas (map (pretty_thm_no_vars ctxt) thms))\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   238
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   239
text \<open>
476
0fb910f62bf9 minor things
Christian Urban <urbanc@in.tum.de>
parents: 475
diff changeset
   240
  Printing functions for @{ML_type typ} are
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   241
\<close>
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   242
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   243
ML %grayML\<open>fun pretty_typ ctxt ty = Syntax.pretty_typ ctxt ty
446
4c32349b9875 added an example to be used for conversions later on
Christian Urban <urbanc@in.tum.de>
parents: 441
diff changeset
   244
fun pretty_typs ctxt tys = 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   245
  Pretty.block (Pretty.commas (map (pretty_typ ctxt) tys))\<close>
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   246
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   247
text \<open>
476
0fb910f62bf9 minor things
Christian Urban <urbanc@in.tum.de>
parents: 475
diff changeset
   248
  respectively @{ML_type ctyp}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   249
\<close>
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   250
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   251
ML %grayML\<open>fun pretty_ctyp ctxt cty = pretty_typ ctxt (Thm.typ_of cty)
446
4c32349b9875 added an example to be used for conversions later on
Christian Urban <urbanc@in.tum.de>
parents: 441
diff changeset
   252
fun pretty_ctyps ctxt ctys = 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   253
  Pretty.block (Pretty.commas (map (pretty_ctyp ctxt) ctys))\<close>
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   254
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   255
text \<open>
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   256
  \begin{readmore}
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   257
  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
   258
  @{ML_type string}s are implemented in @{ML_file "Pure/Syntax/syntax.ML"}. 
467
e11fe0de19a5 soem polishing
Christian Urban <urbanc@in.tum.de>
parents: 466
diff changeset
   259
  The configuration values that change the printing information are declared in 
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   260
  @{ML_file "Pure/Syntax/printer.ML"}.
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   261
  \end{readmore}
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
   262
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   263
  Note that for printing out several ``parcels'' of information that belong
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   264
  together, like a warning message consisting of a term and its type, you
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   265
  should try to print these parcels together in a single string. Therefore do
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   266
  \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
   267
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
   268
@{ML_matchresult_fake [display,gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   269
\<open>pwriteln (Pretty.str "First half,"); 
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   270
pwriteln (Pretty.str "and second half.")\<close>
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   271
\<open>"First half,
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   272
and second half."\<close>}
306
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   273
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   274
  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
   275
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   276
@{ML_response [display,gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   277
\<open>pwriteln (Pretty.str ("First half," ^ "\\n" ^ "and second half."))\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   278
\<open>First half,
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   279
and second half.\<close>}
306
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   280
  
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   281
  To ease this kind of string manipulations, there are a number
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   282
  of library functions in Isabelle. For example,  the function 
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   283
  @{ML_ind cat_lines in Library} concatenates a list of strings 
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   284
  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
   285
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   286
  @{ML_response [display, gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   287
  \<open>pwriteln (Pretty.str (cat_lines ["foo", "bar"]))\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   288
  \<open>foo
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   289
bar\<close>}
306
fe732e890d87 tuned the section about printing several bits of inormation
Christian Urban <urbanc@in.tum.de>
parents: 305
diff changeset
   290
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   291
  Section \ref{sec:pretty} will explain the infrastructure that Isabelle 
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   292
  provides for more elaborate pretty printing. 
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   293
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   294
  \begin{readmore}
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   295
  Most of the basic string functions of Isabelle are defined in 
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   296
  @{ML_file "Pure/library.ML"}.
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   297
  \end{readmore}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   298
\<close>
305
2ac9dc1a95b4 added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
parents: 304
diff changeset
   299
2ac9dc1a95b4 added a comment for printing out information and tuned some examples accordingly
Christian Urban <urbanc@in.tum.de>
parents: 304
diff changeset
   300
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   301
section \<open>Combinators\label{sec:combinators}\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   302
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   303
text \<open>
413
95461cf6fd07 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 412
diff changeset
   304
  For beginners perhaps the most puzzling parts in the existing code of
95461cf6fd07 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 412
diff changeset
   305
  Isabelle are the combinators. At first they seem to greatly obstruct the
95461cf6fd07 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 412
diff changeset
   306
  comprehension of code, but after getting familiar with them and handled with
95461cf6fd07 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 412
diff changeset
   307
  care, they actually ease the understanding and also the programming.
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   308
373
28a49fe024c9 added structure index
Christian Urban <urbanc@in.tum.de>
parents: 372
diff changeset
   309
  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
   310
  identity function defined as
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   311
\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   312
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   313
ML %grayML\<open>fun I x = x\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   314
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   315
text \<open>
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   316
  Another simple combinator is @{ML_ind K in Library}, defined as 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   317
\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   318
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   319
ML %grayML\<open>fun K x = fn _ => x\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   320
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   321
text \<open>
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   322
  @{ML K} ``wraps'' a function around \<open>x\<close> that ignores its argument. As a 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   323
  result, @{ML K} defines a constant function always returning \<open>x\<close>.
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   324
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   325
  The next combinator is reverse application, @{ML_ind  "|>" in Basics}, defined as: 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   326
\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   327
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   328
ML %grayML\<open>fun x |> f = f x\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   329
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   330
text \<open>While just syntactic sugar for the usual function application,
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   331
  the purpose of this combinator is to implement functions in a  
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   332
  ``waterfall fashion''. Consider for example the function\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   333
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   334
ML %linenosgray\<open>fun inc_by_five x =
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   335
  x |> (fn x => x + 1)
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   336
    |> (fn x => (x, x))
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   337
    |> fst
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   338
    |> (fn x => x + 4)\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   339
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   340
text \<open>
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   341
  which increments its argument \<open>x\<close> by 5. It does this by first
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   342
  incrementing the argument by 1 (Line 2); then storing the result in a pair
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   343
  (Line 3); taking the first component of the pair (Line 4) and finally
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   344
  incrementing the first component by 4 (Line 5). This kind of cascading
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   345
  manipulations of values is quite common when dealing with theories. The
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   346
  reverse application allows you to read what happens in a top-down
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   347
  manner. This kind of coding should be familiar, if you have been exposed to
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   348
  Haskell's {\it do}-notation. Writing the function @{ML inc_by_five} using
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   349
  the reverse application is much clearer than writing
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   350
\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   351
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   352
ML %grayML\<open>fun inc_by_five x = fst ((fn x => (x, x)) (x + 1)) + 4\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   353
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   354
text \<open>or\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   355
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   356
ML %grayML\<open>fun inc_by_five x = 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   357
  ((fn x => x + 4) o fst o (fn x => (x, x)) o (fn x => x + 1)) x\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   358
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   359
text \<open>and typographically more economical than\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   360
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   361
ML %grayML\<open>fun inc_by_five x =
257
ce0f60d0351e corrected index references and minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   362
let val y1 = x + 1
ce0f60d0351e corrected index references and minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   363
    val y2 = (y1, y1)
ce0f60d0351e corrected index references and minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   364
    val y3 = fst y2
ce0f60d0351e corrected index references and minor polishing
Christian Urban <urbanc@in.tum.de>
parents: 256
diff changeset
   365
    val y4 = y3 + 4
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   366
in y4 end\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   367
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   368
text \<open>
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   369
  Another reasons to avoid the let-bindings in the code above:
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   370
  it is easy to get the intermediate values wrong and
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   371
  the resulting code is difficult to maintain.
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   372
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   373
  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
   374
  the waterfall fashion might be the following code:
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   375
\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   376
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   377
ML %linenosgray\<open>fun apply_fresh_args f ctxt =
193
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   378
    f |> fastype_of
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   379
      |> binder_types 
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   380
      |> map (pair "z")
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   381
      |> Variable.variant_frees ctxt [f]
ffd93dcc269d polishing to the theorem attributes section
Christian Urban <urbanc@in.tum.de>
parents: 192
diff changeset
   382
      |> map Free  
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   383
      |> curry list_comb f\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   384
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   385
text \<open>
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   386
  This function takes a term and a context as arguments. If the term is of function
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   387
  type, then @{ML \<open>apply_fresh_args\<close>} returns the term with distinct variables 
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   388
  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
   389
  @{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
   390
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   391
  @{ML_response [display,gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   392
\<open>let
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   393
  val trm = @{term "P::nat \<Rightarrow> int \<Rightarrow> unit \<Rightarrow> bool"}
266
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   394
  val ctxt = @{context}
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   395
in 
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   396
  apply_fresh_args trm ctxt
441
Christian Urban <urbanc@in.tum.de>
parents: 440
diff changeset
   397
   |> pretty_term ctxt
Christian Urban <urbanc@in.tum.de>
parents: 440
diff changeset
   398
   |> pwriteln
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   399
end\<close> 
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   400
  \<open>P z za zb\<close>}
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 171
diff changeset
   401
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   402
  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
   403
  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
   404
  term; @{ML_ind binder_types in Term} in the next line produces the list of
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   405
  argument types (in the case above the list \<open>[nat, int, unit]\<close>); Line
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   406
  4 pairs up each type with the string \<open>z\<close>; the function @{ML_ind
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   407
  variant_frees in Variable} generates for each \<open>z\<close> a unique name
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   408
  avoiding the given \<open>f\<close>; the list of name-type pairs is turned into a
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   409
  list of variable terms in Line 6, which in the last line is applied by the
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   410
  function @{ML_ind list_comb in Term} to the original term. In this last step we have
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   411
  to use the function @{ML_ind curry in Library}, because @{ML list_comb}
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   412
  expects the function and the variables list as a pair. 
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   413
  
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   414
  Functions like @{ML apply_fresh_args} are often needed when constructing
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   415
  terms involving fresh variables. For this the infrastructure helps
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   416
  tremendously to avoid any name clashes. Consider for example:
252
f380b13b25a7 added an example
Christian Urban <urbanc@in.tum.de>
parents: 251
diff changeset
   417
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   418
   @{ML_response [display,gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   419
\<open>let
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   420
  val trm = @{term "za::'a \<Rightarrow> 'b \<Rightarrow> 'c"}
266
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   421
  val ctxt = @{context}
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   422
in
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   423
  apply_fresh_args trm ctxt 
441
Christian Urban <urbanc@in.tum.de>
parents: 440
diff changeset
   424
   |> pretty_term ctxt
Christian Urban <urbanc@in.tum.de>
parents: 440
diff changeset
   425
   |> pwriteln
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   426
end\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   427
  \<open>za z zb\<close>}
177
4e2341f6599d polishing
Christian Urban <urbanc@in.tum.de>
parents: 171
diff changeset
   428
  
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   429
  where the \<open>za\<close> is correctly avoided.
266
Christian Urban <urbanc@in.tum.de>
parents: 265
diff changeset
   430
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   431
  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
   432
  It can be used to define the following function
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   433
\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   434
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   435
ML %grayML\<open>val inc_by_six = 
374
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   436
  (fn x => x + 1) #> 
Christian Urban <urbanc@in.tum.de>
parents: 373
diff changeset
   437
  (fn x => x + 2) #> 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   438
  (fn x => x + 3)\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   439
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   440
text \<open>
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   441
  which is the function composed of first the increment-by-one function and
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   442
  then increment-by-two, followed by increment-by-three. Again, the reverse
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   443
  function composition allows you to read the code top-down. This combinator
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   444
  is often used for setup functions inside the
478
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   445
  \isacommand{setup}- or \isacommand{local\_setup}-command. These functions 
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   446
  have to be of type @{ML_type "theory -> theory"}, respectively 
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   447
  @{ML_type "local_theory -> local_theory"}. More than one such setup function 
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   448
  can be composed with @{ML \<open>#>\<close>}. Consider for example the following code, 
478
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   449
  where we store the theorems @{thm [source] conjI}, @{thm [source] conjunct1} 
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   450
  and @{thm [source] conjunct2} under alternative names.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   451
\<close>  
478
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   452
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   453
local_setup %graylinenos \<open>
478
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   454
let  
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   455
  fun my_note name thm = Local_Theory.note ((name, []), [thm]) #> snd
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   456
in
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   457
  my_note @{binding "foo_conjI"} @{thm conjI} #>
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   458
  my_note @{binding "bar_conjunct1"} @{thm conjunct1} #>
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   459
  my_note @{binding "bar_conjunct2"} @{thm conjunct2}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   460
end\<close>
478
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   461
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   462
text \<open>
478
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   463
  The function @{ML_text "my_note"} in line 3 is just a wrapper for the function 
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
   464
  @{ML_ind note in Local_Theory} in the structure @{ML_structure Local_Theory}; 
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   465
  it stores a theorem under a name. 
482
Christian Urban <urbanc@in.tum.de>
parents: 481
diff changeset
   466
  In lines 5 to 6 we call this function to give alternative names for the three
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   467
  theorems. The point of @{ML \<open>#>\<close>} is that you can sequence such function calls. 
478
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   468
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   469
  The remaining combinators we describe in this section add convenience for
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   470
  the ``waterfall method'' of writing functions. The combinator @{ML_ind tap
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   471
  in Basics} allows you to get hold of an intermediate result (to do some
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   472
  side-calculations or print out an intermediate result, for instance). The
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   473
  function
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   474
\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   475
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   476
ML %linenosgray\<open>fun inc_by_three x =
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   477
     x |> (fn x => x + 1)
474
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
   478
       |> tap (fn x => pwriteln (Pretty.str (@{make_string} x)))
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   479
       |> (fn x => x + 2)\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   480
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   481
text \<open>
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   482
  increments the argument first by \<open>1\<close> and then by \<open>2\<close>. In the
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   483
  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
   484
  intermediate result. The function @{ML tap} can only be used for
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   485
  side-calculations, because any value that is computed cannot be merged back
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   486
  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
   487
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   488
  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
   489
  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
   490
  value (as a pair). It is defined as
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   491
\<close>
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   492
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   493
ML %grayML\<open>fun `f = fn x => (f x, x)\<close>
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   494
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   495
text \<open>
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   496
  An example for this combinator is the function
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   497
\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   498
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   499
ML %grayML\<open>fun inc_as_pair x =
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   500
     x |> `(fn x => x + 1)
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   501
       |> (fn (x, y) => (x, y + 1))\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   502
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   503
text \<open>
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   504
  which takes \<open>x\<close> as argument, and then increments \<open>x\<close>, but also keeps
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   505
  \<open>x\<close>. The intermediate result is therefore the pair @{ML \<open>(x + 1, x)\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   506
  for x}. After that, the function increments the right-hand component of the
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   507
  pair. So finally the result will be @{ML \<open>(x + 1, x + 1)\<close> for x}.
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   508
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   509
  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
   510
  defined for functions manipulating pairs. The first applies the function to
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   511
  the first component of the pair, defined as
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   512
\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   513
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   514
ML %grayML\<open>fun (x, y) |>> f = (f x, y)\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   515
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   516
text \<open>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   517
  and the second combinator to the second component, defined as
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   518
\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   519
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   520
ML %grayML\<open>fun (x, y) ||> f = (x, f y)\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   521
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   522
text \<open>
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   523
  These two functions can, for example, be used to avoid explicit \<open>lets\<close> for
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   524
  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
   525
  want to separate a list of integers into two lists according to a
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   526
  threshold. If the threshold is @{ML \<open>5\<close>}, the list @{ML \<open>[1,6,2,5,3,4]\<close>}
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   527
  should be separated as @{ML \<open>([1,2,3,4], [6,5])\<close>}.  Such a function can be
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   528
  implemented as
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   529
\<close>
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   530
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   531
ML %grayML\<open>fun separate i [] = ([], [])
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   532
  | separate i (x::xs) =
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   533
      let 
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   534
        val (los, grs) = separate i xs
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   535
      in
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   536
        if i <= x then (los, x::grs) else (x::los, grs)
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   537
      end\<close>
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   538
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   539
text \<open>
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   540
  where the return value of the recursive call is bound explicitly to
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   541
  the pair @{ML \<open>(los, grs)\<close> for los grs}. However, this function
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   542
  can be implemented more concisely as
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   543
\<close>
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   544
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   545
ML %grayML\<open>fun separate _ [] = ([], [])
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   546
  | separate i (x::xs) =
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   547
      if i <= x 
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   548
      then separate i xs ||> cons x
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   549
      else separate i xs |>> cons x\<close>
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   550
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   551
text \<open>
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   552
  avoiding the explicit \<open>let\<close>. While in this example the gain in
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   553
  conciseness is only small, in more complicated situations the benefit of
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   554
  avoiding \<open>lets\<close> can be substantial.
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   555
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   556
  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
   557
  elements from a pair. This combinator is defined as
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   558
\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   559
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   560
ML %grayML\<open>fun (x, y) |-> f = f x y\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   561
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   562
text \<open>
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   563
  and can be used to write the following roundabout version 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   564
  of the \<open>double\<close> function: 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   565
\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   566
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   567
ML %grayML\<open>fun double x =
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   568
      x |>  (fn x => (x, x))
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   569
        |-> (fn x => fn y => x + y)\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   570
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   571
text \<open>
564
6e2479089226 tuned parsing in document antiquotations for ML
Norbert Schirmer <norbert.schirmer@web.de>
parents: 562
diff changeset
   572
  The combinator @{ML_ind ||>> in Basics} plays a central role whenever your
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   573
  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
   574
  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
   575
  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
   576
  @{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
   577
  accumulate the side-results. Consider the following simple function.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   578
\<close>
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   579
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   580
ML %linenosgray\<open>fun acc_incs x =
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   581
    x |> (fn x => ("", x)) 
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   582
      ||>> (fn x => (x, x + 1))
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   583
      ||>> (fn x => (x, x + 1))
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   584
      ||>> (fn x => (x, x + 1))\<close>
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   585
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   586
text \<open>
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   587
  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
   588
  @{ML ||>>} operates on pairs). Each of the next three lines just increment 
280
a63ca3a9b44a spell checked
griff
parents: 279
diff changeset
   589
  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
   590
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
   591
  @{ML_matchresult [display,gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   592
  \<open>acc_incs 1\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   593
  \<open>(((("", 1), 2), 3), 4)\<close>}
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   594
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   595
  You can continue this chain with:
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   596
  
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
   597
  @{ML_matchresult [display,gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   598
  \<open>acc_incs 1 ||>> (fn x => (x, x + 2))\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   599
  \<open>((((("", 1), 2), 3), 4), 6)\<close>}
215
8d1a344a621e more work on the inductive package
Christian Urban <urbanc@in.tum.de>
parents: 210
diff changeset
   600
483
Christian Urban <urbanc@in.tum.de>
parents: 482
diff changeset
   601
  An example where this combinator is useful is as follows
Christian Urban <urbanc@in.tum.de>
parents: 482
diff changeset
   602
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   603
  @{ML_matchresult [display, gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   604
  \<open>let
483
Christian Urban <urbanc@in.tum.de>
parents: 482
diff changeset
   605
  val ((names1, names2), _) =
Christian Urban <urbanc@in.tum.de>
parents: 482
diff changeset
   606
    @{context}
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   607
    |> Variable.variant_fixes (replicate 4 "x")
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   608
    ||>> Variable.variant_fixes (replicate 5 "x")
483
Christian Urban <urbanc@in.tum.de>
parents: 482
diff changeset
   609
in
Christian Urban <urbanc@in.tum.de>
parents: 482
diff changeset
   610
  (names1, names2)
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   611
end\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   612
  \<open>(["x", "xa", "xb", "xc"], ["xd", "xe", "xf", "xg", "xh"])\<close>}
483
Christian Urban <urbanc@in.tum.de>
parents: 482
diff changeset
   613
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   614
  Its purpose is to create nine variants of the string @{ML \<open>"x"\<close>} so
483
Christian Urban <urbanc@in.tum.de>
parents: 482
diff changeset
   615
  that no variant will clash with another. Suppose for some reason we want
Christian Urban <urbanc@in.tum.de>
parents: 482
diff changeset
   616
  to bind four variants to the lists @{ML_text "name1"} and the
Christian Urban <urbanc@in.tum.de>
parents: 482
diff changeset
   617
  rest to @{ML_text "name2"}. In order to obtain non-clashing
Christian Urban <urbanc@in.tum.de>
parents: 482
diff changeset
   618
  variants we have to thread the context through the function calls
Christian Urban <urbanc@in.tum.de>
parents: 482
diff changeset
   619
  (the context records which variants have been previously created).
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   620
  For the first call we can use @{ML \<open>|>\<close>}, but in the
483
Christian Urban <urbanc@in.tum.de>
parents: 482
diff changeset
   621
  second and any further call to @{ML_ind variant_fixes in Variable} we 
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   622
  have to use @{ML \<open>||>>\<close>} in order to account for the result(s) 
483
Christian Urban <urbanc@in.tum.de>
parents: 482
diff changeset
   623
  obtained by previous calls.
Christian Urban <urbanc@in.tum.de>
parents: 482
diff changeset
   624
  
479
7a84649d8839 a few things added First_Steps
Christian Urban <urbanc@in.tum.de>
parents: 478
diff changeset
   625
  A more realistic example for this combinator is the following code
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   626
\<close>
478
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   627
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   628
ML %grayML\<open>val ((([one_def], [two_def]), [three_def]), ctxt') = 
478
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   629
  @{context}
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   630
  |> Local_Defs.define [((@{binding "One"}, NoSyn), (Binding.empty_atts, @{term "1::nat"}))]
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   631
  ||>> Local_Defs.define [((@{binding "Two"}, NoSyn), (Binding.empty_atts,@{term "2::nat"}))]
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   632
  ||>> Local_Defs.define [((@{binding "Three"}, NoSyn), (Binding.empty_atts,@{term "3::nat"}))]\<close>
478
dfbd535cd1fd some tests
Christian Urban <urbanc@in.tum.de>
parents: 477
diff changeset
   633
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   634
text \<open>
496
Christian Urban <urbanc@in.tum.de>
parents: 489
diff changeset
   635
  where we make three definitions, namely @{term "One \<equiv> 1::nat"}, @{term "Two \<equiv> 2::nat"}
Christian Urban <urbanc@in.tum.de>
parents: 489
diff changeset
   636
  and @{term "Three \<equiv> 3::nat"}. The point of this code is that we augment the initial
479
7a84649d8839 a few things added First_Steps
Christian Urban <urbanc@in.tum.de>
parents: 478
diff changeset
   637
  context with the definitions. The result we are interested in is the
7a84649d8839 a few things added First_Steps
Christian Urban <urbanc@in.tum.de>
parents: 478
diff changeset
   638
  augmented context, that is @{ML_text "ctxt'"}, but also the side-results containing 
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   639
  information about the definitions---the function @{ML_ind define in Local_Defs} returns
479
7a84649d8839 a few things added First_Steps
Christian Urban <urbanc@in.tum.de>
parents: 478
diff changeset
   640
  both as pairs. We can use this information for example to print out the definiens and 
7a84649d8839 a few things added First_Steps
Christian Urban <urbanc@in.tum.de>
parents: 478
diff changeset
   641
  the theorem corresponding to the definitions. For example for the first definition:
7a84649d8839 a few things added First_Steps
Christian Urban <urbanc@in.tum.de>
parents: 478
diff changeset
   642
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   643
  @{ML_response [display, gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   644
  \<open>let 
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   645
  val (one_trm, (_, one_thm)) = one_def
479
7a84649d8839 a few things added First_Steps
Christian Urban <urbanc@in.tum.de>
parents: 478
diff changeset
   646
in
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   647
  (pwriteln (pretty_term ctxt' one_trm),
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   648
   pwriteln (pretty_thm ctxt' one_thm))
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   649
end\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   650
  \<open>One
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   651
One \<equiv> 1\<close>}
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   652
  Recall that @{ML \<open>|>\<close>} is the reverse function application. Recall also that
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   653
  the related reverse function composition is @{ML \<open>#>\<close>}. In fact all the
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   654
  combinators @{ML \<open>|->\<close>}, @{ML \<open>|>>\<close>} , @{ML \<open>||>\<close>} and @{ML \<open>||>>\<close>}
344
83d5bca38bec added structures in the index
Christian Urban <urbanc@in.tum.de>
parents: 343
diff changeset
   655
  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
   656
  @{ML_ind "#->" in Basics}, @{ML_ind "#>>" in Basics}, @{ML_ind "##>" in
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   657
  Basics} and @{ML_ind "##>>" in Basics}. Using @{ML \<open>#->\<close>}, for example, the
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   658
  function \<open>double\<close> can also be written as:
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   659
\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   660
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   661
ML %grayML\<open>val double =
502
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   662
   (fn x => (x, x)) #-> 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   663
   (fn x => fn y => x + y)\<close>
126
fcc0e6e54dca polished
Christian Urban <urbanc@in.tum.de>
parents: 124
diff changeset
   664
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   665
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   666
text \<open>
314
79202e2eab6a polished
Christian Urban <urbanc@in.tum.de>
parents: 313
diff changeset
   667
  When using combinators for writing functions in waterfall fashion, it is
311
ee864694315b polished
Christian Urban <urbanc@in.tum.de>
parents: 310
diff changeset
   668
  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
   669
  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
   670
  apply_fresh_args}, where @{ML curry} is needed for making the function @{ML
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   671
  list_comb}, which works over pairs, to fit with the combinator @{ML \<open>|>\<close>}. Such 
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   672
  plumbing is also needed in situations where a function operates over lists, 
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   673
  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
   674
  @{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
   675
  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
   676
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   677
  @{ML_response [display, gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   678
  \<open>let
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   679
  val ctxt = @{context}
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   680
in
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   681
  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
   682
  |> Syntax.check_terms ctxt
441
Christian Urban <urbanc@in.tum.de>
parents: 440
diff changeset
   683
  |> pretty_terms ctxt
Christian Urban <urbanc@in.tum.de>
parents: 440
diff changeset
   684
  |> pwriteln
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   685
end\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   686
  \<open>m + n, m * n, m - n\<close>}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   687
\<close>
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   688
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   689
text \<open>
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   690
  In this example we obtain three terms (using the function 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   691
  @{ML_ind parse_term in Syntax}) whose variables \<open>m\<close> and \<open>n\<close> 
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   692
  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
   693
  check_terms in Syntax} needs plumbing. This can be done with the function
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   694
  @{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
   695
  Syntax} in the file @{ML_file "Pure/Syntax/syntax.ML"} that is implemented 
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   696
  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
   697
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   698
  @{ML_response [display, gray, linenos]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   699
  \<open>let 
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   700
  val ctxt = @{context}
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   701
in
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   702
  Syntax.parse_term ctxt "m - (n::nat)" 
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   703
  |> singleton (Syntax.check_terms ctxt)
441
Christian Urban <urbanc@in.tum.de>
parents: 440
diff changeset
   704
  |> pretty_term ctxt
Christian Urban <urbanc@in.tum.de>
parents: 440
diff changeset
   705
  |> pwriteln
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   706
end\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   707
  \<open>m - n\<close>}
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   708
   
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   709
  where in Line 5, the function operating over lists fits with the
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   710
  single term generated in Line 4.
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   711
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   712
  \begin{readmore}
196
840b49bfb1cf fixed `str_of_thms' output in example + small changes
griff
parents: 192
diff changeset
   713
  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
   714
  "Pure/library.ML"}
127
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   715
  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
   716
  contains further information about combinators.
74846cb0fff9 updated and added two tentative recipes
Christian Urban <urbanc@in.tum.de>
parents: 126
diff changeset
   717
  \end{readmore}
310
007922777ff1 added some rudimentary inrastructure for producing the ML-code
Christian Urban <urbanc@in.tum.de>
parents: 309
diff changeset
   718
421
620a24bf954a added a section to the introduction; described @{make_string}
Christian Urban <urbanc@in.tum.de>
parents: 420
diff changeset
   719
  \begin{exercise}
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   720
  Find out what the combinator @{ML \<open>K I\<close>} does.
421
620a24bf954a added a section to the introduction; described @{make_string}
Christian Urban <urbanc@in.tum.de>
parents: 420
diff changeset
   721
  \end{exercise}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   722
\<close>
15
9da9ba2b095b added a solution section and some other minor additions
Christian Urban <urbanc@in.tum.de>
parents: 14
diff changeset
   723
10
df09e49b19bf many changes in the FirstSteps section
Christian Urban <urbanc@in.tum.de>
parents: 6
diff changeset
   724
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   725
section \<open>ML-Antiquotations\label{sec:antiquote}\<close>
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   726
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   727
text \<open>
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   728
  Recall from Section \ref{sec:include} that code in Isabelle is always
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   729
  embedded in a theory.  The main advantage of this is that the code can
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   730
  contain references to entities defined on the logical level of Isabelle. By
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   731
  this we mean references to definitions, theorems, terms and so on. These
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   732
  reference are realised in Isabelle with ML-antiquotations, often just called
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   733
  antiquotations.\footnote{Note that there are two kinds of antiquotations in
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   734
  Isabelle, which have very different purposes and infrastructures. The first
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   735
  kind, described in this section, are \emph{\index*{ML-antiquotation}}. They
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   736
  are used to refer to entities (like terms, types etc) from Isabelle's logic
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   737
  layer inside ML-code. The other kind of antiquotations are
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   738
  \emph{document}\index{document antiquotation} antiquotations. They are used
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   739
  only in the text parts of Isabelle and their purpose is to print logical
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   740
  entities inside \LaTeX-documents. Document antiquotations are part of the
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   741
  user level and therefore we are not interested in them in this Tutorial,
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   742
  except in Appendix \ref{rec:docantiquotations} where we show how to
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   743
  implement your own document antiquotations.}  Syntactically antiquotations
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   744
  are indicated by the @{ML_text @}-sign followed by text wrapped in \<open>{\<dots>}\<close>.  For example, one can print out the name of the current theory with
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   745
  the code
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 34
diff changeset
   746
  
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   747
  @{ML_matchresult [display,gray] \<open>Context.theory_name @{theory}\<close> \<open>"First_Steps"\<close>}
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 34
diff changeset
   748
 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   749
  where \<open>@{theory}\<close> is an antiquotation that is substituted with the
49
a0edabf14457 added more material
Christian Urban <urbanc@in.tum.de>
parents: 48
diff changeset
   750
  current theory (remember that we assumed we are inside the theory 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   751
  \<open>First_Steps\<close>). 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
   752
  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
   753
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 87
diff changeset
   754
  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
   755
  determined at ``compile-time'', not at ``run-time''. For example the function
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   756
\<close>
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
   757
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   758
ML %grayML\<open>fun not_current_thyname () = Context.theory_name @{theory}\<close>
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
   759
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   760
text \<open>
89
fee4942c4770 polishing
Christian Urban <urbanc@in.tum.de>
parents: 87
diff changeset
   761
  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
   762
  different theory. Instead, the code above defines the constant function 
441
Christian Urban <urbanc@in.tum.de>
parents: 440
diff changeset
   763
  that always returns the string @{text [quotes] "First_Steps"}, no matter where the
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   764
  function is called. Operationally speaking,  the antiquotation \<open>@{theory}\<close> is 
5
e91f54791e14 minor modifiations to the Intro and FirstSteps chapters
Christian Urban <urbanc@in.tum.de>
parents: 2
diff changeset
   765
  \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
   766
  some data structure and return it. Instead, it is literally
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   767
  replaced with the value representing the theory.
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   768
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   769
  Another important antiquotation is \<open>@{context}\<close>. (What the
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   770
  difference between a theory and a context is will be described in Chapter
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   771
  \ref{chp:advanced}.) A context is for example needed to use the
475
25371f74c768 updated to post-2011-1 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
   772
  function @{ML print_abbrevs in Proof_Context} that list of all currently
557
77ea2de0ca62 updated for Isabelle 2014
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 554
diff changeset
   773
  defined abbreviations. For example
2
978a3c2ed7ce split the document into smaller pieces;
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   774
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
   775
  @{ML_matchresult_fake [display, gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   776
  \<open>Proof_Context.print_abbrevs false @{context}\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   777
\<open>\<dots>
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   778
INTER \<equiv> INFIMUM
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   779
Inter \<equiv> Inf
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   780
\<dots>\<close>}
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   781
557
77ea2de0ca62 updated for Isabelle 2014
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 554
diff changeset
   782
  The precise output of course depends on the abbreviations that are
77ea2de0ca62 updated for Isabelle 2014
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 554
diff changeset
   783
  currently defined; this can change over time.
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   784
  You can also use antiquotations to refer to proved theorems: 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   785
  \<open>@{thm \<dots>}\<close> for a single theorem
39
631d12c25bde substantial changes to the antiquotations (preliminary version)
Christian Urban <urbanc@in.tum.de>
parents: 34
diff changeset
   786
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   787
  @{ML_response [display,gray] \<open>@{thm allI}\<close> \<open>(\<And>x. ?P x) \<Longrightarrow> \<forall>x. ?P x\<close>}
75
Christian Urban <urbanc@in.tum.de>
parents: 73
diff changeset
   788
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   789
  and \<open>@{thms \<dots>}\<close> for more than one
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   790
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   791
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   792
@{ML_response [display,gray] 
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   793
  \<open>@{thms conj_ac}\<close> 
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   794
\<open>["(?P \<and> ?Q) = (?Q \<and> ?P)", 
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   795
  "(?P \<and> ?Q \<and> ?R) = (?Q \<and> ?P \<and> ?R)", 
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   796
  "((?P \<and> ?Q) \<and> ?R) = (?P \<and> ?Q \<and> ?R)"]\<close>}  
132
2d9198bcb850 polished
Christian Urban <urbanc@in.tum.de>
parents: 131
diff changeset
   797
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   798
  The thm-antiquotations can also be used for manipulating theorems. For 
474
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
   799
  example, if you need the version of the theorem @{thm [source] refl} that 
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   800
  has a meta-equality instead of an equality, you can write
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   801
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   802
@{ML_response [display,gray] 
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   803
  \<open>@{thm refl[THEN eq_reflection]}\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   804
  \<open>?x \<equiv> ?x\<close>}
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   805
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   806
  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
   807
  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
   808
  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
   809
  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
   810
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   811
  It is also possible to prove lemmas with the antiquotation \<open>@{lemma \<dots> by \<dots>}\<close>
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   812
  whose first argument is a statement (possibly many of them separated by \<open>and\<close>)
375
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   813
  and the second is a proof. For example
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   814
\<close>
375
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   815
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   816
ML %grayML\<open>val foo_thms = @{lemma "True" and "False \<Longrightarrow> P" by simp_all}\<close>
375
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   817
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   818
text \<open>
377
272ba2cceeb2 added a section about unification and matching
Christian Urban <urbanc@in.tum.de>
parents: 376
diff changeset
   819
  The result can be printed out as follows.
375
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   820
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   821
  @{ML_response [gray,display]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   822
\<open>foo_thms |> pretty_thms_no_vars @{context}
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   823
         |> pwriteln\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   824
  \<open>True, False \<Longrightarrow> P\<close>}
375
92f7328dc5cc added type work and updated to Isabelle and poly 5.3
Christian Urban <urbanc@in.tum.de>
parents: 374
diff changeset
   825
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   826
  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
   827
  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
   828
  simpset.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   829
\<close>
75
Christian Urban <urbanc@in.tum.de>
parents: 73
diff changeset
   830
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   831
ML %grayML\<open>fun get_thm_names_from_ss ctxt =
131
8db9195bb3e9 polished
Christian Urban <urbanc@in.tum.de>
parents: 128
diff changeset
   832
let
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 538
diff changeset
   833
  val simpset = Raw_Simplifier.simpset_of ctxt
458
242e81f4d461 updated to post-2011 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 455
diff changeset
   834
  val {simps,...} = Raw_Simplifier.dest_ss simpset
70
bbb2d5f1d58d deleted the fixme about simpsets
Christian Urban <urbanc@in.tum.de>
parents: 69
diff changeset
   835
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
   836
  map #1 simps
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   837
end\<close>
54
1783211b3494 tuned; added document antiquotation ML_response_fake_both
Christian Urban <urbanc@in.tum.de>
parents: 52
diff changeset
   838
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   839
text \<open>
458
242e81f4d461 updated to post-2011 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 455
diff changeset
   840
  The function @{ML_ind dest_ss in Raw_Simplifier} returns a record containing all
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
   841
  information stored in the simpset, but here 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
   842
  simp-rules. Now you can feed in the current simpset into this function. 
544
501491d56798 updated to simplifier change
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 538
diff changeset
   843
  The current simpset can be referred to using @{ML_ind simpset_of in Raw_Simplifier}.
81
8fda6b452f28 polished
Christian Urban <urbanc@in.tum.de>
parents: 78
diff changeset
   844
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   845
  @{ML_response [display,gray] 
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   846
  \<open>get_thm_names_from_ss @{context}\<close> 
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   847
  \<open>["Euclidean_Division.euclidean_size_int_def",\<dots>\<close>}
10
df09e49b19bf many changes in the FirstSteps section
Christian Urban <urbanc@in.tum.de>
parents: 6
diff changeset
   848
196
840b49bfb1cf fixed `str_of_thms' output in example + small changes
griff
parents: 192
diff changeset
   849
  Again, this way of referencing simpsets makes you independent from additions
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   850
  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
   851
  code.
156
e8f11280c762 polished
Christian Urban <urbanc@in.tum.de>
parents: 153
diff changeset
   852
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   853
  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
   854
  exercise care when introducing new ones, as they can also make your code
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   855
  difficult to read. In the next chapter we describe how to construct
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   856
  terms with the (built-in) antiquotation \<open>@{term \<dots>}\<close>.  A restriction
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   857
  of this antiquotation is that it does not allow you to use schematic
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   858
  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
   859
  this restriction, you can implement your own using the function @{ML_ind
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
   860
  inline in ML_Antiquotation} from the structure @{ML_structure ML_Antiquotation}. The code
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   861
  for the antiquotation \<open>term_pat\<close> is as follows.
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   862
\<close>
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
   863
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   864
ML %linenosgray\<open>val term_pat_setup = 
471
f65b9f14d5de updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 468
diff changeset
   865
let
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   866
  val parser = Args.context -- Scan.lift Args.embedded_inner_syntax
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   867
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   868
  fun term_pat (ctxt, str) =
475
25371f74c768 updated to post-2011-1 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
   869
     str |> Proof_Context.read_term_pattern ctxt
264
311830b43f8f minor tuning
Christian Urban <urbanc@in.tum.de>
parents: 263
diff changeset
   870
         |> ML_Syntax.print_term
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   871
         |> ML_Syntax.atomic
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   872
in
554
638ed040e6f8 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 553
diff changeset
   873
  ML_Antiquotation.inline @{binding "term_pat"} (parser >> term_pat)
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   874
end\<close>
263
195c4444dff7 added section about code maintenance and added an example for antiquotations
Christian Urban <urbanc@in.tum.de>
parents: 262
diff changeset
   875
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   876
text \<open>
474
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
   877
  To use it you also have to install it using \isacommand{setup} like so
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   878
\<close>
474
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
   879
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   880
setup %gray \<open>term_pat_setup\<close>
471
f65b9f14d5de updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 468
diff changeset
   881
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   882
text \<open>
308
c90f4ec30d43 improvements from the workshop
Christian Urban <urbanc@in.tum.de>
parents: 307
diff changeset
   883
  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
   884
  transformed into a term using the function @{ML_ind read_term_pattern in
475
25371f74c768 updated to post-2011-1 Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
   885
  Proof_Context} (Line 5); the next two lines transform the term into a string
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   886
  so that the ML-system can understand it. (All these functions will be explained
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   887
  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
   888
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   889
  @{ML_matchresult [display,gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   890
  \<open>@{term_pat "Suc (?x::nat)"}\<close>
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   891
  \<open>Const ("Nat.Suc", _) $ Var (("x", 0), _)\<close>}
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   892
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   893
  which shows the internal representation of the term \<open>Suc ?x\<close>. Similarly
474
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
   894
  we can write an antiquotation for type patterns. Its code is
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   895
\<close>
377
272ba2cceeb2 added a section about unification and matching
Christian Urban <urbanc@in.tum.de>
parents: 376
diff changeset
   896
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   897
ML %grayML\<open>val type_pat_setup = 
471
f65b9f14d5de updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 468
diff changeset
   898
let
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   899
  val parser = Args.context -- Scan.lift Args.embedded_inner_syntax
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   900
377
272ba2cceeb2 added a section about unification and matching
Christian Urban <urbanc@in.tum.de>
parents: 376
diff changeset
   901
  fun typ_pat (ctxt, str) =
503
3778bddfb824 updated to Isabelle 24 November
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   902
    let
3778bddfb824 updated to Isabelle 24 November
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   903
      val ctxt' = Proof_Context.set_mode Proof_Context.mode_schematic ctxt
3778bddfb824 updated to Isabelle 24 November
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   904
    in 
3778bddfb824 updated to Isabelle 24 November
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   905
      str |> Syntax.read_typ ctxt'
3778bddfb824 updated to Isabelle 24 November
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   906
          |> ML_Syntax.print_typ
3778bddfb824 updated to Isabelle 24 November
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   907
          |> ML_Syntax.atomic
3778bddfb824 updated to Isabelle 24 November
Christian Urban <urbanc@in.tum.de>
parents: 496
diff changeset
   908
      end
377
272ba2cceeb2 added a section about unification and matching
Christian Urban <urbanc@in.tum.de>
parents: 376
diff changeset
   909
in
554
638ed040e6f8 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 553
diff changeset
   910
  ML_Antiquotation.inline @{binding "typ_pat"} (parser >> typ_pat)
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   911
end\<close>
377
272ba2cceeb2 added a section about unification and matching
Christian Urban <urbanc@in.tum.de>
parents: 376
diff changeset
   912
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   913
text \<open>
474
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
   914
  which can be installed with
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   915
\<close>
474
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
   916
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   917
setup %gray \<open>type_pat_setup\<close>
471
f65b9f14d5de updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 468
diff changeset
   918
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   919
text \<open>
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   920
However, a word of warning is in order: new antiquotations should be introduced only after
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
   921
careful deliberations. They can potentially make your code harder, rather than easier, to read.
474
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
   922
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   923
  \begin{readmore} @{url "Norbert/ML/ml_antiquotation.ML"}
554
638ed040e6f8 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 553
diff changeset
   924
  The files @{ML_file "Pure/ML/ml_antiquotation.ML"} and @{ML_file "Pure/ML/ml_antiquotations.ML"}
638ed040e6f8 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 553
diff changeset
   925
  contain the infrastructure and definitions for most antiquotations. Most of the basic operations 
638ed040e6f8 updated
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 553
diff changeset
   926
  on ML-syntax are implemented 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
   927
  \end{readmore}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   928
\<close>
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
   929
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   930
section \<open>Storing Data in Isabelle\label{sec:storing}\<close>
292
41a802bbb7df added more to the ML-antiquotation section
Christian Urban <urbanc@in.tum.de>
parents: 290
diff changeset
   931
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   932
text \<open>
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   933
  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
   934
  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
   935
  wisdom has it that the type-system of ML ensures that  an
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
   936
  @{ML_type "'a list"}, say, can only hold elements of the same type, namely
467
e11fe0de19a5 soem polishing
Christian Urban <urbanc@in.tum.de>
parents: 466
diff changeset
   937
  @{ML_type "'a"} (or whatever is substitued for it). Despite this common 
e11fe0de19a5 soem polishing
Christian Urban <urbanc@in.tum.de>
parents: 466
diff changeset
   938
  wisdom, however, it is possible to implement a
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   939
  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
   940
  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
   941
  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
   942
  in contrast to datatypes, which only allow injection and projection of data for
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   943
  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
   944
  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
   945
  destroy type-safety of ML: storing and accessing the data can only be done
474
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
   946
  in a type-safe manner...though run-time checks are needed for that.
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
   947
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   948
  \begin{readmore}
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   949
  In ML the universal type is implemented as the type @{ML_type
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
   950
  Universal.universal}.
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   951
  \end{readmore}
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   952
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   953
  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
   954
  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
   955
  functions for booleans and integers into and from the type @{ML_type Universal.universal}.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   956
\<close>
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
   957
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   958
ML %grayML\<open>local
325
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   959
  val fn_int  = Universal.tag () : int  Universal.tag
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   960
  val fn_bool = Universal.tag () : bool Universal.tag
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   961
in
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   962
  val inject_int   = Universal.tagInject fn_int;
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   963
  val inject_bool  = Universal.tagInject fn_bool;
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   964
  val project_int  = Universal.tagProject fn_int;
352e31d9dacc started section about storing data
Christian Urban <urbanc@in.tum.de>
parents: 324
diff changeset
   965
  val project_bool = Universal.tagProject fn_bool
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   966
end\<close>
298
8057d65607eb some general polishing
Christian Urban <urbanc@in.tum.de>
parents: 293
diff changeset
   967
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   968
text \<open>
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   969
  Using the injection functions, we can inject the integer @{ML_text "13"} 
330
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
   970
  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
   971
  then store them in a @{ML_type "Universal.universal list"} as follows:
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   972
\<close>
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
   973
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   974
ML %grayML\<open>val foo_list = 
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   975
let
474
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
   976
  val thirteen = inject_int 13
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   977
  val truth_val = inject_bool true
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   978
in
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   979
  [thirteen, truth_val]
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   980
end\<close>
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   981
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
   982
text \<open>
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
   983
  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
   984
  
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   985
  @{ML_response [display, gray]
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   986
\<open>(project_int (nth foo_list 0), 
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   987
project_bool (nth foo_list 1))\<close> 
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   988
\<open>13,
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   989
true\<close>}
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   990
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   991
  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
   992
  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
   993
  a runtime error. 
557
77ea2de0ca62 updated for Isabelle 2014
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 554
diff changeset
   994
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   995
  @{ML_response [display, gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
   996
\<open>project_bool (nth foo_list 0)\<close>  
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
   997
\<open>exception Match raised\<close>}
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   998
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
   999
  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
  1000
  containing a universal type.
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1001
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1002
  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
  1003
  data: theorem lists, simpsets, facts etc.  Roughly speaking, there are two
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1004
  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
  1005
  contexts}. Data such as simpsets are ``global'' and therefore need to be stored
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1006
  in a theory (simpsets need to be maintained across proofs and even across
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1007
  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
  1008
  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
  1009
  maintained inside a proof context, which represents ``local'' data.
467
e11fe0de19a5 soem polishing
Christian Urban <urbanc@in.tum.de>
parents: 466
diff changeset
  1010
  You can think of a theory as the ``longterm memory'' of Isabelle (nothing will
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
  1011
  be deleted from it), and a proof-context as a ``shortterm memory'' (it
467
e11fe0de19a5 soem polishing
Christian Urban <urbanc@in.tum.de>
parents: 466
diff changeset
  1012
  changes according to what is needed at the time).
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1013
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1014
  For theories and proof contexts there are, respectively, the functors 
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
  1015
  @{ML_functor_ind Theory_Data} and @{ML_functor_ind Proof_Data} that help
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1016
  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
  1017
  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
  1018
  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
  1019
  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
  1020
  and therefore it makes sense to store this data inside a theory.  
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
  1021
  Consequently we use the functor @{ML_functor Theory_Data}.
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1022
  The code for the table is:
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1023
\<close>
323
92f6a772e013 some polishing
Christian Urban <urbanc@in.tum.de>
parents: 322
diff changeset
  1024
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1025
ML %linenosgray\<open>structure Data = Theory_Data
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1026
  (type T = thm Symtab.table
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1027
   val empty = Symtab.empty
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1028
   val extend = I
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1029
   val merge = Symtab.merge (K true))\<close>
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1030
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1031
text \<open>
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1032
  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
  1033
  (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
  1034
  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
  1035
  producing an associated @{ML_type thm}. We also have to specify four
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1036
  functions to use this functor: namely how to initialise the data storage
385
78c91a629602 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 378
diff changeset
  1037
  (Line 3), how to extend it (Line 4) and how two
78c91a629602 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 378
diff changeset
  1038
  tables should be merged (Line 5). These functions correspond roughly to the
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1039
  operations performed on theories and we just give some sensible 
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1040
  defaults.\footnote{\bf FIXME: Say more about the
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1041
  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
  1042
  contains functions for accessing the table (@{ML Data.get}) and for updating
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
  1043
  it (@{ML Data.map}). There is also the function @{ML Data.put}, but it is 
385
78c91a629602 updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 378
diff changeset
  1044
  not relevant here. Below we define two
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1045
  auxiliary functions, which help us with accessing the table.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1046
\<close>
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1047
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1048
ML %grayML\<open>val lookup = Symtab.lookup o Data.get
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1049
fun update k v = Data.map (Symtab.update (k, v))\<close>
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1050
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1051
text \<open>
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1052
  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
  1053
  logical connective, we can fill the table as follows.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1054
\<close>
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1055
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1056
setup %gray \<open>
450
102dc5cc1aed a bit closer to the new conventions of naming HOL-constants
Christian Urban <urbanc@in.tum.de>
parents: 446
diff changeset
  1057
  update "conj" @{thm conjI} #>
102dc5cc1aed a bit closer to the new conventions of naming HOL-constants
Christian Urban <urbanc@in.tum.de>
parents: 446
diff changeset
  1058
  update "imp"  @{thm impI}  #>
102dc5cc1aed a bit closer to the new conventions of naming HOL-constants
Christian Urban <urbanc@in.tum.de>
parents: 446
diff changeset
  1059
  update "all"  @{thm allI}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1060
\<close>
326
f76135c6c527 more work on the tutorial
Christian Urban <urbanc@in.tum.de>
parents: 325
diff changeset
  1061
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1062
text \<open>
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1063
  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
  1064
  \emph{current} theory is updated (this is explained further in 
559
ffa5c4ec9611 improvements by Piotr Trojanek
Christian Urban <christian dot urban at kcl dot ac dot uk>
parents: 557
diff changeset
  1065
  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
  1066
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
  1067
  @{ML_response [display, gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1068
\<open>lookup @{theory} "conj"\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1069
\<open>SOME "\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q"\<close>}
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1070
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1071
  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
  1072
  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
  1073
  we can update the table as follows
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1074
\<close>
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1075
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1076
setup %gray \<open>update "conj" @{thm TrueI}\<close>
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1077
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
  1078
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1079
text \<open>
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1080
  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
  1081
  
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
  1082
@{ML_response [display, gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1083
\<open>lookup @{theory} "conj"\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1084
\<open>SOME "True"\<close>}
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1085
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1086
  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
  1087
  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
  1088
  interfere with the multithreaded execution model of Isabelle and also
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1089
  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
  1090
  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
  1091
  integers.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1092
\<close>
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1093
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1094
ML %grayML\<open>structure WrongRefData = Theory_Data
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1095
  (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
  1096
   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
  1097
   val extend = I
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1098
   val merge = fst)\<close>
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1099
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1100
text \<open>
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1101
  We initialise the reference with the empty list. Consequently a first 
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1102
  lookup produces @{ML \<open>ref []\<close> in Unsynchronized}.
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1103
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
  1104
  @{ML_matchresult_fake [display,gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1105
  \<open>WrongRefData.get @{theory}\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1106
  \<open>ref []\<close>}
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1107
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1108
  For updating the reference we use the following function
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1109
\<close>
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1110
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1111
ML %grayML\<open>fun ref_update n = WrongRefData.map 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1112
      (fn r => let val _ = r := n::(!r) in r end)\<close>
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1113
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1114
text \<open>
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1115
  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
  1116
  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
  1117
  \isacommand{setup}. 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1118
\<close>
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1119
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1120
setup %gray \<open>ref_update 1\<close>
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1121
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1122
text \<open>
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1123
  A lookup in the current theory gives then the expected list 
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1124
  @{ML \<open>ref [1]\<close> in Unsynchronized}.
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1125
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
  1126
  @{ML_matchresult_fake [display,gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1127
  \<open>WrongRefData.get @{theory}\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1128
  \<open>ref [1]\<close>}
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1129
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1130
  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
  1131
  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
  1132
  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
  1133
  reference, Isabelle has no control over it. So it is not empty, but still
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1134
  @{ML \<open>ref [1]\<close> in Unsynchronized}. Adding to the trouble, if we execute the
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1135
  \isacommand{setup}-command again, we do not obtain @{ML \<open>ref [1]\<close> in
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1136
  Unsynchronized}, but
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1137
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
  1138
  @{ML_matchresult_fake [display,gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1139
  \<open>WrongRefData.get @{theory}\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1140
  \<open>ref [1, 1]\<close>}
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1141
474
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
  1142
  Now imagine how often you go backwards and forwards in your proof 
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
  1143
  scripts.\footnote{The same problem can be triggered in the Jedit GUI by
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
  1144
  making the parser to go over and over again over the \isacommand{setup} command.} 
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1145
  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
  1146
  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
  1147
  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
  1148
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1149
  \begin{readmore}
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1150
  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
  1151
  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
  1152
  including association lists in @{ML_file "Pure/General/alist.ML"},
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1153
  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
  1154
  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
  1155
  \end{readmore}
327
ce754ad78bc9 more work on the storing section
Christian Urban <urbanc@in.tum.de>
parents: 326
diff changeset
  1156
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1157
  Storing data in a proof context is done in a similar fashion. As mentioned
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
  1158
  before, the corresponding functor is @{ML_functor_ind Proof_Data}. With the
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1159
  following code we can store a list of terms in a proof context.
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1160
\<close>
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1161
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1162
ML %grayML\<open>structure Data = Proof_Data
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1163
  (type T = term list
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1164
   fun init _ = [])\<close>
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1165
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1166
text \<open>
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
  1167
  The init-function we have to specify must produce a list for when a context 
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1168
  is initialised (possibly taking the theory into account from which the 
372
Christian Urban <urbanc@in.tum.de>
parents: 371
diff changeset
  1169
  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
  1170
  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
  1171
  term and printing the list. 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1172
\<close>
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1173
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1174
ML %grayML\<open>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
  1175
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1176
fun print ctxt =
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1177
  case (Data.get ctxt) of
474
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
  1178
    [] => pwriteln (Pretty.str "Empty!")
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1179
  | trms => pwriteln (pretty_terms ctxt trms)\<close>
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1180
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1181
text \<open>
330
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1182
  Next we start with the context generated by the antiquotation 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1183
  \<open>@{context}\<close> and update it in various ways. 
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1184
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
  1185
  @{ML_matchresult_fake [display,gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1186
\<open>let
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1187
  val ctxt0 = @{context}
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1188
  val ctxt1 = ctxt0 |> update @{term "False"}
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1189
                    |> update @{term "True \<and> True"} 
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1190
  val ctxt2 = ctxt0 |> update @{term "1::nat"}
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1191
  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
  1192
in
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1193
  print ctxt0; 
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1194
  print ctxt1;
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1195
  print ctxt2;
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1196
  print ctxt3
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1197
end\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1198
\<open>Empty!
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1199
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
  1200
1
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1201
2, 1\<close>}
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1202
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1203
  Many functions in Isabelle manage and update data in a similar
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
  1204
  fashion. Consequently, such calculations with contexts occur frequently in
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1205
  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
  1206
  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
  1207
  theory. Once we throw away the contexts, we have no access to their
414
5fc2fb34c323 polished
Christian Urban <urbanc@in.tum.de>
parents: 413
diff changeset
  1208
  associated data. This is different for theories, where the command 
328
c0cae24b9d46 updated to new Isabelle; more work on the data section
Christian Urban <urbanc@in.tum.de>
parents: 327
diff changeset
  1209
  \isacommand{setup} registers the data with the current and future 
330
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1210
  theories, and therefore one can access the data potentially 
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1211
  indefinitely.
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1212
484
490fe9483c0d more material
Christian Urban <urbanc@in.tum.de>
parents: 483
diff changeset
  1213
  Move elsewhere
490fe9483c0d more material
Christian Urban <urbanc@in.tum.de>
parents: 483
diff changeset
  1214
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1215
  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
  1216
  for treating theories and proof contexts more uniformly. This type is defined as follows
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1217
\<close>
330
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1218
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1219
ML_val %grayML\<open>datatype generic = 
330
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1220
  Theory of theory 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1221
| Proof of proof\<close>
330
Christian Urban <urbanc@in.tum.de>
parents: 329
diff changeset
  1222
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1223
text \<open>
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1224
  \footnote{\bf FIXME: say more about generic contexts.}
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1225
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1226
  There are two special instances of the data storage mechanism described 
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1227
  above. The first instance implements named theorem lists using the functor
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
  1228
  @{ML_functor_ind Named_Thms}. This is because storing theorems in a list 
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1229
  is such a common task.  To obtain a named theorem list, you just declare
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1230
\<close>
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1231
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1232
ML %grayML\<open>structure FooRules = Named_Thms
481
32323727af96 updated
Christian Urban <urbanc@in.tum.de>
parents: 480 479
diff changeset
  1233
  (val name = @{binding "foo"} 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1234
   val description = "Theorems for foo")\<close>
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1235
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1236
text \<open>
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
  1237
  and set up the @{ML_structure FooRules} with the command
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1238
\<close>
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1239
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1240
setup %gray \<open>FooRules.setup\<close>
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1241
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1242
text \<open>
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1243
  This code declares a data container where the theorems are stored,
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1244
  an attribute \<open>foo\<close> (with the \<open>add\<close> and \<open>del\<close> options
377
272ba2cceeb2 added a section about unification and matching
Christian Urban <urbanc@in.tum.de>
parents: 376
diff changeset
  1245
  for adding and deleting theorems) and an internal ML-interface for retrieving and 
272ba2cceeb2 added a section about unification and matching
Christian Urban <urbanc@in.tum.de>
parents: 376
diff changeset
  1246
  modifying the theorems.
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1247
  Furthermore, the theorems are made available on the user-level under the name 
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1248
  \<open>foo\<close>. For example you can declare three lemmas to be a member of the 
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1249
  theorem list \<open>foo\<close> by:
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1250
\<close>
326
f76135c6c527 more work on the tutorial
Christian Urban <urbanc@in.tum.de>
parents: 325
diff changeset
  1251
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1252
lemma rule1[foo]: "A" sorry
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1253
lemma rule2[foo]: "B" sorry
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1254
lemma rule3[foo]: "C" sorry
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1255
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1256
text \<open>and undeclare the first one by:\<close>
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1257
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1258
declare rule1[foo del]
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1259
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1260
text \<open>You can query the remaining ones with:
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1261
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1262
  \begin{isabelle}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1263
  \isacommand{thm}~\<open>foo\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1264
  \<open>> ?C\<close>\\
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1265
  \<open>> ?B\<close>
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1266
  \end{isabelle}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1267
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1268
  On the ML-level, we can add theorems to the list with @{ML FooRules.add_thm}:
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1269
\<close> 
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1270
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1271
setup %gray \<open>Context.theory_map (FooRules.add_thm @{thm TrueI})\<close>
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1272
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1273
text \<open>
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1274
  The rules in the list can be retrieved using the function 
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1275
  @{ML FooRules.get}:
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1276
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
  1277
  @{ML_response [display,gray] 
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1278
  \<open>FooRules.get @{context}\<close> 
572
438703674711 prefer more result checking in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 569
diff changeset
  1279
  \<open>["True", "?C", "?B"]\<close>}
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1280
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1281
  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
  1282
  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
  1283
  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
  1284
  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
  1285
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1286
  \begin{readmore}
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1287
  For more information about named theorem lists see 
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1288
  @{ML_file "Pure/Tools/named_thms.ML"}.
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1289
  \end{readmore}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1290
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1291
  The second special instance of the data storage mechanism are configuration
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1292
  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
  1293
  resort to the ML-level (and also to avoid references). Assume you want the
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1294
  user to control three values, say \<open>bval\<close> containing a boolean, \<open>ival\<close> containing an integer and \<open>sval\<close> containing a string. These
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1295
  values can be declared by
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1296
\<close>
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1297
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1298
ML %grayML\<open>val bval = Attrib.setup_config_bool @{binding "bval"} (K false)
462
1d1e795bc3ad updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 458
diff changeset
  1299
val ival = Attrib.setup_config_int @{binding "ival"} (K 0)
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1300
val sval = Attrib.setup_config_string @{binding "sval"} (K "some string")\<close>
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1301
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1302
text \<open>
462
1d1e795bc3ad updated to new Isabelle
Christian Urban <urbanc@in.tum.de>
parents: 458
diff changeset
  1303
  where each value needs to be given a default. 
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1304
  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
  1305
  with the command
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1306
\<close>
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1307
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1308
declare [[bval = true, ival = 3]]
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1309
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1310
text \<open>
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1311
  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
  1312
  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
  1313
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
  1314
  @{ML_matchresult [display,gray] 
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1315
  \<open>Config.get @{context} bval\<close> 
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1316
  \<open>true\<close>}
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1317
423
0a25b35178c3 updated to new isabelle
Christian Urban <urbanc@in.tum.de>
parents: 421
diff changeset
  1318
  or directly from a theory using the function @{ML_ind get_global in Config}
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1319
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
  1320
  @{ML_matchresult [display,gray] 
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1321
  \<open>Config.get_global @{theory} bval\<close> 
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1322
  \<open>true\<close>}
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1323
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1324
  It is also possible to manipulate the configuration values
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1325
  from the ML-level with the functions @{ML_ind put in Config}
423
0a25b35178c3 updated to new isabelle
Christian Urban <urbanc@in.tum.de>
parents: 421
diff changeset
  1326
  and @{ML_ind put_global 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
  1327
567
f7c97e64cc2a tuned ML-antiquotations; added intro portions.
Norbert Schirmer <norbert.schirmer@web.de>
parents: 566
diff changeset
  1328
  @{ML_matchresult [display,gray]
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1329
  \<open>let
346
0fea8b7a14a1 tuned the ML-output mechanism; tuned slightly the text
Christian Urban <urbanc@in.tum.de>
parents: 344
diff changeset
  1330
  val ctxt = @{context}
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1331
  val ctxt' = Config.put sval "foo" ctxt
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1332
  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
  1333
in
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1334
  (Config.get ctxt sval, 
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1335
   Config.get ctxt' sval, 
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1336
   Config.get ctxt'' sval)
569
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1337
end\<close>
f875a25aa72d prefer cartouches over " in ML antiquotations
Norbert Schirmer <norbert.schirmer@web.de>
parents: 567
diff changeset
  1338
  \<open>("some string", "foo", "bar")\<close>}
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1339
468
00921ae66622 more polishing
Christian Urban <urbanc@in.tum.de>
parents: 467
diff changeset
  1340
  A concrete example for a configuration value is 
00921ae66622 more polishing
Christian Urban <urbanc@in.tum.de>
parents: 467
diff changeset
  1341
  @{ML_ind simp_trace in Raw_Simplifier}, which switches on trace information
474
683fb6e468b1 a few things updated
Christian Urban <urbanc@in.tum.de>
parents: 471
diff changeset
  1342
  in the simplifier. This can be used for example in the following proof
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1343
\<close>
468
00921ae66622 more polishing
Christian Urban <urbanc@in.tum.de>
parents: 467
diff changeset
  1344
562
daf404920ab9 Accomodate to Isabelle 2018
Norbert Schirmer <norbert.schirmer@web.de>
parents: 559
diff changeset
  1345
468
00921ae66622 more polishing
Christian Urban <urbanc@in.tum.de>
parents: 467
diff changeset
  1346
lemma
00921ae66622 more polishing
Christian Urban <urbanc@in.tum.de>
parents: 467
diff changeset
  1347
  shows "(False \<or> True) \<and> True"
00921ae66622 more polishing
Christian Urban <urbanc@in.tum.de>
parents: 467
diff changeset
  1348
proof (rule conjI)
00921ae66622 more polishing
Christian Urban <urbanc@in.tum.de>
parents: 467
diff changeset
  1349
  show "False \<or> True" using [[simp_trace = true]] by simp
00921ae66622 more polishing
Christian Urban <urbanc@in.tum.de>
parents: 467
diff changeset
  1350
next
00921ae66622 more polishing
Christian Urban <urbanc@in.tum.de>
parents: 467
diff changeset
  1351
  show "True" by simp
00921ae66622 more polishing
Christian Urban <urbanc@in.tum.de>
parents: 467
diff changeset
  1352
qed
00921ae66622 more polishing
Christian Urban <urbanc@in.tum.de>
parents: 467
diff changeset
  1353
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1354
text \<open>
468
00921ae66622 more polishing
Christian Urban <urbanc@in.tum.de>
parents: 467
diff changeset
  1355
  in order to inspect how the simplifier solves the first subgoal.
00921ae66622 more polishing
Christian Urban <urbanc@in.tum.de>
parents: 467
diff changeset
  1356
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1357
  \begin{readmore}
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1358
  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
  1359
  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
  1360
  @{ML_file "Pure/config.ML"}.
329
5dffcab68680 more work
Christian Urban <urbanc@in.tum.de>
parents: 328
diff changeset
  1361
  \end{readmore}
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1362
\<close>
343
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1363
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1364
section \<open>Summary\<close>
343
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1365
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1366
text \<open>
343
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1367
  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
  1368
  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
  1369
  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
  1370
  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
  1371
  contains mechanisms for storing arbitrary data in theory and proof 
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1372
  contexts.
8f73e80c8c6f polished first chapter
Christian Urban <urbanc@in.tum.de>
parents: 339
diff changeset
  1373
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1374
  \begin{conventions}
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1375
  \begin{itemize}
370
2494b5b7a85d added something about show_types references
Christian Urban <urbanc@in.tum.de>
parents: 369
diff changeset
  1376
  \item Print messages that belong together in a single string.
350
364fffa80794 polished
Christian Urban <urbanc@in.tum.de>
parents: 347
diff changeset
  1377
  \item Do not use references in Isabelle code.
347
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1378
  \end{itemize}
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1379
  \end{conventions}
01e71cddf6a3 slightly tuned
Christian Urban <urbanc@in.tum.de>
parents: 346
diff changeset
  1380
565
cecd7a941885 isabelle update_cartouches -t
Norbert Schirmer <norbert.schirmer@web.de>
parents: 564
diff changeset
  1381
\<close>
196
840b49bfb1cf fixed `str_of_thms' output in example + small changes
griff
parents: 192
diff changeset
  1382
end