| author | Christian Urban <urbanc@in.tum.de> | 
| Sun, 26 Apr 2009 23:45:22 +0200 | |
| changeset 242 | 14d5f0cf91dc | 
| parent 240 | d111f5988e49 | 
| child 243 | 6e0f56764ff8 | 
| permissions | -rw-r--r-- | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1 | theory FirstSteps | 
| 25 
e2f9f94b26d4
Antiquotation setup is now contained in theory Base.
 berghofe parents: 
21diff
changeset | 2 | imports Base | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 3 | begin | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 4 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 5 | chapter {* First Steps *}
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 6 | |
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 7 | text {*
 | 
| 242 
14d5f0cf91dc
added note about @{type_names ...}
 Christian Urban <urbanc@in.tum.de> parents: 
240diff
changeset | 8 | Isabelle programming is done in ML. Just like lemmas and proofs, ML-code | 
| 89 | 9 | in Isabelle is 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: 
52diff
changeset | 10 | 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 | 11 | |
| 6 | 12 |   \begin{center}
 | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 13 |   \begin{tabular}{@ {}l}
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
34diff
changeset | 14 |   \isacommand{theory} FirstSteps\\
 | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 15 |   \isacommand{imports} Main\\
 | 
| 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 16 |   \isacommand{begin}\\
 | 
| 6 | 17 | \ldots | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 18 |   \end{tabular}
 | 
| 6 | 19 |   \end{center}
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 20 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 21 | We also generally assume you are working with HOL. The given examples might | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 22 | need to be adapted if you work in a different logic. | 
| 14 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 23 | *} | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 24 | |
| 20 | 25 | section {* Including ML-Code *}
 | 
| 14 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 26 | |
| 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 27 | text {*
 | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 28 | The easiest and quickest way to include code in a theory is | 
| 89 | 29 |   by using the \isacommand{ML}-command. For example:
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 30 | |
| 75 | 31 | \begin{isabelle}
 | 
| 32 | \begin{graybox}
 | |
| 85 
b02904872d6b
better handling of {* and *}
 Christian Urban <urbanc@in.tum.de> parents: 
84diff
changeset | 33 | \isacommand{ML}~@{text "\<verbopen>"}\isanewline
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
34diff
changeset | 34 | \hspace{5mm}@{ML "3 + 4"}\isanewline
 | 
| 85 
b02904872d6b
better handling of {* and *}
 Christian Urban <urbanc@in.tum.de> parents: 
84diff
changeset | 35 | @{text "\<verbclose>"}\isanewline
 | 
| 
b02904872d6b
better handling of {* and *}
 Christian Urban <urbanc@in.tum.de> parents: 
84diff
changeset | 36 | @{text "> 7"}\smallskip
 | 
| 75 | 37 | \end{graybox}
 | 
| 38 | \end{isabelle}
 | |
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 39 | |
| 234 | 40 |   Like normal Isabelle scripts, \isacommand{ML}-commands can be
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 41 | evaluated by using the advance and undo buttons of your Isabelle | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 42 |   environment. The code inside the \isacommand{ML}-command can also contain
 | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 43 | value and function bindings, for example | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 44 | *} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 45 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 46 | ML %gray {* 
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 47 | val r = ref 0 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 48 | fun f n = n + 1 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 49 | *} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 50 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 51 | text {*
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 52 | and even those can be undone when the proof | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 53 | script is retracted. As mentioned in the Introduction, we will drop the | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 54 |   \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"} scaffolding whenever we
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 55 |   show code. The lines prefixed with @{text [quotes] ">"} are not part of the
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 56 | code, rather they indicate what the response is when the code is evaluated. | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 57 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 58 | 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: 
234diff
changeset | 59 | to a separate ML-file. Such files can then be included somewhere inside a | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 60 |   theory by using the command \isacommand{use}. For example
 | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 61 | |
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 62 |   \begin{center}
 | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 63 |   \begin{tabular}{@ {}l}
 | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 64 |   \isacommand{theory} FirstSteps\\
 | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 65 |   \isacommand{imports} Main\\
 | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 66 |   \isacommand{uses}~@{text "(\"file_to_be_included.ML\")"} @{text "\<dots>"}\\
 | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 67 |   \isacommand{begin}\\
 | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 68 | \ldots\\ | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 69 |   \isacommand{use}~@{text "\"file_to_be_included.ML\""}\\
 | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 70 | \ldots | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 71 |   \end{tabular}
 | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 72 |   \end{center}
 | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 73 | |
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 74 |   The \isacommand{uses}-command in the header of the theory is needed in order 
 | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 75 | to indicate the dependency of the theory on the ML-file. Alternatively, the | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 76 | file can be included by just writing in the header | 
| 14 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 77 | |
| 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 78 |   \begin{center}
 | 
| 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 79 |   \begin{tabular}{@ {}l}
 | 
| 54 
1783211b3494
tuned; added document antiquotation ML_response_fake_both
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 80 |   \isacommand{theory} FirstSteps\\
 | 
| 14 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 81 |   \isacommand{imports} Main\\
 | 
| 58 | 82 |   \isacommand{uses} @{text "\"file_to_be_included.ML\""} @{text "\<dots>"}\\
 | 
| 14 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 83 |   \isacommand{begin}\\
 | 
| 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 84 | \ldots | 
| 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 85 |   \end{tabular}
 | 
| 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 86 |   \end{center}
 | 
| 235 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 87 | |
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 88 | Note that no parentheses are given this time. | 
| 14 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 89 | *} | 
| 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 90 | |
| 126 | 91 | section {* Debugging and Printing\label{sec:printing} *}
 | 
| 14 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 92 | |
| 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 93 | text {*
 | 
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 94 | |
| 50 | 95 | During development you might find it necessary to inspect some data | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 96 | in your code. This can be done in a ``quick-and-dirty'' fashion using | 
| 239 
b63c72776f03
replaced "warning" with "writeln"
 Christian Urban <urbanc@in.tum.de> parents: 
235diff
changeset | 97 |   the function @{ML "writeln"}. For example 
 | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 98 | |
| 239 
b63c72776f03
replaced "warning" with "writeln"
 Christian Urban <urbanc@in.tum.de> parents: 
235diff
changeset | 99 |   @{ML_response_fake [display,gray] "writeln \"any string\"" "\"any string\""}
 | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 100 | |
| 58 | 101 |   will print out @{text [quotes] "any string"} inside the response buffer
 | 
| 102 | of Isabelle. This function expects a string as argument. If you develop under PolyML, | |
| 50 | 103 | then there is a convenient, though again ``quick-and-dirty'', method for | 
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
239diff
changeset | 104 |   converting values into strings, namely the function @{ML PolyML.makestring}:
 | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 105 | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
239diff
changeset | 106 |   @{ML_response_fake [display,gray] "writeln (PolyML.makestring 1)" "\"1\""} 
 | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 107 | |
| 196 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 108 |   However, @{ML makestring} only works if the type of what is converted is monomorphic 
 | 
| 78 
ef778679d3e0
added a section about combinators
 Christian Urban <urbanc@in.tum.de> parents: 
75diff
changeset | 109 | and not a function. | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 110 | |
| 239 
b63c72776f03
replaced "warning" with "writeln"
 Christian Urban <urbanc@in.tum.de> parents: 
235diff
changeset | 111 |   The function @{ML "writeln"} should only be used for testing purposes, because any
 | 
| 52 | 112 | output this function generates will be overwritten as soon as an error is | 
| 50 | 113 | raised. For printing anything more serious and elaborate, the | 
| 54 
1783211b3494
tuned; added document antiquotation ML_response_fake_both
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 114 |   function @{ML tracing} is more appropriate. This function writes all output into
 | 
| 89 | 115 | a separate tracing buffer. For example: | 
| 14 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 116 | |
| 161 
83f36a1c62f2
rolled back the changes on the function warning and tracing
 Christian Urban <urbanc@in.tum.de> parents: 
160diff
changeset | 117 |   @{ML_response_fake [display,gray] "tracing \"foo\"" "\"foo\""}
 | 
| 14 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 118 | |
| 58 | 119 |   It is also possible to redirect the ``channel'' where the string @{text "foo"} is 
 | 
| 196 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 120 | printed to a separate file, e.g., to prevent ProofGeneral from choking on massive | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 121 | amounts of trace output. This redirection can be achieved with the code: | 
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 122 | *} | 
| 14 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 123 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 124 | ML{*val strip_specials =
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 125 | let | 
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 126 |   fun strip ("\^A" :: _ :: cs) = strip cs
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 127 | | strip (c :: cs) = c :: strip cs | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 128 | | strip [] = []; | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 129 | in implode o strip o explode end; | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 130 | |
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 131 | fun redirect_tracing stream = | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 132 | Output.tracing_fn := (fn s => | 
| 14 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 133 | (TextIO.output (stream, (strip_specials s)); | 
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 134 | TextIO.output (stream, "\n"); | 
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 135 | TextIO.flushOut stream)) *} | 
| 14 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 136 | |
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 137 | text {*
 | 
| 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 138 |   Calling @{ML "redirect_tracing"} with @{ML "(TextIO.openOut \"foo.bar\")"} 
 | 
| 58 | 139 |   will cause that all tracing information is printed into the file @{text "foo.bar"}.
 | 
| 75 | 140 | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 141 |   You can print out error messages with the function @{ML error}; for example:
 | 
| 75 | 142 | |
| 122 | 143 | @{ML_response_fake [display,gray] "if 0=1 then true else (error \"foo\")" 
 | 
| 144 | "Exception- ERROR \"foo\" raised | |
| 145 | At command \"ML\"."} | |
| 75 | 146 | |
| 192 | 147 |   (FIXME @{ML Toplevel.debug} @{ML Toplevel.profiling})
 | 
| 148 | *} | |
| 149 | ||
| 150 | (* | |
| 207 | 151 | ML {* reset Toplevel.debug *}
 | 
| 152 | ||
| 153 | ML {* fun dodgy_fun () = (raise TYPE ("",[],[]); 1) *}
 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 154 | |
| 207 | 155 | ML {* fun innocent () = dodgy_fun () *}
 | 
| 156 | ML {* exception_trace (fn () => cterm_of @{theory} (Bound 0)) *}
 | |
| 157 | ML {* exception_trace (fn () => innocent ()) *}
 | |
| 192 | 158 | |
| 207 | 159 | ML {* Toplevel.program (fn () => cterm_of @{theory} (Bound 0)) *}
 | 
| 192 | 160 | |
| 207 | 161 | ML {* Toplevel.program (fn () => innocent ()) *}
 | 
| 192 | 162 | *) | 
| 163 | ||
| 164 | text {*
 | |
| 126 | 165 |   Most often you want to inspect data of type @{ML_type term}, @{ML_type cterm} 
 | 
| 166 |   or @{ML_type thm}. Isabelle contains elaborate pretty-printing functions for printing them, 
 | |
| 167 | but for quick-and-dirty solutions they are far too unwieldy. A simple way to transform | |
| 168 |   a term into a string is to use the function @{ML Syntax.string_of_term}.
 | |
| 169 | ||
| 170 |   @{ML_response_fake [display,gray]
 | |
| 171 |   "Syntax.string_of_term @{context} @{term \"1::nat\"}"
 | |
| 172 | "\"\\^E\\^Fterm\\^E\\^E\\^Fconst\\^Fname=HOL.one_class.one\\^E1\\^E\\^F\\^E\\^E\\^F\\^E\""} | |
| 173 | ||
| 174 | This produces a string with some additional information encoded in it. The string | |
| 239 
b63c72776f03
replaced "warning" with "writeln"
 Christian Urban <urbanc@in.tum.de> parents: 
235diff
changeset | 175 |   can be properly printed by using the function @{ML writeln}.
 | 
| 126 | 176 | |
| 177 |   @{ML_response_fake [display,gray]
 | |
| 239 
b63c72776f03
replaced "warning" with "writeln"
 Christian Urban <urbanc@in.tum.de> parents: 
235diff
changeset | 178 |   "writeln (Syntax.string_of_term @{context} @{term \"1::nat\"})"
 | 
| 126 | 179 | "\"1\""} | 
| 180 | ||
| 181 |   A @{ML_type cterm} can be transformed into a string by the following function.
 | |
| 182 | *} | |
| 183 | ||
| 184 | ML{*fun str_of_cterm ctxt t =  
 | |
| 185 | Syntax.string_of_term ctxt (term_of t)*} | |
| 186 | ||
| 187 | text {*
 | |
| 149 | 188 |   In this example the function @{ML term_of} extracts the @{ML_type term} from
 | 
| 189 |   a @{ML_type cterm}.  If there are more than one @{ML_type cterm}s to be
 | |
| 190 |   printed, you can use the function @{ML commas} to separate them.
 | |
| 126 | 191 | *} | 
| 192 | ||
| 193 | ML{*fun str_of_cterms ctxt ts =  
 | |
| 194 | commas (map (str_of_cterm ctxt) ts)*} | |
| 195 | ||
| 196 | text {*
 | |
| 197 | The easiest way to get the string of a theorem is to transform it | |
| 190 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 198 |   into a @{ML_type cterm} using the function @{ML crep_thm}. 
 | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 199 | *} | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 200 | |
