| author | Christian Urban <urbanc@in.tum.de> | 
| Thu, 19 Mar 2009 17:50:28 +0100 | |
| changeset 190 | ca0ac2e75f6d | 
| parent 189 | 069d525f8f1d | 
| child 192 | 2fff636e1fa0 | 
| 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 {*
 | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 8 | |
| 54 
1783211b3494
tuned; added document antiquotation ML_response_fake_both
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 9 | Isabelle programming is done in ML. Just like lemmas and proofs, ML-code | 
| 89 | 10 | 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 | 11 | 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 | 12 | |
| 6 | 13 |   \begin{center}
 | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 14 |   \begin{tabular}{@ {}l}
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
34diff
changeset | 15 |   \isacommand{theory} FirstSteps\\
 | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 16 |   \isacommand{imports} Main\\
 | 
| 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 17 |   \isacommand{begin}\\
 | 
| 6 | 18 | \ldots | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 19 |   \end{tabular}
 | 
| 6 | 20 |   \end{center}
 | 
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 21 | |
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 22 | We also generally assume you are working with HOL. The given examples might | 
| 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 23 | need to be adapted slightly 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 | 24 | *} | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 25 | |
| 20 | 26 | 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 | 27 | |
| 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 28 | text {*
 | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 29 | The easiest and quickest way to include code in a theory is | 
| 89 | 30 |   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 | 31 | |
| 75 | 32 | \begin{isabelle}
 | 
| 33 | \begin{graybox}
 | |
| 85 
b02904872d6b
better handling of {* and *}
 Christian Urban <urbanc@in.tum.de> parents: 
84diff
changeset | 34 | \isacommand{ML}~@{text "\<verbopen>"}\isanewline
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
34diff
changeset | 35 | \hspace{5mm}@{ML "3 + 4"}\isanewline
 | 
| 85 
b02904872d6b
better handling of {* and *}
 Christian Urban <urbanc@in.tum.de> parents: 
84diff
changeset | 36 | @{text "\<verbclose>"}\isanewline
 | 
| 
b02904872d6b
better handling of {* and *}
 Christian Urban <urbanc@in.tum.de> parents: 
84diff
changeset | 37 | @{text "> 7"}\smallskip
 | 
| 75 | 38 | \end{graybox}
 | 
| 39 | \end{isabelle}
 | |
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 40 | |
| 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 |   Like normal Isabelle proof scripts, \isacommand{ML}-commands can be
 | 
| 
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 | 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 | 43 |   environment. The code inside the \isacommand{ML}-command can also contain
 | 
| 
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 | 44 | value and function bindings, and even those can be undone when the proof | 
| 160 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 45 | 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 | 46 |   \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 | 47 |   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 | 48 | 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 | 49 | |
| 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 | 50 | Once a portion of code is relatively stable, you usually want to export it | 
| 
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 | 51 | to a separate ML-file. Such files can then be included in a theory by using | 
| 
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 | 52 |   the \isacommand{uses}-command in the header of the theory, like:
 | 
| 14 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 53 | |
| 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 54 |   \begin{center}
 | 
| 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 55 |   \begin{tabular}{@ {}l}
 | 
| 54 
1783211b3494
tuned; added document antiquotation ML_response_fake_both
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 56 |   \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 | 57 |   \isacommand{imports} Main\\
 | 
| 58 | 58 |   \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 | 59 |   \isacommand{begin}\\
 | 
| 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 60 | \ldots | 
| 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 61 |   \end{tabular}
 | 
| 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 62 |   \end{center}
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
34diff
changeset | 63 | |
| 14 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 64 | *} | 
| 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 65 | |
| 126 | 66 | 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 | 67 | |
| 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 68 | text {*
 | 
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 69 | |
| 50 | 70 | 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 | 71 | in your code. This can be done in a ``quick-and-dirty'' fashion using | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 72 |   the function @{ML "warning"}. For example 
 | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 73 | |
| 161 
83f36a1c62f2
rolled back the changes on the function warning and tracing
 Christian Urban <urbanc@in.tum.de> parents: 
160diff
changeset | 74 |   @{ML_response_fake [display,gray] "warning \"any string\"" "\"any string\""}
 | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 75 | |
| 58 | 76 |   will print out @{text [quotes] "any string"} inside the response buffer
 | 
| 77 | of Isabelle. This function expects a string as argument. If you develop under PolyML, | |
| 50 | 78 | then there is a convenient, though again ``quick-and-dirty'', method for | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 79 |   converting values into strings, namely the function @{ML makestring}:
 | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 80 | |
| 161 
83f36a1c62f2
rolled back the changes on the function warning and tracing
 Christian Urban <urbanc@in.tum.de> parents: 
160diff
changeset | 81 |   @{ML_response_fake [display,gray] "warning (makestring 1)" "\"1\""} 
 | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 82 | |
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 83 |   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 | 84 | and not a function. | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 85 | |
| 52 | 86 |   The function @{ML "warning"} should only be used for testing purposes, because any
 | 
| 87 | output this function generates will be overwritten as soon as an error is | |
| 50 | 88 | 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 | 89 |   function @{ML tracing} is more appropriate. This function writes all output into
 | 
| 89 | 90 | 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 | 91 | |
| 161 
83f36a1c62f2
rolled back the changes on the function warning and tracing
 Christian Urban <urbanc@in.tum.de> parents: 
160diff
changeset | 92 |   @{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 | 93 | |
| 58 | 94 |   It is also possible to redirect the ``channel'' where the string @{text "foo"} is 
 | 
| 78 
ef778679d3e0
added a section about combinators
 Christian Urban <urbanc@in.tum.de> parents: 
75diff
changeset | 95 | 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 | 96 | 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 | 97 | *} | 
| 14 
1c17e99f6f66
added a paragraph about "uses" and started a paragraph about tracing
 Christian Urban <urbanc@in.tum.de> parents: 
13diff
changeset | 98 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 99 | ML{*val strip_specials =
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 100 | 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 | 101 |   fun strip ("\^A" :: _ :: cs) = strip cs
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 102 | | strip (c :: cs) = c :: strip cs | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 103 | | strip [] = []; | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 104 | in implode o strip o explode end; | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 105 | |
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 106 | fun redirect_tracing stream = | 
| 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 107 | 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 | 108 | (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 | 109 | TextIO.output (stream, "\n"); | 
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 110 | 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 | 111 | |
| 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 | 112 | 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 | 113 |   Calling @{ML "redirect_tracing"} with @{ML "(TextIO.openOut \"foo.bar\")"} 
 | 
| 58 | 114 |   will cause that all tracing information is printed into the file @{text "foo.bar"}.
 | 
| 75 | 115 | |
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 116 |   You can print out error messages with the function @{ML error}; for example:
 | 
| 75 | 117 | |
| 122 | 118 | @{ML_response_fake [display,gray] "if 0=1 then true else (error \"foo\")" 
 | 
| 119 | "Exception- ERROR \"foo\" raised | |
| 120 | At command \"ML\"."} | |
| 75 | 121 | |
| 126 | 122 |   Most often you want to inspect data of type @{ML_type term}, @{ML_type cterm} 
 | 
| 123 |   or @{ML_type thm}. Isabelle contains elaborate pretty-printing functions for printing them, 
 | |
| 124 | but for quick-and-dirty solutions they are far too unwieldy. A simple way to transform | |
| 125 |   a term into a string is to use the function @{ML Syntax.string_of_term}.
 | |
| 126 | ||
| 127 |   @{ML_response_fake [display,gray]
 | |
| 128 |   "Syntax.string_of_term @{context} @{term \"1::nat\"}"
 | |
| 129 | "\"\\^E\\^Fterm\\^E\\^E\\^Fconst\\^Fname=HOL.one_class.one\\^E1\\^E\\^F\\^E\\^E\\^F\\^E\""} | |
| 130 | ||
| 131 | This produces a string with some additional information encoded in it. The string | |
| 132 |   can be properly printed by using the function @{ML warning}.
 | |
| 133 | ||
| 134 |   @{ML_response_fake [display,gray]
 | |
| 135 |   "warning (Syntax.string_of_term @{context} @{term \"1::nat\"})"
 | |
| 136 | "\"1\""} | |
| 137 | ||
| 138 |   A @{ML_type cterm} can be transformed into a string by the following function.
 | |
| 139 | *} | |
| 140 | ||
| 141 | ML{*fun str_of_cterm ctxt t =  
 | |
| 142 | Syntax.string_of_term ctxt (term_of t)*} | |
| 143 | ||
| 144 | text {*
 | |
| 149 | 145 |   In this example the function @{ML term_of} extracts the @{ML_type term} from
 | 
| 146 |   a @{ML_type cterm}.  If there are more than one @{ML_type cterm}s to be
 | |
| 147 |   printed, you can use the function @{ML commas} to separate them.
 | |
| 126 | 148 | *} | 
| 149 | ||
| 150 | ML{*fun str_of_cterms ctxt ts =  
 | |
| 151 | commas (map (str_of_cterm ctxt) ts)*} | |
| 152 | ||
| 153 | text {*
 | |
| 154 | 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 | 155 |   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 | 156 | *} | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 157 | |
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 158 | ML{*fun str_of_thm_raw ctxt thm =
 | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 159 | str_of_cterm ctxt (#prop (crep_thm thm))*} | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 160 | |
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 161 | text {*
 | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 162 |   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 | 163 |   @{text "?Q"} and so on. 
 | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 164 | |
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 165 |   @{ML_response_fake [display, gray]
 | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 166 |   "warning (str_of_thm_raw @{context} @{thm conjI})"
 | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 167 | "\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q"} | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 168 | |
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 169 | 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 | 170 | 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 | 171 |   function @{ML Variable.import_thms}.
 | 
| 126 | 172 | *} | 
| 173 | ||
| 158 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 174 | 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 | 175 | let | 
| 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 176 | 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 | 177 | in | 
| 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 178 | thm' | 
| 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 179 | end | 
| 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 180 | |
| 
d7944bdf7b3f
removed infix_conv and moved function no_vars into the FirstSteps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
157diff
changeset | 181 | fun str_of_thm ctxt thm = | 
| 171 
18f90044c777
simplified antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
163diff
changeset | 182 | str_of_cterm ctxt (#prop (crep_thm (no_vars ctxt thm)))*} | 
| 126 | 183 | |
| 184 | text {* 
 | |
| 190 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 185 |   Theorem @{thm [source] conjI} looks now as follows
 | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 186 | |
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 187 |   @{ML_response_fake [display, gray]
 | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 188 |   "warning (str_of_thm_raw @{context} @{thm conjI})"
 | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 189 | "\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q"} | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 190 | |
| 126 | 191 |   Again the function @{ML commas} helps with printing more than one theorem. 
 | 
| 192 | *} | |
| 193 | ||
| 194 | ML{*fun str_of_thms ctxt thms =  
 | |
| 190 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 195 | commas (map (str_of_thm ctxt) thms) | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 196 | |
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 197 | fun str_of_thms_raw ctxt thms = | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 198 | commas (map (str_of_thm_raw ctxt) thms)*} | 
| 126 | 199 | |
| 188 | 200 | text {*
 | 
| 201 | (FIXME @{ML Toplevel.debug} @{ML Toplevel.profiling} @{ML Toplevel.debug})
 | |
| 202 | *} | |
| 203 | ||
| 126 | 204 | section {* Combinators\label{sec:combinators} *}
 | 
| 205 | ||
| 206 | text {*
 | |
| 131 | 207 | For beginners perhaps the most puzzling parts in the existing code of Isabelle are | 
| 126 | 208 | the combinators. At first they seem to greatly obstruct the | 
| 209 | comprehension of the code, but after getting familiar with them, they | |
| 210 | actually ease the understanding and also the programming. | |
| 211 | ||
| 212 |   The simplest combinator is @{ML I}, which is just the identity function defined as
 | |
| 213 | *} | |
| 214 | ||
| 215 | ML{*fun I x = x*}
 | |
| 216 | ||
| 217 | text {* Another simple combinator is @{ML K}, defined as *}
 | |
| 218 | ||
| 219 | ML{*fun K x = fn _ => x*}
 | |
| 220 | ||
| 221 | text {*
 | |
| 222 |   @{ML K} ``wraps'' a function around the argument @{text "x"}. However, this 
 | |
| 223 |   function ignores its argument. As a result, @{ML K} defines a constant function 
 | |
| 224 |   always returning @{text x}.
 | |
| 225 | ||
| 226 |   The next combinator is reverse application, @{ML "|>"}, defined as: 
 | |
| 227 | *} | |
| 228 | ||
| 229 | ML{*fun x |> f = f x*}
 | |
| 230 | ||
| 231 | text {* While just syntactic sugar for the usual function application,
 | |
| 232 | the purpose of this combinator is to implement functions in a | |
| 233 | ``waterfall fashion''. Consider for example the function *} | |
| 234 | ||
| 235 | ML %linenosgray{*fun inc_by_five x =
 | |
| 236 | x |> (fn x => x + 1) | |
| 237 | |> (fn x => (x, x)) | |
| 238 | |> fst | |
| 239 | |> (fn x => x + 4)*} | |
| 240 | ||
| 241 | text {*
 | |
| 242 |   which increments its argument @{text x} by 5. It does this by first incrementing 
 | |
| 243 | the argument by 1 (Line 2); then storing the result in a pair (Line 3); taking | |
| 244 | the first component of the pair (Line 4) and finally incrementing the first | |
| 245 | component by 4 (Line 5). This kind of cascading manipulations of values is quite | |
| 246 | common when dealing with theories (for example by adding a definition, followed by | |
| 247 | lemmas and so on). The reverse application allows you to read what happens in | |
| 248 | a top-down manner. This kind of coding should also be familiar, | |
| 149 | 249 |   if you have been exposed to Haskell's do-notation. Writing the function @{ML inc_by_five} using 
 | 
| 126 | 250 | the reverse application is much clearer than writing | 
| 251 | *} | |
| 252 | ||
| 253 | ML{*fun inc_by_five x = fst ((fn x => (x, x)) (x + 1)) + 4*}
 | |
| 254 | ||
| 255 | text {* or *}
 | |
| 256 | ||
| 257 | ML{*fun inc_by_five x = 
 | |
| 258 | ((fn x => x + 4) o fst o (fn x => (x, x)) o (fn x => x + 1)) x*} | |
| 259 | ||
| 260 | text {* and typographically more economical than *}
 | |
| 261 | ||
| 262 | ML{*fun inc_by_five x =
 | |
| 263 | let val y1 = x + 1 | |
| 264 | val y2 = (y1, y1) | |
| 265 | val y3 = fst y2 | |
| 266 | val y4 = y3 + 4 | |
| 267 | in y4 end*} | |
| 268 | ||
| 269 | text {* 
 | |
| 270 | Another reason why the let-bindings in the code above are better to be | |
| 271 | avoided: it is more than easy to get the intermediate values wrong, not to | |
| 272 | mention the nightmares the maintenance of this code causes! | |
| 273 | ||
| 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 | 274 | In the context of Isabelle, a ``real world'' example for a function written 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 | 275 | the waterfall fashion might be the following code: | 
| 177 | 276 | *} | 
| 126 | 277 | |
| 177 | 278 | ML %linenosgray{*fun apply_fresh_args pred ctxt =
 | 
| 279 | pred |> fastype_of | |
| 280 | |> binder_types | |
| 281 | |> map (pair "z") | |
| 282 | |> Variable.variant_frees ctxt [pred] | |
| 283 | |> map Free | |
| 284 | |> (curry list_comb) pred *} | |
| 126 | 285 | |
| 177 | 286 | text {*
 | 
| 184 | 287 | This code extracts the argument types of a given | 
| 288 | predicate and then generates for each argument type a distinct variable; finally it | |
| 183 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 289 | applies the generated variables to the predicate. For example | 
| 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 290 | |
| 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 291 |   @{ML_response_fake [display,gray]
 | 
| 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 292 | "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 | 293 |  |> Syntax.string_of_term @{context}
 | 
| 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 294 | |> warning" | 
| 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 295 | "P z za zb"} | 
| 177 | 296 | |
| 184 | 297 |   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 | 298 |   coded: in Line 2, the function @{ML fastype_of} calculates the type of the
 | 
| 184 | 299 |   predicate; @{ML binder_types} in the next line produces the list of argument
 | 
| 300 |   types (in the case above the list @{text "[nat, int, unit]"}); Line 4 
 | |
| 301 |   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 | 302 |   function @{ML variant_frees in Variable} generates for each @{text "z"} a
 | 
| 184 | 303 |   unique name avoiding the given @{text pred}; the list of name-type pairs is turned
 | 
| 304 | into a list of variable terms in Line 6, which in the last line is applied | |
| 183 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 305 |   by the function @{ML list_comb} to the predicate. We have to use the
 | 
| 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 306 |   function @{ML curry}, because @{ML list_comb} expects the predicate and the
 | 
| 184 | 307 | variables list as a pair. | 
| 177 | 308 | |
| 309 |   The combinator @{ML "#>"} is the reverse function composition. It can be
 | |
| 310 | used to define the following function | |
| 126 | 311 | *} | 
| 312 | ||
| 313 | ML{*val inc_by_six = 
 | |
| 314 | (fn x => x + 1) | |
| 315 | #> (fn x => x + 2) | |
| 316 | #> (fn x => x + 3)*} | |
| 317 | ||
| 318 | text {* 
 | |
| 319 | which is the function composed of first the increment-by-one function and then | |
| 320 | increment-by-two, followed by increment-by-three. Again, the reverse function | |
| 321 | composition allows you to read the code top-down. | |
| 322 | ||
| 323 | The remaining combinators described in this section add convenience for the | |
| 324 |   ``waterfall method'' of writing functions. The combinator @{ML tap} allows
 | |
| 325 | you to get hold of an intermediate result (to do some side-calculations for | |
| 326 | instance). The function | |
| 327 | ||
| 328 | *} | |
| 329 | ||
| 330 | ML %linenosgray{*fun inc_by_three x =
 | |
| 331 | x |> (fn x => x + 1) | |
| 332 | |> tap (fn x => tracing (makestring x)) | |
| 333 | |> (fn x => x + 2)*} | |
| 334 | ||
| 335 | text {* 
 | |
| 336 |   increments the argument first by @{text "1"} and then by @{text "2"}. In the
 | |
| 337 |   middle (Line 3), however, it uses @{ML tap} for printing the ``plus-one''
 | |
| 338 |   intermediate result inside the tracing buffer. The function @{ML tap} can
 | |
| 339 | only be used for side-calculations, because any value that is computed | |
| 340 | cannot be merged back into the ``main waterfall''. To do this, you can use | |
| 341 | the next combinator. | |
| 342 | ||
| 343 |   The combinator @{ML "`"} is similar to @{ML tap}, but applies a function to the value
 | |
| 344 | and returns the result together with the value (as a pair). For example | |
| 345 | the function | |
| 346 | *} | |
| 347 | ||
| 348 | ML{*fun inc_as_pair x =
 | |
| 349 | x |> `(fn x => x + 1) | |
| 350 | |> (fn (x, y) => (x, y + 1))*} | |
| 351 | ||
| 352 | text {*
 | |
| 353 |   takes @{text x} as argument, and then increments @{text x}, but also keeps
 | |
| 354 |   @{text x}. The intermediate result is therefore the pair @{ML "(x + 1, x)"
 | |
| 355 | for x}. After that, the function increments the right-hand component of the | |
| 356 |   pair. So finally the result will be @{ML "(x + 1, x + 1)" for x}.
 | |
| 357 | ||
| 358 |   The combinators @{ML "|>>"} and @{ML "||>"} are defined for 
 | |
| 359 | functions manipulating pairs. The first applies the function to | |
| 360 | the first component of the pair, defined as | |
| 361 | *} | |
| 362 | ||
| 363 | ML{*fun (x, y) |>> f = (f x, y)*}
 | |
| 364 | ||
| 365 | text {*
 | |
| 366 | and the second combinator to the second component, defined as | |
| 367 | *} | |
| 368 | ||
| 369 | ML{*fun (x, y) ||> f = (x, f y)*}
 | |
| 370 | ||
| 371 | text {*
 | |
| 372 |   With the combinator @{ML "|->"} you can re-combine the elements from a pair.
 | |
| 373 | This combinator is defined as | |
| 374 | *} | |
| 375 | ||
| 376 | ML{*fun (x, y) |-> f = f x y*}
 | |
| 377 | ||
| 378 | text {* and can be used to write the following roundabout version 
 | |
| 379 |   of the @{text double} function: 
 | |
| 380 | *} | |
| 381 | ||
| 382 | ML{*fun double x =
 | |
| 383 | x |> (fn x => (x, x)) | |
| 384 | |-> (fn x => fn y => x + y)*} | |
| 385 | ||
| 386 | text {*
 | |
| 387 |   Recall that @{ML "|>"} is the reverse function applications. Recall also that the related 
 | |
| 388 |   reverse function composition is @{ML "#>"}. In fact all the combinators @{ML "|->"},
 | |
| 389 |   @{ML "|>>"} and @{ML "||>"} described above have related combinators for function
 | |
| 390 |   composition, namely @{ML "#->"}, @{ML "#>>"} and @{ML "##>"}. Using @{ML "#->"}, 
 | |
| 391 |   for example, the function @{text double} can also be written as:
 | |
| 392 | *} | |
| 393 | ||
| 394 | ML{*val double =
 | |
| 395 | (fn x => (x, x)) | |
| 396 | #-> (fn x => fn y => x + y)*} | |
| 397 | ||
| 398 | text {*
 | |
| 399 | ||
| 400 | (FIXME: find a good exercise for combinators) | |
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 401 | |
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 402 |   \begin{readmore}
 | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 403 |   The most frequently used combinator are defined in the files @{ML_file "Pure/library.ML"}
 | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 404 |   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 | 405 | contains further information about combinators. | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 406 |   \end{readmore}
 | 
| 126 | 407 | |
| 15 
9da9ba2b095b
added a solution section and some other minor additions
 Christian Urban <urbanc@in.tum.de> parents: 
14diff
changeset | 408 | *} | 
| 
9da9ba2b095b
added a solution section and some other minor additions
 Christian Urban <urbanc@in.tum.de> parents: 
14diff
changeset | 409 | |
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 410 | |
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 411 | section {* Antiquotations *}
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 412 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 413 | text {*
 | 
| 49 | 414 | The main advantage of embedding all code in a theory is that the code can | 
| 58 | 415 | contain references to entities defined on the logical level of Isabelle. By | 
| 416 | this we mean definitions, theorems, terms and so on. This kind of reference is | |
| 417 | realised with antiquotations. For example, one can print out the name of the current | |
| 49 | 418 | theory by typing | 
| 419 | ||
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
34diff
changeset | 420 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
71diff
changeset | 421 |   @{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 | 422 | |
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 423 |   where @{text "@{theory}"} is an antiquotation that is substituted with the
 | 
| 49 | 424 | current theory (remember that we assumed we are inside the theory | 
| 89 | 425 |   @{text FirstSteps}). The name of this theory can be extracted using
 | 
| 49 | 426 |   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 | 427 | |
| 89 | 428 | 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 | 429 | 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 | 430 | *} | 
| 5 
e91f54791e14
minor modifiations to the Intro and FirstSteps chapters
 Christian Urban <urbanc@in.tum.de> parents: 
2diff
changeset | 431 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 432 | 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 | 433 | |
| 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 434 | text {*
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 435 | |
| 89 | 436 |   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 | 437 | different theory. Instead, the code above defines the constant function | 
| 58 | 438 |   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 | 439 |   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 | 440 |   \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 | 441 | 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 | 442 | 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 | 443 | |
| 132 | 444 | 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 | 445 |   @{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 | 446 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
71diff
changeset | 447 |   @{ML_response_fake [display,gray] "@{thm allI}" "(\<And>x. ?P x) \<Longrightarrow> \<forall>x. ?P x"}
 | 
| 75 | 448 | |
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 449 |   and @{text "@{thms \<dots>}"} for more than one
 | 
| 132 | 450 | |
| 451 | @{ML_response_fake [display,gray] "@{thms conj_ac}" 
 | |
| 452 | "(?P \<and> ?Q) = (?Q \<and> ?P) | |
| 453 | (?P \<and> ?Q \<and> ?R) = (?Q \<and> ?P \<and> ?R) | |
| 454 | ((?P \<and> ?Q) \<and> ?R) = (?P \<and> ?Q \<and> ?R)"} | |
| 455 | ||
| 149 | 456 | You can also refer to the current simpset. To illustrate this we implement the | 
| 132 | 457 | function that extracts the theorem names stored in a simpset. | 
| 131 | 458 | *} | 
| 75 | 459 | |
| 149 | 460 | ML{*fun get_thm_names_from_ss simpset =
 | 
| 131 | 461 | 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 | 462 |   val {simps,...} = MetaSimplifier.dest_ss simpset
 | 
| 70 
bbb2d5f1d58d
deleted the fixme about simpsets
 Christian Urban <urbanc@in.tum.de> parents: 
69diff
changeset | 463 | 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 | 464 | map #1 simps | 
| 131 | 465 | end*} | 
| 54 
1783211b3494
tuned; added document antiquotation ML_response_fake_both
 Christian Urban <urbanc@in.tum.de> parents: 
52diff
changeset | 466 | |
| 131 | 467 | 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 | 468 |   The function @{ML dest_ss in MetaSimplifier} returns a record containing all
 | 
| 184 | 469 | information stored in the simpset. We are only interested in the names of the | 
| 470 | simp-rules. So now you can feed in the current simpset into this function. | |
| 471 |   The simpset can be referred to using the antiquotation @{text "@{simpset}"}.
 | |
| 81 | 472 | |
| 131 | 473 |   @{ML_response_fake [display,gray] 
 | 
| 149 | 474 |   "get_thm_names_from_ss @{simpset}" 
 | 
| 475 | "[\"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 | 476 | |
| 156 | 477 | Again, this way or referencing simpsets makes you independent from additions | 
| 478 | of lemmas to the simpset by the user that potentially cause loops. | |
| 479 | ||
| 75 | 480 | While antiquotations have many applications, they were originally introduced in order | 
| 89 | 481 | to avoid explicit bindings for 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 | 482 | *} | 
| 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 483 | |
| 69 
19106a9975c1
highligted the background of ML-code
 Christian Urban <urbanc@in.tum.de> parents: 
68diff
changeset | 484 | 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 | 485 | |
| 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 486 | text {*
 | 
| 78 
ef778679d3e0
added a section about combinators
 Christian Urban <urbanc@in.tum.de> parents: 
75diff
changeset | 487 | These bindings are difficult to maintain and also can be accidentally | 
| 156 | 488 | overwritten by the user. This often broke Isabelle | 
| 49 | 489 | packages. Antiquotations solve this problem, since they are ``linked'' | 
| 89 | 490 | statically at compile-time. However, this static linkage also limits their | 
| 491 | usefulness in cases where data needs to be build up dynamically. In the | |
| 149 | 492 | course of this chapter you will learn more about these antiquotations: | 
| 122 | 493 | they can simplify Isabelle programming since one can directly access all | 
| 89 | 494 | kinds of logical elements from th ML-level. | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 495 | *} | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 496 | |
| 183 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 497 | text {*
 | 
| 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 498 |   (FIXME: say something about @{text "@{binding \<dots>}"}
 | 
| 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 499 | *} | 
| 
8bb4eaa2ec92
a simplification suggested by Stefan and some polishing
 Christian Urban <urbanc@in.tum.de> parents: 
178diff
changeset | 500 | |
| 15 
9da9ba2b095b
added a solution section and some other minor additions
 Christian Urban <urbanc@in.tum.de> parents: 
14diff
changeset | 501 | section {* Terms and Types *}
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 502 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 503 | text {*
 | 
| 122 | 504 | One way to construct terms of Isabelle is by using the antiquotation | 
| 89 | 505 |   \mbox{@{text "@{term \<dots>}"}}. For example:
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 506 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
71diff
changeset | 507 |   @{ML_response [display,gray] 
 | 
| 75 | 508 | "@{term \"(a::nat) + b = c\"}" 
 | 
| 509 | "Const (\"op =\", \<dots>) $ | |
| 149 | 510 | (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 | 511 | |
| 50 | 512 |   This will show the term @{term "(a::nat) + b = c"}, but printed using the internal
 | 
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 513 | representation of this term. This internal representation corresponds to the | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
34diff
changeset | 514 |   datatype @{ML_type "term"}.
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 515 | |
| 34 | 516 | The internal representation of terms uses the usual de Bruijn index mechanism where bound | 
| 11 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 517 |   variables are represented by the constructor @{ML Bound}. The index in @{ML Bound} refers to
 | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 518 |   the number of Abstractions (@{ML Abs}) we have to skip until we hit the @{ML Abs} that
 | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 519 | binds the corresponding variable. However, in Isabelle the names of bound variables are | 
| 149 | 520 | kept at abstractions for printing purposes, and so should be treated only as ``comments''. | 
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 521 |   Application in Isabelle is realised with the term-constructor @{ML $}.
 | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 522 | |
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 523 |   \begin{readmore}
 | 
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 524 |   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 | 525 |   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 | 526 |   \end{readmore}
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 527 | |
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 528 | 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 | 529 | 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 | 530 | 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 | 531 | |
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 532 |   \begin{exercise}
 | 
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 533 | Look at the internal term representation of the following terms, and | 
| 89 | 534 | 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 | 535 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 536 |   \begin{itemize}
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 537 |   \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 | 538 |   \item @{term "\<lambda>(x,y). P y x"}  
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 539 |   \item @{term "{ [x::int] | x. x \<le> -2 }"}  
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 540 |   \end{itemize}
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 541 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 542 | 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 | 543 | may omit parts of it by default. If you want to see all of it, you | 
| 122 | 544 | can use the following ML-function to set the printing depth to a higher | 
| 545 | value: | |
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 546 | |
| 75 | 547 |   @{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 | 548 |   \end{exercise}
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
34diff
changeset | 549 | |
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 550 |   The antiquotation @{text "@{prop \<dots>}"} constructs terms of propositional type, 
 | 
| 50 | 551 |   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 | 552 | Consider for example the pairs | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 553 | |
| 126 | 554 | @{ML_response [display,gray] "(@{term \"P x\"}, @{prop \"P x\"})" 
 | 
| 555 | "(Free (\"P\", \<dots>) $ Free (\"x\", \<dots>), | |
| 149 | 556 | 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 | 557 | |
| 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 | 558 | 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 | 559 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
71diff
changeset | 560 |   @{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 | 561 | "(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 | 562 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
71diff
changeset | 563 | 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 | 564 | |
| 89 | 565 |   Types can be constructed using the antiquotation @{text "@{typ \<dots>}"}. For example:
 | 
| 39 
631d12c25bde
substantial changes to the antiquotations (preliminary version)
 Christian Urban <urbanc@in.tum.de> parents: 
34diff
changeset | 566 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
71diff
changeset | 567 |   @{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 | 568 | |
| 19 
34b93dbf8c3c
some tuning in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
15diff
changeset | 569 |   \begin{readmore}
 | 
| 
34b93dbf8c3c
some tuning in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
15diff
changeset | 570 |   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 | 571 | definition and many useful operations are implemented | 
| 
ef778679d3e0
added a section about combinators
 Christian Urban <urbanc@in.tum.de> parents: 
75diff
changeset | 572 |   in @{ML_file "Pure/type.ML"}.
 | 
| 19 
34b93dbf8c3c
some tuning in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
15diff
changeset | 573 |   \end{readmore}
 | 
| 47 
4daf913fdbe1
hakked latex so that it does not display ML {* *}; general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
43diff
changeset | 574 | *} | 
| 19 
34b93dbf8c3c
some tuning in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
15diff
changeset | 575 | |
| 
34b93dbf8c3c
some tuning in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
15diff
changeset | 576 | |
| 156 | 577 | 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 | 578 | |
| 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 579 | text {*
 | 
| 81 | 580 | 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 | 581 | 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 | 582 | 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 | 583 |   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 | 584 |   @{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 | 585 | |
| 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 | 586 | *} | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 587 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 588 | ML{*fun make_imp P Q =
 | 
| 131 | 589 | let | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 590 |   val x = Free ("x", @{typ nat})
 | 
| 131 | 591 | in | 
| 592 | Logic.all x (Logic.mk_implies (P $ x, Q $ x)) | |
| 593 | 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 | 594 | |
| 
02f76f1b6e7b
added positions to anti-quotations; removed old antiquotation_setup; tuned the text a bit
 Christian Urban <urbanc@in.tum.de> parents: 
42diff
changeset | 595 | text {*
 | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 596 |   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 | 597 |   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 | 598 | *} | 
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 599 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 600 | 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 | 601 | |
| 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 | 602 | text {*
 | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 603 |   To see this apply @{text "@{term S}"} and @{text "@{term T}"} 
 | 
| 75 | 604 |   to both functions. With @{ML make_imp} we 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 | 605 | the given arguments | 
| 65 
c8e9a4f97916
tuned and added a section about creating keyword files
 Christian Urban <urbanc@in.tum.de> parents: 
60diff
changeset | 606 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 607 |   @{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 | 608 | "Const \<dots> $ | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 609 | Abs (\"x\", Type (\"nat\",[]), | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 610 | 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 | 611 | |
| 81 | 612 |   whereas with @{ML make_wrong_imp} we 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 | 613 |   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 | 614 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 615 |   @{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 | 616 | "Const \<dots> $ | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 617 | Abs (\"x\", \<dots>, | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 618 | Const \<dots> $ (Const \<dots> $ (Free (\"P\",\<dots>) $ \<dots>)) $ | 
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 619 | (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 | 620 | |
| 190 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 621 |   (FIXME: handy functions for constructing terms: @{ML list_comb}, @{ML lambda}, 
 | 
| 
ca0ac2e75f6d
more one the simple-inductive chapter
 Christian Urban <urbanc@in.tum.de> parents: 
189diff
changeset | 622 |   @{ML subst_free})
 | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 623 | *} | 
| 49 | 624 | |
| 625 | ||
| 626 | text {*
 | |
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 627 |   \begin{readmore}
 | 
| 89 | 628 |   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 | 629 |   @{ML_file "HOL/Tools/hologic.ML"} that make such manual constructions of terms 
 | 
| 49 | 630 |   and types easier.\end{readmore}
 | 
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 631 | |
| 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 632 | 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 | 633 | *} | 
| 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 634 | |
| 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 635 | text {*
 | 
| 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 636 | |
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 637 |   \begin{exercise}\label{fun:revsum}
 | 
| 58 | 638 |   Write a function @{text "rev_sum : term -> term"} that takes a
 | 
| 122 | 639 |   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 | 640 |   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 | 641 |   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 | 642 | 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 | 643 |   \end{exercise}
 | 
| 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 644 | |
| 15 
9da9ba2b095b
added a solution section and some other minor additions
 Christian Urban <urbanc@in.tum.de> parents: 
14diff
changeset | 645 |   \begin{exercise}\label{fun:makesum}
 | 
| 11 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 646 | Write a function which takes two terms representing natural numbers | 
| 75 | 647 |   in unary notation (like @{term "Suc (Suc (Suc 0))"}), and produce the
 | 
| 11 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 648 | number representing their sum. | 
| 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 649 |   \end{exercise}
 | 
| 
733614e236a3
tuned and updated antquote_setup.ML
 Christian Urban <urbanc@in.tum.de> parents: 
10diff
changeset | 650 | |
| 122 | 651 | There are a few subtle issues with constants. They usually crop up when | 
| 149 | 652 | pattern matching terms or types, or when constructing them. While it is perfectly ok | 
| 122 | 653 |   to write the function @{text is_true} as follows
 | 
| 654 | *} | |
| 655 | ||
| 656 | ML{*fun is_true @{term True} = true
 | |
| 657 | | is_true _ = false*} | |
| 658 | ||
| 659 | text {*
 | |
| 660 |   this does not work for picking out @{text "\<forall>"}-quantified terms. Because
 | |
| 661 | the function | |
| 662 | *} | |
| 663 | ||
| 664 | ML{*fun is_all (@{term All} $ _) = true
 | |
| 665 | | is_all _ = false*} | |
| 666 | ||
| 667 | text {* 
 | |
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 668 |   will not correctly match the formula @{prop "\<forall>x::nat. P x"}: 
 | 
| 122 | 669 | |
| 670 |   @{ML_response [display,gray] "is_all @{term \"\<forall>x::nat. P x\"}" "false"}
 | |
| 671 | ||
| 672 |   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 | 673 |   fixes the type of the constant @{term "All"} to be @{typ "('a \<Rightarrow> bool) \<Rightarrow> bool"} for 
 | 
| 122 | 674 |   an arbitrary, but fixed type @{typ "'a"}. A properly working alternative 
 | 
| 675 | for this function is | |
| 676 | *} | |
| 677 | ||
| 678 | ML{*fun is_all (Const ("All", _) $ _) = true
 | |
| 679 | | is_all _ = false*} | |
| 680 | ||
| 681 | text {* 
 | |
| 682 | because now | |
| 683 | ||
| 684 | @{ML_response [display,gray] "is_all @{term \"\<forall>x::nat. P x\"}" "true"}
 | |
| 685 | ||
| 149 | 686 | matches correctly (the first wildcard in the pattern matches any type and the | 
| 687 | second any term). | |
| 122 | 688 | |
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 689 | However there is still a problem: consider the similar function that | 
| 131 | 690 |   attempts to pick out @{text "Nil"}-terms:
 | 
| 122 | 691 | *} | 
| 692 | ||
| 693 | ML{*fun is_nil (Const ("Nil", _)) = true
 | |
| 694 | | is_nil _ = false *} | |
| 695 | ||
| 696 | text {* 
 | |
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 697 |   Unfortunately, also this function does \emph{not} work as expected, since
 | 
| 122 | 698 | |
| 699 |   @{ML_response [display,gray] "is_nil @{term \"Nil\"}" "false"}
 | |
| 700 | ||
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 701 | The problem is that on the ML-level the name of a constant is more | 
| 149 | 702 |   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 | 703 |   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 | 704 |   by @{ML "Const (\"All\", some_type)" for some_type}. However, if you look at
 | 
| 122 | 705 | |
| 706 |   @{ML_response [display,gray] "@{term \"Nil\"}" "Const (\"List.list.Nil\", \<dots>)"}
 | |
| 707 | ||
| 131 | 708 |   the name of the constant @{text "Nil"} depends on the theory in which the
 | 
| 128 | 709 |   term constructor is defined (@{text "List"}) and also in which datatype
 | 
| 710 |   (@{text "list"}). Even worse, some constants have a name involving
 | |
| 711 |   type-classes. Consider for example the constants for @{term "zero"} and
 | |
| 131 | 712 |   \mbox{@{text "(op *)"}}:
 | 
| 122 | 713 | |
| 714 |   @{ML_response [display,gray] "(@{term \"0::nat\"}, @{term \"op *\"})" 
 | |
| 126 | 715 | "(Const (\"HOL.zero_class.zero\", \<dots>), | 
| 122 | 716 | Const (\"HOL.times_class.times\", \<dots>))"} | 
| 717 | ||
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 718 | 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 | 719 |   @{ML "Const (\"List.list.Nil\", some_type)" for some_type}, for referring to or
 | 
| 122 | 720 |   matching against @{text "Nil"}, this would make the code rather brittle. 
 | 
| 721 | The reason is that the theory and the name of the datatype can easily change. | |
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 722 | To make the code more robust, it is better to use the antiquotation | 
| 122 | 723 |   @{text "@{const_name \<dots>}"}. With this antiquotation you can harness the 
 | 
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 724 | variable parts of the constant's name. Therefore a functions for | 
| 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 725 | 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 | 726 | be written as follows. | 
| 122 | 727 | *} | 
| 728 | ||
| 729 | ML{*fun is_nil_or_all (Const (@{const_name "Nil"}, _)) = true
 | |
| 730 |   | is_nil_or_all (Const (@{const_name "All"}, _) $ _) = true
 | |
| 731 | | is_nil_or_all _ = false *} | |
| 732 | ||
| 733 | text {*
 | |
| 131 | 734 | Occasional 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 | 735 |   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 | 736 |   @{ML Long_Name.base_name}. For example:
 | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 737 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 738 |   @{ML_response [display,gray] "Sign.extern_const @{theory} \"List.list.Nil\"" "\"Nil\""}
 | 
| 122 | 739 | |
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 740 |   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 | 741 |   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 | 742 | strips off all qualifiers. | 
| 122 | 743 | |
| 744 |   \begin{readmore}
 | |
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 745 |   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 | 746 |   functions about signatures in @{ML_file "Pure/sign.ML"}.
 | 
| 122 | 747 |   \end{readmore}
 | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 748 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 749 | 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 | 750 | 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 | 751 | 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 | 752 | 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 | 753 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 754 | *} | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 755 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 756 | 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 | 757 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 758 | 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 | 759 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 760 | 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 | 761 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 762 | text {*
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 763 |   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 | 764 | 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 | 765 |   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 | 766 | *} | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 767 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 768 | 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 | 769 | (case t of | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 770 |      @{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 | 771 | | 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 | 772 | | _ => t)*} | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 773 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 774 | text {*
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 775 | An example as follows: | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 776 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 777 | @{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 | 778 | "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 | 779 | "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 | 780 | $ 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 | 781 | |
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 782 | (FIXME: readmore about types) | 
| 122 | 783 | *} | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 784 | |
| 122 | 785 | |
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 786 | section {* Type-Checking *}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 787 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 788 | text {* 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 789 | |
| 131 | 790 |   You can freely construct and manipulate @{ML_type "term"}s and @{ML_type
 | 
| 791 | typ}es, since they are just arbitrary unchecked trees. However, you | |
| 792 | eventually want to see if a term is well-formed, or type-checks, relative to | |
| 793 |   a theory.  Type-checking is done via the function @{ML cterm_of}, which
 | |
| 794 |   converts a @{ML_type term} into a @{ML_type cterm}, a \emph{certified}
 | |
| 795 |   term. Unlike @{ML_type term}s, which are just trees, @{ML_type "cterm"}s are
 | |
| 796 | abstract objects that are guaranteed to be type-correct, and they can only | |
| 797 | be constructed via ``official interfaces''. | |
| 798 | ||
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 799 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 800 | 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 | 801 |   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 | 802 | For example you can write: | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 803 | |
| 149 | 804 |   @{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 | 805 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 806 | This can also be written with an antiquotation: | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 807 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 808 |   @{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 | 809 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 810 | Attempting to obtain the certified term for | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 811 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 812 |   @{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 | 813 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 814 | 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 | 815 | example that type-checks is: | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 816 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 817 | @{ML_response_fake [display,gray] 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 818 | "let | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 819 |   val natT = @{typ \"nat\"}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 820 |   val zero = @{term \"0::nat\"}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 821 | in | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 822 |   cterm_of @{theory} 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 823 |       (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 | 824 | end" "0 + 0"} | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 825 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 826 | In Isabelle not just terms need to be certified, but also types. For example, | 
| 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 | 827 |   you obtain the certified type for the Isablle type @{typ "nat \<Rightarrow> bool"} on 
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 828 | 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 | 829 | |
| 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 830 |   @{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 | 831 |   "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 | 832 | "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 | 833 | |
| 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 834 |   \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 | 835 |   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 | 836 |   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 | 837 |   \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 | 838 | |
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 839 |   \begin{exercise}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 840 |   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 | 841 | result that type-checks. | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 842 |   \end{exercise}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 843 | |
| 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 | 844 | Remember Isabelle foolows the Church-style typing for terms, i.e.~a term contains | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 845 | 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 | 846 | 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 | 847 |   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 | 848 | type of a term. Consider for example: | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 849 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 850 |   @{ML_response [display,gray] 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 851 |   "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 | 852 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 853 | To calculate the type, this function traverses the whole term and will | 
| 128 | 854 | detect any typing inconsistency. For examle changing the type of the variable | 
| 149 | 855 |   @{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 | 856 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 857 |   @{ML_response_fake [display,gray] 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 858 |   "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 | 859 | "*** 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 | 860 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 861 | Since the complete traversal might sometimes be too costly and | 
| 149 | 862 |   not necessary, there is the function @{ML fastype_of}, which 
 | 
| 863 | also returns the type of a term. | |
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 864 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 865 |   @{ML_response [display,gray] 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 866 |   "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 | 867 | |
| 177 | 868 | 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 | 869 | can see this in the following example | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 870 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 871 |    @{ML_response [display,gray] 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 872 |   "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 | 873 | |
| 149 | 874 | where no error is detected. | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 875 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 876 | 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 | 877 | 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 | 878 | 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 | 879 |   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 | 880 | complete type. An example is as follows: | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 881 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 882 |   @{ML_response_fake [display,gray] 
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 883 | "let | 
| 126 | 884 |   val c = Const (@{const_name \"plus\"}, dummyT) 
 | 
| 885 |   val o = @{term \"1::nat\"} 
 | |
| 886 | val v = Free (\"x\", dummyT) | |
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 887 | in | 
| 126 | 888 |   Syntax.check_term @{context} (c $ o $ v)
 | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 889 | end" | 
| 126 | 890 | "Const (\"HOL.plus_class.plus\", \"nat \<Rightarrow> nat \<Rightarrow> nat\") $ | 
| 891 | 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 | 892 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 893 |   Instead of giving explicitly the type for the constant @{text "plus"} and the free 
 | 
| 149 | 894 |   variable @{text "x"}, the type-inference filles in the missing information.
 | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 895 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 896 |   \begin{readmore}
 | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 897 |   See @{ML_file "Pure/Syntax/syntax.ML"} where more functions about reading,
 | 
| 160 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 898 | checking and pretty-printing of terms are defined. Fuctions related to the | 
| 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 899 |   type inference are implemented in @{ML_file "Pure/type.ML"} and 
 | 
| 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 900 |   @{ML_file "Pure/type_infer.ML"}. 
 | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 901 |   \end{readmore}
 | 
| 162 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 902 | |
| 
3fb9f820a294
some additions to the simplifier section and general tuning
 Christian Urban <urbanc@in.tum.de> parents: 
161diff
changeset | 903 | (FIXME: say something about sorts) | 
| 124 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 904 | *} | 
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 905 | |
| 
0b9fa606a746
added to the first-steps section
 Christian Urban <urbanc@in.tum.de> parents: 
123diff
changeset | 906 | |
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 907 | section {* Theorems *}
 | 
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 908 | |
| 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 909 | text {*
 | 
| 50 | 910 |   Just like @{ML_type cterm}s, theorems are abstract objects of type @{ML_type thm} 
 | 
| 122 | 911 | that can only be build 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 | 912 | 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 | 913 |   \cite{GordonMilnerWadsworth79}.
 | 
| 107 
258ce361ba1b
polished and more material in the tactic chapter
 Christian Urban <urbanc@in.tum.de> parents: 
104diff
changeset | 914 | |
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 915 | |
| 78 
ef778679d3e0
added a section about combinators
 Christian Urban <urbanc@in.tum.de> parents: 
75diff
changeset | 916 | 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 | 917 | statement: | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 918 | *} | 
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 919 | |
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 920 | lemma | 
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 921 | 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 | 922 | 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 | 923 | shows "Q t" (*<*)oops(*>*) | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 924 | |
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 925 | text {*
 | 
| 185 | 926 | The corresponding ML-code is as follows: | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 927 | |
| 72 
7b8c4fe235aa
added an antiquotation option [gray] for gray boxes around displays
 Christian Urban <urbanc@in.tum.de> parents: 
71diff
changeset | 928 | @{ML_response_fake [display,gray]
 | 
| 42 
cd612b489504
tuned mostly antiquotation and text
 Christian Urban <urbanc@in.tum.de> parents: 
41diff
changeset | 929 | "let | 
| 138 
e4d8dfb7e34a
included comment from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
136diff
changeset | 930 |   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 | 931 |   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 | 932 | |
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 933 | val Pt_implies_Qt = | 
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 934 | assume assm1 | 
| 138 
e4d8dfb7e34a
included comment from Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
136diff
changeset | 935 |         |> forall_elim @{cterm \"t::nat\"};
 | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 936 | |
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 937 | 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 | 938 | in | 
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 939 | Qt | 
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 940 | |> implies_intr assm2 | 
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 941 | |> implies_intr assm1 | 
| 48 
609f9ef73494
fixed FIXME's in fake responses
 Christian Urban <urbanc@in.tum.de> parents: 
47diff
changeset | 942 | 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 | 943 | |
| 21 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 944 | This code-snippet constructs the following proof: | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 945 | |
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 946 | \[ | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 947 |   \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 | 948 |     {\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 | 949 |        {\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 | 950 |           {\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 | 951 |                  {\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 | 952 | & | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 953 |            \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 | 954 | } | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 955 | } | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 956 | } | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 957 | \] | 
| 
2356e5c70d98
added a proof and tuned the rest
 Christian Urban <urbanc@in.tum.de> parents: 
20diff
changeset | 958 | |
| 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 | 959 | 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 | 960 | yet stored in Isabelle's theorem database. So it cannot be referenced later | 
| 128 | 961 |   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 | 962 | |
| 13 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 963 |   \begin{readmore}
 | 
| 50 | 964 |   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 | 965 |   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 | 966 |   @{ML_file "Pure/thm.ML"}. 
 | 
| 
2b07da8b310d
polished and added a subdirectory for the recipes
 Christian Urban <urbanc@in.tum.de> parents: 
12diff
changeset | 967 |   \end{readmore}
 | 
| 12 
2f1736cb8f26
various changes by Alex and Christian
 Christian Urban <urbanc@in.tum.de> parents: 
11diff
changeset | 968 | |
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 969 | (FIXME: how to add case-names to goal states - maybe in the next section) | 
| 10 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 970 | *} | 
| 
df09e49b19bf
many changes in the FirstSteps section
 Christian Urban <urbanc@in.tum.de> parents: 
6diff
changeset | 971 | |
| 123 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 972 | section {* Theorem Attributes *}
 | 
| 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 973 | |
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 974 | text {*
 | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 975 |   Theorem attributes are @{text "[simp]"}, @{text "[OF \<dots>]"}, @{text
 | 
| 136 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 976 |   "[symmetric]"} and so on. Such attributes are \emph{neither} tags \emph{nor} flags
 | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 977 | annotated to theorems, but functions that do further processing once a | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 978 | theorem is proven. In particular, it is not possible to find out | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 979 | what are all theorems that have a given attribute in common, unless of course | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 980 | the function behind the attribute stores the theorems in a retrivable | 
| 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 | 981 | datastructure. | 
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 982 | |
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 983 | If you want to print out all currently known attributes a theorem | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 984 | can have, you can use the function: | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 985 | |
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 986 |   @{ML_response_fake [display,gray] "Attrib.print_attributes @{theory}" 
 | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 987 | "COMP: direct composition with rules (no lifting) | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 988 | HOL.dest: declaration of Classical destruction rule | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 989 | HOL.elim: declaration of Classical elimination rule | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 990 | \<dots>"} | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 991 | |
| 136 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 992 | 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 | 993 |   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 | 994 | 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 | 995 | this attribute is | 
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 996 | *} | 
| 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 997 | |
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 998 | 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 | 999 | |
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1000 | text {* 
 | 
| 136 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1001 |   where the function @{ML "Thm.rule_attribute"} expects a function taking a 
 | 
| 149 | 1002 |   context (which we ignore in the code above) and a theorem (@{text thm}), and 
 | 
| 1003 |   returns another theorem (namely @{text thm} resolved with the lemma 
 | |
| 1004 |   @{thm [source] sym}: @{thm sym[no_vars]}). The function @{ML "Thm.rule_attribute"} then returns 
 | |
| 156 | 1005 | an attribute. | 
| 136 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1006 | |
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1007 | Before we can use the attribute, we need to set it up. This can be done | 
| 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 | 1008 |   using the function @{ML Attrib.setup} as follows.
 | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1009 | *} | 
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1010 | |
| 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 | 1011 | setup %gray {* Attrib.setup @{binding "my_sym"} 
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1012 | (Scan.succeed my_symmetric) "applying the sym rule"*} | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1013 | |
| 136 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1014 | text {*
 | 
| 149 | 1015 |   The attribute does not expect any further arguments (unlike @{text "[OF
 | 
| 1016 | \<dots>]"}, for example, which can take a list of theorems as argument). Therefore | |
| 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 | 1017 |   we use the parser @{ML Scan.succeed}. Later on we will also consider
 | 
| 149 | 1018 |   attributes taking further arguments. An example for the attribute @{text
 | 
| 1019 | "[my_sym]"} is the proof | |
| 136 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1020 | *} | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1021 | |
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1022 | 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 | 1023 | |
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1024 | text {*
 | 
| 136 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1025 |   which stores the theorem @{thm test} under the name @{thm [source] test}. We 
 | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1026 | can also use the attribute when referring to this theorem. | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1027 | |
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1028 |   \begin{isabelle}
 | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1029 |   \isacommand{thm}~@{text "test[my_sym]"}\\
 | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1030 |   @{text "> "}~@{thm test[my_sym]}
 | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1031 |   \end{isabelle}
 | 
| 
58277de8493c
to be in sync with Sascha
 Christian Urban <urbanc@in.tum.de> parents: 
134diff
changeset | 1032 | |
| 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 | 1033 |   The purpose of @{ML Thm.rule_attribute} is to directly manipulate theorems.
 | 
| 
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 | 1034 | Another usage of attributes is to add and delete theorems from stored 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 | 1035 |   For example the attribute @{text "[simp]"} adds or deletes a theorem from 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 | 1036 |   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 | 1037 | 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 | 1038 | of theorems. | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1039 | *} | 
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1040 | |
| 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 | 1041 | ML{*val my_thms = ref ([]: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 | 1042 | |
| 
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 | 1043 | 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 | 1044 | A word of warning: such references must not be used in any code that | 
| 
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 | 1045 | is meant to be more than just for testing purposes! Here it is only used | 
| 
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 | 1046 | to illustrate matters. We will show later how to store data properly without | 
| 
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 | 1047 |   using references. The function @{ML Thm.declaration_attribute} expects us to 
 | 
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1048 | provide two functions that add and delete theorems from this list. | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1049 | 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 | 1050 | *} | 
| 
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 | 1051 | |
| 
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 | 1052 | ML{*fun my_thms_add thm 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 | 1053 | (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 | 1054 | |
| 
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 | 1055 | fun my_thms_del thm 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 | 1056 | (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 | 1057 | |
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1058 | 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 | 1059 | These functions take a theorem and a context and, for what we are explaining | 
| 156 | 1060 | 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 | 1061 |   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 | 1062 |   adds a theorem if it is not already included in the list, and @{ML
 | 
| 
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 | 1063 |   Thm.del_thm} deletes one. Both functions use the predicate @{ML
 | 
| 
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 | 1064 | Thm.eq_thm_prop} which compares theorems according to their proved | 
| 
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 | 1065 | propositions (modulo alpha-equivalence). | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1066 | |
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1067 | |
| 156 | 1068 | You can turn both functions into attributes using | 
| 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 | 1069 | *} | 
| 
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 | 1070 | |
| 
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 | 1071 | ML{*val my_add = Thm.declaration_attribute my_thms_add
 | 
| 
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 | 1072 | val my_del = Thm.declaration_attribute 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 | 1073 | |
| 
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 | 1074 | 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 | 1075 | 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 | 1076 | *} | 
| 
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 | 1077 | |
| 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 | 1078 | setup %gray {* Attrib.setup @{binding "my_thms"}
 | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1079 | (Attrib.add_del my_add my_del) "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 | 1080 | |
| 
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 | 1081 | 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 | 1082 |   Now if you prove the lemma attaching the attribute @{text "[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 | 1083 | *} | 
| 
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 | 1084 | |
| 
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 | 1085 | 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 | 1086 | |
| 
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 | 1087 | 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 | 1088 | then you can see the lemma is added to the initially empty 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 | 1089 | |
| 
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 | 1090 |   @{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 | 1091 | "!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 | 1092 | |
| 160 
cc9359bfacf4
redefined the functions warning and tracing in order to properly match more antiquotations
 Christian Urban <urbanc@in.tum.de> parents: 
158diff
changeset | 1093 |   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 | 1094 | *} | 
| 
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 | 1095 | |
| 
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 | 1096 | declare test[my_thms] trueI_2[my_thms add] | 
| 
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 | 1097 | |
| 
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 | 1098 | 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 | 1099 |   The @{text "add"} is the default operation and does not need
 | 
| 
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 | 1100 | to be given. This declaration will cause the theorem list to be | 
| 
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 | 1101 | updated 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 | 1102 | |
| 
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 | 1103 |   @{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 | 1104 | "!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 | 1105 | "[\"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 | 1106 | |
| 
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 | 1107 |   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 | 1108 |   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 | 1109 | 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 | 1110 | *} | 
| 
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 | 1111 | |
| 
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 | 1112 | 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 | 1113 | |
| 156 | 1114 | 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 | 1115 | |
| 
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 | 1116 |   @{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 | 1117 | "!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 | 1118 | "[\"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 | 1119 | |
| 
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 | 1120 |   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 | 1121 | 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 | 1122 | 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 | 1123 | |
| 156 | 1124 | 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 | 1125 |   \emph{not} the received way of doing such things. The received way is to 
 | 
| 
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 | 1126 |   start a ``data slot'' in a context by using the functor @{text GenericDataFun}:
 | 
| 
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 | 1127 | *} | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1128 | |
| 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 | 1129 | ML {*structure Data = GenericDataFun
 | 
| 
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 | 1130 | (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 | 1131 | 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 | 1132 | 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 | 1133 | 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 | 1134 | |
| 
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 | 1135 | text {*
 | 
| 156 | 1136 |   To use this data slot, you only have to change @{ML my_thms_add} and
 | 
| 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 | 1137 |   @{ML my_thms_del} to:
 | 
| 
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 | 1138 | *} | 
| 
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 | 1139 | |
| 
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 | 1140 | ML{*val thm_add = Data.map o 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 | 1141 | val thm_del = Data.map o Thm.del_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 | 1142 | |
| 
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 | 1143 | 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 | 1144 |   where @{ML Data.map} updates the data appropriately in the context. Since storing
 | 
| 
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 | 1145 |   theorems in a special list is such a common task, there is the functor @{text 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 | 1146 | which does most of the work for you. To obtain such a named theorem lists, you just | 
| 
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 | 1147 | declare | 
| 
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 | 1148 | *} | 
| 
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 | 1149 | |
| 
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 | 1150 | 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 | 1151 | (val name = "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 | 1152 | val description = "Rules for 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 | 1153 | |
| 
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 | 1154 | 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 | 1155 |   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 | 1156 | *} | 
| 
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 | 1157 | |
| 177 | 1158 | 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 | 1159 | |
| 
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 | 1160 | 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 | 1161 | 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 | 1162 |   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 | 1163 | 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 | 1164 | modify the theorems. | 
| 133 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1165 | |
| 157 
76cdc8f562fc
added more to the simplifier section
 Christian Urban <urbanc@in.tum.de> parents: 
156diff
changeset | 1166 | Furthermore, the facts are made available on the user-level under the dynamic | 
| 156 | 1167 |   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 | 1168 |   @{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 | 1169 | *} | 
| 
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 | 1170 | |
| 
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 | 1171 | 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 | 1172 | 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 | 1173 | 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 | 1174 | |
| 
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 | 1175 | 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 | 1176 | |
| 
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 | 1177 | 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 | 1178 | |
| 
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 | 1179 | 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 | 1180 | |
| 
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 | 1181 |   \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 | 1182 |   \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 | 1183 |   @{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 | 1184 |   @{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 | 1185 |   \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 | 1186 | |
| 156 | 1187 |   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 | 1188 |   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 | 1189 | |
| 
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 | 1190 |   @{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 | 1191 | |
| 
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 | 1192 |   \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 | 1193 |   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 | 1194 |   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 | 1195 | 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 | 1196 |   \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 | 1197 | |
| 156 | 1198 |   (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 | 1199 | |
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1200 | |
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1201 |   \begin{readmore}
 | 
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1202 |   FIXME: @{ML_file "Pure/more_thm.ML"} @{ML_file "Pure/Isar/attrib.ML"}
 | 
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1203 |   \end{readmore}
 | 
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1204 | *} | 
| 
3e94ccc0f31e
polishing and start of the section about attributes
 Christian Urban <urbanc@in.tum.de> parents: 
132diff
changeset | 1205 | |
| 127 
74846cb0fff9
updated and added two tentative recipes
 Christian Urban <urbanc@in.tum.de> parents: 
126diff
changeset | 1206 | |
| 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 | 1207 | 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 | 1208 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1209 | 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 | 1210 |   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 | 1211 | a theorem attribute known to Isabelle. What happens behind the scenes | 
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 1212 |   is that \isacommand{setup} expects a functions of type 
 | 
| 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 1213 |   @{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 | 1214 | 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 | 1215 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1216 | 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 | 1217 | 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 | 1218 |   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 | 1219 |   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 | 1220 |   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 | 1221 | *} | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1222 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 1223 | 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 | 1224 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1225 | 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 | 1226 |   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 | 1227 | 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 | 1228 |   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 | 1229 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1230 |   \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 | 1231 |   \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 | 1232 |   @{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 | 1233 |   \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 | 1234 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1235 |   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 | 1236 | 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 | 1237 | 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 | 1238 |   \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 | 1239 | *} | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1240 | |
| 186 
371e4375c994
made the Ackermann function example safer and included suggestions from MW
 Christian Urban <urbanc@in.tum.de> parents: 
185diff
changeset | 1241 | 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 | 1242 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1243 | 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 | 1244 | 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 | 1245 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1246 |   \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 | 1247 |   \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 | 1248 |   @{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 | 1249 |   \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 | 1250 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1251 |   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 | 1252 | |
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1253 |   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 | 1254 |   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 | 1255 |   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 | 1256 |   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 | 1257 | 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 | 1258 | *} | 
| 
fb8f22dd8ad0
adapted to latest Attrib.setup changes and more work on the simple induct chapter
 Christian Urban <urbanc@in.tum.de> parents: 
177diff
changeset | 1259 | |
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1260 | 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 | 1261 | |
| 126 | 1262 | text {*
 | 
| 1263 | There are theories, proof contexts and local theories (in this order, if you | |
| 1264 | want to order them). | |
| 1265 | ||
| 1266 | In contrast to an ordinary theory, which simply consists of a type | |
| 1267 | signature, as well as tables for constants, axioms and theorems, a local | |
| 1268 | theory also contains additional context information, such as locally fixed | |
| 1269 | variables and local assumptions that may be used by the package. The type | |
| 1270 |   @{ML_type local_theory} is identical to the type of \emph{proof contexts}
 | |
| 1271 |   @{ML_type "Proof.context"}, although not every proof context constitutes a
 | |
| 1272 | 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 | 1273 | *} | 
| 126 | 1274 | |
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1275 | 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 | 1276 | |
| 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 1277 | text {* @{ML PureThy.add_thms_dynamic} *}
 | 
| 
460bc2ee14e9
some polishing in the first-steps chapter
 Christian Urban <urbanc@in.tum.de> parents: 
122diff
changeset | 1278 | |
| 100 | 1279 | |
| 75 | 1280 | |
| 126 | 1281 | (* FIXME: some code below *) | 
| 89 | 1282 | |
| 1283 | (*<*) | |
| 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 | 1284 | (* | 
| 89 | 1285 | 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 | 1286 |  Sign.add_consts_i [(Binding"bar", @{typ "nat"},NoSyn)] 
 | 
| 89 | 1287 | *} | 
| 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 | 1288 | *) | 
| 89 | 1289 | lemma "bar = (1::nat)" | 
| 1290 | oops | |
| 1291 | ||
| 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 | 1292 | (* | 
| 89 | 1293 | setup {*   
 | 
| 1294 |   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 | 1295 |  #> PureThy.add_defs false [((@{binding "foo_def"}, 
 | 
| 89 | 1296 |        Logic.mk_equals (Const ("FirstSteps.foo", @{typ "nat"}), @{term "1::nat"})), [])] 
 | 
| 1297 | #> snd | |
| 1298 | *} | |
| 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 | 1299 | *) | 
| 
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 | 1300 | (* | 
| 89 | 1301 | lemma "foo = (1::nat)" | 
| 1302 | apply(simp add: foo_def) | |
| 1303 | done | |
| 1304 | ||
| 1305 | 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 | 1306 | *) | 
| 89 | 1307 | (*>*) | 
| 1308 | ||
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1309 | section {* Pretty-Printing (TBD) *}
 | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1310 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1311 | 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 | 1312 |   @{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 | 1313 |   @{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 | 1314 |   @{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 | 1315 |   @{ML Pretty.chunks}
 | 
| 153 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1316 | *} | 
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1317 | |
| 
c22b507e1407
general polishing; added versioning to the document
 Christian Urban <urbanc@in.tum.de> parents: 
151diff
changeset | 1318 | section {* Misc (TBD) *}
 | 
| 92 | 1319 | |
| 1320 | ML {*DatatypePackage.get_datatype @{theory} "List.list"*}
 | |
| 1321 | ||
| 2 
978a3c2ed7ce
split the document into smaller pieces;
 Christian Urban <urbanc@in.tum.de> parents: diff
changeset | 1322 | end |