ProgTutorial/FirstSteps.thy
changeset 315 de49d5780f57
parent 314 79202e2eab6a
child 316 74f0a06f751f
equal deleted inserted replaced
314:79202e2eab6a 315:de49d5780f57
   113 section {* Debugging and Printing\label{sec:printing} *}
   113 section {* Debugging and Printing\label{sec:printing} *}
   114 
   114 
   115 text {*
   115 text {*
   116   During development you might find it necessary to inspect some data in your
   116   During development you might find it necessary to inspect some data in your
   117   code. This can be done in a ``quick-and-dirty'' fashion using the function
   117   code. This can be done in a ``quick-and-dirty'' fashion using the function
   118   @{ML [index] "writeln"}. For example
   118   @{ML_ind [index] "writeln"}. For example
   119 
   119 
   120   @{ML_response_fake [display,gray] "writeln \"any string\"" "\"any string\""}
   120   @{ML_response_fake [display,gray] "writeln \"any string\"" "\"any string\""}
   121 
   121 
   122   will print out @{text [quotes] "any string"} inside the response buffer of
   122   will print out @{text [quotes] "any string"} inside the response buffer of
   123   Isabelle.  This function expects a string as argument. If you develop under
   123   Isabelle.  This function expects a string as argument. If you develop under
   124   PolyML, then there is a convenient, though again ``quick-and-dirty'', method
   124   PolyML, then there is a convenient, though again ``quick-and-dirty'', method
   125   for converting values into strings, namely the function 
   125   for converting values into strings, namely the function 
   126   @{ML [index] makestring in PolyML}:
   126   @{ML_ind [index] makestring in PolyML}:
   127 
   127 
   128   @{ML_response_fake [display,gray] "writeln (PolyML.makestring 1)" "\"1\""} 
   128   @{ML_response_fake [display,gray] "writeln (PolyML.makestring 1)" "\"1\""} 
   129 
   129 
   130   However, @{ML makestring in PolyML} only works if the type of what
   130   However, @{ML makestring in PolyML} only works if the type of what
   131   is converted is monomorphic and not a function.
   131   is converted is monomorphic and not a function.
   132 
   132 
   133   The function @{ML "writeln"} should only be used for testing purposes,
   133   The function @{ML "writeln"} should only be used for testing purposes,
   134   because any output this function generates will be overwritten as soon as an
   134   because any output this function generates will be overwritten as soon as an
   135   error is raised. For printing anything more serious and elaborate, the
   135   error is raised. For printing anything more serious and elaborate, the
   136   function @{ML [index] tracing} is more appropriate. This function writes all
   136   function @{ML_ind [index] tracing} is more appropriate. This function writes all
   137   output into a separate tracing buffer. For example:
   137   output into a separate tracing buffer. For example:
   138 
   138 
   139   @{ML_response_fake [display,gray] "tracing \"foo\"" "\"foo\""}
   139   @{ML_response_fake [display,gray] "tracing \"foo\"" "\"foo\""}
   140 
   140 
   141   It is also possible to redirect the ``channel'' where the string @{text
   141   It is also possible to redirect the ``channel'' where the string @{text
   164 
   164 
   165   @{ML [display,gray] "redirect_tracing (TextIO.openOut \"foo.bar\")"}
   165   @{ML [display,gray] "redirect_tracing (TextIO.openOut \"foo.bar\")"}
   166 
   166 
   167   will cause that all tracing information is printed into the file @{text "foo.bar"}.
   167   will cause that all tracing information is printed into the file @{text "foo.bar"}.
   168 
   168 
   169   You can print out error messages with the function @{ML [index] error}; for 
   169   You can print out error messages with the function @{ML_ind [index] error}; for 
   170   example:
   170   example:
   171 
   171 
   172 @{ML_response_fake [display,gray] "if 0=1 then true else (error \"foo\")" 
   172 @{ML_response_fake [display,gray] "if 0=1 then true else (error \"foo\")" 
   173 "Exception- ERROR \"foo\" raised
   173 "Exception- ERROR \"foo\" raised
   174 At command \"ML\"."}
   174 At command \"ML\"."}
   175 
   175 
   176   This function raises the exception @{text ERROR}, which will then 
   176   This function raises the exception @{text ERROR}, which will then 
   177   be displayed by the infrastructure.
   177   be displayed by the infrastructure.
   178 
   178 
   179 
   179 
   180   (FIXME Mention how to work with @{ML [index] debug in Toplevel} and 
   180   (FIXME Mention how to work with @{ML_ind [index] debug in Toplevel} and 
   181   @{ML [index] profiling in Toplevel}.)
   181   @{ML_ind [index] profiling in Toplevel}.)
   182 *}
   182 *}
   183 
   183 
   184 (*
   184 (*
   185 ML {* reset Toplevel.debug *}
   185 ML {* reset Toplevel.debug *}
   186 
   186 
   199   Most often you want to inspect data of Isabelle's most basic data
   199   Most often you want to inspect data of Isabelle's most basic data
   200   structures, namely @{ML_type term}, @{ML_type cterm} and @{ML_type
   200   structures, namely @{ML_type term}, @{ML_type cterm} and @{ML_type
   201   thm}. Isabelle contains elaborate pretty-printing functions for printing
   201   thm}. Isabelle contains elaborate pretty-printing functions for printing
   202   them (see Section \ref{sec:pretty}), but for quick-and-dirty solutions they
   202   them (see Section \ref{sec:pretty}), but for quick-and-dirty solutions they
   203   are a bit unwieldy. One way to transform a term into a string is to use the
   203   are a bit unwieldy. One way to transform a term into a string is to use the
   204   function @{ML [index] string_of_term in Syntax} in structure @{ML_struct
   204   function @{ML_ind [index] string_of_term in Syntax} in structure @{ML_struct
   205   Syntax}, which we bind for more convenience to the toplevel.
   205   Syntax}, which we bind for more convenience to the toplevel.
   206 *}
   206 *}
   207 
   207 
   208 ML{*val string_of_term = Syntax.string_of_term*}
   208 ML{*val string_of_term = Syntax.string_of_term*}
   209 
   209 
   214   "string_of_term @{context} @{term \"1::nat\"}"
   214   "string_of_term @{context} @{term \"1::nat\"}"
   215   "\"\\^E\\^Fterm\\^E\\^E\\^Fconst\\^Fname=HOL.one_class.one\\^E1\\^E\\^F\\^E\\^E\\^F\\^E\""}
   215   "\"\\^E\\^Fterm\\^E\\^E\\^Fconst\\^Fname=HOL.one_class.one\\^E1\\^E\\^F\\^E\\^E\\^F\\^E\""}
   216 
   216 
   217   This produces a string corrsponding to the term @{term [show_types] "1::nat"} with some
   217   This produces a string corrsponding to the term @{term [show_types] "1::nat"} with some
   218   additional information encoded in it. The string can be properly printed by
   218   additional information encoded in it. The string can be properly printed by
   219   using either the function @{ML [index] writeln} or @{ML [index] tracing}:
   219   using either the function @{ML_ind [index] writeln} or @{ML_ind [index] tracing}:
   220 
   220 
   221   @{ML_response_fake [display,gray]
   221   @{ML_response_fake [display,gray]
   222   "writeln (string_of_term @{context} @{term \"1::nat\"})"
   222   "writeln (string_of_term @{context} @{term \"1::nat\"})"
   223   "\"1\""}
   223   "\"1\""}
   224 
   224 
   227   @{ML_response_fake [display,gray]
   227   @{ML_response_fake [display,gray]
   228   "tracing (string_of_term @{context} @{term \"1::nat\"})"
   228   "tracing (string_of_term @{context} @{term \"1::nat\"})"
   229   "\"1\""}
   229   "\"1\""}
   230 
   230 
   231   If there are more than one @{ML_type term}s to be printed, you can use the 
   231   If there are more than one @{ML_type term}s to be printed, you can use the 
   232   function @{ML [index] commas} to separate them.
   232   function @{ML_ind [index] commas} to separate them.
   233 *} 
   233 *} 
   234 
   234 
   235 ML{*fun string_of_terms ctxt ts =  
   235 ML{*fun string_of_terms ctxt ts =  
   236   commas (map (string_of_term ctxt) ts)*}
   236   commas (map (string_of_term ctxt) ts)*}
   237 
   237 
   241 
   241 
   242 ML{*fun string_of_cterm ctxt ct =  
   242 ML{*fun string_of_cterm ctxt ct =  
   243   string_of_term ctxt (term_of ct)*}
   243   string_of_term ctxt (term_of ct)*}
   244 
   244 
   245 text {*
   245 text {*
   246   In this example the function @{ML [index] term_of} extracts the @{ML_type
   246   In this example the function @{ML_ind [index] term_of} extracts the @{ML_type
   247   term} from a @{ML_type cterm}.  More than one @{ML_type cterm}s can again be
   247   term} from a @{ML_type cterm}.  More than one @{ML_type cterm}s can again be
   248   printed with @{ML [index] commas}.
   248   printed with @{ML_ind [index] commas}.
   249 *} 
   249 *} 
   250 
   250 
   251 ML{*fun string_of_cterms ctxt cts =  
   251 ML{*fun string_of_cterms ctxt cts =  
   252   commas (map (string_of_cterm ctxt) cts)*}
   252   commas (map (string_of_cterm ctxt) cts)*}
   253 
   253 
   254 text {*
   254 text {*
   255   The easiest way to get the string of a theorem is to transform it
   255   The easiest way to get the string of a theorem is to transform it
   256   into a @{ML_type cterm} using the function @{ML [index] crep_thm}. 
   256   into a @{ML_type cterm} using the function @{ML_ind [index] crep_thm}. 
   257 *}
   257 *}
   258 
   258 
   259 ML{*fun string_of_thm ctxt thm =
   259 ML{*fun string_of_thm ctxt thm =
   260   string_of_cterm ctxt (#prop (crep_thm thm))*}
   260   string_of_cterm ctxt (#prop (crep_thm thm))*}
   261 
   261 
   270   "tracing (string_of_thm @{context} @{thm conjI})"
   270   "tracing (string_of_thm @{context} @{thm conjI})"
   271   "\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q"}
   271   "\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q"}
   272 
   272 
   273   However, in order to improve the readability when printing theorems, we
   273   However, in order to improve the readability when printing theorems, we
   274   convert these schematic variables into free variables using the function
   274   convert these schematic variables into free variables using the function
   275   @{ML [index] import in Variable}. This is similar to statements like @{text
   275   @{ML_ind [index] import in Variable}. This is similar to statements like @{text
   276   "conjI[no_vars]"} from Isabelle's user-level.
   276   "conjI[no_vars]"} from Isabelle's user-level.
   277 *}
   277 *}
   278 
   278 
   279 ML{*fun no_vars ctxt thm =
   279 ML{*fun no_vars ctxt thm =
   280 let 
   280 let 
   320 "tracing (\"First half,\" ^ \"\\n\" ^ \"and second half.\")"
   320 "tracing (\"First half,\" ^ \"\\n\" ^ \"and second half.\")"
   321 "First half,
   321 "First half,
   322 and second half."}
   322 and second half."}
   323   
   323   
   324   To ease this kind of string manipulations, there are a number
   324   To ease this kind of string manipulations, there are a number
   325   of library functions. For example,  the function @{ML [index] cat_lines}
   325   of library functions. For example,  the function @{ML_ind [index] cat_lines}
   326   concatenates a list of strings and inserts newlines. 
   326   concatenates a list of strings and inserts newlines. 
   327 
   327 
   328   @{ML_response [display, gray]
   328   @{ML_response [display, gray]
   329   "cat_lines [\"foo\", \"bar\"]"
   329   "cat_lines [\"foo\", \"bar\"]"
   330   "\"foo\\nbar\""}
   330   "\"foo\\nbar\""}
   340   For beginners perhaps the most puzzling parts in the existing code of Isabelle are
   340   For beginners perhaps the most puzzling parts in the existing code of Isabelle are
   341   the combinators. At first they seem to greatly obstruct the
   341   the combinators. At first they seem to greatly obstruct the
   342   comprehension of the code, but after getting familiar with them, they
   342   comprehension of the code, but after getting familiar with them, they
   343   actually ease the understanding and also the programming.
   343   actually ease the understanding and also the programming.
   344 
   344 
   345   The simplest combinator is @{ML [index] I}, which is just the identity function defined as
   345   The simplest combinator is @{ML_ind [index] I}, which is just the identity function defined as
   346 *}
   346 *}
   347 
   347 
   348 ML{*fun I x = x*}
   348 ML{*fun I x = x*}
   349 
   349 
   350 text {* Another simple combinator is @{ML [index] K}, defined as *}
   350 text {* Another simple combinator is @{ML_ind [index] K}, defined as *}
   351 
   351 
   352 ML{*fun K x = fn _ => x*}
   352 ML{*fun K x = fn _ => x*}
   353 
   353 
   354 text {*
   354 text {*
   355   @{ML [index] K} ``wraps'' a function around the argument @{text "x"}. However, this 
   355   @{ML_ind [index] K} ``wraps'' a function around the argument @{text "x"}. However, this 
   356   function ignores its argument. As a result, @{ML K} defines a constant function 
   356   function ignores its argument. As a result, @{ML K} defines a constant function 
   357   always returning @{text x}.
   357   always returning @{text x}.
   358 
   358 
   359   The next combinator is reverse application, @{ML [index] "|>"}, defined as: 
   359   The next combinator is reverse application, @{ML_ind [index] "|>"}, defined as: 
   360 *}
   360 *}
   361 
   361 
   362 ML{*fun x |> f = f x*}
   362 ML{*fun x |> f = f x*}
   363 
   363 
   364 text {* While just syntactic sugar for the usual function application,
   364 text {* While just syntactic sugar for the usual function application,
   432    |> tracing
   432    |> tracing
   433 end" 
   433 end" 
   434   "P z za zb"}
   434   "P z za zb"}
   435 
   435 
   436   You can read off this behaviour from how @{ML apply_fresh_args} is
   436   You can read off this behaviour from how @{ML apply_fresh_args} is
   437   coded: in Line 2, the function @{ML [index] fastype_of} calculates the type of the
   437   coded: in Line 2, the function @{ML_ind [index] fastype_of} calculates the type of the
   438   term; @{ML [index] binder_types} in the next line produces the list of argument
   438   term; @{ML_ind [index] binder_types} in the next line produces the list of argument
   439   types (in the case above the list @{text "[nat, int, unit]"}); Line 4 
   439   types (in the case above the list @{text "[nat, int, unit]"}); Line 4 
   440   pairs up each type with the string  @{text "z"}; the
   440   pairs up each type with the string  @{text "z"}; the
   441   function @{ML [index] variant_frees in Variable} generates for each @{text "z"} a
   441   function @{ML_ind [index] variant_frees in Variable} generates for each @{text "z"} a
   442   unique name avoiding the given @{text f}; the list of name-type pairs is turned
   442   unique name avoiding the given @{text f}; the list of name-type pairs is turned
   443   into a list of variable terms in Line 6, which in the last line is applied
   443   into a list of variable terms in Line 6, which in the last line is applied
   444   by the function @{ML [index] list_comb} to the term. In this last step we have to 
   444   by the function @{ML_ind [index] list_comb} to the term. In this last step we have to 
   445   use the function @{ML [index] curry}, because @{ML [index] list_comb} expects the 
   445   use the function @{ML_ind [index] curry}, because @{ML_ind [index] list_comb} expects the 
   446   function and the variables list as a pair. This kind of functions is often needed when
   446   function and the variables list as a pair. This kind of functions is often needed when
   447   constructing terms with fresh variables. The infrastructure helps tremendously 
   447   constructing terms with fresh variables. The infrastructure helps tremendously 
   448   to avoid any name clashes. Consider for example: 
   448   to avoid any name clashes. Consider for example: 
   449 
   449 
   450    @{ML_response_fake [display,gray]
   450    @{ML_response_fake [display,gray]
   458 end"
   458 end"
   459   "za z zb"}
   459   "za z zb"}
   460   
   460   
   461   where the @{text "za"} is correctly avoided.
   461   where the @{text "za"} is correctly avoided.
   462 
   462 
   463   The combinator @{ML [index] "#>"} is the reverse function composition. It can be
   463   The combinator @{ML_ind [index] "#>"} is the reverse function composition. It can be
   464   used to define the following function
   464   used to define the following function
   465 *}
   465 *}
   466 
   466 
   467 ML{*val inc_by_six = 
   467 ML{*val inc_by_six = 
   468       (fn x => x + 1)
   468       (fn x => x + 1)
   473   which is the function composed of first the increment-by-one function and then
   473   which is the function composed of first the increment-by-one function and then
   474   increment-by-two, followed by increment-by-three. Again, the reverse function 
   474   increment-by-two, followed by increment-by-three. Again, the reverse function 
   475   composition allows you to read the code top-down.
   475   composition allows you to read the code top-down.
   476 
   476 
   477   The remaining combinators described in this section add convenience for the
   477   The remaining combinators described in this section add convenience for the
   478   ``waterfall method'' of writing functions. The combinator @{ML [index] tap} allows
   478   ``waterfall method'' of writing functions. The combinator @{ML_ind [index] tap} allows
   479   you to get hold of an intermediate result (to do some side-calculations for
   479   you to get hold of an intermediate result (to do some side-calculations for
   480   instance). The function
   480   instance). The function
   481 
   481 
   482  *}
   482  *}
   483 
   483 
   486        |> tap (fn x => tracing (PolyML.makestring x))
   486        |> tap (fn x => tracing (PolyML.makestring x))
   487        |> (fn x => x + 2)*}
   487        |> (fn x => x + 2)*}
   488 
   488 
   489 text {* 
   489 text {* 
   490   increments the argument first by @{text "1"} and then by @{text "2"}. In the
   490   increments the argument first by @{text "1"} and then by @{text "2"}. In the
   491   middle (Line 3), however, it uses @{ML [index] tap} for printing the ``plus-one''
   491   middle (Line 3), however, it uses @{ML_ind [index] tap} for printing the ``plus-one''
   492   intermediate result inside the tracing buffer. The function @{ML [index] tap} can
   492   intermediate result inside the tracing buffer. The function @{ML_ind [index] tap} can
   493   only be used for side-calculations, because any value that is computed
   493   only be used for side-calculations, because any value that is computed
   494   cannot be merged back into the ``main waterfall''. To do this, you can use
   494   cannot be merged back into the ``main waterfall''. To do this, you can use
   495   the next combinator.
   495   the next combinator.
   496 
   496 
   497   The combinator @{ML [index] "`"} (a backtick) is similar to @{ML tap}, but applies a
   497   The combinator @{ML_ind [index] "`"} (a backtick) is similar to @{ML tap}, but applies a
   498   function to the value and returns the result together with the value (as a
   498   function to the value and returns the result together with the value (as a
   499   pair). For example the function 
   499   pair). For example the function 
   500 *}
   500 *}
   501 
   501 
   502 ML{*fun inc_as_pair x =
   502 ML{*fun inc_as_pair x =
   507   takes @{text x} as argument, and then increments @{text x}, but also keeps
   507   takes @{text x} as argument, and then increments @{text x}, but also keeps
   508   @{text x}. The intermediate result is therefore the pair @{ML "(x + 1, x)"
   508   @{text x}. The intermediate result is therefore the pair @{ML "(x + 1, x)"
   509   for x}. After that, the function increments the right-hand component of the
   509   for x}. After that, the function increments the right-hand component of the
   510   pair. So finally the result will be @{ML "(x + 1, x + 1)" for x}.
   510   pair. So finally the result will be @{ML "(x + 1, x + 1)" for x}.
   511 
   511 
   512   The combinators @{ML [index] "|>>"} and @{ML [index] "||>"} are defined for 
   512   The combinators @{ML_ind [index] "|>>"} and @{ML_ind [index] "||>"} are defined for 
   513   functions manipulating pairs. The first applies the function to
   513   functions manipulating pairs. The first applies the function to
   514   the first component of the pair, defined as
   514   the first component of the pair, defined as
   515 *}
   515 *}
   516 
   516 
   517 ML{*fun (x, y) |>> f = (f x, y)*}
   517 ML{*fun (x, y) |>> f = (f x, y)*}
   554 text {*
   554 text {*
   555   avoiding the explicit @{text "let"}. While in this example the gain in
   555   avoiding the explicit @{text "let"}. While in this example the gain in
   556   conciseness is only small, in more complicated situations the benefit of
   556   conciseness is only small, in more complicated situations the benefit of
   557   avoiding @{text "lets"} can be substantial.
   557   avoiding @{text "lets"} can be substantial.
   558 
   558 
   559   With the combinator @{ML [index] "|->"} you can re-combine the elements from a pair.
   559   With the combinator @{ML_ind [index] "|->"} you can re-combine the elements from a pair.
   560   This combinator is defined as
   560   This combinator is defined as
   561 *}
   561 *}
   562 
   562 
   563 ML{*fun (x, y) |-> f = f x y*}
   563 ML{*fun (x, y) |-> f = f x y*}
   564 
   564 
   570 ML{*fun double x =
   570 ML{*fun double x =
   571       x |>  (fn x => (x, x))
   571       x |>  (fn x => (x, x))
   572         |-> (fn x => fn y => x + y)*}
   572         |-> (fn x => fn y => x + y)*}
   573 
   573 
   574 text {* 
   574 text {* 
   575   The combinator @{ML [index] ||>>} plays a central rôle whenever your task is to update a 
   575   The combinator @{ML_ind [index] ||>>} plays a central rôle whenever your task is to update a 
   576   theory and the update also produces a side-result (for example a theorem). Functions 
   576   theory and the update also produces a side-result (for example a theorem). Functions 
   577   for such tasks return a pair whose second component is the theory and the fist 
   577   for such tasks return a pair whose second component is the theory and the fist 
   578   component is the side-result. Using @{ML [index] ||>>}, you can do conveniently the update 
   578   component is the side-result. Using @{ML_ind [index] ||>>}, you can do conveniently the update 
   579   and also accumulate the side-results. Consider the following simple function. 
   579   and also accumulate the side-results. Consider the following simple function. 
   580 *}
   580 *}
   581 
   581 
   582 ML %linenosgray{*fun acc_incs x =
   582 ML %linenosgray{*fun acc_incs x =
   583     x |> (fn x => ("", x)) 
   583     x |> (fn x => ("", x)) 
   585       ||>> (fn x => (x, x + 1))
   585       ||>> (fn x => (x, x + 1))
   586       ||>> (fn x => (x, x + 1))*}
   586       ||>> (fn x => (x, x + 1))*}
   587 
   587 
   588 text {*
   588 text {*
   589   The purpose of Line 2 is to just pair up the argument with a dummy value (since
   589   The purpose of Line 2 is to just pair up the argument with a dummy value (since
   590   @{ML [index] "||>>"} operates on pairs). Each of the next three lines just increment 
   590   @{ML_ind [index] "||>>"} operates on pairs). Each of the next three lines just increment 
   591   the value by one, but also nest the intermediate results to the left. For example 
   591   the value by one, but also nest the intermediate results to the left. For example 
   592 
   592 
   593   @{ML_response [display,gray]
   593   @{ML_response [display,gray]
   594   "acc_incs 1"
   594   "acc_incs 1"
   595   "((((\"\", 1), 2), 3), 4)"}
   595   "((((\"\", 1), 2), 3), 4)"}
   602 
   602 
   603   (FIXME: maybe give a ``real world'' example for this combinator)
   603   (FIXME: maybe give a ``real world'' example for this combinator)
   604 *}
   604 *}
   605 
   605 
   606 text {*
   606 text {*
   607   Recall that @{ML [index] "|>"} is the reverse function application. Recall also that
   607   Recall that @{ML_ind [index] "|>"} is the reverse function application. Recall also that
   608   the related 
   608   the related 
   609   reverse function composition is @{ML [index] "#>"}. In fact all the combinators 
   609   reverse function composition is @{ML_ind [index] "#>"}. In fact all the combinators 
   610   @{ML [index] "|->"}, @{ML [index] "|>>"} , @{ML [index] "||>"} and @{ML [index] 
   610   @{ML_ind [index] "|->"}, @{ML_ind [index] "|>>"} , @{ML_ind [index] "||>"} and @{ML_ind [index] 
   611   "||>>"} described above have related combinators for 
   611   "||>>"} described above have related combinators for 
   612   function composition, namely @{ML [index] "#->"}, @{ML [index] "#>>"}, 
   612   function composition, namely @{ML_ind [index] "#->"}, @{ML_ind [index] "#>>"}, 
   613   @{ML [index] "##>"} and @{ML [index] "##>>"}. 
   613   @{ML_ind [index] "##>"} and @{ML_ind [index] "##>>"}. 
   614   Using @{ML [index] "#->"}, for example, the function @{text double} can also be written as:
   614   Using @{ML_ind [index] "#->"}, for example, the function @{text double} can also be written as:
   615 *}
   615 *}
   616 
   616 
   617 ML{*val double =
   617 ML{*val double =
   618             (fn x => (x, x))
   618             (fn x => (x, x))
   619         #-> (fn x => fn y => x + y)*}
   619         #-> (fn x => fn y => x + y)*}
   625   together. We have already seen such plumbing in the function @{ML
   625   together. We have already seen such plumbing in the function @{ML
   626   apply_fresh_args}, where @{ML curry} is needed for making the function @{ML
   626   apply_fresh_args}, where @{ML curry} is needed for making the function @{ML
   627   list_comb} that works over pairs to fit with the combinator @{ML "|>"}. Such 
   627   list_comb} that works over pairs to fit with the combinator @{ML "|>"}. Such 
   628   plumbing is also needed in situations where a functions operate over lists, 
   628   plumbing is also needed in situations where a functions operate over lists, 
   629   but one calculates only with a single entity. An example is the function 
   629   but one calculates only with a single entity. An example is the function 
   630   @{ML [index] check_terms in Syntax}, whose purpose is to type-check a list 
   630   @{ML_ind [index] check_terms in Syntax}, whose purpose is to type-check a list 
   631   of terms.
   631   of terms.
   632 
   632 
   633   @{ML_response_fake [display, gray]
   633   @{ML_response_fake [display, gray]
   634   "let
   634   "let
   635   val ctxt = @{context}
   635   val ctxt = @{context}
   688   @{ML_response [display,gray] "Context.theory_name @{theory}" "\"FirstSteps\""}
   688   @{ML_response [display,gray] "Context.theory_name @{theory}" "\"FirstSteps\""}
   689  
   689  
   690   where @{text "@{theory}"} is an antiquotation that is substituted with the
   690   where @{text "@{theory}"} is an antiquotation that is substituted with the
   691   current theory (remember that we assumed we are inside the theory 
   691   current theory (remember that we assumed we are inside the theory 
   692   @{text FirstSteps}). The name of this theory can be extracted using
   692   @{text FirstSteps}). The name of this theory can be extracted using
   693   the function @{ML [index] theory_name in Context}. 
   693   the function @{ML_ind [index] theory_name in Context}. 
   694 
   694 
   695   Note, however, that antiquotations are statically linked, that is their value is
   695   Note, however, that antiquotations are statically linked, that is their value is
   696   determined at ``compile-time'', not ``run-time''. For example the function
   696   determined at ``compile-time'', not ``run-time''. For example the function
   697 *}
   697 *}
   698 
   698 
   736 in
   736 in
   737   map #1 simps
   737   map #1 simps
   738 end*}
   738 end*}
   739 
   739 
   740 text {*
   740 text {*
   741   The function @{ML [index] dest_ss in MetaSimplifier} returns a record containing all
   741   The function @{ML_ind [index] dest_ss in MetaSimplifier} returns a record containing all
   742   information stored in the simpset, but we are only interested in the names of the
   742   information stored in the simpset, but we are only interested in the names of the
   743   simp-rules. Now you can feed in the current simpset into this function. 
   743   simp-rules. Now you can feed in the current simpset into this function. 
   744   The current simpset can be referred to using the antiquotation @{text "@{simpset}"}.
   744   The current simpset can be referred to using the antiquotation @{text "@{simpset}"}.
   745 
   745 
   746   @{ML_response_fake [display,gray] 
   746   @{ML_response_fake [display,gray] 
   750   Again, this way of referencing simpsets makes you independent from additions
   750   Again, this way of referencing simpsets makes you independent from additions
   751   of lemmas to the simpset by the user that can potentially cause loops in your
   751   of lemmas to the simpset by the user that can potentially cause loops in your
   752   code.
   752   code.
   753 
   753 
   754   On the ML-level of Isabelle, you often have to work with qualified names.
   754   On the ML-level of Isabelle, you often have to work with qualified names.
   755   These are strings with some additional information, such as positional information
   755   These are strings with some additional information, such as positional
   756   and qualifiers. Such qualified names can be generated with the antiquotation 
   756   information and qualifiers. Such qualified names can be generated with the
   757   @{text "@{binding \<dots>}"}. For example
   757   antiquotation @{text "@{binding \<dots>}"}. For example
   758 
   758 
   759   @{ML_response [display,gray]
   759   @{ML_response [display,gray]
   760   "@{binding \"name\"}"
   760   "@{binding \"name\"}"
   761   "name"}
   761   "name"}
   762 
   762 
   763   An example where a qualified name is needed is the function 
   763   An example where a qualified name is needed is the function 
   764   @{ML [index] define in LocalTheory}.  This function is used below to define 
   764   @{ML_ind [index] define in LocalTheory}.  This function is used below to define 
   765   the constant @{term "TrueConj"} as the conjunction @{term "True \<and> True"}.
   765   the constant @{term "TrueConj"} as the conjunction @{term "True \<and> True"}.
   766 *}
   766 *}
   767   
   767   
   768 local_setup %gray {* 
   768 local_setup %gray {* 
   769   snd o LocalTheory.define Thm.internalK 
   769   snd o LocalTheory.define Thm.internalK 
   781   (FIXME give a better example why bindings are important; maybe
   781   (FIXME give a better example why bindings are important; maybe
   782   give a pointer to \isacommand{local\_setup}; if not, then explain
   782   give a pointer to \isacommand{local\_setup}; if not, then explain
   783   why @{ML snd} is needed)
   783   why @{ML snd} is needed)
   784 
   784 
   785   It is also possible to define your own antiquotations. But you should
   785   It is also possible to define your own antiquotations. But you should
   786   exercise care when introducing new ones, as they can also make your
   786   exercise care when introducing new ones, as they can also make your code
   787   code also difficult to read. In the next section we will see how the (build in) 
   787   also difficult to read. In the next section we will see how the (build in)
   788   antiquotation @{text "@{term \<dots>}"} can be used to construct terms.
   788   antiquotation @{text "@{term \<dots>}"} can be used to construct terms.  A
   789   A restriction of this antiquotation is that it does not allow you to
   789   restriction of this antiquotation is that it does not allow you to use
   790   use schematic variables. If you want to have an antiquotation that does not
   790   schematic variables. If you want to have an antiquotation that does not have
   791   have this restriction, you can implement your own using the 
   791   this restriction, you can implement your own using the function @{ML_ind [index]
   792   function @{ML [index] ML_Antiquote.inline}. The code is as follows.
   792   ML_Antiquote.inline}. The code is as follows.
   793 *}
   793 *}
   794 
   794 
   795 ML %linenosgray{*ML_Antiquote.inline "term_pat"
   795 ML %linenosgray{*ML_Antiquote.inline "term_pat"
   796   (Args.context -- Scan.lift Args.name_source >>
   796   (Args.context -- Scan.lift Args.name_source >>
   797      (fn (ctxt, s) =>
   797      (fn (ctxt, s) =>
   799          |> ML_Syntax.print_term
   799          |> ML_Syntax.print_term
   800          |> ML_Syntax.atomic))*}
   800          |> ML_Syntax.atomic))*}
   801 
   801 
   802 text {*
   802 text {*
   803   The parser in Line 2 provides us with a context and a string; this string is
   803   The parser in Line 2 provides us with a context and a string; this string is
   804   transformed into a term using the function @{ML [index] read_term_pattern in
   804   transformed into a term using the function @{ML_ind [index] read_term_pattern in
   805   ProofContext} (Line 4); the next two lines print the term so that the
   805   ProofContext} (Line 4); the next two lines print the term so that the
   806   ML-system can understand them. An example of this antiquotation is as
   806   ML-system can understand them. An example of this antiquotation is as
   807   follows.
   807   follows.
   808 
   808 
   809   @{ML_response_fake [display,gray]
   809   @{ML_response_fake [display,gray]
   877 
   877 
   878   Isabelle makes a distinction between \emph{free} variables (term-constructor
   878   Isabelle makes a distinction between \emph{free} variables (term-constructor
   879   @{ML Free} and written on the user level in blue colour) and
   879   @{ML Free} and written on the user level in blue colour) and
   880   \emph{schematic} variables (term-constructor @{ML Var} and written with a
   880   \emph{schematic} variables (term-constructor @{ML Var} and written with a
   881   leading question mark). Consider the following two examples
   881   leading question mark). Consider the following two examples
   882 
       
   883 
       
   884   
   882   
   885   @{ML_response_fake [display, gray]
   883   @{ML_response_fake [display, gray]
   886 "let
   884 "let
   887   val v1 = Var ((\"x\", 3), @{typ bool})
   885   val v1 = Var ((\"x\", 3), @{typ bool})
   888   val v2 = Var ((\"x1\", 3), @{typ bool})
   886   val v2 = Var ((\"x1\", 3), @{typ bool})
  1045   Abs (\"x\", \<dots>,
  1043   Abs (\"x\", \<dots>,
  1046     Const \<dots> $ (Const \<dots> $ (Free (\"P\",\<dots>) $ \<dots>)) $ 
  1044     Const \<dots> $ (Const \<dots> $ (Free (\"P\",\<dots>) $ \<dots>)) $ 
  1047                 (Const \<dots> $ (Free (\"Q\",\<dots>) $ \<dots>)))"}
  1045                 (Const \<dots> $ (Free (\"Q\",\<dots>) $ \<dots>)))"}
  1048 
  1046 
  1049   There are a number of handy functions that are frequently used for 
  1047   There are a number of handy functions that are frequently used for 
  1050   constructing terms. One is the function @{ML [index] list_comb}, which takes a term
  1048   constructing terms. One is the function @{ML_ind [index] list_comb}, which takes a term
  1051   and a list of terms as arguments, and produces as output the term
  1049   and a list of terms as arguments, and produces as output the term
  1052   list applied to the term. For example
  1050   list applied to the term. For example
  1053 
  1051 
  1054 @{ML_response_fake [display,gray]
  1052 @{ML_response_fake [display,gray]
  1055 "let
  1053 "let
  1058 in
  1056 in
  1059   list_comb (trm, args)
  1057   list_comb (trm, args)
  1060 end"
  1058 end"
  1061 "Free (\"P\", \"nat\") $ Const (\"True\", \"bool\") $ Const (\"False\", \"bool\")"}
  1059 "Free (\"P\", \"nat\") $ Const (\"True\", \"bool\") $ Const (\"False\", \"bool\")"}
  1062 
  1060 
  1063   Another handy function is @{ML [index] lambda}, which abstracts a variable 
  1061   Another handy function is @{ML_ind [index] lambda}, which abstracts a variable 
  1064   in a term. For example
  1062   in a term. For example
  1065 
  1063 
  1066   @{ML_response_fake [display,gray]
  1064   @{ML_response_fake [display,gray]
  1067 "let
  1065 "let
  1068   val x_nat = @{term \"x::nat\"}
  1066   val x_nat = @{term \"x::nat\"}
  1091   then the variable @{text "Free (\"x\", \"int\")"} is \emph{not} abstracted. 
  1089   then the variable @{text "Free (\"x\", \"int\")"} is \emph{not} abstracted. 
  1092   This is a fundamental principle
  1090   This is a fundamental principle
  1093   of Church-style typing, where variables with the same name still differ, if they 
  1091   of Church-style typing, where variables with the same name still differ, if they 
  1094   have different type.
  1092   have different type.
  1095 
  1093 
  1096   There is also the function @{ML [index] subst_free} with which terms can be
  1094   There is also the function @{ML_ind [index] subst_free} with which terms can be
  1097   replaced by other terms. For example below, we will replace in @{term
  1095   replaced by other terms. For example below, we will replace in @{term
  1098   "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0 x"} the subterm @{term "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0"} by
  1096   "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0 x"} the subterm @{term "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0"} by
  1099   @{term y}, and @{term x} by @{term True}.
  1097   @{term y}, and @{term x} by @{term True}.
  1100 
  1098 
  1101   @{ML_response_fake [display,gray]
  1099   @{ML_response_fake [display,gray]
  1118 in
  1116 in
  1119   subst_free [sub] trm
  1117   subst_free [sub] trm
  1120 end"
  1118 end"
  1121   "Free (\"x\", \"nat\")"}
  1119   "Free (\"x\", \"nat\")"}
  1122 
  1120 
  1123   Similarly the function @{ML [index] subst_bounds}, replaces lose bound 
  1121   Similarly the function @{ML_ind [index] subst_bounds}, replaces lose bound 
  1124   variables with terms. To see how this function works, let us implement a
  1122   variables with terms. To see how this function works, let us implement a
  1125   function that strips off the outermost quantifiers in a term.
  1123   function that strips off the outermost quantifiers in a term.
  1126 *}
  1124 *}
  1127 
  1125 
  1128 ML{*fun strip_alls (Const ("All", _) $ Abs (n, T, t)) =
  1126 ML{*fun strip_alls (Const ("All", _) $ Abs (n, T, t)) =
  1136   @{ML_response_fake [display, gray]
  1134   @{ML_response_fake [display, gray]
  1137   "strip_alls @{term \"\<forall>x y. x = (y::bool)\"}"
  1135   "strip_alls @{term \"\<forall>x y. x = (y::bool)\"}"
  1138 "([Free (\"x\", \"bool\"), Free (\"y\", \"bool\")],
  1136 "([Free (\"x\", \"bool\"), Free (\"y\", \"bool\")],
  1139   Const (\"op =\", \<dots>) $ Bound 1 $ Bound 0)"}
  1137   Const (\"op =\", \<dots>) $ Bound 1 $ Bound 0)"}
  1140 
  1138 
  1141   With the function @{ML subst_bounds}, you can replace the lose 
  1139   After calling @{ML strip_alls}, you obtain a term with lose bound variables. With
  1142   @{ML [index] Bound}s with the stripped off variables.
  1140   the function @{ML subst_bounds}, you can replace these lose @{ML_ind [index]
       
  1141   Bound}s with the stripped off variables.
  1143 
  1142 
  1144   @{ML_response_fake [display, gray, linenos]
  1143   @{ML_response_fake [display, gray, linenos]
  1145   "let
  1144   "let
  1146   val (vrs, trm) = strip_alls @{term \"\<forall>x y. x = (y::bool)\"}
  1145   val (vrs, trm) = strip_alls @{term \"\<forall>x y. x = (y::bool)\"}
  1147 in 
  1146 in 
  1155   returned. The reason is that the head of the list the function @{ML subst_bounds}
  1154   returned. The reason is that the head of the list the function @{ML subst_bounds}
  1156   takes is the replacement for @{ML "Bound 0"}, the next element for @{ML "Bound 1"}
  1155   takes is the replacement for @{ML "Bound 0"}, the next element for @{ML "Bound 1"}
  1157   and so on. 
  1156   and so on. 
  1158 
  1157 
  1159   There are many convenient functions that construct specific HOL-terms. For
  1158   There are many convenient functions that construct specific HOL-terms. For
  1160   example @{ML [index] mk_eq in HOLogic} constructs an equality out of two terms.
  1159   example @{ML_ind [index] mk_eq in HOLogic} constructs an equality out of two terms.
  1161   The types needed in this equality are calculated from the type of the
  1160   The types needed in this equality are calculated from the type of the
  1162   arguments. For example
  1161   arguments. For example
  1163 
  1162 
  1164 @{ML_response_fake [gray,display]
  1163 @{ML_response_fake [gray,display]
  1165 "let
  1164 "let
  1171   "True = False"}
  1170   "True = False"}
  1172 *}
  1171 *}
  1173 
  1172 
  1174 text {*
  1173 text {*
  1175   \begin{readmore}
  1174   \begin{readmore}
  1176   Most of the HOL-specific operations on terms and types are defined 
  1175   There are many functions in @{ML_file "Pure/term.ML"}, @{ML_file
  1177   in @{ML_file "HOL/Tools/hologic.ML"}.
  1176   "Pure/logic.ML"} and @{ML_file "HOL/Tools/hologic.ML"} that make such manual
       
  1177   constructions of terms and types easier.
  1178   \end{readmore}
  1178   \end{readmore}
  1179 
  1179 
  1180   When constructing terms manually, there are a few subtle issues with
  1180   When constructing terms manually, there are a few subtle issues with
  1181   constants. They usually crop up when pattern matching terms or types, or
  1181   constants. They usually crop up when pattern matching terms or types, or
  1182   when constructing them. While it is perfectly ok to write the function
  1182   when constructing them. While it is perfectly ok to write the function
  1259 ML{*fun is_nil_or_all (Const (@{const_name "Nil"}, _)) = true
  1259 ML{*fun is_nil_or_all (Const (@{const_name "Nil"}, _)) = true
  1260   | is_nil_or_all (Const (@{const_name "All"}, _) $ _) = true
  1260   | is_nil_or_all (Const (@{const_name "All"}, _) $ _) = true
  1261   | is_nil_or_all _ = false *}
  1261   | is_nil_or_all _ = false *}
  1262 
  1262 
  1263 text {*
  1263 text {*
  1264   \begin{readmore}
       
  1265   There are many functions in @{ML_file "Pure/term.ML"}, @{ML_file "Pure/logic.ML"} and
       
  1266   @{ML_file "HOL/Tools/hologic.ML"} that make such manual constructions of terms 
       
  1267   and types easier.
       
  1268   \end{readmore}
       
  1269 
       
  1270   The antiquotation for properly referencing type constants is @{text "@{type_name \<dots>}"}.
  1264   The antiquotation for properly referencing type constants is @{text "@{type_name \<dots>}"}.
  1271   For example
  1265   For example
  1272 
  1266 
  1273   @{ML_response [display,gray]
  1267   @{ML_response [display,gray]
  1274   "@{type_name \"list\"}" "\"List.list\""}
  1268   "@{type_name \"list\"}" "\"List.list\""}
  1275 
  1269 
  1276   (FIXME: Explain the following better.)
  1270   (FIXME: Explain the following better.)
  1277 
  1271 
  1278   Occasionally you have to calculate what the ``base'' name of a given
  1272   Occasionally you have to calculate what the ``base'' name of a given
  1279   constant is. For this you can use the function @{ML [index] "Sign.extern_const"} or
  1273   constant is. For this you can use the function @{ML_ind [index] "Sign.extern_const"} or
  1280   @{ML [index] Long_Name.base_name}. For example:
  1274   @{ML_ind [index] Long_Name.base_name}. For example:
  1281 
  1275 
  1282   @{ML_response [display,gray] "Sign.extern_const @{theory} \"List.list.Nil\"" "\"Nil\""}
  1276   @{ML_response [display,gray] "Sign.extern_const @{theory} \"List.list.Nil\"" "\"Nil\""}
  1283 
  1277 
  1284   The difference between both functions is that @{ML extern_const in Sign} returns
  1278   The difference between both functions is that @{ML extern_const in Sign} returns
  1285   the smallest name that is still unique, whereas @{ML base_name in Long_Name} always
  1279   the smallest name that is still unique, whereas @{ML base_name in Long_Name} always
  1297 
  1291 
  1298 *} 
  1292 *} 
  1299 
  1293 
  1300 ML{*fun make_fun_type ty1 ty2 = Type ("fun", [ty1, ty2]) *}
  1294 ML{*fun make_fun_type ty1 ty2 = Type ("fun", [ty1, ty2]) *}
  1301 
  1295 
  1302 text {* This can be equally written with the combinator @{ML [index] "-->"} as: *}
  1296 text {* This can be equally written with the combinator @{ML_ind [index] "-->"} as: *}
  1303 
  1297 
  1304 ML{*fun make_fun_type ty1 ty2 = ty1 --> ty2 *}
  1298 ML{*fun make_fun_type ty1 ty2 = ty1 --> ty2 *}
  1305 
  1299 
  1306 text {*
  1300 text {*
  1307   If you want to construct a function type with more than one argument
  1301   If you want to construct a function type with more than one argument
  1308   type, then you can use @{ML [index] "--->"}.
  1302   type, then you can use @{ML_ind [index] "--->"}.
  1309 *}
  1303 *}
  1310 
  1304 
  1311 ML{*fun make_fun_types tys ty = tys ---> ty *}
  1305 ML{*fun make_fun_types tys ty = tys ---> ty *}
  1312 
  1306 
  1313 text {*
  1307 text {*
  1314   A handy function for manipulating terms is @{ML [index] map_types}: it takes a 
  1308   A handy function for manipulating terms is @{ML_ind [index] map_types}: it takes a 
  1315   function and applies it to every type in a term. You can, for example,
  1309   function and applies it to every type in a term. You can, for example,
  1316   change every @{typ nat} in a term into an @{typ int} using the function:
  1310   change every @{typ nat} in a term into an @{typ int} using the function:
  1317 *}
  1311 *}
  1318 
  1312 
  1319 ML{*fun nat_to_int ty =
  1313 ML{*fun nat_to_int ty =
  1329 "map_types nat_to_int @{term \"a = (1::nat)\"}" 
  1323 "map_types nat_to_int @{term \"a = (1::nat)\"}" 
  1330 "Const (\"op =\", \"int \<Rightarrow> int \<Rightarrow> bool\")
  1324 "Const (\"op =\", \"int \<Rightarrow> int \<Rightarrow> bool\")
  1331            $ Free (\"a\", \"int\") $ Const (\"HOL.one_class.one\", \"int\")"}
  1325            $ Free (\"a\", \"int\") $ Const (\"HOL.one_class.one\", \"int\")"}
  1332 
  1326 
  1333   If you want to obtain the list of free type-variables of a term, you
  1327   If you want to obtain the list of free type-variables of a term, you
  1334   can use the function @{ML [index] add_tfrees in Term} 
  1328   can use the function @{ML_ind [index] add_tfrees in Term} 
  1335   (similarly @{ML [index] add_tvars in Term} for the schematic type-variables). 
  1329   (similarly @{ML_ind [index] add_tvars in Term} for the schematic type-variables). 
  1336   One would expect that such functions
  1330   One would expect that such functions
  1337   take a term as input and return a list of types. But their type is actually 
  1331   take a term as input and return a list of types. But their type is actually 
  1338 
  1332 
  1339   @{text[display] "Term.term -> (string * Term.sort) list -> (string * Term.sort) list"}
  1333   @{text[display] "Term.term -> (string * Term.sort) list -> (string * Term.sort) list"}
  1340 
  1334 
  1399 text {* 
  1393 text {* 
  1400   
  1394   
  1401   You can freely construct and manipulate @{ML_type "term"}s and @{ML_type
  1395   You can freely construct and manipulate @{ML_type "term"}s and @{ML_type
  1402   typ}es, since they are just arbitrary unchecked trees. However, you
  1396   typ}es, since they are just arbitrary unchecked trees. However, you
  1403   eventually want to see if a term is well-formed, or type-checks, relative to
  1397   eventually want to see if a term is well-formed, or type-checks, relative to
  1404   a theory.  Type-checking is done via the function @{ML [index] cterm_of}, which
  1398   a theory.  Type-checking is done via the function @{ML_ind [index] cterm_of}, which
  1405   converts a @{ML_type [index] term} into a @{ML_type [index] cterm}, a \emph{certified}
  1399   converts a @{ML_type [index] term} into a @{ML_type [index] cterm}, a \emph{certified}
  1406   term. Unlike @{ML_type term}s, which are just trees, @{ML_type "cterm"}s are
  1400   term. Unlike @{ML_type term}s, which are just trees, @{ML_type "cterm"}s are
  1407   abstract objects that are guaranteed to be type-correct, and they can only
  1401   abstract objects that are guaranteed to be type-correct, and they can only
  1408   be constructed via ``official interfaces''.
  1402   be constructed via ``official interfaces''.
  1409 
  1403 
  1451   \begin{readmore}
  1445   \begin{readmore}
  1452   For functions related to @{ML_type cterm}s and @{ML_type ctyp}s see 
  1446   For functions related to @{ML_type cterm}s and @{ML_type ctyp}s see 
  1453   the file @{ML_file "Pure/thm.ML"}.
  1447   the file @{ML_file "Pure/thm.ML"}.
  1454   \end{readmore}
  1448   \end{readmore}
  1455 
  1449 
  1456   \begin{exercise}
       
  1457   Check that the function defined in Exercise~\ref{fun:revsum} returns a
       
  1458   result that type-checks. See what happens to the  solutions of this 
       
  1459   exercise given in \ref{ch:solutions} when they receive an ill-typed term
       
  1460   as input.
       
  1461   \end{exercise}
       
  1462 
       
  1463   Remember Isabelle follows the Church-style typing for terms, i.e., a term contains 
  1450   Remember Isabelle follows the Church-style typing for terms, i.e., a term contains 
  1464   enough typing information (constants, free variables and abstractions all have typing 
  1451   enough typing information (constants, free variables and abstractions all have typing 
  1465   information) so that it is always clear what the type of a term is. 
  1452   information) so that it is always clear what the type of a term is. 
  1466   Given a well-typed term, the function @{ML [index] type_of} returns the 
  1453   Given a well-typed term, the function @{ML_ind [index] type_of} returns the 
  1467   type of a term. Consider for example:
  1454   type of a term. Consider for example:
  1468 
  1455 
  1469   @{ML_response [display,gray] 
  1456   @{ML_response [display,gray] 
  1470   "type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"}
  1457   "type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"}
  1471 
  1458 
  1476   @{ML_response_fake [display,gray] 
  1463   @{ML_response_fake [display,gray] 
  1477   "type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::int\"})" 
  1464   "type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::int\"})" 
  1478   "*** Exception- TYPE (\"type_of: type mismatch in application\" \<dots>"}
  1465   "*** Exception- TYPE (\"type_of: type mismatch in application\" \<dots>"}
  1479 
  1466 
  1480   Since the complete traversal might sometimes be too costly and
  1467   Since the complete traversal might sometimes be too costly and
  1481   not necessary, there is the function @{ML [index] fastype_of}, which 
  1468   not necessary, there is the function @{ML_ind [index] fastype_of}, which 
  1482   also returns the type of a term.
  1469   also returns the type of a term.
  1483 
  1470 
  1484   @{ML_response [display,gray] 
  1471   @{ML_response [display,gray] 
  1485   "fastype_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"}
  1472   "fastype_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"}
  1486 
  1473 
  1493   where no error is detected.
  1480   where no error is detected.
  1494 
  1481 
  1495   Sometimes it is a bit inconvenient to construct a term with 
  1482   Sometimes it is a bit inconvenient to construct a term with 
  1496   complete typing annotations, especially in cases where the typing 
  1483   complete typing annotations, especially in cases where the typing 
  1497   information is redundant. A short-cut is to use the ``place-holder'' 
  1484   information is redundant. A short-cut is to use the ``place-holder'' 
  1498   type @{ML [index] dummyT} and then let type-inference figure out the 
  1485   type @{ML_ind [index] dummyT} and then let type-inference figure out the 
  1499   complete type. An example is as follows:
  1486   complete type. An example is as follows:
  1500 
  1487 
  1501   @{ML_response_fake [display,gray] 
  1488   @{ML_response_fake [display,gray] 
  1502 "let
  1489 "let
  1503   val c = Const (@{const_name \"plus\"}, dummyT) 
  1490   val c = Const (@{const_name \"plus\"}, dummyT) 
  1518   type-inference are implemented in @{ML_file "Pure/type.ML"} and 
  1505   type-inference are implemented in @{ML_file "Pure/type.ML"} and 
  1519   @{ML_file "Pure/type_infer.ML"}. 
  1506   @{ML_file "Pure/type_infer.ML"}. 
  1520   \end{readmore}
  1507   \end{readmore}
  1521 
  1508 
  1522   (FIXME: say something about sorts)
  1509   (FIXME: say something about sorts)
       
  1510 
       
  1511   \begin{exercise}
       
  1512   Check that the function defined in Exercise~\ref{fun:revsum} returns a
       
  1513   result that type-checks. See what happens to the  solutions of this 
       
  1514   exercise given in \ref{ch:solutions} when they receive an ill-typed term
       
  1515   as input.
       
  1516   \end{exercise}
  1523 *}
  1517 *}
  1524 
  1518 
  1525 
  1519 
  1526 section {* Theorems *}
  1520 section {* Theorems *}
  1527 
  1521 
  1583   For the functions @{text "assume"}, @{text "forall_elim"} etc 
  1577   For the functions @{text "assume"}, @{text "forall_elim"} etc 
  1584   see \isccite{sec:thms}. The basic functions for theorems are defined in
  1578   see \isccite{sec:thms}. The basic functions for theorems are defined in
  1585   @{ML_file "Pure/thm.ML"}. 
  1579   @{ML_file "Pure/thm.ML"}. 
  1586   \end{readmore}
  1580   \end{readmore}
  1587 
  1581 
  1588   (FIXME: handy functions working on theorems, like @{ML [index] rulify in ObjectLogic} and so on) 
  1582   (FIXME: handy functions working on theorems, like @{ML_ind [index] rulify in ObjectLogic} and so on) 
  1589 
  1583 
  1590   (FIXME: how to add case-names to goal states - maybe in the 
  1584   (FIXME: how to add case-names to goal states - maybe in the 
  1591   next section)
  1585   next section)
  1592 
  1586 
  1593   (FIXME: example for how to add theorem styles)
  1587   (FIXME: example for how to add theorem styles)
  1610 
  1604 
  1611 ML {*
  1605 ML {*
  1612 strip_assums_all ([], @{term "\<And>x y. A x y"})
  1606 strip_assums_all ([], @{term "\<And>x y. A x y"})
  1613 *}
  1607 *}
  1614 
  1608 
  1615 setup {*
  1609 setup %gray {*
  1616   TermStyle.add_style "no_all_prem1" (style_parm_premise 1) #>
  1610   TermStyle.add_style "no_all_prem1" (style_parm_premise 1) #>
  1617   TermStyle.add_style "no_all_prem2" (style_parm_premise 2)
  1611   TermStyle.add_style "no_all_prem2" (style_parm_premise 2)
  1618 *}
  1612 *}
  1619 
  1613 
  1620 
  1614 
  1627   @{ML_type "theory -> theory"}: the input theory is the current theory and the 
  1621   @{ML_type "theory -> theory"}: the input theory is the current theory and the 
  1628   output the theory where the theory attribute has been stored.
  1622   output the theory where the theory attribute has been stored.
  1629   
  1623   
  1630   This is a fundamental principle in Isabelle. A similar situation occurs 
  1624   This is a fundamental principle in Isabelle. A similar situation occurs 
  1631   for example with declaring constants. The function that declares a 
  1625   for example with declaring constants. The function that declares a 
  1632   constant on the ML-level is @{ML [index] add_consts_i in Sign}. 
  1626   constant on the ML-level is @{ML_ind [index] add_consts_i in Sign}. 
  1633   If you write\footnote{Recall that ML-code  needs to be 
  1627   If you write\footnote{Recall that ML-code  needs to be 
  1634   enclosed in \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"}.} 
  1628   enclosed in \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"}.} 
  1635 *}  
  1629 *}  
  1636 
  1630 
  1637 ML{*Sign.add_consts_i [(@{binding "BAR"}, @{typ "nat"}, NoSyn)] @{theory} *}
  1631 ML{*Sign.add_consts_i [(@{binding "BAR"}, @{typ "nat"}, NoSyn)] @{theory} *}
  1705 *}
  1699 *}
  1706 
  1700 
  1707 ML{*val my_symmetric = Thm.rule_attribute (fn _ => fn thm => thm RS @{thm sym})*}
  1701 ML{*val my_symmetric = Thm.rule_attribute (fn _ => fn thm => thm RS @{thm sym})*}
  1708 
  1702 
  1709 text {* 
  1703 text {* 
  1710   where the function @{ML [index] rule_attribute in Thm} expects a function taking a 
  1704   where the function @{ML_ind [index] rule_attribute in Thm} expects a function taking a 
  1711   context (which we ignore in the code above) and a theorem (@{text thm}), and 
  1705   context (which we ignore in the code above) and a theorem (@{text thm}), and 
  1712   returns another theorem (namely @{text thm} resolved with the theorem 
  1706   returns another theorem (namely @{text thm} resolved with the theorem 
  1713   @{thm [source] sym}: @{thm sym[no_vars]}).\footnote{The function @{ML [index] RS} 
  1707   @{thm [source] sym}: @{thm sym[no_vars]}).\footnote{The function @{ML_ind [index] "RS"} 
  1714   is explained in Section~\ref{sec:simpletacs}.} The function 
  1708   is explained in Section~\ref{sec:simpletacs}.} The function 
  1715   @{ML rule_attribute in Thm} then returns  an attribute.
  1709   @{ML rule_attribute in Thm} then returns  an attribute.
  1716 
  1710 
  1717   Before we can use the attribute, we need to set it up. This can be done
  1711   Before we can use the attribute, we need to set it up. This can be done
  1718   using the Isabelle command \isacommand{attribute\_setup} as follows:
  1712   using the Isabelle command \isacommand{attribute\_setup} as follows:
  1745   \begin{isabelle}
  1739   \begin{isabelle}
  1746   \isacommand{thm}~@{text "test[my_sym]"}\\
  1740   \isacommand{thm}~@{text "test[my_sym]"}\\
  1747   @{text "> "}~@{thm test[my_sym]}
  1741   @{text "> "}~@{thm test[my_sym]}
  1748   \end{isabelle}
  1742   \end{isabelle}
  1749 
  1743 
  1750   An alternative for setting up an attribute is the function @{ML [index] setup in Attrib}.
  1744   An alternative for setting up an attribute is the function @{ML_ind [index] setup in Attrib}.
  1751   So instead of using \isacommand{attribute\_setup}, you can also set up the
  1745   So instead of using \isacommand{attribute\_setup}, you can also set up the
  1752   attribute as follows:
  1746   attribute as follows:
  1753 *}
  1747 *}
  1754 
  1748 
  1755 ML{*Attrib.setup @{binding "my_sym"} (Scan.succeed my_symmetric)
  1749 ML{*Attrib.setup @{binding "my_sym"} (Scan.succeed my_symmetric)
  1810   \end{isabelle}
  1804   \end{isabelle}
  1811 
  1805 
  1812   you get an exception indicating that the theorem @{thm [source] sym}
  1806   you get an exception indicating that the theorem @{thm [source] sym}
  1813   does not resolve with meta-equations. 
  1807   does not resolve with meta-equations. 
  1814 
  1808 
  1815   The purpose of @{ML [index] rule_attribute in Thm} is to directly manipulate theorems.
  1809   The purpose of @{ML_ind [index] rule_attribute in Thm} is to directly manipulate theorems.
  1816   Another usage of theorem attributes is to add and delete theorems from stored data.
  1810   Another usage of theorem attributes is to add and delete theorems from stored data.
  1817   For example the theorem attribute @{text "[simp]"} adds or deletes a theorem from the
  1811   For example the theorem attribute @{text "[simp]"} adds or deletes a theorem from the
  1818   current simpset. For these applications, you can use @{ML [index] declaration_attribute in Thm}. 
  1812   current simpset. For these applications, you can use @{ML_ind [index] declaration_attribute in Thm}. 
  1819   To illustrate this function, let us introduce a reference containing a list
  1813   To illustrate this function, let us introduce a reference containing a list
  1820   of theorems.
  1814   of theorems.
  1821 *}
  1815 *}
  1822 
  1816 
  1823 ML{*val my_thms = ref ([] : thm list)*}
  1817 ML{*val my_thms = ref ([] : thm list)*}
  1842 
  1836 
  1843 text {*
  1837 text {*
  1844   These functions take a theorem and a context and, for what we are explaining
  1838   These functions take a theorem and a context and, for what we are explaining
  1845   here it is sufficient that they just return the context unchanged. They change
  1839   here it is sufficient that they just return the context unchanged. They change
  1846   however the reference @{ML my_thms}, whereby the function 
  1840   however the reference @{ML my_thms}, whereby the function 
  1847   @{ML [index] add_thm in Thm} adds a theorem if it is not already included in 
  1841   @{ML_ind [index] add_thm in Thm} adds a theorem if it is not already included in 
  1848   the list, and @{ML [index] del_thm in Thm} deletes one (both functions use the 
  1842   the list, and @{ML_ind [index] del_thm in Thm} deletes one (both functions use the 
  1849   predicate @{ML [index] eq_thm_prop in Thm}, which compares theorems according to 
  1843   predicate @{ML_ind [index] eq_thm_prop in Thm}, which compares theorems according to 
  1850   their proved propositions modulo alpha-equivalence).
  1844   their proved propositions modulo alpha-equivalence).
  1851 
  1845 
  1852   You can turn functions @{ML my_thm_add} and @{ML my_thm_del} into 
  1846   You can turn functions @{ML my_thm_add} and @{ML my_thm_del} into 
  1853   attributes with the code
  1847   attributes with the code
  1854 *}
  1848 *}
  1862 
  1856 
  1863 attribute_setup %gray my_thms = {* Attrib.add_del my_add my_del *} 
  1857 attribute_setup %gray my_thms = {* Attrib.add_del my_add my_del *} 
  1864   "maintaining a list of my_thms - rough test only!" 
  1858   "maintaining a list of my_thms - rough test only!" 
  1865 
  1859 
  1866 text {*
  1860 text {*
  1867   The parser @{ML [index] add_del in Attrib} is a predefined parser for 
  1861   The parser @{ML_ind [index] add_del in Attrib} is a predefined parser for 
  1868   adding and deleting lemmas. Now if you prove the next lemma 
  1862   adding and deleting lemmas. Now if you prove the next lemma 
  1869   and attach to it the attribute @{text "[my_thms]"}
  1863   and attach to it the attribute @{text "[my_thms]"}
  1870 *}
  1864 *}
  1871 
  1865 
  1872 lemma trueI_2[my_thms]: "True" by simp
  1866 lemma trueI_2[my_thms]: "True" by simp
  1890   @{ML_response_fake [display,gray]
  1884   @{ML_response_fake [display,gray]
  1891   "!my_thms"
  1885   "!my_thms"
  1892   "[\"True\", \"Suc (Suc 0) = 2\"]"}
  1886   "[\"True\", \"Suc (Suc 0) = 2\"]"}
  1893 
  1887 
  1894   The theorem @{thm [source] trueI_2} only appears once, since the 
  1888   The theorem @{thm [source] trueI_2} only appears once, since the 
  1895   function @{ML [index] add_thm in Thm} tests for duplicates, before extending
  1889   function @{ML_ind [index] add_thm in Thm} tests for duplicates, before extending
  1896   the list. Deletion from the list works as follows:
  1890   the list. Deletion from the list works as follows:
  1897 *}
  1891 *}
  1898 
  1892 
  1899 declare test[my_thms del]
  1893 declare test[my_thms del]
  1900 
  1894 
  1902 
  1896 
  1903   @{ML_response_fake [display,gray]
  1897   @{ML_response_fake [display,gray]
  1904   "!my_thms"
  1898   "!my_thms"
  1905   "[\"True\"]"}
  1899   "[\"True\"]"}
  1906 
  1900 
  1907   We used in this example two functions declared as @{ML [index] declaration_attribute in Thm}, 
  1901   We used in this example two functions declared as @{ML_ind [index] declaration_attribute in Thm}, 
  1908   but there can be any number of them. We just have to change the parser for reading
  1902   but there can be any number of them. We just have to change the parser for reading
  1909   the arguments accordingly. 
  1903   the arguments accordingly. 
  1910 
  1904 
  1911   However, as said at the beginning of this example, using references for storing theorems is
  1905   However, as said at the beginning of this example, using references for storing theorems is
  1912   \emph{not} the received way of doing such things. The received way is to 
  1906   \emph{not} the received way of doing such things. The received way is to 
  2131 ML {* LocalTheory.set_group *}
  2125 ML {* LocalTheory.set_group *}
  2132 *)
  2126 *)
  2133 
  2127 
  2134 section {* Storing Theorems\label{sec:storing} (TBD) *}
  2128 section {* Storing Theorems\label{sec:storing} (TBD) *}
  2135 
  2129 
  2136 text {* @{ML [index] add_thms_dynamic in PureThy} *}
  2130 text {* @{ML_ind [index] add_thms_dynamic in PureThy} *}
  2137 
  2131 
  2138 local_setup {* 
  2132 local_setup %gray {* 
  2139   LocalTheory.note Thm.theoremK 
  2133   LocalTheory.note Thm.theoremK 
  2140     ((@{binding "allI_alt"}, []), [@{thm allI}]) #> snd *}
  2134     ((@{binding "allI_alt"}, []), [@{thm allI}]) #> snd *}
  2141 
  2135 
  2142 
  2136 
  2143 (* FIXME: some code below *)
  2137 (* FIXME: some code below *)
  2190   where the result indicates that we transformed a string with length 4. Once
  2184   where the result indicates that we transformed a string with length 4. Once
  2191   you have a pretty type, you can, for example, control where linebreaks may
  2185   you have a pretty type, you can, for example, control where linebreaks may
  2192   occur in case the text wraps over a line, or with how much indentation a
  2186   occur in case the text wraps over a line, or with how much indentation a
  2193   text should be printed.  However, if you want to actually output the
  2187   text should be printed.  However, if you want to actually output the
  2194   formatted text, you have to transform the pretty type back into a @{ML_type
  2188   formatted text, you have to transform the pretty type back into a @{ML_type
  2195   string}. This can be done with the function @{ML [index] string_of in Pretty}. In what
  2189   string}. This can be done with the function @{ML_ind [index] string_of in Pretty}. In what
  2196   follows we will use the following wrapper function for printing a pretty
  2190   follows we will use the following wrapper function for printing a pretty
  2197   type:
  2191   type:
  2198 *}
  2192 *}
  2199 
  2193 
  2200 ML{*fun pprint prt = tracing (Pretty.string_of prt)*}
  2194 ML{*fun pprint prt = tracing (Pretty.string_of prt)*}
  2235 aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo
  2229 aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo
  2236 oooooooooooooobaaaaaaaaaaaar"}
  2230 oooooooooooooobaaaaaaaaaaaar"}
  2237 
  2231 
  2238   However by using pretty types you have the ability to indicate a possible
  2232   However by using pretty types you have the ability to indicate a possible
  2239   line break for example at each space. You can achieve this with the function
  2233   line break for example at each space. You can achieve this with the function
  2240   @{ML [index] breaks in Pretty}, which expects a list of pretty types and inserts a
  2234   @{ML_ind [index] breaks in Pretty}, which expects a list of pretty types and inserts a
  2241   possible line break in between every two elements in this list. To print
  2235   possible line break in between every two elements in this list. To print
  2242   this list of pretty types as a single string, we concatenate them 
  2236   this list of pretty types as a single string, we concatenate them 
  2243   with the function @{ML [index] blk in Pretty} as follows:
  2237   with the function @{ML_ind [index] blk in Pretty} as follows:
  2244 
  2238 
  2245 
  2239 
  2246 @{ML_response_fake [display,gray]
  2240 @{ML_response_fake [display,gray]
  2247 "let
  2241 "let
  2248   val ptrs = map Pretty.str (space_explode \" \" test_str)
  2242   val ptrs = map Pretty.str (space_explode \" \" test_str)
  2253 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar 
  2247 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar 
  2254 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
  2248 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
  2255 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
  2249 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
  2256 
  2250 
  2257   Here the layout of @{ML test_str} is much more pleasing to the 
  2251   Here the layout of @{ML test_str} is much more pleasing to the 
  2258   eye. The @{ML "0"} in @{ML [index] blk in Pretty} stands for no indentation
  2252   eye. The @{ML "0"} in @{ML_ind [index] blk in Pretty} stands for no indentation
  2259   of the printed string. You can increase the indentation and obtain
  2253   of the printed string. You can increase the indentation and obtain
  2260   
  2254   
  2261 @{ML_response_fake [display,gray]
  2255 @{ML_response_fake [display,gray]
  2262 "let
  2256 "let
  2263   val ptrs = map Pretty.str (space_explode \" \" test_str)
  2257   val ptrs = map Pretty.str (space_explode \" \" test_str)
  2269    fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
  2263    fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
  2270    fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
  2264    fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
  2271 
  2265 
  2272   where starting from the second line the indent is 3. If you want
  2266   where starting from the second line the indent is 3. If you want
  2273   that every line starts with the same indent, you can use the
  2267   that every line starts with the same indent, you can use the
  2274   function @{ML [index] indent in Pretty} as follows:
  2268   function @{ML_ind [index] indent in Pretty} as follows:
  2275 
  2269 
  2276 @{ML_response_fake [display,gray]
  2270 @{ML_response_fake [display,gray]
  2277 "let
  2271 "let
  2278   val ptrs = map Pretty.str (space_explode \" \" test_str)
  2272   val ptrs = map Pretty.str (space_explode \" \" test_str)
  2279 in
  2273 in
  2284           fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
  2278           fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
  2285           fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
  2279           fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
  2286 
  2280 
  2287   If you want to print out a list of items separated by commas and 
  2281   If you want to print out a list of items separated by commas and 
  2288   have the linebreaks handled properly, you can use the function 
  2282   have the linebreaks handled properly, you can use the function 
  2289   @{ML [index] commas in Pretty}. For example
  2283   @{ML_ind [index] commas in Pretty}. For example
  2290 
  2284 
  2291 @{ML_response_fake [display,gray]
  2285 @{ML_response_fake [display,gray]
  2292 "let
  2286 "let
  2293   val ptrs = map (Pretty.str o string_of_int) (99998 upto 100020)
  2287   val ptrs = map (Pretty.str o string_of_int) (99998 upto 100020)
  2294 in
  2288 in
  2299 100016, 100017, 100018, 100019, 100020"}
  2293 100016, 100017, 100018, 100019, 100020"}
  2300 
  2294 
  2301   where @{ML upto} generates a list of integers. You can print out this
  2295   where @{ML upto} generates a list of integers. You can print out this
  2302   list as a ``set'', that means enclosed inside @{text [quotes] "{"} and
  2296   list as a ``set'', that means enclosed inside @{text [quotes] "{"} and
  2303   @{text [quotes] "}"}, and separated by commas using the function
  2297   @{text [quotes] "}"}, and separated by commas using the function
  2304   @{ML [index] enum in Pretty}. For example
  2298   @{ML_ind [index] enum in Pretty}. For example
  2305 *}
  2299 *}
  2306 
  2300 
  2307 text {*
  2301 text {*
  2308   
  2302   
  2309 @{ML_response_fake [display,gray]
  2303 @{ML_response_fake [display,gray]
  2370 end*}
  2364 end*}
  2371 
  2365 
  2372 text {*
  2366 text {*
  2373   In Line 3 we define a function that inserts possible linebreaks in places
  2367   In Line 3 we define a function that inserts possible linebreaks in places
  2374   where a space is. In Lines 4 and 5 we pretty-print the term and its type
  2368   where a space is. In Lines 4 and 5 we pretty-print the term and its type
  2375   using the functions @{ML [index] pretty_term in Syntax} and @{ML [index]
  2369   using the functions @{ML_ind [index] pretty_term in Syntax} and @{ML_ind [index]
  2376   pretty_typ in Syntax}. We also use the function @{ML [index] quote in
  2370   pretty_typ in Syntax}. We also use the function @{ML_ind [index] quote in
  2377   Pretty} in order to enclose the term and type inside quotation marks. In
  2371   Pretty} in order to enclose the term and type inside quotation marks. In
  2378   Line 9 we add a period right after the type without the possibility of a
  2372   Line 9 we add a period right after the type without the possibility of a
  2379   line break, because we do not want that a line break occurs there.
  2373   line break, because we do not want that a line break occurs there.
  2380 
  2374 
  2381 
  2375 
  2458 ML {* fun my_hook interactive state =
  2452 ML {* fun my_hook interactive state =
  2459          (interactive ? Proof.goal_message (fn () => Pretty.str  
  2453          (interactive ? Proof.goal_message (fn () => Pretty.str  
  2460 "foo")) state
  2454 "foo")) state
  2461 *}
  2455 *}
  2462 
  2456 
  2463 setup {* Context.theory_map (Specification.add_theorem_hook my_hook) *}
  2457 setup %gray {* Context.theory_map (Specification.add_theorem_hook my_hook) *}
  2464 
  2458 
  2465 lemma "False"
  2459 lemma "False"
  2466 oops
  2460 oops
  2467 
  2461 
  2468 
  2462