| 194 | 201 | ML{*fun str_of_thm ctxt thm =
 | 
| 190 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 202 | str_of_cterm ctxt (#prop (crep_thm thm))*} | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 203 | |
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 204 | text {*
 | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 205 |   Theorems also include schematic variables, such as @{text "?P"}, 
 | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 206 |   @{text "?Q"} and so on. 
 | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 207 | |
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 208 |   @{ML_response_fake [display, gray]
 | 
| 239 
b63c72776f03
replaced "warning" with "writeln"
 Christian Urban <urbanc@in.tum.de> parents: 
235diff
changeset | 209 |   "writeln (str_of_thm @{context} @{thm conjI})"
 | 
| 190 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 210 | "\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q"} | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 211 | |
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 212 | In order to improve the readability of theorems we convert | 
| 158 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 213 | these schematic variables into free variables using the | 
| 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 214 |   function @{ML Variable.import_thms}.
 | 
| 126 | 215 | *} | 
| 216 | ||
| 158 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 217 | ML{*fun no_vars ctxt thm =
 | 
| 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 218 | let | 
| 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 219 | val ((_, [thm']), _) = Variable.import_thms true [thm] ctxt | 
| 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 220 | in | 
| 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 221 | thm' | 
| 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 222 | end | 
| 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 223 | |
| 194 | 224 | fun str_of_thm_no_vars ctxt thm = | 
| 171 
18f90044c777
simplified antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
163diff
changeset | 225 | str_of_cterm ctxt (#prop (crep_thm (no_vars ctxt thm)))*} | 
| 126 | 226 | |
| 227 | text {* 
 | |
| 207 | 228 |   Theorem @{thm [source] conjI} is now printed as follows:
 | 
| 190 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 229 | |
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 230 |   @{ML_response_fake [display, gray]
 | 
| 239 
b63c72776f03
replaced "warning" with "writeln"
 Christian Urban <urbanc@in.tum.de> parents: 
235diff
changeset | 231 |   "writeln (str_of_thm_no_vars @{context} @{thm conjI})"
 | 
| 196 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 232 | "\<lbrakk>P; Q\<rbrakk> \<Longrightarrow> P \<and> Q"} | 
| 190 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 233 | |
| 126 | 234 |   Again the function @{ML commas} helps with printing more than one theorem. 
 | 
| 235 | *} | |
| 236 | ||
| 237 | ML{*fun str_of_thms ctxt thms =  
 | |
| 190 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 238 | commas (map (str_of_thm ctxt) thms) | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 239 | |
| 194 | 240 | fun str_of_thms_no_vars ctxt thms = | 
| 241 | commas (map (str_of_thm_no_vars ctxt) thms) *} | |
| 126 | 242 | |
| 243 | section {* Combinators\label{sec:combinators} *}
 | |
| 244 | ||
| 245 | text {*
 | |
| 235 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 246 | (FIXME: Calling convention) | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 247 | |
| 131 | 248 | For beginners perhaps the most puzzling parts in the existing code of Isabelle are | 
| 126 | 249 | the combinators. At first they seem to greatly obstruct the | 
| 250 | comprehension of the code, but after getting familiar with them, they | |
| 251 | actually ease the understanding and also the programming. | |
| 252 | ||
| 253 |   The simplest combinator is @{ML I}, which is just the identity function defined as
 | |
| 254 | *} | |
| 255 | ||
| 256 | ML{*fun I x = x*}
 | |
| 257 | ||
| 258 | text {* Another simple combinator is @{ML K}, defined as *}
 | |
| 259 | ||
| 260 | ML{*fun K x = fn _ => x*}
 | |
| 261 | ||
| 262 | text {*
 | |
| 263 |   @{ML K} ``wraps'' a function around the argument @{text "x"}. However, this 
 | |
| 264 |   function ignores its argument. As a result, @{ML K} defines a constant function 
 | |
| 265 |   always returning @{text x}.
 | |
| 266 | ||
| 267 |   The next combinator is reverse application, @{ML "|>"}, defined as: 
 | |
| 268 | *} | |
| 269 | ||
| 270 | ML{*fun x |> f = f x*}
 | |
| 271 | ||
| 272 | text {* While just syntactic sugar for the usual function application,
 | |
| 273 | the purpose of this combinator is to implement functions in a | |
| 274 | ``waterfall fashion''. Consider for example the function *} | |
| 275 | ||
| 276 | ML %linenosgray{*fun inc_by_five x =
 | |
| 277 | x |> (fn x => x + 1) | |
| 278 | |> (fn x => (x, x)) | |
| 279 | |> fst | |
| 280 | |> (fn x => x + 4)*} | |
| 281 | ||
| 282 | text {*
 | |
| 196 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 283 |   which increments its argument @{text x} by 5. It proceeds by first incrementing 
 | 
| 126 | 284 | the argument by 1 (Line 2); then storing the result in a pair (Line 3); taking | 
| 285 | the first component of the pair (Line 4) and finally incrementing the first | |
| 286 | component by 4 (Line 5). This kind of cascading manipulations of values is quite | |
| 287 | common when dealing with theories (for example by adding a definition, followed by | |
| 288 | lemmas and so on). The reverse application allows you to read what happens in | |
| 289 | a top-down manner. This kind of coding should also be familiar, | |
| 149 | 290 |   if you have been exposed to Haskell's do-notation. Writing the function @{ML inc_by_five} using 
 | 
| 126 | 291 | the reverse application is much clearer than writing | 
| 292 | *} | |
| 293 | ||
| 294 | ML{*fun inc_by_five x = fst ((fn x => (x, x)) (x + 1)) + 4*}
 | |
| 295 | ||
| 296 | text {* or *}
 | |
| 297 | ||
| 298 | ML{*fun inc_by_five x = 
 | |
| 299 | ((fn x => x + 4) o fst o (fn x => (x, x)) o (fn x => x + 1)) x*} | |
| 300 | ||
| 301 | text {* and typographically more economical than *}
 | |
| 302 | ||
| 303 | ML{*fun inc_by_five x =
 | |
| 304 | let val y1 = x + 1 | |
| 305 | val y2 = (y1, y1) | |
| 306 | val y3 = fst y2 | |
| 307 | val y4 = y3 + 4 | |
| 308 | in y4 end*} | |
| 309 | ||
| 310 | text {* 
 | |
| 311 | Another reason why the let-bindings in the code above are better to be | |
| 312 | avoided: it is more than easy to get the intermediate values wrong, not to | |
| 313 | mention the nightmares the maintenance of this code causes! | |
| 314 | ||
| 215 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 315 | 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: 
177diff
changeset | 316 | the waterfall fashion might be the following code: | 
| 177 | 317 | *} | 
| 126 | 318 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 319 | ML %linenosgray{*fun apply_fresh_args f ctxt =
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 320 | f |> fastype_of | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 321 | |> binder_types | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 322 | |> map (pair "z") | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 323 | |> Variable.variant_frees ctxt [f] | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 324 | |> map Free | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 325 | |> (curry list_comb) f *} | 
| 126 | 326 | |
| 177 | 327 | text {*
 | 
| 194 | 328 |   This code extracts the argument types of a given function @{text "f"} and then generates 
 | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 329 | for each argument type a distinct variable; finally it applies the generated | 
| 194 | 330 | variables to the function. For example: | 
| 183 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 331 | |
| 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 332 |   @{ML_response_fake [display,gray]
 | 
| 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 333 | "apply_fresh_args @{term \"P::nat \<Rightarrow> int \<Rightarrow> unit \<Rightarrow> bool\"} @{context} 
 | 
| 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 334 |  |> Syntax.string_of_term @{context}
 | 
| 239 
b63c72776f03
replaced "warning" with "writeln"
 Christian Urban <urbanc@in.tum.de> parents: 
235diff
changeset | 335 | |> writeln" | 
| 183 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 336 | "P z za zb"} | 
| 177 | 337 | |
| 184 | 338 |   You can read off this behaviour from how @{ML apply_fresh_args} is
 | 
| 183 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 339 |   coded: in Line 2, the function @{ML fastype_of} calculates the type of the
 | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 340 |   function; @{ML binder_types} in the next line produces the list of argument
 | 
| 184 | 341 |   types (in the case above the list @{text "[nat, int, unit]"}); Line 4 
 | 
| 342 |   pairs up each type with the string  @{text "z"}; the
 | |
| 183 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 343 |   function @{ML variant_frees in Variable} generates for each @{text "z"} a
 | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 344 |   unique name avoiding the given @{text f}; the list of name-type pairs is turned
 | 
| 184 | 345 | into a list of variable terms in Line 6, which in the last line is applied | 
| 194 | 346 |   by the function @{ML list_comb} to the function. In this last step we have to 
 | 
| 347 |   use the function @{ML curry}, because @{ML list_comb} expects the function and the
 | |
| 184 | 348 | variables list as a pair. | 
| 177 | 349 | |
| 350 |   The combinator @{ML "#>"} is the reverse function composition. It can be
 | |
| 351 | used to define the following function | |
| 126 | 352 | *} | 
| 353 | ||
| 354 | ML{*val inc_by_six = 
 | |
| 355 | (fn x => x + 1) | |
| 356 | #> (fn x => x + 2) | |
| 357 | #> (fn x => x + 3)*} | |
| 358 | ||
| 359 | text {* 
 | |
| 360 | which is the function composed of first the increment-by-one function and then | |
| 361 | increment-by-two, followed by increment-by-three. Again, the reverse function | |
| 362 | composition allows you to read the code top-down. | |
| 363 | ||
| 364 | The remaining combinators described in this section add convenience for the | |
| 365 |   ``waterfall method'' of writing functions. The combinator @{ML tap} allows
 | |
| 366 | you to get hold of an intermediate result (to do some side-calculations for | |
| 367 | instance). The function | |
| 368 | ||
| 369 | *} | |
| 370 | ||
| 371 | ML %linenosgray{*fun inc_by_three x =
 | |
| 372 | x |> (fn x => x + 1) | |
| 240 
d111f5988e49
replaced explode by Symbol.explode
 Christian Urban <urbanc@in.tum.de> parents: 
239diff
changeset | 373 | |> tap (fn x => tracing (PolyML.makestring x)) | 
| 126 | 374 | |> (fn x => x + 2)*} | 
| 375 | ||
| 376 | text {* 
 | |
| 377 |   increments the argument first by @{text "1"} and then by @{text "2"}. In the
 | |
| 378 |   middle (Line 3), however, it uses @{ML tap} for printing the ``plus-one''
 | |
| 379 |   intermediate result inside the tracing buffer. The function @{ML tap} can
 | |
| 380 | only be used for side-calculations, because any value that is computed | |
| 381 | cannot be merged back into the ``main waterfall''. To do this, you can use | |
| 382 | the next combinator. | |
| 383 | ||
| 196 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 384 |   The combinator @{ML "`"} (a backtick) is similar to @{ML tap}, but applies a
 | 
| 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 385 | function to the value and returns the result together with the value (as a | 
| 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 386 | pair). For example the function | 
| 126 | 387 | *} | 
| 388 | ||
| 389 | ML{*fun inc_as_pair x =
 | |
| 390 | x |> `(fn x => x + 1) | |
| 391 | |> (fn (x, y) => (x, y + 1))*} | |
| 392 | ||
| 393 | text {*
 | |
| 394 |   takes @{text x} as argument, and then increments @{text x}, but also keeps
 | |
| 395 |   @{text x}. The intermediate result is therefore the pair @{ML "(x + 1, x)"
 | |
| 396 | for x}. After that, the function increments the right-hand component of the | |
| 397 |   pair. So finally the result will be @{ML "(x + 1, x + 1)" for x}.
 | |
| 398 | ||
| 399 |   The combinators @{ML "|>>"} and @{ML "||>"} are defined for 
 | |
| 400 | functions manipulating pairs. The first applies the function to | |
| 401 | the first component of the pair, defined as | |
| 402 | *} | |
| 403 | ||
| 404 | ML{*fun (x, y) |>> f = (f x, y)*}
 | |
| 405 | ||
| 406 | text {*
 | |
| 407 | and the second combinator to the second component, defined as | |
| 408 | *} | |
| 409 | ||
| 410 | ML{*fun (x, y) ||> f = (x, f y)*}
 | |
| 411 | ||
| 412 | text {*
 | |
| 413 |   With the combinator @{ML "|->"} you can re-combine the elements from a pair.
 | |
| 414 | This combinator is defined as | |
| 415 | *} | |
| 416 | ||
| 417 | ML{*fun (x, y) |-> f = f x y*}
 | |
| 418 | ||
| 215 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 419 | text {* 
 | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 420 | and can be used to write the following roundabout version | 
| 126 | 421 |   of the @{text double} function: 
 | 
| 422 | *} | |
| 423 | ||
| 424 | ML{*fun double x =
 | |
| 425 | x |> (fn x => (x, x)) | |
| 426 | |-> (fn x => fn y => x + y)*} | |
| 427 | ||
| 215 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 428 | text {* 
 | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 429 |   The combinator @{ML ||>>} plays a central rôle whenever your task is to update a 
 | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 430 | theory and the update also produces a side-result (for example a theorem). Functions | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 431 | for such tasks return a pair whose second component is the theory and the fist | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 432 |   component is the side-result. Using @{ML ||>>}, you can do conveniently the update 
 | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 433 | and also accumulate the side-results. Considder the following simple function. | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 434 | *} | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 435 | |
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 436 | ML %linenosgray{*fun acc_incs x =
 | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 437 |     x |> (fn x => ("", x)) 
 | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 438 | ||>> (fn x => (x, x + 1)) | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 439 | ||>> (fn x => (x, x + 1)) | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 440 | ||>> (fn x => (x, x + 1))*} | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 441 | |
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 442 | text {*
 | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 443 | The purpose of Line 2 is to just pair up the argument with a dummy value (since | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 444 |   @{ML "||>>"} operates on pairs). Each of the next three lines just increment 
 | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 445 | the value by one, but also nest the intrermediate results to the left. For example | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 446 | |
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 447 |   @{ML_response [display,gray]
 | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 448 | "acc_incs 1" | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 449 | "((((\"\", 1), 2), 3), 4)"} | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 450 | |
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 451 | You can continue this chain with: | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 452 | |
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 453 |   @{ML_response [display,gray]
 | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 454 | "acc_incs 1 ||>> (fn x => (x, x + 2))" | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 455 | "(((((\"\", 1), 2), 3), 4), 6)"} | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 456 | |
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 457 | (FIXME: maybe give a ``real world'' example) | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 458 | *} | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 459 | |
| 126 | 460 | text {*
 | 
| 196 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 461 |   Recall that @{ML "|>"} is the reverse function application. Recall also that
 | 
| 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 462 | the related | 
| 126 | 463 |   reverse function composition is @{ML "#>"}. In fact all the combinators @{ML "|->"},
 | 
| 215 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 464 |   @{ML "|>>"} , @{ML "||>"} and @{ML "||>>"} described above have related combinators for 
 | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 465 |   function composition, namely @{ML "#->"}, @{ML "#>>"}, @{ML "##>"} and @{ML "##>>"}. 
 | 
| 
8d1a344a621e
more work on the inductive package
 Christian Urban <urbanc@in.tum.de> parents: 
210diff
changeset | 466 |   Using @{ML "#->"}, for example, the function @{text double} can also be written as:
 | 
| 126 | 467 | *} | 
| 468 | ||
| 469 | ML{*val double =
 | |
| 470 | (fn x => (x, x)) | |
| 471 | #-> (fn x => fn y => x + y)*} | |
| 472 | ||
| 473 | text {*
 | |
| 474 | ||
| 475 | (FIXME: find a good exercise for combinators) | |
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 476 | |
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 477 |   \begin{readmore}
 | 
| 196 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 478 |   The most frequently used combinators are defined in the files @{ML_file
 | 
| 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 479 | "Pure/library.ML"} | 
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 480 |   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: 
126diff
changeset | 481 | contains further information about combinators. | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 482 |   \end{readmore}
 | 
| 126 | 483 | |
| 15 
9da9ba2b095b
added a solution section and some other minor additions
 Christian Urban <urbanc@in.tum.de> parents: 
14diff
changeset | 484 | *} | 
| 
9da9ba2b095b
added a solution section and some other minor additions
 Christian Urban <urbanc@in.tum.de> parents: 
14diff
changeset | 485 | |
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 486 | |
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 487 | section {* Antiquotations *}
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 488 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 489 | text {*
 | 
| 49 | 490 | The main advantage of embedding all code in a theory is that the code can | 
| 58 | 491 | contain references to entities defined on the logical level of Isabelle. By | 
| 492 | this we mean definitions, theorems, terms and so on. This kind of reference is | |
| 493 | realised with antiquotations. For example, one can print out the name of the current | |
| 49 | 494 | theory by typing | 
| 495 | ||
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
34diff
changeset | 496 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
71diff
changeset | 497 |   @{ML_response [display,gray] "Context.theory_name @{theory}" "\"FirstSteps\""}
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
34diff
changeset | 498 | |
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 499 |   where @{text "@{theory}"} is an antiquotation that is substituted with the
 | 
| 49 | 500 | current theory (remember that we assumed we are inside the theory | 
| 89 | 501 |   @{text FirstSteps}). The name of this theory can be extracted using
 | 
| 49 | 502 |   the function @{ML "Context.theory_name"}. 
 | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 503 | |
| 89 | 504 | Note, however, that antiquotations are statically linked, that is their value is | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 505 | determined at ``compile-time'', not ``run-time''. For example the function | 
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 506 | *} | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 507 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 508 | ML{*fun not_current_thyname () = Context.theory_name @{theory} *}
 | 
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 509 | |
| 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 510 | text {*
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 511 | |
| 89 | 512 |   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: 
2diff
changeset | 513 | different theory. Instead, the code above defines the constant function | 
| 58 | 514 |   that always returns the string @{text [quotes] "FirstSteps"}, no matter where the
 | 
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 515 |   function is called. Operationally speaking,  the antiquotation @{text "@{theory}"} is 
 | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 516 |   \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: 
2diff
changeset | 517 | some data structure and return it. Instead, it is literally | 
| 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 518 | replaced with the value representing the theory name. | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 519 | |
| 132 | 520 | In a similar way you can use antiquotations to refer to proved theorems: | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 521 |   @{text "@{thm \<dots>}"} for a single theorem
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
34diff
changeset | 522 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
71diff
changeset | 523 |   @{ML_response_fake [display,gray] "@{thm allI}" "(\<And>x. ?P x) \<Longrightarrow> \<forall>x. ?P x"}
 | 
| 75 | 524 | |
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 525 |   and @{text "@{thms \<dots>}"} for more than one
 | 
| 132 | 526 | |
| 527 | @{ML_response_fake [display,gray] "@{thms conj_ac}" 
 | |
| 528 | "(?P \<and> ?Q) = (?Q \<and> ?P) | |
| 529 | (?P \<and> ?Q \<and> ?R) = (?Q \<and> ?P \<and> ?R) | |
| 530 | ((?P \<and> ?Q) \<and> ?R) = (?P \<and> ?Q \<and> ?R)"} | |
| 531 | ||
| 149 | 532 | You can also refer to the current simpset. To illustrate this we implement the | 
| 132 | 533 | function that extracts the theorem names stored in a simpset. | 
| 131 | 534 | *} | 
| 75 | 535 | |
| 149 | 536 | ML{*fun get_thm_names_from_ss simpset =
 | 
| 131 | 537 | let | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 538 |   val {simps,...} = MetaSimplifier.dest_ss simpset
 | 
| 70 
bbb2d5f1d58d
deleted the fixme about simpsets
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 539 | 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: 
162diff
changeset | 540 | map #1 simps | 
| 131 | 541 | end*} | 
| 54 
1783211b3494
tuned; added document antiquotation ML_response_fake_both
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 542 | |
| 131 | 543 | text {*
 | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 544 |   The function @{ML dest_ss in MetaSimplifier} returns a record containing all
 | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 545 | information stored in the simpset, but we are only interested in the names of the | 
| 184 | 546 | simp-rules. So now you can feed in the current simpset into this function. | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 547 |   The current simpset can be referred to using the antiquotation @{text "@{simpset}"}.
 | 
| 81 | 548 | |
| 131 | 549 |   @{ML_response_fake [display,gray] 
 | 
| 149 | 550 |   "get_thm_names_from_ss @{simpset}" 
 | 
| 551 | "[\"Nat.of_nat_eq_id\", \"Int.of_int_eq_id\", \"Nat.One_nat_def\", \<dots>]"} | |
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 552 | |
| 196 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 553 | Again, this way of referencing simpsets makes you independent from additions | 
| 156 | 554 | of lemmas to the simpset by the user that potentially cause loops. | 
| 555 | ||
| 192 | 556 | On the ML-level of Isabelle, you often have to work with qualified names; | 
| 196 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 557 | these are strings with some additional information, such as positional information | 
| 192 | 558 | and qualifiers. Such bindings can be generated with the antiquotation | 
| 196 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 559 |   @{text "@{binding \<dots>}"}.
 | 
| 192 | 560 | |
| 561 |   @{ML_response [display,gray]
 | |
| 562 |   "@{binding \"name\"}"
 | |
| 563 | "name"} | |
| 564 | ||
| 196 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 565 |   An example where a binding is needed is the function @{ML define in
 | 
| 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 566 |   LocalTheory}.  Below, this function is used to define the constant @{term
 | 
| 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 567 | "TrueConj"} as the conjunction | 
| 192 | 568 |   @{term "True \<and> True"}.
 | 
| 569 | *} | |
| 570 | ||
| 571 | local_setup %gray {* 
 | |
| 572 | snd o LocalTheory.define Thm.internalK | |
| 573 |     ((@{binding "TrueConj"}, NoSyn), 
 | |
| 574 |      (Attrib.empty_binding, @{term "True \<and> True"})) *}
 | |
| 575 | ||
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 576 | text {* 
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 577 | Now querying the definition you obtain: | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 578 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 579 |   \begin{isabelle}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 580 |   \isacommand{thm}~@{text "TrueConj_def"}\\
 | 
| 225 
7859fc59249a
section for further material about simple inductive
 Christian Urban <urbanc@in.tum.de> parents: 
215diff
changeset | 581 |   @{text "> "}~@{thm TrueConj_def}
 | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 582 |   \end{isabelle}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 583 | |
| 194 | 584 | (FIXME give a better example why bindings are important; maybe | 
| 585 |   give a pointer to \isacommand{local\_setup})
 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 586 | |
| 196 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 587 | While antiquotations have many applications, they were originally introduced | 
| 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 588 | in order to avoid explicit bindings of theorems such as: | 
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 589 | *} | 
| 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 590 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 591 | ML{*val allI = thm "allI" *}
 | 
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 592 | |
| 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 593 | text {*
 | 
| 196 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 594 | Such bindings are difficult to maintain and can be overwritten by the | 
| 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 595 | user accidentally. This often broke Isabelle | 
| 49 | 596 | packages. Antiquotations solve this problem, since they are ``linked'' | 
| 89 | 597 | statically at compile-time. However, this static linkage also limits their | 
| 201 | 598 | usefulness in cases where data needs to be built up dynamically. In the | 
| 196 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 599 | course of this chapter you will learn more about antiquotations: | 
| 122 | 600 | they can simplify Isabelle programming since one can directly access all | 
| 196 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 601 | kinds of logical elements from the ML-level. | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 602 | *} | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 603 | |
| 15 
9da9ba2b095b
added a solution section and some other minor additions
 Christian Urban <urbanc@in.tum.de> parents: 
14diff
changeset | 604 | section {* Terms and Types *}
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 605 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 606 | text {*
 | 
| 197 | 607 | One way to construct Isabelle terms, is by using the antiquotation | 
| 89 | 608 |   \mbox{@{text "@{term \<dots>}"}}. For example:
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 609 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
71diff
changeset | 610 |   @{ML_response [display,gray] 
 | 
| 75 | 611 | "@{term \"(a::nat) + b = c\"}" 
 | 
| 612 | "Const (\"op =\", \<dots>) $ | |
| 149 | 613 | (Const (\"HOL.plus_class.plus\", \<dots>) $ \<dots> $ \<dots>) $ \<dots>"} | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 614 | |
| 207 | 615 |   will show the term @{term "(a::nat) + b = c"}, but printed using the internal
 | 
| 198 | 616 |   representation corresponding to the data type @{ML_type "term"}.
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 617 | |
| 197 | 618 | This internal representation uses the usual de Bruijn index mechanism---where | 
| 619 |   bound variables are represented by the constructor @{ML Bound}.  The index in
 | |
| 620 |   @{ML Bound} refers to the number of Abstractions (@{ML Abs}) we have to skip
 | |
| 621 |   until we hit the @{ML Abs} that binds the corresponding variable. Note that
 | |
| 622 | the names of bound variables are kept at abstractions for printing purposes, | |
| 623 | and so should be treated only as ``comments''. Application in Isabelle is | |
| 624 |   realised with the term-constructor @{ML $}.
 | |
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 625 | |
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 626 |   \begin{readmore}
 | 
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 627 |   Terms are described in detail in \isccite{sec:terms}. Their
 | 
| 78 
ef778679d3e0
added a section about combinators
 Christian Urban <urbanc@in.tum.de> parents: 
75diff
changeset | 628 |   definition and many useful operations are implemented in @{ML_file "Pure/term.ML"}.
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 629 |   \end{readmore}
 | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 630 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 631 | Constructing terms via antiquotations has the advantage that only typable | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 632 | terms can be constructed. For example | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 633 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 634 |   @{ML_response_fake_both [display,gray]
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 635 |   "@{term \"(x::nat) x\"}"
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 636 | "Type unification failed \<dots>"} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 637 | |
| 194 | 638 | raises a typing error, while it perfectly ok to construct the term | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 639 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 640 |   @{ML [display,gray] "Free (\"x\", @{typ nat}) $ Free (\"x\", @{typ nat})"}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 641 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 642 | with the raw ML-constructors. | 
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 643 | Sometimes the internal representation of terms can be surprisingly different | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 644 | from what you see at the user-level, because the layers of | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 645 | parsing/type-checking/pretty printing can be quite elaborate. | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 646 | |
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 647 |   \begin{exercise}
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 648 | Look at the internal term representation of the following terms, and | 
| 89 | 649 | find out why they are represented like this: | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 650 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 651 |   \begin{itemize}
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 652 |   \item @{term "case x of 0 \<Rightarrow> 0 | Suc y \<Rightarrow> y"}  
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 653 |   \item @{term "\<lambda>(x,y). P y x"}  
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 654 |   \item @{term "{ [x::int] | x. x \<le> -2 }"}  
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 655 |   \end{itemize}
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 656 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 657 | Hint: The third term is already quite big, and the pretty printer | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 658 | may omit parts of it by default. If you want to see all of it, you | 
| 122 | 659 | can use the following ML-function to set the printing depth to a higher | 
| 660 | value: | |
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 661 | |
| 75 | 662 |   @{ML [display,gray] "print_depth 50"}
 | 
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
71diff
changeset | 663 |   \end{exercise}
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
34diff
changeset | 664 | |
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 665 |   The antiquotation @{text "@{prop \<dots>}"} constructs terms of propositional type, 
 | 
| 50 | 666 |   inserting the invisible @{text "Trueprop"}-coercions whenever necessary. 
 | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
66diff
changeset | 667 | Consider for example the pairs | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 668 | |
| 126 | 669 | @{ML_response [display,gray] "(@{term \"P x\"}, @{prop \"P x\"})" 
 | 
| 670 | "(Free (\"P\", \<dots>) $ Free (\"x\", \<dots>), | |
| 149 | 671 | Const (\"Trueprop\", \<dots>) $ (Free (\"P\", \<dots>) $ Free (\"x\", \<dots>)))"} | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 672 | |
| 108 
8bea3f74889d
added to the tactical chapter; polished; added the tabularstar environment (which is just tabular*)
 Christian Urban <urbanc@in.tum.de> parents: 
107diff
changeset | 673 | where a coercion is inserted in the second component and | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 674 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
71diff
changeset | 675 |   @{ML_response [display,gray] "(@{term \"P x \<Longrightarrow> Q x\"}, @{prop \"P x \<Longrightarrow> Q x\"})" 
 | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 676 | "(Const (\"==>\", \<dots>) $ \<dots> $ \<dots>, Const (\"==>\", \<dots>) $ \<dots> $ \<dots>)"} | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 677 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
71diff
changeset | 678 | where it is not (since it is already constructed by a meta-implication). | 
| 19 
34b93dbf8c3c
some tuning in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
15diff
changeset | 679 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 680 | As already seen above, types can be constructed using the antiquotation | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 681 |   @{text "@{typ \<dots>}"}. For example:
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
34diff
changeset | 682 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
71diff
changeset | 683 |   @{ML_response_fake [display,gray] "@{typ \"bool \<Rightarrow> nat\"}" "bool \<Rightarrow> nat"}
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
34diff
changeset | 684 | |
| 19 
34b93dbf8c3c
some tuning in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
15diff
changeset | 685 |   \begin{readmore}
 | 
| 
34b93dbf8c3c
some tuning in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
15diff
changeset | 686 |   Types are described in detail in \isccite{sec:types}. Their
 | 
| 78 
ef778679d3e0
added a section about combinators
 Christian Urban <urbanc@in.tum.de> parents: 
75diff
changeset | 687 | definition and many useful operations are implemented | 
| 
ef778679d3e0
added a section about combinators
 Christian Urban <urbanc@in.tum.de> parents: 
75diff
changeset | 688 |   in @{ML_file "Pure/type.ML"}.
 | 
| 19 
34b93dbf8c3c
some tuning in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
15diff
changeset | 689 |   \end{readmore}
 | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 690 | *} | 
| 19 
34b93dbf8c3c
some tuning in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
15diff
changeset | 691 | |
| 
34b93dbf8c3c
some tuning in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
15diff
changeset | 692 | |
| 156 | 693 | section {* Constructing Terms and Types Manually\label{sec:terms_types_manually} *} 
 | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 694 | |
| 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 695 | text {*
 | 
| 81 | 696 | While antiquotations are very convenient for constructing terms, they can | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 697 | only construct fixed terms (remember they are ``linked'' at compile-time). | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 698 | However, you often need to construct terms dynamically. For example, a | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 699 |   function that returns the implication @{text "\<And>(x::nat). P x \<Longrightarrow> Q x"} taking
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 700 |   @{term P} and @{term Q} as arguments can only be written as:
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 701 | |
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 702 | *} | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 703 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 704 | ML{*fun make_imp P Q =
 | 
| 131 | 705 | let | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 706 |   val x = Free ("x", @{typ nat})
 | 
| 131 | 707 | in | 
| 708 | Logic.all x (Logic.mk_implies (P $ x, Q $ x)) | |
| 709 | end *} | |
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 710 | |
| 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 711 | text {*
 | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 712 |   The reason is that you cannot pass the arguments @{term P} and @{term Q} 
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 713 |   into an antiquotation. For example the following does \emph{not} work.
 | 
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 714 | *} | 
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 715 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 716 | ML{*fun make_wrong_imp P Q = @{prop "\<And>(x::nat). P x \<Longrightarrow> Q x"} *}
 | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 717 | |
| 43 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 718 | text {*
 | 
| 194 | 719 |   To see this, apply @{text "@{term S}"} and @{text "@{term T}"} 
 | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 720 |   to both functions. With @{ML make_imp} you obtain the intended term involving 
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 721 | the given arguments | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 722 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 723 |   @{ML_response [display,gray] "make_imp @{term S} @{term T}" 
 | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 724 | "Const \<dots> $ | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 725 | Abs (\"x\", Type (\"nat\",[]), | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 726 | Const \<dots> $ (Free (\"S\",\<dots>) $ \<dots>) $ (Free (\"T\",\<dots>) $ \<dots>))"} | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
66diff
changeset | 727 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 728 |   whereas with @{ML make_wrong_imp} you obtain a term involving the @{term "P"} 
 | 
| 68 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
66diff
changeset | 729 |   and @{text "Q"} from the antiquotation.
 | 
| 
e7519207c2b7
added more to the "new command section" and tuning
 Christian Urban <urbanc@in.tum.de> parents: 
66diff
changeset | 730 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 731 |   @{ML_response [display,gray] "make_wrong_imp @{term S} @{term T}" 
 | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 732 | "Const \<dots> $ | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 733 | Abs (\"x\", \<dots>, | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 734 | Const \<dots> $ (Const \<dots> $ (Free (\"P\",\<dots>) $ \<dots>)) $ | 
| 192 | 735 | (Const \<dots> $ (Free (\"Q\",\<dots>) $ \<dots>)))"} | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 736 | |
| 192 | 737 | There are a number of handy functions that are frequently used for | 
| 204 | 738 |   constructing terms. One is the function @{ML list_comb}, which takes a term
 | 
| 199 | 739 | and a list of terms as arguments, and produces as output the term | 
| 192 | 740 | list applied to the term. For example | 
| 741 | ||
| 742 | @{ML_response_fake [display,gray]
 | |
| 743 | "list_comb (@{term \"P::nat\"}, [@{term \"True\"}, @{term \"False\"}])"
 | |
| 744 | "Free (\"P\", \"nat\") $ Const (\"True\", \"bool\") $ Const (\"False\", \"bool\")"} | |
| 745 | ||
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 746 |   Another handy function is @{ML lambda}, which abstracts a variable 
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 747 | in a term. For example | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 748 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 749 |   @{ML_response_fake [display,gray]
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 750 |   "lambda @{term \"x::nat\"} @{term \"(P::nat\<Rightarrow>bool) x\"}"
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 751 | "Abs (\"x\", \"nat\", Free (\"P\", \"bool \<Rightarrow> bool\") $ Bound 0)"} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 752 | |
| 228 | 753 |   In this example, @{ML lambda} produces a de Bruijn index (i.e.~@{ML "Bound  0"}), 
 | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 754 | and an abstraction. It also records the type of the abstracted | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 755 | variable and for printing purposes also its name. Note that because of the | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 756 |   typing annotation on @{text "P"}, the variable @{text "x"} in @{text "P x"}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 757 | is of the same type as the abstracted variable. If it is of different type, | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 758 | as in | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 759 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 760 |   @{ML_response_fake [display,gray]
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 761 |   "lambda @{term \"x::nat\"} @{term \"(P::bool\<Rightarrow>bool) x\"}"
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 762 | "Abs (\"x\", \"nat\", Free (\"P\", \"bool \<Rightarrow> bool\") $ Free (\"x\", \"bool\"))"} | 
| 49 | 763 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 764 |   then the variable @{text "Free (\"x\", \"bool\")"} is \emph{not} abstracted. 
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 765 | This is a fundamental principle | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 766 | of Church-style typing, where variables with the same name still differ, if they | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 767 | have different type. | 
| 192 | 768 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 769 |   There is also the function @{ML subst_free} with which terms can 
 | 
| 194 | 770 | be replaced by other terms. For example below, we will replace in | 
| 771 |   @{term "(f::nat\<Rightarrow>nat\<Rightarrow>nat) 0 x"} 
 | |
| 772 |   the subterm @{term "(f::nat\<Rightarrow>nat\<Rightarrow>nat) 0"} by @{term y}, and @{term x} by @{term True}.
 | |
| 49 | 773 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 774 |   @{ML_response_fake [display,gray]
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 775 | "subst_free [(@{term \"(f::nat\<Rightarrow>nat\<Rightarrow>nat) 0\"}, @{term \"y::nat\<Rightarrow>nat\"}),
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 776 |             (@{term \"x::nat\"}, @{term \"True\"})] 
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 777 |   @{term \"((f::nat\<Rightarrow>nat\<Rightarrow>nat) 0) x\"}"
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 778 | "Free (\"y\", \"nat \<Rightarrow> nat\") $ Const (\"True\", \"bool\")"} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 779 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 780 |   As can be seen, @{ML subst_free} does not take typability into account.
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 781 | However it takes alpha-equivalence into account: | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 782 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 783 |   @{ML_response_fake [display, gray]
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 784 |   "subst_free [(@{term \"(\<lambda>y::nat. y)\"}, @{term \"x::nat\"})] 
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 785 |   @{term \"(\<lambda>x::nat. x)\"}"
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 786 | "Free (\"x\", \"nat\")"} | 
| 192 | 787 | |
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 788 |   \begin{readmore}
 | 
| 89 | 789 |   There are many functions in @{ML_file "Pure/term.ML"}, @{ML_file "Pure/logic.ML"} and
 | 
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 790 |   @{ML_file "HOL/Tools/hologic.ML"} that make such manual constructions of terms 
 | 
| 49 | 791 |   and types easier.\end{readmore}
 | 
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 792 | |
| 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 793 | Have a look at these files and try to solve the following two exercises: | 
| 11 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 794 | |
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 795 |   \begin{exercise}\label{fun:revsum}
 | 
| 58 | 796 |   Write a function @{text "rev_sum : term -> term"} that takes a
 | 
| 122 | 797 |   term of the form @{text "t\<^isub>1 + t\<^isub>2 + \<dots> + t\<^isub>n"} (whereby @{text "n"} might be zero)
 | 
| 11 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 798 |   and returns the reversed sum @{text "t\<^isub>n + \<dots> + t\<^isub>2 + t\<^isub>1"}. Assume
 | 
| 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 799 |   the @{text "t\<^isub>i"} can be arbitrary expressions and also note that @{text "+"} 
 | 
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 800 | associates to the left. Try your function on some examples. | 
| 11 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 801 |   \end{exercise}
 | 
| 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 802 | |
| 15 
9da9ba2b095b
added a solution section and some other minor additions
 Christian Urban <urbanc@in.tum.de> parents: 
14diff
changeset | 803 |   \begin{exercise}\label{fun:makesum}
 | 
| 11 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 804 | Write a function which takes two terms representing natural numbers | 
| 199 | 805 |   in unary notation (like @{term "Suc (Suc (Suc 0))"}), and produces the
 | 
| 11 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 806 | number representing their sum. | 
| 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 807 |   \end{exercise}
 | 
| 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 808 | |
| 122 | 809 | There are a few subtle issues with constants. They usually crop up when | 
| 149 | 810 | pattern matching terms or types, or when constructing them. While it is perfectly ok | 
| 122 | 811 |   to write the function @{text is_true} as follows
 | 
| 812 | *} | |
| 813 | ||
| 814 | ML{*fun is_true @{term True} = true
 | |
| 815 | | is_true _ = false*} | |
| 816 | ||
| 817 | text {*
 | |
| 818 |   this does not work for picking out @{text "\<forall>"}-quantified terms. Because
 | |
| 819 | the function | |
| 820 | *} | |
| 821 | ||
| 822 | ML{*fun is_all (@{term All} $ _) = true
 | |
| 823 | | is_all _ = false*} | |
| 824 | ||
| 825 | text {* 
 | |
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 826 |   will not correctly match the formula @{prop "\<forall>x::nat. P x"}: 
 | 
| 122 | 827 | |
| 828 |   @{ML_response [display,gray] "is_all @{term \"\<forall>x::nat. P x\"}" "false"}
 | |
| 829 | ||
| 830 |   The problem is that the @{text "@term"}-antiquotation in the pattern 
 | |
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 831 |   fixes the type of the constant @{term "All"} to be @{typ "('a \<Rightarrow> bool) \<Rightarrow> bool"} for 
 | 
| 122 | 832 |   an arbitrary, but fixed type @{typ "'a"}. A properly working alternative 
 | 
| 833 | for this function is | |
| 834 | *} | |
| 835 | ||
| 836 | ML{*fun is_all (Const ("All", _) $ _) = true
 | |
| 837 | | is_all _ = false*} | |
| 838 | ||
| 839 | text {* 
 | |
| 840 | because now | |
| 841 | ||
| 842 | @{ML_response [display,gray] "is_all @{term \"\<forall>x::nat. P x\"}" "true"}
 | |
| 843 | ||
| 149 | 844 | matches correctly (the first wildcard in the pattern matches any type and the | 
| 845 | second any term). | |
| 122 | 846 | |
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 847 | However there is still a problem: consider the similar function that | 
| 131 | 848 |   attempts to pick out @{text "Nil"}-terms:
 | 
| 122 | 849 | *} | 
| 850 | ||
| 851 | ML{*fun is_nil (Const ("Nil", _)) = true
 | |
| 852 | | is_nil _ = false *} | |
| 853 | ||
| 854 | text {* 
 | |
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 855 |   Unfortunately, also this function does \emph{not} work as expected, since
 | 
| 122 | 856 | |
| 857 |   @{ML_response [display,gray] "is_nil @{term \"Nil\"}" "false"}
 | |
| 858 | ||
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 859 | The problem is that on the ML-level the name of a constant is more | 
| 149 | 860 |   subtle than you might expect. The function @{ML is_all} worked correctly,
 | 
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 861 |   because @{term "All"} is such a fundamental constant, which can be referenced
 | 
| 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 862 |   by @{ML "Const (\"All\", some_type)" for some_type}. However, if you look at
 | 
| 122 | 863 | |
| 864 |   @{ML_response [display,gray] "@{term \"Nil\"}" "Const (\"List.list.Nil\", \<dots>)"}
 | |
| 865 | ||
| 131 | 866 |   the name of the constant @{text "Nil"} depends on the theory in which the
 | 
| 198 | 867 |   term constructor is defined (@{text "List"}) and also in which data type
 | 
| 128 | 868 |   (@{text "list"}). Even worse, some constants have a name involving
 | 
| 869 |   type-classes. Consider for example the constants for @{term "zero"} and
 | |
| 131 | 870 |   \mbox{@{text "(op *)"}}:
 | 
| 122 | 871 | |
| 872 |   @{ML_response [display,gray] "(@{term \"0::nat\"}, @{term \"op *\"})" 
 | |
| 126 | 873 | "(Const (\"HOL.zero_class.zero\", \<dots>), | 
| 122 | 874 | Const (\"HOL.times_class.times\", \<dots>))"} | 
| 875 | ||
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 876 | While you could use the complete name, for example | 
| 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 877 |   @{ML "Const (\"List.list.Nil\", some_type)" for some_type}, for referring to or
 | 
| 122 | 878 |   matching against @{text "Nil"}, this would make the code rather brittle. 
 | 
| 198 | 879 | The reason is that the theory and the name of the data type can easily change. | 
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 880 | To make the code more robust, it is better to use the antiquotation | 
| 122 | 881 |   @{text "@{const_name \<dots>}"}. With this antiquotation you can harness the 
 | 
| 200 | 882 | variable parts of the constant's name. Therefore a function for | 
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 883 | matching against constants that have a polymorphic type should | 
| 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 884 | be written as follows. | 
| 122 | 885 | *} | 
| 886 | ||
| 887 | ML{*fun is_nil_or_all (Const (@{const_name "Nil"}, _)) = true
 | |
| 888 |   | is_nil_or_all (Const (@{const_name "All"}, _) $ _) = true
 | |
| 889 | | is_nil_or_all _ = false *} | |
| 890 | ||
| 891 | text {*
 | |
| 242 
14d5f0cf91dc
added note about @{type_names ...}
 Christian Urban <urbanc@in.tum.de> parents: 
240diff
changeset | 892 |   The antiquotation for properly referencing type constants is  is @{text "@{type_name \<dots>}"}.
 | 
| 
14d5f0cf91dc
added note about @{type_names ...}
 Christian Urban <urbanc@in.tum.de> parents: 
240diff
changeset | 893 | For example | 
| 
14d5f0cf91dc
added note about @{type_names ...}
 Christian Urban <urbanc@in.tum.de> parents: 
240diff
changeset | 894 | |
| 
14d5f0cf91dc
added note about @{type_names ...}
 Christian Urban <urbanc@in.tum.de> parents: 
240diff
changeset | 895 |   @{ML_response [display,gray]
 | 
| 
14d5f0cf91dc
added note about @{type_names ...}
 Christian Urban <urbanc@in.tum.de> parents: 
240diff
changeset | 896 |   "@{type_name \"list\"}" "\"List.list\""}
 | 
| 
14d5f0cf91dc
added note about @{type_names ...}
 Christian Urban <urbanc@in.tum.de> parents: 
240diff
changeset | 897 | |
| 200 | 898 | Occasionally you have to calculate what the ``base'' name of a given | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 899 |   constant is. For this you can use the function @{ML Sign.extern_const} or
 | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 900 |   @{ML Long_Name.base_name}. For example:
 | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 901 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 902 |   @{ML_response [display,gray] "Sign.extern_const @{theory} \"List.list.Nil\"" "\"Nil\""}
 | 
| 122 | 903 | |
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 904 |   The difference between both functions is that @{ML extern_const in Sign} returns
 | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 905 |   the smallest name that is still unique, whereas @{ML base_name in Long_Name} always
 | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 906 | strips off all qualifiers. | 
| 122 | 907 | |
| 908 |   \begin{readmore}
 | |
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 909 |   Functions about naming are implemented in @{ML_file "Pure/General/name_space.ML"};
 | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 910 |   functions about signatures in @{ML_file "Pure/sign.ML"}.
 | 
| 122 | 911 |   \end{readmore}
 | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 912 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 913 | Although types of terms can often be inferred, there are many | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 914 | situations where you need to construct types manually, especially | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 915 | when defining constants. For example the function returning a function | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 916 | type is as follows: | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 917 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 918 | *} | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 919 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 920 | ML{*fun make_fun_type tau1 tau2 = Type ("fun", [tau1, tau2]) *}
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 921 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 922 | text {* This can be equally written with the combinator @{ML "-->"} as: *}
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 923 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 924 | ML{*fun make_fun_type tau1 tau2 = tau1 --> tau2 *}
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 925 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 926 | text {*
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 927 |   A handy function for manipulating terms is @{ML map_types}: it takes a 
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 928 | function and applies it to every type in a term. You can, for example, | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 929 |   change every @{typ nat} in a term into an @{typ int} using the function:
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 930 | *} | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 931 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 932 | ML{*fun nat_to_int t =
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 933 | (case t of | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 934 |      @{typ nat} => @{typ int}
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 935 | | Type (s, ts) => Type (s, map nat_to_int ts) | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 936 | | _ => t)*} | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 937 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 938 | text {*
 | 
| 200 | 939 | Here is an example: | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 940 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 941 | @{ML_response_fake [display,gray] 
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 942 | "map_types nat_to_int @{term \"a = (1::nat)\"}" 
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 943 | "Const (\"op =\", \"int \<Rightarrow> int \<Rightarrow> bool\") | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 944 | $ Free (\"a\", \"int\") $ Const (\"HOL.one_class.one\", \"int\")"} | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 945 | |
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 946 | (FIXME: a readmore about types) | 
| 122 | 947 | *} | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 948 | |
| 122 | 949 | |
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 950 | section {* Type-Checking *}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 951 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 952 | text {* 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 953 | |
| 131 | 954 |   You can freely construct and manipulate @{ML_type "term"}s and @{ML_type
 | 
| 955 | typ}es, since they are just arbitrary unchecked trees. However, you | |
| 956 | eventually want to see if a term is well-formed, or type-checks, relative to | |
| 957 |   a theory.  Type-checking is done via the function @{ML cterm_of}, which
 | |
| 958 |   converts a @{ML_type term} into a @{ML_type cterm}, a \emph{certified}
 | |
| 959 |   term. Unlike @{ML_type term}s, which are just trees, @{ML_type "cterm"}s are
 | |
| 960 | abstract objects that are guaranteed to be type-correct, and they can only | |
| 961 | be constructed via ``official interfaces''. | |
| 962 | ||
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 963 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 964 | Type-checking is always relative to a theory context. For now we use | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 965 |   the @{ML "@{theory}"} antiquotation to get hold of the current theory.
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 966 | For example you can write: | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 967 | |
| 149 | 968 |   @{ML_response_fake [display,gray] "cterm_of @{theory} @{term \"(a::nat) + b = c\"}" "a + b = c"}
 | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 969 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 970 | This can also be written with an antiquotation: | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 971 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 972 |   @{ML_response_fake [display,gray] "@{cterm \"(a::nat) + b = c\"}" "a + b = c"}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 973 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 974 | Attempting to obtain the certified term for | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 975 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 976 |   @{ML_response_fake_both [display,gray] "@{cterm \"1 + True\"}" "Type unification failed \<dots>"}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 977 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 978 | yields an error (since the term is not typable). A slightly more elaborate | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 979 | example that type-checks is: | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 980 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 981 | @{ML_response_fake [display,gray] 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 982 | "let | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 983 |   val natT = @{typ \"nat\"}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 984 |   val zero = @{term \"0::nat\"}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 985 | in | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 986 |   cterm_of @{theory} 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 987 |       (Const (@{const_name plus}, natT --> natT --> natT) $ zero $ zero)
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 988 | end" "0 + 0"} | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 989 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 990 | In Isabelle not just terms need to be certified, but also types. For example, | 
| 198 | 991 |   you obtain the certified type for the Isabelle type @{typ "nat \<Rightarrow> bool"} on 
 | 
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 992 | the ML-level as follows: | 
| 160 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 993 | |
| 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 994 |   @{ML_response_fake [display,gray]
 | 
| 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 995 |   "ctyp_of @{theory} (@{typ nat} --> @{typ bool})"
 | 
| 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 996 | "nat \<Rightarrow> bool"} | 
| 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 997 | |
| 207 | 998 | or with the antiquotation: | 
| 999 | ||
| 1000 |   @{ML_response_fake [display,gray]
 | |
| 1001 |   "@{ctyp \"nat \<Rightarrow> bool\"}"
 | |
| 1002 | "nat \<Rightarrow> bool"} | |
| 1003 | ||
| 160 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 1004 |   \begin{readmore}
 | 
| 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 1005 |   For functions related to @{ML_type cterm}s and @{ML_type ctyp}s see 
 | 
| 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 1006 |   the file @{ML_file "Pure/thm.ML"}.
 | 
| 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 1007 |   \end{readmore}
 | 
| 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 1008 | |
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1009 |   \begin{exercise}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1010 |   Check that the function defined in Exercise~\ref{fun:revsum} returns a
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1011 | result that type-checks. | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1012 |   \end{exercise}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1013 | |
| 200 | 1014 | Remember Isabelle follows the Church-style typing for terms, i.e., a term contains | 
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1015 | enough typing information (constants, free variables and abstractions all have typing | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1016 | information) so that it is always clear what the type of a term is. | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1017 |   Given a well-typed term, the function @{ML type_of} returns the 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1018 | type of a term. Consider for example: | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1019 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1020 |   @{ML_response [display,gray] 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1021 |   "type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1022 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1023 | To calculate the type, this function traverses the whole term and will | 
| 197 | 1024 | detect any typing inconsistency. For example changing the type of the variable | 
| 149 | 1025 |   @{term "x"} from @{typ "nat"} to @{typ "int"} will result in the error message: 
 | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1026 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1027 |   @{ML_response_fake [display,gray] 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1028 |   "type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::int\"})" 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1029 | "*** Exception- TYPE (\"type_of: type mismatch in application\" \<dots>"} | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1030 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1031 | Since the complete traversal might sometimes be too costly and | 
| 149 | 1032 |   not necessary, there is the function @{ML fastype_of}, which 
 | 
| 1033 | also returns the type of a term. | |
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1034 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1035 |   @{ML_response [display,gray] 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1036 |   "fastype_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1037 | |
| 177 | 1038 | However, efficiency is gained on the expense of skipping some tests. You | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1039 | can see this in the following example | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1040 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1041 |    @{ML_response [display,gray] 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1042 |   "fastype_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::int\"})" "bool"}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1043 | |
| 149 | 1044 | where no error is detected. | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1045 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1046 | Sometimes it is a bit inconvenient to construct a term with | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1047 | complete typing annotations, especially in cases where the typing | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1048 | information is redundant. A short-cut is to use the ``place-holder'' | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1049 |   type @{ML "dummyT"} and then let type-inference figure out the 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1050 | complete type. An example is as follows: | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1051 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1052 |   @{ML_response_fake [display,gray] 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1053 | "let | 
| 126 | 1054 |   val c = Const (@{const_name \"plus\"}, dummyT) 
 | 
| 1055 |   val o = @{term \"1::nat\"} 
 | |
| 1056 | val v = Free (\"x\", dummyT) | |
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1057 | in | 
| 126 | 1058 |   Syntax.check_term @{context} (c $ o $ v)
 | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1059 | end" | 
| 126 | 1060 | "Const (\"HOL.plus_class.plus\", \"nat \<Rightarrow> nat \<Rightarrow> nat\") $ | 
| 1061 | Const (\"HOL.one_class.one\", \"nat\") $ Free (\"x\", \"nat\")"} | |
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1062 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1063 |   Instead of giving explicitly the type for the constant @{text "plus"} and the free 
 | 
| 200 | 1064 |   variable @{text "x"}, type-inference fills in the missing information.
 | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1065 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1066 |   \begin{readmore}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1067 |   See @{ML_file "Pure/Syntax/syntax.ML"} where more functions about reading,
 | 
| 200 | 1068 | checking and pretty-printing of terms are defined. Functions related to | 
| 1069 |   type-inference are implemented in @{ML_file "Pure/type.ML"} and 
 | |
| 160 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 1070 |   @{ML_file "Pure/type_infer.ML"}. 
 | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1071 |   \end{readmore}
 | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1072 | |
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 1073 | (FIXME: say something about sorts) | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1074 | *} | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1075 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1076 | |
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1077 | section {* Theorems *}
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1078 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1079 | text {*
 | 
| 50 | 1080 |   Just like @{ML_type cterm}s, theorems are abstract objects of type @{ML_type thm} 
 | 
| 201 | 1081 | that can only be built by going through interfaces. As a consequence, every proof | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1082 | in Isabelle is correct by construction. This follows the tradition of the LCF approach | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 1083 |   \cite{GordonMilnerWadsworth79}.
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 1084 | |
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1085 | |
| 78 
ef778679d3e0
added a section about combinators
 Christian Urban <urbanc@in.tum.de> parents: 
75diff
changeset | 1086 | To see theorems in ``action'', let us give a proof on the ML-level for the following | 
| 
ef778679d3e0
added a section about combinators
 Christian Urban <urbanc@in.tum.de> parents: 
75diff
changeset | 1087 | statement: | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1088 | *} | 
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1089 | |
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1090 | lemma | 
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1091 | assumes assm\<^isub>1: "\<And>(x::nat). P x \<Longrightarrow> Q x" | 
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1092 | and assm\<^isub>2: "P t" | 
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 1093 | shows "Q t" (*<*)oops(*>*) | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1094 | |
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1095 | text {*
 | 
| 185 | 1096 | The corresponding ML-code is as follows: | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1097 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
71diff
changeset | 1098 | @{ML_response_fake [display,gray]
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 1099 | "let | 
| 138 
e4d8dfb7e34a
included comment from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
136diff
changeset | 1100 |   val assm1 = @{cprop \"\<And>(x::nat). P x \<Longrightarrow> Q x\"}
 | 
| 
e4d8dfb7e34a
included comment from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
136diff
changeset | 1101 |   val assm2 = @{cprop \"(P::nat\<Rightarrow>bool) t\"}
 | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1102 | |
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1103 | val Pt_implies_Qt = | 
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1104 | assume assm1 | 
| 138 
e4d8dfb7e34a
included comment from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
136diff
changeset | 1105 |         |> forall_elim @{cterm \"t::nat\"};
 | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1106 | |
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1107 | val Qt = implies_elim Pt_implies_Qt (assume assm2); | 
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1108 | in | 
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1109 | Qt | 
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1110 | |> implies_intr assm2 | 
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1111 | |> implies_intr assm1 | 
| 48 
609f9ef73494
fixed FIXME's in fake responses
 Christian Urban <urbanc@in.tum.de> parents: 
47diff
changeset | 1112 | end" "\<lbrakk>\<And>x. P x \<Longrightarrow> Q x; P t\<rbrakk> \<Longrightarrow> Q t"} | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 1113 | |
| 21 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 1114 | This code-snippet constructs the following proof: | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 1115 | |
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 1116 | \[ | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 1117 |   \infer[(@{text "\<Longrightarrow>"}$-$intro)]{\vdash @{prop "(\<And>x. P x \<Longrightarrow> Q x) \<Longrightarrow> P t \<Longrightarrow> Q t"}}
 | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 1118 |     {\infer[(@{text "\<Longrightarrow>"}$-$intro)]{@{prop "\<And>x. P x \<Longrightarrow> Q x"} \vdash @{prop "P t \<Longrightarrow> Q t"}}
 | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 1119 |        {\infer[(@{text "\<Longrightarrow>"}$-$elim)]{@{prop "\<And>x. P x \<Longrightarrow> Q x"}, @{prop "P t"} \vdash @{prop "Q t"}}
 | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 1120 |           {\infer[(@{text "\<And>"}$-$elim)]{@{prop "\<And>x. P x \<Longrightarrow> Q x"} \vdash @{prop "P t \<Longrightarrow> Q t"}}
 | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 1121 |                  {\infer[(assume)]{@{prop "\<And>x. P x \<Longrightarrow> Q x"} \vdash @{prop "\<And>x. P x \<Longrightarrow> Q x"}}{}}
 | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 1122 | & | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 1123 |            \infer[(assume)]{@{prop "P t"} \vdash @{prop "P t"}}{}
 | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 1124 | } | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 1125 | } | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 1126 | } | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 1127 | \] | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 1128 | |
| 102 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 1129 | However, while we obtained a theorem as result, this theorem is not | 
| 
5e309df58557
general cleaning up; deleted antiquotation ML_text; adjusted pathnames of various files in the distribution
 Christian Urban <urbanc@in.tum.de> parents: 
101diff
changeset | 1130 | yet stored in Isabelle's theorem database. So it cannot be referenced later | 
| 128 | 1131 |   on. How to store theorems will be explained in Section~\ref{sec:storing}.
 | 
| 21 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 1132 | |
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 1133 |   \begin{readmore}
 | 
| 50 | 1134 |   For the functions @{text "assume"}, @{text "forall_elim"} etc 
 | 
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 1135 |   see \isccite{sec:thms}. The basic functions for theorems are defined in
 | 
| 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 1136 |   @{ML_file "Pure/thm.ML"}. 
 | 
| 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 1137 |   \end{readmore}
 | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 1138 | |
| 207 | 1139 |   (FIXME: handy functions working on theorems, like @{ML ObjectLogic.rulify} and so on) 
 | 
| 1140 | ||
| 230 
8def50824320
added material about OuterKeyword.keyword and OuterParse.reserved
 Christian Urban <urbanc@in.tum.de> parents: 
229diff
changeset | 1141 | (FIXME: how to add case-names to goal states - maybe in the | 
| 207 | 1142 | next section) | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1143 | *} | 
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 1144 | |
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 1145 | section {* Theorem Attributes *}
 | 
| 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 1146 | |
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 1147 | text {*
 | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1148 |   Theorem attributes are @{text "[symmetric]"}, @{text "[THEN \<dots>]"}, @{text
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1149 |   "[simp]"} and so on. Such attributes are \emph{neither} tags \emph{nor} flags
 | 
| 136 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1150 | annotated to theorems, but functions that do further processing once a | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1151 | theorem is proved. In particular, it is not possible to find out | 
| 136 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1152 | what are all theorems that have a given attribute in common, unless of course | 
| 197 | 1153 | the function behind the attribute stores the theorems in a retrievable | 
| 1154 | data structure. | |
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 1155 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1156 | If you want to print out all currently known attributes a theorem can have, | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1157 | you can use the Isabelle command | 
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 1158 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1159 |   \begin{isabelle}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1160 |   \isacommand{print\_attributes}\\
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1161 |   @{text "> COMP:  direct composition with rules (no lifting)"}\\
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1162 |   @{text "> HOL.dest:  declaration of Classical destruction rule"}\\
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1163 |   @{text "> HOL.elim:  declaration of Classical elimination rule"}\\
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1164 |   @{text "> \<dots>"}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1165 |   \end{isabelle}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1166 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1167 | The theorem attributes fall roughly into two categories: the first category manipulates | 
| 207 | 1168 |   the proved theorem (for example @{text "[symmetric]"} and @{text "[THEN \<dots>]"}), and the second
 | 
| 1169 |   stores the proved theorem somewhere as data (for example @{text "[simp]"}, which adds
 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1170 | the theorem to the current simpset). | 
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 1171 | |
| 136 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1172 | To explain how to write your own attribute, let us start with an extremely simple | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1173 |   version of the attribute @{text "[symmetric]"}. The purpose of this attribute is
 | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1174 | to produce the ``symmetric'' version of an equation. The main function behind | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1175 | this attribute is | 
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 1176 | *} | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 1177 | |
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1178 | ML{*val my_symmetric = Thm.rule_attribute (fn _ => fn thm => thm RS @{thm sym})*}
 | 
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1179 | |
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1180 | text {* 
 | 
| 136 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1181 |   where the function @{ML "Thm.rule_attribute"} expects a function taking a 
 | 
| 149 | 1182 |   context (which we ignore in the code above) and a theorem (@{text thm}), and 
 | 
| 207 | 1183 |   returns another theorem (namely @{text thm} resolved with the theorem 
 | 
| 1184 |   @{thm [source] sym}: @{thm sym[no_vars]}).\footnote{The function @{ML RS} is explained
 | |
| 1185 |   later on in Section~\ref{sec:simpletacs}.} The function 
 | |
| 1186 |   @{ML "Thm.rule_attribute"} then returns 
 | |
| 156 | 1187 | an attribute. | 
| 136 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1188 | |
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1189 | Before we can use the attribute, we need to set it up. This can be done | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1190 |   using the Isabelle command \isacommand{attribute\_setup} as follows:
 | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1191 | *} | 
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1192 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1193 | attribute_setup %gray my_sym = {* Scan.succeed my_symmetric *} 
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1194 | "applying the sym rule" | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1195 | |
| 136 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1196 | text {*
 | 
| 207 | 1197 |   Inside the @{text "\<verbopen> \<dots> \<verbclose>"}, we have to specify a parser
 | 
| 1198 | for the theorem attribute. Since the attribute does not expect any further | |
| 1199 |   arguments (unlike @{text "[THEN \<dots>]"}, for example), we use the parser @{ML
 | |
| 1200 | Scan.succeed}. Later on we will also consider attributes taking further | |
| 1201 |   arguments. An example for the attribute @{text "[my_sym]"} is the proof
 | |
| 136 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1202 | *} | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1203 | |
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1204 | lemma test[my_sym]: "2 = Suc (Suc 0)" by simp | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1205 | |
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1206 | text {*
 | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1207 |   which stores the theorem @{thm test} under the name @{thm [source] test}. You
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1208 | can see this, if you query the lemma: | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1209 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1210 |   \begin{isabelle}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1211 |   \isacommand{thm}~@{text "test"}\\
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1212 |   @{text "> "}~@{thm test}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1213 |   \end{isabelle}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1214 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1215 | We can also use the attribute when referring to this theorem: | 
| 136 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1216 | |
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1217 |   \begin{isabelle}
 | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1218 |   \isacommand{thm}~@{text "test[my_sym]"}\\
 | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1219 |   @{text "> "}~@{thm test[my_sym]}
 | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1220 |   \end{isabelle}
 | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1221 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1222 | As an example of a slightly more complicated theorem attribute, we implement | 
| 207 | 1223 |   our own version of @{text "[THEN \<dots>]"}. This attribute will take a list of theorems
 | 
| 194 | 1224 | as argument and resolve the proved theorem with this list (one theorem | 
| 1225 | after another). The code for this attribute is | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1226 | *} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1227 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1228 | ML{*fun MY_THEN thms = 
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1229 | Thm.rule_attribute (fn _ => fn thm => foldl ((op RS) o swap) thm thms)*} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1230 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1231 | text {* 
 | 
| 207 | 1232 |   where @{ML swap} swaps the components of a pair. The setup of this theorem
 | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1233 |   attribute uses the parser @{ML Attrib.thms}, which parses a list of
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1234 | theorems. | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1235 | *} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1236 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1237 | attribute_setup %gray MY_THEN = {* Attrib.thms >> MY_THEN *} 
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1238 | "resolving the list of theorems with the proved theorem" | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1239 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1240 | text {* 
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1241 | You can, for example, use this theorem attribute to turn an equation into a | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1242 | meta-equation: | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1243 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1244 |   \begin{isabelle}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1245 |   \isacommand{thm}~@{text "test[MY_THEN eq_reflection]"}\\
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1246 |   @{text "> "}~@{thm test[MY_THEN eq_reflection]}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1247 |   \end{isabelle}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1248 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1249 | If you need the symmetric version as a meta-equation, you can write | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1250 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1251 |   \begin{isabelle}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1252 |   \isacommand{thm}~@{text "test[MY_THEN sym eq_reflection]"}\\
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1253 |   @{text "> "}~@{thm test[MY_THEN sym eq_reflection]}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1254 |   \end{isabelle}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1255 | |
| 194 | 1256 | It is also possible to combine different theorem attributes, as in: | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1257 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1258 |   \begin{isabelle}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1259 |   \isacommand{thm}~@{text "test[my_sym, MY_THEN eq_reflection]"}\\
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1260 |   @{text "> "}~@{thm test[my_sym, MY_THEN eq_reflection]}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1261 |   \end{isabelle}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1262 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1263 | However, here also a weakness of the concept | 
| 194 | 1264 | of theorem attributes shows through: since theorem attributes can be | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1265 | arbitrary functions, they do not in general commute. If you try | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1266 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1267 |   \begin{isabelle}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1268 |   \isacommand{thm}~@{text "test[MY_THEN eq_reflection, my_sym]"}\\
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1269 |   @{text "> "}~@{text "exception THM 1 raised: RSN: no unifiers"}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1270 |   \end{isabelle}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1271 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1272 |   you get an exception indicating that the theorem @{thm [source] sym}
 | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1273 | does not resolve with meta-equations. | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1274 | |
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1275 |   The purpose of @{ML Thm.rule_attribute} is to directly manipulate theorems.
 | 
| 194 | 1276 | Another usage of theorem attributes is to add and delete theorems from stored data. | 
| 1277 |   For example the theorem attribute @{text "[simp]"} adds or deletes a theorem from the
 | |
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1278 |   current simpset. For these applications, you can use @{ML Thm.declaration_attribute}. 
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1279 | To illustrate this function, let us introduce a reference containing a list | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1280 | of theorems. | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1281 | *} | 
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1282 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1283 | ML{*val my_thms = ref ([] : thm list)*}
 | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1284 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1285 | text {* 
 | 
| 207 | 1286 | The purpose of this reference is that we are going to add and delete theorems | 
| 1287 | to the referenced list. However, a word of warning: such references must not | |
| 1288 | be used in any code that is meant to be more than just for testing purposes! | |
| 1289 | Here it is only used to illustrate matters. We will show later how to store | |
| 1290 | data properly without using references. | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1291 | |
| 207 | 1292 | We need to provide two functions that add and delete theorems from this list. | 
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1293 | For this we use the two functions: | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1294 | *} | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1295 | |
| 207 | 1296 | ML{*fun my_thm_add thm ctxt =
 | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1297 | (my_thms := Thm.add_thm thm (!my_thms); ctxt) | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1298 | |
| 207 | 1299 | fun my_thm_del thm ctxt = | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1300 | (my_thms := Thm.del_thm thm (!my_thms); ctxt)*} | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1301 | |
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1302 | text {*
 | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1303 | These functions take a theorem and a context and, for what we are explaining | 
| 156 | 1304 | here it is sufficient that they just return the context unchanged. They change | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1305 |   however the reference @{ML my_thms}, whereby the function @{ML Thm.add_thm}
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1306 |   adds a theorem if it is not already included in the list, and @{ML
 | 
| 194 | 1307 |   Thm.del_thm} deletes one (both functions use the predicate @{ML
 | 
| 1308 | Thm.eq_thm_prop}, which compares theorems according to their proved | |
| 1309 | propositions modulo alpha-equivalence). | |
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1310 | |
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1311 | |
| 207 | 1312 |   You can turn functions @{ML my_thm_add} and @{ML my_thm_del} into 
 | 
| 194 | 1313 | attributes with the code | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1314 | *} | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1315 | |
| 207 | 1316 | ML{*val my_add = Thm.declaration_attribute my_thm_add
 | 
| 1317 | val my_del = Thm.declaration_attribute my_thm_del *} | |
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1318 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1319 | text {* 
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1320 | and set up the attributes as follows | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1321 | *} | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1322 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1323 | attribute_setup %gray my_thms = {* Attrib.add_del my_add my_del *} 
 | 
| 207 | 1324 | "maintaining a list of my_thms - rough test only!" | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1325 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1326 | text {*
 | 
| 207 | 1327 |   The parser @{ML Attrib.add_del} is a pre-defined parser for 
 | 
| 194 | 1328 | adding and deleting lemmas. Now if you prove the next lemma | 
| 207 | 1329 |   and attach to it the attribute @{text "[my_thms]"}
 | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1330 | *} | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1331 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1332 | lemma trueI_2[my_thms]: "True" by simp | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1333 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1334 | text {*
 | 
| 194 | 1335 | then you can see it is added to the initially empty list. | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1336 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1337 |   @{ML_response_fake [display,gray]
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1338 | "!my_thms" "[\"True\"]"} | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1339 | |
| 160 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 1340 |   You can also add theorems using the command \isacommand{declare}.
 | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1341 | *} | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1342 | |
| 207 | 1343 | declare test[my_thms] trueI_2[my_thms add] | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1344 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1345 | text {*
 | 
| 207 | 1346 |   With this attribute, the @{text "add"} operation is the default and does 
 | 
| 1347 | not need to be explicitly given. These three declarations will cause the | |
| 194 | 1348 | theorem list to be updated as: | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1349 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1350 |   @{ML_response_fake [display,gray]
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1351 | "!my_thms" | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1352 | "[\"True\", \"Suc (Suc 0) = 2\"]"} | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1353 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1354 |   The theorem @{thm [source] trueI_2} only appears once, since the 
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1355 |   function @{ML Thm.add_thm} tests for duplicates, before extending
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1356 | the list. Deletion from the list works as follows: | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1357 | *} | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1358 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1359 | declare test[my_thms del] | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1360 | |
| 156 | 1361 | text {* After this, the theorem list is again: 
 | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1362 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1363 |   @{ML_response_fake [display,gray]
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1364 | "!my_thms" | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1365 | "[\"True\"]"} | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1366 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1367 |   We used in this example two functions declared as @{ML Thm.declaration_attribute}, 
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1368 | but there can be any number of them. We just have to change the parser for reading | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1369 | the arguments accordingly. | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1370 | |
| 156 | 1371 | However, as said at the beginning of this example, using references for storing theorems is | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1372 |   \emph{not} the received way of doing such things. The received way is to 
 | 
| 207 | 1373 |   start a ``data slot'', below called @{text MyThmsData}, generated by the functor 
 | 
| 194 | 1374 |   @{text GenericDataFun}:
 | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1375 | *} | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1376 | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1377 | ML {*structure MyThmsData = GenericDataFun
 | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1378 | (type T = thm list | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1379 | val empty = [] | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1380 | val extend = I | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1381 | fun merge _ = Thm.merge_thms) *} | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1382 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1383 | text {*
 | 
| 207 | 1384 |   The type @{text "T"} of this data slot is @{ML_type "thm list"}.\footnote{FIXME: give a pointer
 | 
| 1385 | to where data slots are explained properly.} | |
| 1386 |   To use this data slot, you only have to change @{ML my_thm_add} and
 | |
| 1387 |   @{ML my_thm_del} to:
 | |
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1388 | *} | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1389 | |
| 207 | 1390 | ML{*val my_thm_add = MyThmsData.map o Thm.add_thm
 | 
| 1391 | val my_thm_del = MyThmsData.map o Thm.del_thm*} | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1392 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1393 | text {*
 | 
| 194 | 1394 |   where @{ML MyThmsData.map} updates the data appropriately. The
 | 
| 1395 | corresponding theorem addtributes are | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1396 | *} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1397 | |
| 207 | 1398 | ML{*val my_add = Thm.declaration_attribute my_thm_add
 | 
| 1399 | val my_del = Thm.declaration_attribute my_thm_del *} | |
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1400 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1401 | text {*
 | 
| 194 | 1402 | and the setup is as follows | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1403 | *} | 
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1404 | |
| 207 | 1405 | attribute_setup %gray my_thms2 = {* Attrib.add_del my_add my_del *} 
 | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1406 | "properly maintaining a list of my_thms" | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1407 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1408 | text {*
 | 
| 207 | 1409 | Initially, the data slot is empty | 
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1410 | |
| 194 | 1411 |   @{ML_response_fake [display,gray]
 | 
| 1412 |   "MyThmsData.get (Context.Proof @{context})"
 | |
| 1413 | "[]"} | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1414 | |
| 194 | 1415 | but if you prove | 
| 1416 | *} | |
| 1417 | ||
| 1418 | lemma three[my_thms2]: "3 = Suc (Suc (Suc 0))" by simp | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1419 | |
| 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1420 | text {*
 | 
| 207 | 1421 | then the lemma is recorded. | 
| 194 | 1422 | |
| 1423 |   @{ML_response_fake [display,gray]
 | |
| 1424 |   "MyThmsData.get (Context.Proof @{context})"
 | |
| 1425 | "[\"3 = Suc (Suc (Suc 0))\"]"} | |
| 1426 | ||
| 207 | 1427 |   With theorem attribute @{text my_thms2} you can also nicely see why it 
 | 
| 1428 | is important to | |
| 194 | 1429 |   store data in a ``data slot'' and \emph{not} in a reference. Backtrack
 | 
| 207 | 1430 |   to the point just before the lemma @{thm [source] three} was proved and 
 | 
| 1431 |   check the the content of @{ML_struct "MyThmsData"}: it should be empty. 
 | |
| 1432 | The addition has been properly retracted. Now consider the proof: | |
| 194 | 1433 | *} | 
| 1434 | ||
| 1435 | lemma four[my_thms]: "4 = Suc (Suc (Suc (Suc 0)))" by simp | |
| 193 
ffd93dcc269d
polishing to the  theorem attributes section
 Christian Urban <urbanc@in.tum.de> parents: 
192diff
changeset | 1436 | |
| 194 | 1437 | text {*
 | 
| 1438 |   Checking the content of @{ML my_thms} gives
 | |
| 1439 | ||
| 1440 |   @{ML_response_fake [display,gray]
 | |
| 1441 | "!my_thms" | |
| 1442 | "[\"4 = Suc (Suc (Suc (Suc 0)))\", \"True\"]"} | |
| 1443 | ||
| 207 | 1444 |   as expected, but if you backtrack before the lemma @{thm [source] four}, the
 | 
| 194 | 1445 |   content of @{ML my_thms} is unchanged. The backtracking mechanism
 | 
| 207 | 1446 | of Isabelle is completely oblivious about what to do with references, but | 
| 1447 | properly treats ``data slots''! | |
| 194 | 1448 | |
| 207 | 1449 | Since storing theorems in a list is such a common task, there is the special | 
| 194 | 1450 |   functor @{text NamedThmsFun}, which does most of the work for you. To obtain
 | 
| 207 | 1451 | a named theorem lists, you just declare | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1452 | *} | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1453 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1454 | ML{*structure FooRules = NamedThmsFun 
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1455 | (val name = "foo" | 
| 194 | 1456 | val description = "Rules for foo") *} | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1457 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1458 | text {*
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1459 |   and set up the @{ML_struct FooRules} with the command
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1460 | *} | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1461 | |
| 177 | 1462 | setup %gray {* FooRules.setup *}
 | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1463 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1464 | text {*
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1465 | This code declares a data slot where the theorems are stored, | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1466 |   an attribute @{text foo} (with the @{text add} and @{text del} options
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1467 | for adding and deleting theorems) and an internal ML interface to retrieve and | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1468 | modify the theorems. | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1469 | |
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1470 | Furthermore, the facts are made available on the user-level under the dynamic | 
| 156 | 1471 |   fact name @{text foo}. For example you can declare three lemmas to be of the kind
 | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1472 |   @{text foo} by:
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1473 | *} | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1474 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1475 | lemma rule1[foo]: "A" sorry | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1476 | lemma rule2[foo]: "B" sorry | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1477 | lemma rule3[foo]: "C" sorry | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1478 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1479 | text {* and undeclare the first one by: *}
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1480 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1481 | declare rule1[foo del] | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1482 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1483 | text {* and query the remaining ones with:
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1484 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1485 |   \begin{isabelle}
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1486 |   \isacommand{thm}~@{text "foo"}\\
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1487 |   @{text "> ?C"}\\
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1488 |   @{text "> ?B"}
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1489 |   \end{isabelle}
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1490 | |
| 156 | 1491 |   On the ML-level the rules marked with @{text "foo"} can be retrieved
 | 
| 151 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1492 |   using the function @{ML FooRules.get}:
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1493 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1494 |   @{ML_response_fake [display,gray] "FooRules.get @{context}" "[\"?C\",\"?B\"]"}
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1495 | |
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1496 |   \begin{readmore}
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1497 |   For more information see @{ML_file "Pure/Tools/named_thms.ML"} and also
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1498 |   the recipe in Section~\ref{recipe:storingdata} about storing arbitrary
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1499 | data. | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1500 |   \end{readmore}
 | 
| 
7e0bf13bf743
added more material to the attribute section; merged the recipe about named theorems into the main body; added a solution to an exercise in the conversion section
 Christian Urban <urbanc@in.tum.de> parents: 
149diff
changeset | 1501 | |
| 156 | 1502 |   (FIXME What are: @{text "theory_attributes"}, @{text "proof_attributes"}?)
 | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1503 | |
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1504 | |
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1505 |   \begin{readmore}
 | 
| 207 | 1506 |   FIXME: @{ML_file "Pure/more_thm.ML"}; parsers for attributes is in 
 | 
| 1507 |   @{ML_file "Pure/Isar/attrib.ML"}...also explained in the chapter about
 | |
| 1508 | parsing. | |
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1509 |   \end{readmore}
 | 
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1510 | *} | 
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1511 | |
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 1512 | |
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1513 | section {* Setups (TBD) *}
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1514 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1515 | text {*
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1516 |   In the previous section we used \isacommand{setup} in order to make
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1517 | a theorem attribute known to Isabelle. What happens behind the scenes | 
| 202 | 1518 |   is that \isacommand{setup} expects a function of type 
 | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 1519 |   @{ML_type "theory -> theory"}: the input theory is the current theory and the 
 | 
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1520 | output the theory where the theory attribute has been stored. | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1521 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1522 | This is a fundamental principle in Isabelle. A similar situation occurs | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1523 | for example with declaring constants. The function that declares a | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 1524 |   constant on the ML-level is @{ML Sign.add_consts_i}. 
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 1525 |   If you write\footnote{Recall that ML-code  needs to be 
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 1526 |   enclosed in \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"}.} 
 | 
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1527 | *} | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1528 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 1529 | ML{*Sign.add_consts_i [(@{binding "BAR"}, @{typ "nat"}, NoSyn)] @{theory} *}
 | 
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1530 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1531 | text {*
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1532 |   for declaring the constant @{text "BAR"} with type @{typ nat} and 
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1533 | run the code, then you indeed obtain a theory as result. But if you | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 1534 |   query the constant on the Isabelle level using the command \isacommand{term}
 | 
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1535 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1536 |   \begin{isabelle}
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1537 |   \isacommand{term}~@{text [quotes] "BAR"}\\
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1538 |   @{text "> \"BAR\" :: \"'a\""}
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1539 |   \end{isabelle}
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1540 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1541 |   you do not obtain a constant of type @{typ nat}, but a free variable (printed in 
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1542 | blue) of polymorphic type. The problem is that the ML-expression above did | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1543 | not register the declaration with the current theory. This is what the command | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1544 |   \isacommand{setup} is for. The constant is properly declared with
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1545 | *} | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1546 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 1547 | setup %gray {* Sign.add_consts_i [(@{binding "BAR"}, @{typ "nat"}, NoSyn)] *}
 | 
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1548 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1549 | text {* 
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1550 | Now | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1551 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1552 |   \begin{isabelle}
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1553 |   \isacommand{term}~@{text [quotes] "BAR"}\\
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1554 |   @{text "> \"BAR\" :: \"nat\""}
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1555 |   \end{isabelle}
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1556 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1557 |   returns a (black) constant with the type @{typ nat}.
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1558 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1559 |   A similar command is \isacommand{local\_setup}, which expects a function
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1560 |   of type @{ML_type "local_theory -> local_theory"}. Later on we will also
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1561 |   use the commands \isacommand{method\_setup} for installing methods in the
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1562 |   current theory and \isacommand{simproc\_setup} for adding new simprocs to
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1563 | the current simpset. | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1564 | *} | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1565 | |
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1566 | section {* Theories, Contexts and Local Theories (TBD) *}
 | 
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 1567 | |
| 126 | 1568 | text {*
 | 
| 1569 | There are theories, proof contexts and local theories (in this order, if you | |
| 1570 | want to order them). | |
| 1571 | ||
| 1572 | In contrast to an ordinary theory, which simply consists of a type | |
| 1573 | signature, as well as tables for constants, axioms and theorems, a local | |
| 202 | 1574 | theory contains additional context information, such as locally fixed | 
| 126 | 1575 | variables and local assumptions that may be used by the package. The type | 
| 1576 |   @{ML_type local_theory} is identical to the type of \emph{proof contexts}
 | |
| 1577 |   @{ML_type "Proof.context"}, although not every proof context constitutes a
 | |
| 1578 | valid local theory. | |
| 178 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1579 | *} | 
| 126 | 1580 | |
| 235 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 1581 | (* | 
| 229 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1582 | ML{*signature UNIVERSAL_TYPE =
 | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1583 | sig | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1584 | type t | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1585 | |
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1586 |   val embed: unit -> ('a -> t) * (t -> 'a option)
 | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1587 | end*} | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1588 | |
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1589 | ML{*structure U:> UNIVERSAL_TYPE =
 | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1590 | struct | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1591 | type t = exn | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1592 | |
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1593 | fun 'a embed () = | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1594 | let | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1595 | exception E of 'a | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1596 | fun project (e: t): 'a option = | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1597 | case e of | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1598 | E a => SOME a | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1599 | | _ => NONE | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1600 | in | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1601 | (E, project) | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1602 | end | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1603 | end*} | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1604 | |
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1605 | text {*
 | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1606 | The idea is that type t is the universal type and that each call to embed | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1607 | returns a new pair of functions (inject, project), where inject embeds a | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1608 | value into the universal type and project extracts the value from the | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1609 | universal type. A pair (inject, project) returned by embed works together in | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1610 | that project u will return SOME v if and only if u was created by inject | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1611 | v. If u was created by a different function inject', then project returns | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1612 | NONE. | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1613 | |
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1614 | in library.ML | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1615 | *} | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1616 | |
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1617 | ML_val{*structure Object = struct type T = exn end; *}
 | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1618 | |
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1619 | ML{*functor Test (U: UNIVERSAL_TYPE): sig end =
 | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1620 | struct | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1621 | val (intIn: int -> U.t, intOut) = U.embed () | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1622 | val r: U.t ref = ref (intIn 13) | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1623 | val s1 = | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1624 | case intOut (!r) of | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1625 | NONE => "NONE" | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1626 | | SOME i => Int.toString i | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1627 | val (realIn: real -> U.t, realOut) = U.embed () | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1628 | val () = r := realIn 13.0 | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1629 | val s2 = | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1630 | case intOut (!r) of | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1631 | NONE => "NONE" | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1632 | | SOME i => Int.toString i | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1633 | val s3 = | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1634 | case realOut (!r) of | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1635 | NONE => "NONE" | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1636 | | SOME x => Real.toString x | 
| 239 
b63c72776f03
replaced "warning" with "writeln"
 Christian Urban <urbanc@in.tum.de> parents: 
235diff
changeset | 1637 | val () = writeln (concat [s1, " ", s2, " ", s3, "\n"]) | 
| 229 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1638 | end*} | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1639 | |
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1640 | ML_val{*structure t = Test(U) *} 
 | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1641 | |
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1642 | ML_val{*structure Datatab = TableFun(type key = int val ord = int_ord);*}
 | 
| 
abc7f90188af
permutation example uses now recent infrastructure
 Christian Urban <urbanc@in.tum.de> parents: 
228diff
changeset | 1643 | |
| 235 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 1644 | ML {* LocalTheory.restore *}
 | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 1645 | ML {* LocalTheory.set_group *}
 | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 1646 | *) | 
| 
dc955603d813
explained uses and use commands
 Christian Urban <urbanc@in.tum.de> parents: 
234diff
changeset | 1647 | |
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1648 | section {* Storing Theorems\label{sec:storing} (TBD) *}
 | 
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 1649 | |
| 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 1650 | text {* @{ML PureThy.add_thms_dynamic} *}
 | 
| 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 1651 | |
| 100 | 1652 | |
| 75 | 1653 | |
| 126 | 1654 | (* FIXME: some code below *) | 
| 89 | 1655 | |
| 1656 | (*<*) | |
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1657 | (* | 
| 89 | 1658 | setup {*
 | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1659 |  Sign.add_consts_i [(Binding"bar", @{typ "nat"},NoSyn)] 
 | 
| 89 | 1660 | *} | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1661 | *) | 
| 89 | 1662 | lemma "bar = (1::nat)" | 
| 1663 | oops | |
| 1664 | ||
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1665 | (* | 
| 89 | 1666 | setup {*   
 | 
| 1667 |   Sign.add_consts_i [("foo", @{typ "nat"},NoSyn)] 
 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 1668 |  #> PureThy.add_defs false [((@{binding "foo_def"}, 
 | 
| 89 | 1669 |        Logic.mk_equals (Const ("FirstSteps.foo", @{typ "nat"}), @{term "1::nat"})), [])] 
 | 
| 1670 | #> snd | |
| 1671 | *} | |
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1672 | *) | 
| 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1673 | (* | 
| 89 | 1674 | lemma "foo = (1::nat)" | 
| 1675 | apply(simp add: foo_def) | |
| 1676 | done | |
| 1677 | ||
| 1678 | thm foo_def | |
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1679 | *) | 
| 89 | 1680 | (*>*) | 
| 1681 | ||
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1682 | section {* Pretty-Printing (TBD) *}
 | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1683 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1684 | text {*
 | 
| 210 
db8e302f44c8
more work on the simple inductive section
 Christian Urban <urbanc@in.tum.de> parents: 
207diff
changeset | 1685 | Isabelle has a pretty sphisticated pretty printing module. | 
| 
db8e302f44c8
more work on the simple inductive section
 Christian Urban <urbanc@in.tum.de> parents: 
207diff
changeset | 1686 | *} | 
| 
db8e302f44c8
more work on the simple inductive section
 Christian Urban <urbanc@in.tum.de> parents: 
207diff
changeset | 1687 | |
| 
db8e302f44c8
more work on the simple inductive section
 Christian Urban <urbanc@in.tum.de> parents: 
207diff
changeset | 1688 | text {*
 | 
| 163 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1689 |   @{ML Pretty.big_list},
 | 
| 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1690 |   @{ML Pretty.brk},
 | 
| 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1691 |   @{ML Pretty.block},
 | 
| 
2319cff107f0
removed rep_ss, and used dest_ss instead; some very slight changes to simple_inductive
 Christian Urban <urbanc@in.tum.de> parents: 
162diff
changeset | 1692 |   @{ML Pretty.chunks}
 | 
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1693 | *} | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1694 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1695 | section {* Misc (TBD) *}
 | 
| 92 | 1696 | |
| 1697 | ML {*DatatypePackage.get_datatype @{theory} "List.list"*}
 | |
| 1698 | ||
| 196 
840b49bfb1cf
fixed `str_of_thms' output in example + small changes
 griff parents: 
192diff
changeset | 1699 | end |