ProgTutorial/FirstSteps.thy
changeset 316 74f0a06f751f
parent 315 de49d5780f57
child 317 d69214e47ef9
equal deleted inserted replaced
315:de49d5780f57 316:74f0a06f751f
     1 theory FirstSteps
     1 theory FirstSteps
     2 imports Base
     2 imports Base
     3 uses "FirstSteps.ML"
     3 uses "FirstSteps.ML"
     4 begin
     4 begin
     5 
       
     6 
       
     7 
     5 
     8 chapter {* First Steps *}
     6 chapter {* First Steps *}
     9 
     7 
    10 text {*
     8 text {*
    11   Isabelle programming is done in ML. Just like lemmas and proofs, ML-code for
     9   Isabelle programming is done in ML. Just like lemmas and proofs, ML-code for
   113 section {* Debugging and Printing\label{sec:printing} *}
   111 section {* Debugging and Printing\label{sec:printing} *}
   114 
   112 
   115 text {*
   113 text {*
   116   During development you might find it necessary to inspect some data in your
   114   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
   115   code. This can be done in a ``quick-and-dirty'' fashion using the function
   118   @{ML_ind [index] "writeln"}. For example
   116   @{ML_ind  "writeln"}. For example
   119 
   117 
   120   @{ML_response_fake [display,gray] "writeln \"any string\"" "\"any string\""}
   118   @{ML_response_fake [display,gray] "writeln \"any string\"" "\"any string\""}
   121 
   119 
   122   will print out @{text [quotes] "any string"} inside the response buffer of
   120   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
   121   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
   122   PolyML, then there is a convenient, though again ``quick-and-dirty'', method
   125   for converting values into strings, namely the function 
   123   for converting values into strings, namely the function 
   126   @{ML_ind [index] makestring in PolyML}:
   124   @{ML_ind  makestring in PolyML}:
   127 
   125 
   128   @{ML_response_fake [display,gray] "writeln (PolyML.makestring 1)" "\"1\""} 
   126   @{ML_response_fake [display,gray] "writeln (PolyML.makestring 1)" "\"1\""} 
   129 
   127 
   130   However, @{ML makestring in PolyML} only works if the type of what
   128   However, @{ML makestring in PolyML} only works if the type of what
   131   is converted is monomorphic and not a function.
   129   is converted is monomorphic and not a function.
   132 
   130 
   133   The function @{ML "writeln"} should only be used for testing purposes,
   131   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
   132   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
   133   error is raised. For printing anything more serious and elaborate, the
   136   function @{ML_ind [index] tracing} is more appropriate. This function writes all
   134   function @{ML_ind  tracing} is more appropriate. This function writes all
   137   output into a separate tracing buffer. For example:
   135   output into a separate tracing buffer. For example:
   138 
   136 
   139   @{ML_response_fake [display,gray] "tracing \"foo\"" "\"foo\""}
   137   @{ML_response_fake [display,gray] "tracing \"foo\"" "\"foo\""}
   140 
   138 
   141   It is also possible to redirect the ``channel'' where the string @{text
   139   It is also possible to redirect the ``channel'' where the string @{text
   164 
   162 
   165   @{ML [display,gray] "redirect_tracing (TextIO.openOut \"foo.bar\")"}
   163   @{ML [display,gray] "redirect_tracing (TextIO.openOut \"foo.bar\")"}
   166 
   164 
   167   will cause that all tracing information is printed into the file @{text "foo.bar"}.
   165   will cause that all tracing information is printed into the file @{text "foo.bar"}.
   168 
   166 
   169   You can print out error messages with the function @{ML_ind [index] error}; for 
   167   You can print out error messages with the function @{ML_ind  error}; for 
   170   example:
   168   example:
   171 
   169 
   172 @{ML_response_fake [display,gray] "if 0=1 then true else (error \"foo\")" 
   170 @{ML_response_fake [display,gray] "if 0=1 then true else (error \"foo\")" 
   173 "Exception- ERROR \"foo\" raised
   171 "Exception- ERROR \"foo\" raised
   174 At command \"ML\"."}
   172 At command \"ML\"."}
   175 
   173 
   176   This function raises the exception @{text ERROR}, which will then 
   174   This function raises the exception @{text ERROR}, which will then 
   177   be displayed by the infrastructure.
   175   be displayed by the infrastructure.
   178 
   176 
   179 
   177 
   180   (FIXME Mention how to work with @{ML_ind [index] debug in Toplevel} and 
   178   (FIXME Mention how to work with @{ML_ind  debug in Toplevel} and 
   181   @{ML_ind [index] profiling in Toplevel}.)
   179   @{ML_ind  profiling in Toplevel}.)
   182 *}
   180 *}
   183 
   181 
   184 (*
   182 (*
   185 ML {* reset Toplevel.debug *}
   183 ML {* reset Toplevel.debug *}
   186 
   184 
   199   Most often you want to inspect data of Isabelle's most basic data
   197   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
   198   structures, namely @{ML_type term}, @{ML_type cterm} and @{ML_type
   201   thm}. Isabelle contains elaborate pretty-printing functions for printing
   199   thm}. Isabelle contains elaborate pretty-printing functions for printing
   202   them (see Section \ref{sec:pretty}), but for quick-and-dirty solutions they
   200   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
   201   are a bit unwieldy. One way to transform a term into a string is to use the
   204   function @{ML_ind [index] string_of_term in Syntax} in structure @{ML_struct
   202   function @{ML_ind  string_of_term in Syntax} in structure @{ML_struct
   205   Syntax}, which we bind for more convenience to the toplevel.
   203   Syntax}, which we bind for more convenience to the toplevel.
   206 *}
   204 *}
   207 
   205 
   208 ML{*val string_of_term = Syntax.string_of_term*}
   206 ML{*val string_of_term = Syntax.string_of_term*}
   209 
   207 
   214   "string_of_term @{context} @{term \"1::nat\"}"
   212   "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\""}
   213   "\"\\^E\\^Fterm\\^E\\^E\\^Fconst\\^Fname=HOL.one_class.one\\^E1\\^E\\^F\\^E\\^E\\^F\\^E\""}
   216 
   214 
   217   This produces a string corrsponding to the term @{term [show_types] "1::nat"} with some
   215   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
   216   additional information encoded in it. The string can be properly printed by
   219   using either the function @{ML_ind [index] writeln} or @{ML_ind [index] tracing}:
   217   using either the function @{ML_ind  writeln} or @{ML_ind  tracing}:
   220 
   218 
   221   @{ML_response_fake [display,gray]
   219   @{ML_response_fake [display,gray]
   222   "writeln (string_of_term @{context} @{term \"1::nat\"})"
   220   "writeln (string_of_term @{context} @{term \"1::nat\"})"
   223   "\"1\""}
   221   "\"1\""}
   224 
   222 
   227   @{ML_response_fake [display,gray]
   225   @{ML_response_fake [display,gray]
   228   "tracing (string_of_term @{context} @{term \"1::nat\"})"
   226   "tracing (string_of_term @{context} @{term \"1::nat\"})"
   229   "\"1\""}
   227   "\"1\""}
   230 
   228 
   231   If there are more than one @{ML_type term}s to be printed, you can use the 
   229   If there are more than one @{ML_type term}s to be printed, you can use the 
   232   function @{ML_ind [index] commas} to separate them.
   230   function @{ML_ind  commas} to separate them.
   233 *} 
   231 *} 
   234 
   232 
   235 ML{*fun string_of_terms ctxt ts =  
   233 ML{*fun string_of_terms ctxt ts =  
   236   commas (map (string_of_term ctxt) ts)*}
   234   commas (map (string_of_term ctxt) ts)*}
   237 
   235 
   241 
   239 
   242 ML{*fun string_of_cterm ctxt ct =  
   240 ML{*fun string_of_cterm ctxt ct =  
   243   string_of_term ctxt (term_of ct)*}
   241   string_of_term ctxt (term_of ct)*}
   244 
   242 
   245 text {*
   243 text {*
   246   In this example the function @{ML_ind [index] term_of} extracts the @{ML_type
   244   In this example the function @{ML_ind  term_of} extracts the @{ML_type
   247   term} from a @{ML_type cterm}.  More than one @{ML_type cterm}s can again be
   245   term} from a @{ML_type cterm}.  More than one @{ML_type cterm}s can again be
   248   printed with @{ML_ind [index] commas}.
   246   printed with @{ML_ind  commas}.
   249 *} 
   247 *} 
   250 
   248 
   251 ML{*fun string_of_cterms ctxt cts =  
   249 ML{*fun string_of_cterms ctxt cts =  
   252   commas (map (string_of_cterm ctxt) cts)*}
   250   commas (map (string_of_cterm ctxt) cts)*}
   253 
   251 
   254 text {*
   252 text {*
   255   The easiest way to get the string of a theorem is to transform it
   253   The easiest way to get the string of a theorem is to transform it
   256   into a @{ML_type cterm} using the function @{ML_ind [index] crep_thm}. 
   254   into a @{ML_type cterm} using the function @{ML_ind  crep_thm}. 
   257 *}
   255 *}
   258 
   256 
   259 ML{*fun string_of_thm ctxt thm =
   257 ML{*fun string_of_thm ctxt thm =
   260   string_of_cterm ctxt (#prop (crep_thm thm))*}
   258   string_of_cterm ctxt (#prop (crep_thm thm))*}
   261 
   259 
   270   "tracing (string_of_thm @{context} @{thm conjI})"
   268   "tracing (string_of_thm @{context} @{thm conjI})"
   271   "\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q"}
   269   "\<lbrakk>?P; ?Q\<rbrakk> \<Longrightarrow> ?P \<and> ?Q"}
   272 
   270 
   273   However, in order to improve the readability when printing theorems, we
   271   However, in order to improve the readability when printing theorems, we
   274   convert these schematic variables into free variables using the function
   272   convert these schematic variables into free variables using the function
   275   @{ML_ind [index] import in Variable}. This is similar to statements like @{text
   273   @{ML_ind  import in Variable}. This is similar to statements like @{text
   276   "conjI[no_vars]"} from Isabelle's user-level.
   274   "conjI[no_vars]"} from Isabelle's user-level.
   277 *}
   275 *}
   278 
   276 
   279 ML{*fun no_vars ctxt thm =
   277 ML{*fun no_vars ctxt thm =
   280 let 
   278 let 
   320 "tracing (\"First half,\" ^ \"\\n\" ^ \"and second half.\")"
   318 "tracing (\"First half,\" ^ \"\\n\" ^ \"and second half.\")"
   321 "First half,
   319 "First half,
   322 and second half."}
   320 and second half."}
   323   
   321   
   324   To ease this kind of string manipulations, there are a number
   322   To ease this kind of string manipulations, there are a number
   325   of library functions. For example,  the function @{ML_ind [index] cat_lines}
   323   of library functions. For example,  the function @{ML_ind  cat_lines}
   326   concatenates a list of strings and inserts newlines. 
   324   concatenates a list of strings and inserts newlines. 
   327 
   325 
   328   @{ML_response [display, gray]
   326   @{ML_response [display, gray]
   329   "cat_lines [\"foo\", \"bar\"]"
   327   "cat_lines [\"foo\", \"bar\"]"
   330   "\"foo\\nbar\""}
   328   "\"foo\\nbar\""}
   340   For beginners perhaps the most puzzling parts in the existing code of Isabelle are
   338   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
   339   the combinators. At first they seem to greatly obstruct the
   342   comprehension of the code, but after getting familiar with them, they
   340   comprehension of the code, but after getting familiar with them, they
   343   actually ease the understanding and also the programming.
   341   actually ease the understanding and also the programming.
   344 
   342 
   345   The simplest combinator is @{ML_ind [index] I}, which is just the identity function defined as
   343   The simplest combinator is @{ML_ind  I}, which is just the identity function defined as
   346 *}
   344 *}
   347 
   345 
   348 ML{*fun I x = x*}
   346 ML{*fun I x = x*}
   349 
   347 
   350 text {* Another simple combinator is @{ML_ind [index] K}, defined as *}
   348 text {* Another simple combinator is @{ML_ind  K}, defined as *}
   351 
   349 
   352 ML{*fun K x = fn _ => x*}
   350 ML{*fun K x = fn _ => x*}
   353 
   351 
   354 text {*
   352 text {*
   355   @{ML_ind [index] K} ``wraps'' a function around the argument @{text "x"}. However, this 
   353   @{ML_ind  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 
   354   function ignores its argument. As a result, @{ML K} defines a constant function 
   357   always returning @{text x}.
   355   always returning @{text x}.
   358 
   356 
   359   The next combinator is reverse application, @{ML_ind [index] "|>"}, defined as: 
   357   The next combinator is reverse application, @{ML_ind  "|>"}, defined as: 
   360 *}
   358 *}
   361 
   359 
   362 ML{*fun x |> f = f x*}
   360 ML{*fun x |> f = f x*}
   363 
   361 
   364 text {* While just syntactic sugar for the usual function application,
   362 text {* While just syntactic sugar for the usual function application,
   432    |> tracing
   430    |> tracing
   433 end" 
   431 end" 
   434   "P z za zb"}
   432   "P z za zb"}
   435 
   433 
   436   You can read off this behaviour from how @{ML apply_fresh_args} is
   434   You can read off this behaviour from how @{ML apply_fresh_args} is
   437   coded: in Line 2, the function @{ML_ind [index] fastype_of} calculates the type of the
   435   coded: in Line 2, the function @{ML_ind  fastype_of} calculates the type of the
   438   term; @{ML_ind [index] binder_types} in the next line produces the list of argument
   436   term; @{ML_ind  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 
   437   types (in the case above the list @{text "[nat, int, unit]"}); Line 4 
   440   pairs up each type with the string  @{text "z"}; the
   438   pairs up each type with the string  @{text "z"}; the
   441   function @{ML_ind [index] variant_frees in Variable} generates for each @{text "z"} a
   439   function @{ML_ind  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
   440   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
   441   into a list of variable terms in Line 6, which in the last line is applied
   444   by the function @{ML_ind [index] list_comb} to the term. In this last step we have to 
   442   by the function @{ML_ind  list_comb} to the term. In this last step we have to 
   445   use the function @{ML_ind [index] curry}, because @{ML_ind [index] list_comb} expects the 
   443   use the function @{ML_ind  curry}, because @{ML_ind  list_comb} expects the 
   446   function and the variables list as a pair. This kind of functions is often needed when
   444   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 
   445   constructing terms with fresh variables. The infrastructure helps tremendously 
   448   to avoid any name clashes. Consider for example: 
   446   to avoid any name clashes. Consider for example: 
   449 
   447 
   450    @{ML_response_fake [display,gray]
   448    @{ML_response_fake [display,gray]
   458 end"
   456 end"
   459   "za z zb"}
   457   "za z zb"}
   460   
   458   
   461   where the @{text "za"} is correctly avoided.
   459   where the @{text "za"} is correctly avoided.
   462 
   460 
   463   The combinator @{ML_ind [index] "#>"} is the reverse function composition. It can be
   461   The combinator @{ML_ind  "#>"} is the reverse function composition. It can be
   464   used to define the following function
   462   used to define the following function
   465 *}
   463 *}
   466 
   464 
   467 ML{*val inc_by_six = 
   465 ML{*val inc_by_six = 
   468       (fn x => x + 1)
   466       (fn x => x + 1)
   473   which is the function composed of first the increment-by-one function and then
   471   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 
   472   increment-by-two, followed by increment-by-three. Again, the reverse function 
   475   composition allows you to read the code top-down.
   473   composition allows you to read the code top-down.
   476 
   474 
   477   The remaining combinators described in this section add convenience for the
   475   The remaining combinators described in this section add convenience for the
   478   ``waterfall method'' of writing functions. The combinator @{ML_ind [index] tap} allows
   476   ``waterfall method'' of writing functions. The combinator @{ML_ind  tap} allows
   479   you to get hold of an intermediate result (to do some side-calculations for
   477   you to get hold of an intermediate result (to do some side-calculations for
   480   instance). The function
   478   instance). The function
   481 
   479 
   482  *}
   480  *}
   483 
   481 
   486        |> tap (fn x => tracing (PolyML.makestring x))
   484        |> tap (fn x => tracing (PolyML.makestring x))
   487        |> (fn x => x + 2)*}
   485        |> (fn x => x + 2)*}
   488 
   486 
   489 text {* 
   487 text {* 
   490   increments the argument first by @{text "1"} and then by @{text "2"}. In the
   488   increments the argument first by @{text "1"} and then by @{text "2"}. In the
   491   middle (Line 3), however, it uses @{ML_ind [index] tap} for printing the ``plus-one''
   489   middle (Line 3), however, it uses @{ML_ind  tap} for printing the ``plus-one''
   492   intermediate result inside the tracing buffer. The function @{ML_ind [index] tap} can
   490   intermediate result inside the tracing buffer. The function @{ML_ind  tap} can
   493   only be used for side-calculations, because any value that is computed
   491   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
   492   cannot be merged back into the ``main waterfall''. To do this, you can use
   495   the next combinator.
   493   the next combinator.
   496 
   494 
   497   The combinator @{ML_ind [index] "`"} (a backtick) is similar to @{ML tap}, but applies a
   495   The combinator @{ML_ind  "`"} (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
   496   function to the value and returns the result together with the value (as a
   499   pair). For example the function 
   497   pair). For example the function 
   500 *}
   498 *}
   501 
   499 
   502 ML{*fun inc_as_pair x =
   500 ML{*fun inc_as_pair x =
   507   takes @{text x} as argument, and then increments @{text x}, but also keeps
   505   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)"
   506   @{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
   507   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}.
   508   pair. So finally the result will be @{ML "(x + 1, x + 1)" for x}.
   511 
   509 
   512   The combinators @{ML_ind [index] "|>>"} and @{ML_ind [index] "||>"} are defined for 
   510   The combinators @{ML_ind  "|>>"} and @{ML_ind  "||>"} are defined for 
   513   functions manipulating pairs. The first applies the function to
   511   functions manipulating pairs. The first applies the function to
   514   the first component of the pair, defined as
   512   the first component of the pair, defined as
   515 *}
   513 *}
   516 
   514 
   517 ML{*fun (x, y) |>> f = (f x, y)*}
   515 ML{*fun (x, y) |>> f = (f x, y)*}
   554 text {*
   552 text {*
   555   avoiding the explicit @{text "let"}. While in this example the gain in
   553   avoiding the explicit @{text "let"}. While in this example the gain in
   556   conciseness is only small, in more complicated situations the benefit of
   554   conciseness is only small, in more complicated situations the benefit of
   557   avoiding @{text "lets"} can be substantial.
   555   avoiding @{text "lets"} can be substantial.
   558 
   556 
   559   With the combinator @{ML_ind [index] "|->"} you can re-combine the elements from a pair.
   557   With the combinator @{ML_ind  "|->"} you can re-combine the elements from a pair.
   560   This combinator is defined as
   558   This combinator is defined as
   561 *}
   559 *}
   562 
   560 
   563 ML{*fun (x, y) |-> f = f x y*}
   561 ML{*fun (x, y) |-> f = f x y*}
   564 
   562 
   570 ML{*fun double x =
   568 ML{*fun double x =
   571       x |>  (fn x => (x, x))
   569       x |>  (fn x => (x, x))
   572         |-> (fn x => fn y => x + y)*}
   570         |-> (fn x => fn y => x + y)*}
   573 
   571 
   574 text {* 
   572 text {* 
   575   The combinator @{ML_ind [index] ||>>} plays a central rôle whenever your task is to update a 
   573   The combinator @{ML_ind  ||>>} 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 
   574   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 
   575   for such tasks return a pair whose second component is the theory and the fist 
   578   component is the side-result. Using @{ML_ind [index] ||>>}, you can do conveniently the update 
   576   component is the side-result. Using @{ML_ind  ||>>}, you can do conveniently the update 
   579   and also accumulate the side-results. Consider the following simple function. 
   577   and also accumulate the side-results. Consider the following simple function. 
   580 *}
   578 *}
   581 
   579 
   582 ML %linenosgray{*fun acc_incs x =
   580 ML %linenosgray{*fun acc_incs x =
   583     x |> (fn x => ("", x)) 
   581     x |> (fn x => ("", x)) 
   585       ||>> (fn x => (x, x + 1))
   583       ||>> (fn x => (x, x + 1))
   586       ||>> (fn x => (x, x + 1))*}
   584       ||>> (fn x => (x, x + 1))*}
   587 
   585 
   588 text {*
   586 text {*
   589   The purpose of Line 2 is to just pair up the argument with a dummy value (since
   587   The purpose of Line 2 is to just pair up the argument with a dummy value (since
   590   @{ML_ind [index] "||>>"} operates on pairs). Each of the next three lines just increment 
   588   @{ML_ind  "||>>"} 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 
   589   the value by one, but also nest the intermediate results to the left. For example 
   592 
   590 
   593   @{ML_response [display,gray]
   591   @{ML_response [display,gray]
   594   "acc_incs 1"
   592   "acc_incs 1"
   595   "((((\"\", 1), 2), 3), 4)"}
   593   "((((\"\", 1), 2), 3), 4)"}
   602 
   600 
   603   (FIXME: maybe give a ``real world'' example for this combinator)
   601   (FIXME: maybe give a ``real world'' example for this combinator)
   604 *}
   602 *}
   605 
   603 
   606 text {*
   604 text {*
   607   Recall that @{ML_ind [index] "|>"} is the reverse function application. Recall also that
   605   Recall that @{ML_ind  "|>"} is the reverse function application. Recall also that
   608   the related 
   606   the related 
   609   reverse function composition is @{ML_ind [index] "#>"}. In fact all the combinators 
   607   reverse function composition is @{ML_ind  "#>"}. In fact all the combinators 
   610   @{ML_ind [index] "|->"}, @{ML_ind [index] "|>>"} , @{ML_ind [index] "||>"} and @{ML_ind [index] 
   608   @{ML_ind  "|->"}, @{ML_ind  "|>>"} , @{ML_ind  "||>"} and @{ML_ind  
   611   "||>>"} described above have related combinators for 
   609   "||>>"} described above have related combinators for 
   612   function composition, namely @{ML_ind [index] "#->"}, @{ML_ind [index] "#>>"}, 
   610   function composition, namely @{ML_ind  "#->"}, @{ML_ind  "#>>"}, 
   613   @{ML_ind [index] "##>"} and @{ML_ind [index] "##>>"}. 
   611   @{ML_ind  "##>"} and @{ML_ind  "##>>"}. 
   614   Using @{ML_ind [index] "#->"}, for example, the function @{text double} can also be written as:
   612   Using @{ML_ind  "#->"}, for example, the function @{text double} can also be written as:
   615 *}
   613 *}
   616 
   614 
   617 ML{*val double =
   615 ML{*val double =
   618             (fn x => (x, x))
   616             (fn x => (x, x))
   619         #-> (fn x => fn y => x + y)*}
   617         #-> (fn x => fn y => x + y)*}
   625   together. We have already seen such plumbing in the function @{ML
   623   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
   624   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 
   625   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, 
   626   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 
   627   but one calculates only with a single entity. An example is the function 
   630   @{ML_ind [index] check_terms in Syntax}, whose purpose is to type-check a list 
   628   @{ML_ind  check_terms in Syntax}, whose purpose is to type-check a list 
   631   of terms.
   629   of terms.
   632 
   630 
   633   @{ML_response_fake [display, gray]
   631   @{ML_response_fake [display, gray]
   634   "let
   632   "let
   635   val ctxt = @{context}
   633   val ctxt = @{context}
   688   @{ML_response [display,gray] "Context.theory_name @{theory}" "\"FirstSteps\""}
   686   @{ML_response [display,gray] "Context.theory_name @{theory}" "\"FirstSteps\""}
   689  
   687  
   690   where @{text "@{theory}"} is an antiquotation that is substituted with the
   688   where @{text "@{theory}"} is an antiquotation that is substituted with the
   691   current theory (remember that we assumed we are inside the theory 
   689   current theory (remember that we assumed we are inside the theory 
   692   @{text FirstSteps}). The name of this theory can be extracted using
   690   @{text FirstSteps}). The name of this theory can be extracted using
   693   the function @{ML_ind [index] theory_name in Context}. 
   691   the function @{ML_ind  theory_name in Context}. 
   694 
   692 
   695   Note, however, that antiquotations are statically linked, that is their value is
   693   Note, however, that antiquotations are statically linked, that is their value is
   696   determined at ``compile-time'', not ``run-time''. For example the function
   694   determined at ``compile-time'', not ``run-time''. For example the function
   697 *}
   695 *}
   698 
   696 
   736 in
   734 in
   737   map #1 simps
   735   map #1 simps
   738 end*}
   736 end*}
   739 
   737 
   740 text {*
   738 text {*
   741   The function @{ML_ind [index] dest_ss in MetaSimplifier} returns a record containing all
   739   The function @{ML_ind  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
   740   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. 
   741   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}"}.
   742   The current simpset can be referred to using the antiquotation @{text "@{simpset}"}.
   745 
   743 
   746   @{ML_response_fake [display,gray] 
   744   @{ML_response_fake [display,gray] 
   759   @{ML_response [display,gray]
   757   @{ML_response [display,gray]
   760   "@{binding \"name\"}"
   758   "@{binding \"name\"}"
   761   "name"}
   759   "name"}
   762 
   760 
   763   An example where a qualified name is needed is the function 
   761   An example where a qualified name is needed is the function 
   764   @{ML_ind [index] define in LocalTheory}.  This function is used below to define 
   762   @{ML_ind  define in LocalTheory}.  This function is used below to define 
   765   the constant @{term "TrueConj"} as the conjunction @{term "True \<and> True"}.
   763   the constant @{term "TrueConj"} as the conjunction @{term "True \<and> True"}.
   766 *}
   764 *}
   767   
   765   
   768 local_setup %gray {* 
   766 local_setup %gray {* 
   769   snd o LocalTheory.define Thm.internalK 
   767   snd o LocalTheory.define Thm.internalK 
   786   exercise care when introducing new ones, as they can also make your code
   784   exercise care when introducing new ones, as they can also make your code
   787   also difficult to read. In the next section we will see how the (build in)
   785   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.  A
   786   antiquotation @{text "@{term \<dots>}"} can be used to construct terms.  A
   789   restriction of this antiquotation is that it does not allow you to use
   787   restriction of this antiquotation is that it does not allow you to use
   790   schematic variables. If you want to have an antiquotation that does not have
   788   schematic variables. If you want to have an antiquotation that does not have
   791   this restriction, you can implement your own using the function @{ML_ind [index]
   789   this restriction, you can implement your own using the function @{ML_ind 
   792   ML_Antiquote.inline}. The code is as follows.
   790   ML_Antiquote.inline}. The code is as follows.
   793 *}
   791 *}
   794 
   792 
   795 ML %linenosgray{*ML_Antiquote.inline "term_pat"
   793 ML %linenosgray{*ML_Antiquote.inline "term_pat"
   796   (Args.context -- Scan.lift Args.name_source >>
   794   (Args.context -- Scan.lift Args.name_source >>
   799          |> ML_Syntax.print_term
   797          |> ML_Syntax.print_term
   800          |> ML_Syntax.atomic))*}
   798          |> ML_Syntax.atomic))*}
   801 
   799 
   802 text {*
   800 text {*
   803   The parser in Line 2 provides us with a context and a string; this string is
   801   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_ind [index] read_term_pattern in
   802   transformed into a term using the function @{ML_ind  read_term_pattern in
   805   ProofContext} (Line 4); the next two lines print the term so that the
   803   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
   804   ML-system can understand them. An example of this antiquotation is as
   807   follows.
   805   follows.
   808 
   806 
   809   @{ML_response_fake [display,gray]
   807   @{ML_response_fake [display,gray]
   843 "@{term \"(a::nat) + b = c\"}" 
   841 "@{term \"(a::nat) + b = c\"}" 
   844 "Const (\"op =\", \<dots>) $ 
   842 "Const (\"op =\", \<dots>) $ 
   845   (Const (\"HOL.plus_class.plus\", \<dots>) $ \<dots> $ \<dots>) $ \<dots>"}
   843   (Const (\"HOL.plus_class.plus\", \<dots>) $ \<dots> $ \<dots>) $ \<dots>"}
   846 
   844 
   847   will show the term @{term "(a::nat) + b = c"}, but printed using the internal
   845   will show the term @{term "(a::nat) + b = c"}, but printed using the internal
   848   representation corresponding to the datatype @{ML_type [index] "term"} defined as follows: 
   846   representation corresponding to the datatype @{ML_type  "term"} defined as follows: 
   849 *}  
   847 *}  
   850 
   848 
   851 ML_val %linenosgray{*datatype term =
   849 ML_val %linenosgray{*datatype term =
   852   Const of string * typ 
   850   Const of string * typ 
   853 | Free of string * typ 
   851 | Free of string * typ 
  1043   Abs (\"x\", \<dots>,
  1041   Abs (\"x\", \<dots>,
  1044     Const \<dots> $ (Const \<dots> $ (Free (\"P\",\<dots>) $ \<dots>)) $ 
  1042     Const \<dots> $ (Const \<dots> $ (Free (\"P\",\<dots>) $ \<dots>)) $ 
  1045                 (Const \<dots> $ (Free (\"Q\",\<dots>) $ \<dots>)))"}
  1043                 (Const \<dots> $ (Free (\"Q\",\<dots>) $ \<dots>)))"}
  1046 
  1044 
  1047   There are a number of handy functions that are frequently used for 
  1045   There are a number of handy functions that are frequently used for 
  1048   constructing terms. One is the function @{ML_ind [index] list_comb}, which takes a term
  1046   constructing terms. One is the function @{ML_ind  list_comb}, which takes a term
  1049   and a list of terms as arguments, and produces as output the term
  1047   and a list of terms as arguments, and produces as output the term
  1050   list applied to the term. For example
  1048   list applied to the term. For example
  1051 
  1049 
  1052 @{ML_response_fake [display,gray]
  1050 @{ML_response_fake [display,gray]
  1053 "let
  1051 "let
  1056 in
  1054 in
  1057   list_comb (trm, args)
  1055   list_comb (trm, args)
  1058 end"
  1056 end"
  1059 "Free (\"P\", \"nat\") $ Const (\"True\", \"bool\") $ Const (\"False\", \"bool\")"}
  1057 "Free (\"P\", \"nat\") $ Const (\"True\", \"bool\") $ Const (\"False\", \"bool\")"}
  1060 
  1058 
  1061   Another handy function is @{ML_ind [index] lambda}, which abstracts a variable 
  1059   Another handy function is @{ML_ind  lambda}, which abstracts a variable 
  1062   in a term. For example
  1060   in a term. For example
  1063 
  1061 
  1064   @{ML_response_fake [display,gray]
  1062   @{ML_response_fake [display,gray]
  1065 "let
  1063 "let
  1066   val x_nat = @{term \"x::nat\"}
  1064   val x_nat = @{term \"x::nat\"}
  1089   then the variable @{text "Free (\"x\", \"int\")"} is \emph{not} abstracted. 
  1087   then the variable @{text "Free (\"x\", \"int\")"} is \emph{not} abstracted. 
  1090   This is a fundamental principle
  1088   This is a fundamental principle
  1091   of Church-style typing, where variables with the same name still differ, if they 
  1089   of Church-style typing, where variables with the same name still differ, if they 
  1092   have different type.
  1090   have different type.
  1093 
  1091 
  1094   There is also the function @{ML_ind [index] subst_free} with which terms can be
  1092   There is also the function @{ML_ind  subst_free} with which terms can be
  1095   replaced by other terms. For example below, we will replace in @{term
  1093   replaced by other terms. For example below, we will replace in @{term
  1096   "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0 x"} the subterm @{term "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0"} by
  1094   "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0 x"} the subterm @{term "(f::nat \<Rightarrow> nat \<Rightarrow> nat) 0"} by
  1097   @{term y}, and @{term x} by @{term True}.
  1095   @{term y}, and @{term x} by @{term True}.
  1098 
  1096 
  1099   @{ML_response_fake [display,gray]
  1097   @{ML_response_fake [display,gray]
  1116 in
  1114 in
  1117   subst_free [sub] trm
  1115   subst_free [sub] trm
  1118 end"
  1116 end"
  1119   "Free (\"x\", \"nat\")"}
  1117   "Free (\"x\", \"nat\")"}
  1120 
  1118 
  1121   Similarly the function @{ML_ind [index] subst_bounds}, replaces lose bound 
  1119   Similarly the function @{ML_ind  subst_bounds}, replaces lose bound 
  1122   variables with terms. To see how this function works, let us implement a
  1120   variables with terms. To see how this function works, let us implement a
  1123   function that strips off the outermost quantifiers in a term.
  1121   function that strips off the outermost quantifiers in a term.
  1124 *}
  1122 *}
  1125 
  1123 
  1126 ML{*fun strip_alls (Const ("All", _) $ Abs (n, T, t)) =
  1124 ML{*fun strip_alls (Const ("All", _) $ Abs (n, T, t)) =
  1135   "strip_alls @{term \"\<forall>x y. x = (y::bool)\"}"
  1133   "strip_alls @{term \"\<forall>x y. x = (y::bool)\"}"
  1136 "([Free (\"x\", \"bool\"), Free (\"y\", \"bool\")],
  1134 "([Free (\"x\", \"bool\"), Free (\"y\", \"bool\")],
  1137   Const (\"op =\", \<dots>) $ Bound 1 $ Bound 0)"}
  1135   Const (\"op =\", \<dots>) $ Bound 1 $ Bound 0)"}
  1138 
  1136 
  1139   After calling @{ML strip_alls}, you obtain a term with lose bound variables. With
  1137   After calling @{ML strip_alls}, you obtain a term with lose bound variables. With
  1140   the function @{ML subst_bounds}, you can replace these lose @{ML_ind [index]
  1138   the function @{ML subst_bounds}, you can replace these lose @{ML_ind 
  1141   Bound}s with the stripped off variables.
  1139   Bound}s with the stripped off variables.
  1142 
  1140 
  1143   @{ML_response_fake [display, gray, linenos]
  1141   @{ML_response_fake [display, gray, linenos]
  1144   "let
  1142   "let
  1145   val (vrs, trm) = strip_alls @{term \"\<forall>x y. x = (y::bool)\"}
  1143   val (vrs, trm) = strip_alls @{term \"\<forall>x y. x = (y::bool)\"}
  1154   returned. The reason is that the head of the list the function @{ML subst_bounds}
  1152   returned. The reason is that the head of the list the function @{ML subst_bounds}
  1155   takes is the replacement for @{ML "Bound 0"}, the next element for @{ML "Bound 1"}
  1153   takes is the replacement for @{ML "Bound 0"}, the next element for @{ML "Bound 1"}
  1156   and so on. 
  1154   and so on. 
  1157 
  1155 
  1158   There are many convenient functions that construct specific HOL-terms. For
  1156   There are many convenient functions that construct specific HOL-terms. For
  1159   example @{ML_ind [index] mk_eq in HOLogic} constructs an equality out of two terms.
  1157   example @{ML_ind  mk_eq in HOLogic} constructs an equality out of two terms.
  1160   The types needed in this equality are calculated from the type of the
  1158   The types needed in this equality are calculated from the type of the
  1161   arguments. For example
  1159   arguments. For example
  1162 
  1160 
  1163 @{ML_response_fake [gray,display]
  1161 @{ML_response_fake [gray,display]
  1164 "let
  1162 "let
  1268   "@{type_name \"list\"}" "\"List.list\""}
  1266   "@{type_name \"list\"}" "\"List.list\""}
  1269 
  1267 
  1270   (FIXME: Explain the following better.)
  1268   (FIXME: Explain the following better.)
  1271 
  1269 
  1272   Occasionally you have to calculate what the ``base'' name of a given
  1270   Occasionally you have to calculate what the ``base'' name of a given
  1273   constant is. For this you can use the function @{ML_ind [index] "Sign.extern_const"} or
  1271   constant is. For this you can use the function @{ML_ind  "Sign.extern_const"} or
  1274   @{ML_ind [index] Long_Name.base_name}. For example:
  1272   @{ML_ind  Long_Name.base_name}. For example:
  1275 
  1273 
  1276   @{ML_response [display,gray] "Sign.extern_const @{theory} \"List.list.Nil\"" "\"Nil\""}
  1274   @{ML_response [display,gray] "Sign.extern_const @{theory} \"List.list.Nil\"" "\"Nil\""}
  1277 
  1275 
  1278   The difference between both functions is that @{ML extern_const in Sign} returns
  1276   The difference between both functions is that @{ML extern_const in Sign} returns
  1279   the smallest name that is still unique, whereas @{ML base_name in Long_Name} always
  1277   the smallest name that is still unique, whereas @{ML base_name in Long_Name} always
  1291 
  1289 
  1292 *} 
  1290 *} 
  1293 
  1291 
  1294 ML{*fun make_fun_type ty1 ty2 = Type ("fun", [ty1, ty2]) *}
  1292 ML{*fun make_fun_type ty1 ty2 = Type ("fun", [ty1, ty2]) *}
  1295 
  1293 
  1296 text {* This can be equally written with the combinator @{ML_ind [index] "-->"} as: *}
  1294 text {* This can be equally written with the combinator @{ML_ind  "-->"} as: *}
  1297 
  1295 
  1298 ML{*fun make_fun_type ty1 ty2 = ty1 --> ty2 *}
  1296 ML{*fun make_fun_type ty1 ty2 = ty1 --> ty2 *}
  1299 
  1297 
  1300 text {*
  1298 text {*
  1301   If you want to construct a function type with more than one argument
  1299   If you want to construct a function type with more than one argument
  1302   type, then you can use @{ML_ind [index] "--->"}.
  1300   type, then you can use @{ML_ind  "--->"}.
  1303 *}
  1301 *}
  1304 
  1302 
  1305 ML{*fun make_fun_types tys ty = tys ---> ty *}
  1303 ML{*fun make_fun_types tys ty = tys ---> ty *}
  1306 
  1304 
  1307 text {*
  1305 text {*
  1308   A handy function for manipulating terms is @{ML_ind [index] map_types}: it takes a 
  1306   A handy function for manipulating terms is @{ML_ind  map_types}: it takes a 
  1309   function and applies it to every type in a term. You can, for example,
  1307   function and applies it to every type in a term. You can, for example,
  1310   change every @{typ nat} in a term into an @{typ int} using the function:
  1308   change every @{typ nat} in a term into an @{typ int} using the function:
  1311 *}
  1309 *}
  1312 
  1310 
  1313 ML{*fun nat_to_int ty =
  1311 ML{*fun nat_to_int ty =
  1323 "map_types nat_to_int @{term \"a = (1::nat)\"}" 
  1321 "map_types nat_to_int @{term \"a = (1::nat)\"}" 
  1324 "Const (\"op =\", \"int \<Rightarrow> int \<Rightarrow> bool\")
  1322 "Const (\"op =\", \"int \<Rightarrow> int \<Rightarrow> bool\")
  1325            $ Free (\"a\", \"int\") $ Const (\"HOL.one_class.one\", \"int\")"}
  1323            $ Free (\"a\", \"int\") $ Const (\"HOL.one_class.one\", \"int\")"}
  1326 
  1324 
  1327   If you want to obtain the list of free type-variables of a term, you
  1325   If you want to obtain the list of free type-variables of a term, you
  1328   can use the function @{ML_ind [index] add_tfrees in Term} 
  1326   can use the function @{ML_ind  add_tfrees in Term} 
  1329   (similarly @{ML_ind [index] add_tvars in Term} for the schematic type-variables). 
  1327   (similarly @{ML_ind  add_tvars in Term} for the schematic type-variables). 
  1330   One would expect that such functions
  1328   One would expect that such functions
  1331   take a term as input and return a list of types. But their type is actually 
  1329   take a term as input and return a list of types. But their type is actually 
  1332 
  1330 
  1333   @{text[display] "Term.term -> (string * Term.sort) list -> (string * Term.sort) list"}
  1331   @{text[display] "Term.term -> (string * Term.sort) list -> (string * Term.sort) list"}
  1334 
  1332 
  1393 text {* 
  1391 text {* 
  1394   
  1392   
  1395   You can freely construct and manipulate @{ML_type "term"}s and @{ML_type
  1393   You can freely construct and manipulate @{ML_type "term"}s and @{ML_type
  1396   typ}es, since they are just arbitrary unchecked trees. However, you
  1394   typ}es, since they are just arbitrary unchecked trees. However, you
  1397   eventually want to see if a term is well-formed, or type-checks, relative to
  1395   eventually want to see if a term is well-formed, or type-checks, relative to
  1398   a theory.  Type-checking is done via the function @{ML_ind [index] cterm_of}, which
  1396   a theory.  Type-checking is done via the function @{ML_ind  cterm_of}, which
  1399   converts a @{ML_type [index] term} into a @{ML_type [index] cterm}, a \emph{certified}
  1397   converts a @{ML_type  term} into a @{ML_type  cterm}, a \emph{certified}
  1400   term. Unlike @{ML_type term}s, which are just trees, @{ML_type "cterm"}s are
  1398   term. Unlike @{ML_type term}s, which are just trees, @{ML_type "cterm"}s are
  1401   abstract objects that are guaranteed to be type-correct, and they can only
  1399   abstract objects that are guaranteed to be type-correct, and they can only
  1402   be constructed via ``official interfaces''.
  1400   be constructed via ``official interfaces''.
  1403 
  1401 
  1404 
  1402 
  1448   \end{readmore}
  1446   \end{readmore}
  1449 
  1447 
  1450   Remember Isabelle follows the Church-style typing for terms, i.e., a term contains 
  1448   Remember Isabelle follows the Church-style typing for terms, i.e., a term contains 
  1451   enough typing information (constants, free variables and abstractions all have typing 
  1449   enough typing information (constants, free variables and abstractions all have typing 
  1452   information) so that it is always clear what the type of a term is. 
  1450   information) so that it is always clear what the type of a term is. 
  1453   Given a well-typed term, the function @{ML_ind [index] type_of} returns the 
  1451   Given a well-typed term, the function @{ML_ind  type_of} returns the 
  1454   type of a term. Consider for example:
  1452   type of a term. Consider for example:
  1455 
  1453 
  1456   @{ML_response [display,gray] 
  1454   @{ML_response [display,gray] 
  1457   "type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"}
  1455   "type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"}
  1458 
  1456 
  1463   @{ML_response_fake [display,gray] 
  1461   @{ML_response_fake [display,gray] 
  1464   "type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::int\"})" 
  1462   "type_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::int\"})" 
  1465   "*** Exception- TYPE (\"type_of: type mismatch in application\" \<dots>"}
  1463   "*** Exception- TYPE (\"type_of: type mismatch in application\" \<dots>"}
  1466 
  1464 
  1467   Since the complete traversal might sometimes be too costly and
  1465   Since the complete traversal might sometimes be too costly and
  1468   not necessary, there is the function @{ML_ind [index] fastype_of}, which 
  1466   not necessary, there is the function @{ML_ind  fastype_of}, which 
  1469   also returns the type of a term.
  1467   also returns the type of a term.
  1470 
  1468 
  1471   @{ML_response [display,gray] 
  1469   @{ML_response [display,gray] 
  1472   "fastype_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"}
  1470   "fastype_of (@{term \"f::nat \<Rightarrow> bool\"} $ @{term \"x::nat\"})" "bool"}
  1473 
  1471 
  1480   where no error is detected.
  1478   where no error is detected.
  1481 
  1479 
  1482   Sometimes it is a bit inconvenient to construct a term with 
  1480   Sometimes it is a bit inconvenient to construct a term with 
  1483   complete typing annotations, especially in cases where the typing 
  1481   complete typing annotations, especially in cases where the typing 
  1484   information is redundant. A short-cut is to use the ``place-holder'' 
  1482   information is redundant. A short-cut is to use the ``place-holder'' 
  1485   type @{ML_ind [index] dummyT} and then let type-inference figure out the 
  1483   type @{ML_ind  dummyT} and then let type-inference figure out the 
  1486   complete type. An example is as follows:
  1484   complete type. An example is as follows:
  1487 
  1485 
  1488   @{ML_response_fake [display,gray] 
  1486   @{ML_response_fake [display,gray] 
  1489 "let
  1487 "let
  1490   val c = Const (@{const_name \"plus\"}, dummyT) 
  1488   val c = Const (@{const_name \"plus\"}, dummyT) 
  1577   For the functions @{text "assume"}, @{text "forall_elim"} etc 
  1575   For the functions @{text "assume"}, @{text "forall_elim"} etc 
  1578   see \isccite{sec:thms}. The basic functions for theorems are defined in
  1576   see \isccite{sec:thms}. The basic functions for theorems are defined in
  1579   @{ML_file "Pure/thm.ML"}. 
  1577   @{ML_file "Pure/thm.ML"}. 
  1580   \end{readmore}
  1578   \end{readmore}
  1581 
  1579 
  1582   (FIXME: handy functions working on theorems, like @{ML_ind [index] rulify in ObjectLogic} and so on) 
  1580   (FIXME: handy functions working on theorems, like @{ML_ind  rulify in ObjectLogic} and so on) 
  1583 
  1581 
  1584   (FIXME: how to add case-names to goal states - maybe in the 
  1582   (FIXME: how to add case-names to goal states - maybe in the 
  1585   next section)
  1583   next section)
  1586 
  1584 
  1587   (FIXME: example for how to add theorem styles)
  1585   (FIXME: example for how to add theorem styles)
  1621   @{ML_type "theory -> theory"}: the input theory is the current theory and the 
  1619   @{ML_type "theory -> theory"}: the input theory is the current theory and the 
  1622   output the theory where the theory attribute has been stored.
  1620   output the theory where the theory attribute has been stored.
  1623   
  1621   
  1624   This is a fundamental principle in Isabelle. A similar situation occurs 
  1622   This is a fundamental principle in Isabelle. A similar situation occurs 
  1625   for example with declaring constants. The function that declares a 
  1623   for example with declaring constants. The function that declares a 
  1626   constant on the ML-level is @{ML_ind [index] add_consts_i in Sign}. 
  1624   constant on the ML-level is @{ML_ind  add_consts_i in Sign}. 
  1627   If you write\footnote{Recall that ML-code  needs to be 
  1625   If you write\footnote{Recall that ML-code  needs to be 
  1628   enclosed in \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"}.} 
  1626   enclosed in \isacommand{ML}~@{text "\<verbopen> \<dots> \<verbclose>"}.} 
  1629 *}  
  1627 *}  
  1630 
  1628 
  1631 ML{*Sign.add_consts_i [(@{binding "BAR"}, @{typ "nat"}, NoSyn)] @{theory} *}
  1629 ML{*Sign.add_consts_i [(@{binding "BAR"}, @{typ "nat"}, NoSyn)] @{theory} *}
  1699 *}
  1697 *}
  1700 
  1698 
  1701 ML{*val my_symmetric = Thm.rule_attribute (fn _ => fn thm => thm RS @{thm sym})*}
  1699 ML{*val my_symmetric = Thm.rule_attribute (fn _ => fn thm => thm RS @{thm sym})*}
  1702 
  1700 
  1703 text {* 
  1701 text {* 
  1704   where the function @{ML_ind [index] rule_attribute in Thm} expects a function taking a 
  1702   where the function @{ML_ind  rule_attribute in Thm} expects a function taking a 
  1705   context (which we ignore in the code above) and a theorem (@{text thm}), and 
  1703   context (which we ignore in the code above) and a theorem (@{text thm}), and 
  1706   returns another theorem (namely @{text thm} resolved with the theorem 
  1704   returns another theorem (namely @{text thm} resolved with the theorem 
  1707   @{thm [source] sym}: @{thm sym[no_vars]}).\footnote{The function @{ML_ind [index] "RS"} 
  1705   @{thm [source] sym}: @{thm sym[no_vars]}; the function @{ML_ind "RS"} 
  1708   is explained in Section~\ref{sec:simpletacs}.} The function 
  1706   is explained in Section~\ref{sec:simpletacs}). The function 
  1709   @{ML rule_attribute in Thm} then returns  an attribute.
  1707   @{ML rule_attribute in Thm} then returns  an attribute.
  1710 
  1708 
  1711   Before we can use the attribute, we need to set it up. This can be done
  1709   Before we can use the attribute, we need to set it up. This can be done
  1712   using the Isabelle command \isacommand{attribute\_setup} as follows:
  1710   using the Isabelle command \isacommand{attribute\_setup} as follows:
  1713 *}
  1711 *}
  1739   \begin{isabelle}
  1737   \begin{isabelle}
  1740   \isacommand{thm}~@{text "test[my_sym]"}\\
  1738   \isacommand{thm}~@{text "test[my_sym]"}\\
  1741   @{text "> "}~@{thm test[my_sym]}
  1739   @{text "> "}~@{thm test[my_sym]}
  1742   \end{isabelle}
  1740   \end{isabelle}
  1743 
  1741 
  1744   An alternative for setting up an attribute is the function @{ML_ind [index] setup in Attrib}.
  1742   An alternative for setting up an attribute is the function @{ML_ind  setup in Attrib}.
  1745   So instead of using \isacommand{attribute\_setup}, you can also set up the
  1743   So instead of using \isacommand{attribute\_setup}, you can also set up the
  1746   attribute as follows:
  1744   attribute as follows:
  1747 *}
  1745 *}
  1748 
  1746 
  1749 ML{*Attrib.setup @{binding "my_sym"} (Scan.succeed my_symmetric)
  1747 ML{*Attrib.setup @{binding "my_sym"} (Scan.succeed my_symmetric)
  1804   \end{isabelle}
  1802   \end{isabelle}
  1805 
  1803 
  1806   you get an exception indicating that the theorem @{thm [source] sym}
  1804   you get an exception indicating that the theorem @{thm [source] sym}
  1807   does not resolve with meta-equations. 
  1805   does not resolve with meta-equations. 
  1808 
  1806 
  1809   The purpose of @{ML_ind [index] rule_attribute in Thm} is to directly manipulate theorems.
  1807   The purpose of @{ML_ind  rule_attribute in Thm} is to directly manipulate theorems.
  1810   Another usage of theorem attributes is to add and delete theorems from stored data.
  1808   Another usage of theorem attributes is to add and delete theorems from stored data.
  1811   For example the theorem attribute @{text "[simp]"} adds or deletes a theorem from the
  1809   For example the theorem attribute @{text "[simp]"} adds or deletes a theorem from the
  1812   current simpset. For these applications, you can use @{ML_ind [index] declaration_attribute in Thm}. 
  1810   current simpset. For these applications, you can use @{ML_ind  declaration_attribute in Thm}. 
  1813   To illustrate this function, let us introduce a reference containing a list
  1811   To illustrate this function, let us introduce a reference containing a list
  1814   of theorems.
  1812   of theorems.
  1815 *}
  1813 *}
  1816 
  1814 
  1817 ML{*val my_thms = ref ([] : thm list)*}
  1815 ML{*val my_thms = ref ([] : thm list)*}
  1836 
  1834 
  1837 text {*
  1835 text {*
  1838   These functions take a theorem and a context and, for what we are explaining
  1836   These functions take a theorem and a context and, for what we are explaining
  1839   here it is sufficient that they just return the context unchanged. They change
  1837   here it is sufficient that they just return the context unchanged. They change
  1840   however the reference @{ML my_thms}, whereby the function 
  1838   however the reference @{ML my_thms}, whereby the function 
  1841   @{ML_ind [index] add_thm in Thm} adds a theorem if it is not already included in 
  1839   @{ML_ind  add_thm in Thm} adds a theorem if it is not already included in 
  1842   the list, and @{ML_ind [index] del_thm in Thm} deletes one (both functions use the 
  1840   the list, and @{ML_ind  del_thm in Thm} deletes one (both functions use the 
  1843   predicate @{ML_ind [index] eq_thm_prop in Thm}, which compares theorems according to 
  1841   predicate @{ML_ind  eq_thm_prop in Thm}, which compares theorems according to 
  1844   their proved propositions modulo alpha-equivalence).
  1842   their proved propositions modulo alpha-equivalence).
  1845 
  1843 
  1846   You can turn functions @{ML my_thm_add} and @{ML my_thm_del} into 
  1844   You can turn functions @{ML my_thm_add} and @{ML my_thm_del} into 
  1847   attributes with the code
  1845   attributes with the code
  1848 *}
  1846 *}
  1856 
  1854 
  1857 attribute_setup %gray my_thms = {* Attrib.add_del my_add my_del *} 
  1855 attribute_setup %gray my_thms = {* Attrib.add_del my_add my_del *} 
  1858   "maintaining a list of my_thms - rough test only!" 
  1856   "maintaining a list of my_thms - rough test only!" 
  1859 
  1857 
  1860 text {*
  1858 text {*
  1861   The parser @{ML_ind [index] add_del in Attrib} is a predefined parser for 
  1859   The parser @{ML_ind  add_del in Attrib} is a predefined parser for 
  1862   adding and deleting lemmas. Now if you prove the next lemma 
  1860   adding and deleting lemmas. Now if you prove the next lemma 
  1863   and attach to it the attribute @{text "[my_thms]"}
  1861   and attach to it the attribute @{text "[my_thms]"}
  1864 *}
  1862 *}
  1865 
  1863 
  1866 lemma trueI_2[my_thms]: "True" by simp
  1864 lemma trueI_2[my_thms]: "True" by simp
  1884   @{ML_response_fake [display,gray]
  1882   @{ML_response_fake [display,gray]
  1885   "!my_thms"
  1883   "!my_thms"
  1886   "[\"True\", \"Suc (Suc 0) = 2\"]"}
  1884   "[\"True\", \"Suc (Suc 0) = 2\"]"}
  1887 
  1885 
  1888   The theorem @{thm [source] trueI_2} only appears once, since the 
  1886   The theorem @{thm [source] trueI_2} only appears once, since the 
  1889   function @{ML_ind [index] add_thm in Thm} tests for duplicates, before extending
  1887   function @{ML_ind  add_thm in Thm} tests for duplicates, before extending
  1890   the list. Deletion from the list works as follows:
  1888   the list. Deletion from the list works as follows:
  1891 *}
  1889 *}
  1892 
  1890 
  1893 declare test[my_thms del]
  1891 declare test[my_thms del]
  1894 
  1892 
  1896 
  1894 
  1897   @{ML_response_fake [display,gray]
  1895   @{ML_response_fake [display,gray]
  1898   "!my_thms"
  1896   "!my_thms"
  1899   "[\"True\"]"}
  1897   "[\"True\"]"}
  1900 
  1898 
  1901   We used in this example two functions declared as @{ML_ind [index] declaration_attribute in Thm}, 
  1899   We used in this example two functions declared as @{ML_ind  declaration_attribute in Thm}, 
  1902   but there can be any number of them. We just have to change the parser for reading
  1900   but there can be any number of them. We just have to change the parser for reading
  1903   the arguments accordingly. 
  1901   the arguments accordingly. 
  1904 
  1902 
  1905   However, as said at the beginning of this example, using references for storing theorems is
  1903   However, as said at the beginning of this example, using references for storing theorems is
  1906   \emph{not} the received way of doing such things. The received way is to 
  1904   \emph{not} the received way of doing such things. The received way is to 
  2125 ML {* LocalTheory.set_group *}
  2123 ML {* LocalTheory.set_group *}
  2126 *)
  2124 *)
  2127 
  2125 
  2128 section {* Storing Theorems\label{sec:storing} (TBD) *}
  2126 section {* Storing Theorems\label{sec:storing} (TBD) *}
  2129 
  2127 
  2130 text {* @{ML_ind [index] add_thms_dynamic in PureThy} *}
  2128 text {* @{ML_ind  add_thms_dynamic in PureThy} *}
  2131 
  2129 
  2132 local_setup %gray {* 
  2130 local_setup %gray {* 
  2133   LocalTheory.note Thm.theoremK 
  2131   LocalTheory.note Thm.theoremK 
  2134     ((@{binding "allI_alt"}, []), [@{thm allI}]) #> snd *}
  2132     ((@{binding "allI_alt"}, []), [@{thm allI}]) #> snd *}
  2135 
  2133 
  2171   sophisticated, Isabelle includes an infrastructure for properly formatting text.
  2169   sophisticated, Isabelle includes an infrastructure for properly formatting text.
  2172   This infrastructure is loosely based on a paper by Oppen~\cite{Oppen80}. Most of
  2170   This infrastructure is loosely based on a paper by Oppen~\cite{Oppen80}. Most of
  2173   its functions do not operate on @{ML_type string}s, but on instances of the
  2171   its functions do not operate on @{ML_type string}s, but on instances of the
  2174   type:
  2172   type:
  2175 
  2173 
  2176   @{ML_type [display, gray, index] "Pretty.T"}
  2174   @{ML_type_ind [display, gray] "Pretty.T"}
  2177 
  2175 
  2178   The function @{ML str in Pretty} transforms a (plain) string into such a pretty 
  2176   The function @{ML str in Pretty} transforms a (plain) string into such a pretty 
  2179   type. For example
  2177   type. For example
  2180 
  2178 
  2181   @{ML_response_fake [display,gray]
  2179   @{ML_response_fake [display,gray]
  2184   where the result indicates that we transformed a string with length 4. Once
  2182   where the result indicates that we transformed a string with length 4. Once
  2185   you have a pretty type, you can, for example, control where linebreaks may
  2183   you have a pretty type, you can, for example, control where linebreaks may
  2186   occur in case the text wraps over a line, or with how much indentation a
  2184   occur in case the text wraps over a line, or with how much indentation a
  2187   text should be printed.  However, if you want to actually output the
  2185   text should be printed.  However, if you want to actually output the
  2188   formatted text, you have to transform the pretty type back into a @{ML_type
  2186   formatted text, you have to transform the pretty type back into a @{ML_type
  2189   string}. This can be done with the function @{ML_ind [index] string_of in Pretty}. In what
  2187   string}. This can be done with the function @{ML_ind  string_of in Pretty}. In what
  2190   follows we will use the following wrapper function for printing a pretty
  2188   follows we will use the following wrapper function for printing a pretty
  2191   type:
  2189   type:
  2192 *}
  2190 *}
  2193 
  2191 
  2194 ML{*fun pprint prt = tracing (Pretty.string_of prt)*}
  2192 ML{*fun pprint prt = tracing (Pretty.string_of prt)*}
  2229 aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo
  2227 aaaaaaar fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar fo
  2230 oooooooooooooobaaaaaaaaaaaar"}
  2228 oooooooooooooobaaaaaaaaaaaar"}
  2231 
  2229 
  2232   However by using pretty types you have the ability to indicate a possible
  2230   However by using pretty types you have the ability to indicate a possible
  2233   line break for example at each space. You can achieve this with the function
  2231   line break for example at each space. You can achieve this with the function
  2234   @{ML_ind [index] breaks in Pretty}, which expects a list of pretty types and inserts a
  2232   @{ML_ind  breaks in Pretty}, which expects a list of pretty types and inserts a
  2235   possible line break in between every two elements in this list. To print
  2233   possible line break in between every two elements in this list. To print
  2236   this list of pretty types as a single string, we concatenate them 
  2234   this list of pretty types as a single string, we concatenate them 
  2237   with the function @{ML_ind [index] blk in Pretty} as follows:
  2235   with the function @{ML_ind  blk in Pretty} as follows:
  2238 
  2236 
  2239 
  2237 
  2240 @{ML_response_fake [display,gray]
  2238 @{ML_response_fake [display,gray]
  2241 "let
  2239 "let
  2242   val ptrs = map Pretty.str (space_explode \" \" test_str)
  2240   val ptrs = map Pretty.str (space_explode \" \" test_str)
  2247 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar 
  2245 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar 
  2248 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
  2246 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
  2249 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
  2247 fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
  2250 
  2248 
  2251   Here the layout of @{ML test_str} is much more pleasing to the 
  2249   Here the layout of @{ML test_str} is much more pleasing to the 
  2252   eye. The @{ML "0"} in @{ML_ind [index] blk in Pretty} stands for no indentation
  2250   eye. The @{ML "0"} in @{ML_ind  blk in Pretty} stands for no indentation
  2253   of the printed string. You can increase the indentation and obtain
  2251   of the printed string. You can increase the indentation and obtain
  2254   
  2252   
  2255 @{ML_response_fake [display,gray]
  2253 @{ML_response_fake [display,gray]
  2256 "let
  2254 "let
  2257   val ptrs = map Pretty.str (space_explode \" \" test_str)
  2255   val ptrs = map Pretty.str (space_explode \" \" test_str)
  2263    fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
  2261    fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
  2264    fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
  2262    fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
  2265 
  2263 
  2266   where starting from the second line the indent is 3. If you want
  2264   where starting from the second line the indent is 3. If you want
  2267   that every line starts with the same indent, you can use the
  2265   that every line starts with the same indent, you can use the
  2268   function @{ML_ind [index] indent in Pretty} as follows:
  2266   function @{ML_ind  indent in Pretty} as follows:
  2269 
  2267 
  2270 @{ML_response_fake [display,gray]
  2268 @{ML_response_fake [display,gray]
  2271 "let
  2269 "let
  2272   val ptrs = map Pretty.str (space_explode \" \" test_str)
  2270   val ptrs = map Pretty.str (space_explode \" \" test_str)
  2273 in
  2271 in
  2278           fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
  2276           fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar
  2279           fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
  2277           fooooooooooooooobaaaaaaaaaaaar fooooooooooooooobaaaaaaaaaaaar"}
  2280 
  2278 
  2281   If you want to print out a list of items separated by commas and 
  2279   If you want to print out a list of items separated by commas and 
  2282   have the linebreaks handled properly, you can use the function 
  2280   have the linebreaks handled properly, you can use the function 
  2283   @{ML_ind [index] commas in Pretty}. For example
  2281   @{ML_ind  commas in Pretty}. For example
  2284 
  2282 
  2285 @{ML_response_fake [display,gray]
  2283 @{ML_response_fake [display,gray]
  2286 "let
  2284 "let
  2287   val ptrs = map (Pretty.str o string_of_int) (99998 upto 100020)
  2285   val ptrs = map (Pretty.str o string_of_int) (99998 upto 100020)
  2288 in
  2286 in
  2293 100016, 100017, 100018, 100019, 100020"}
  2291 100016, 100017, 100018, 100019, 100020"}
  2294 
  2292 
  2295   where @{ML upto} generates a list of integers. You can print out this
  2293   where @{ML upto} generates a list of integers. You can print out this
  2296   list as a ``set'', that means enclosed inside @{text [quotes] "{"} and
  2294   list as a ``set'', that means enclosed inside @{text [quotes] "{"} and
  2297   @{text [quotes] "}"}, and separated by commas using the function
  2295   @{text [quotes] "}"}, and separated by commas using the function
  2298   @{ML_ind [index] enum in Pretty}. For example
  2296   @{ML_ind  enum in Pretty}. For example
  2299 *}
  2297 *}
  2300 
  2298 
  2301 text {*
  2299 text {*
  2302   
  2300   
  2303 @{ML_response_fake [display,gray]
  2301 @{ML_response_fake [display,gray]
  2364 end*}
  2362 end*}
  2365 
  2363 
  2366 text {*
  2364 text {*
  2367   In Line 3 we define a function that inserts possible linebreaks in places
  2365   In Line 3 we define a function that inserts possible linebreaks in places
  2368   where a space is. In Lines 4 and 5 we pretty-print the term and its type
  2366   where a space is. In Lines 4 and 5 we pretty-print the term and its type
  2369   using the functions @{ML_ind [index] pretty_term in Syntax} and @{ML_ind [index]
  2367   using the functions @{ML_ind  pretty_term in Syntax} and @{ML_ind 
  2370   pretty_typ in Syntax}. We also use the function @{ML_ind [index] quote in
  2368   pretty_typ in Syntax}. We also use the function @{ML_ind quote in
  2371   Pretty} in order to enclose the term and type inside quotation marks. In
  2369   Pretty} in order to enclose the term and type inside quotation marks. In
  2372   Line 9 we add a period right after the type without the possibility of a
  2370   Line 9 we add a period right after the type without the possibility of a
  2373   line break, because we do not want that a line break occurs there.
  2371   line break, because we do not want that a line break occurs there.
  2374 
  2372 
  2375 
  2373