CookBook/Recipes/Antiquotes.thy
changeset 74 f6f8f8ba1eb1
parent 72 7b8c4fe235aa
child 114 13fd0a83d3c3
equal deleted inserted replaced
73:bcbcf5c839ae 74:f6f8f8ba1eb1
    13   {\bf Solution:} This can be achieved using document antiquotations.\smallskip
    13   {\bf Solution:} This can be achieved using document antiquotations.\smallskip
    14 
    14 
    15   Document antiquotations can be used for ensuring consistent type-setting of
    15   Document antiquotations can be used for ensuring consistent type-setting of
    16   various entities in a document. They can also be used for sophisticated
    16   various entities in a document. They can also be used for sophisticated
    17   \LaTeX-hacking. If you type @{text "Ctrl-c Ctrl-a h A"} inside ProofGeneral, you
    17   \LaTeX-hacking. If you type @{text "Ctrl-c Ctrl-a h A"} inside ProofGeneral, you
    18   obtain a list of all available document antiquotations and their options.
    18   obtain a list of all currently available document antiquotations and their options.
    19 
    19 
    20   Below we give the code for two additional antiquotations that can be used to typeset
    20   Below we give the code for two additional antiquotations that can be used to typeset
    21   ML-code and also to check whether the given code actually compiles. This
    21   ML-code and also to check whether the given code actually compiles. This
    22   provides a sanity check for the code and also allows one to keep documents
    22   provides a sanity check for the code and also allows one to keep documents
    23   in sync with other code, for example Isabelle.
    23   in sync with other code, for example Isabelle.
    24 
    24 
    25   We first describe the antiquotation @{text "@{ML_checked \"expr\"}"}. This
    25   We first describe the antiquotation @{text "ML_checked"} with the syntax:
    26   antiquotation takes a piece of code as argument. The code is checked
    26  
    27   by sending the ML-expression @{text [quotes] "val _ = \<dots>"} containing the
    27   @{text [display] "@{ML_checked \"a_piece_of_code\"}"}
    28   given argument to the ML-compiler (i.e.~the function @{ML "ML_Context.eval_in"}
    28 
    29   in Line 4 below). The complete code of the antiquotation is as follows:
    29   The code is checked by sending the ML-expression @{text [quotes] "val _ =
       
    30   a_piece_of_code"} to the ML-compiler (i.e.~the function @{ML
       
    31   "ML_Context.eval_in"} in Line 4 below). The complete code of the
       
    32   antiquotation is as follows:
    30 
    33 
    31 *}
    34 *}
    32 
    35 
    33 ML%linenumbers{*fun ml_val code_txt = "val _ = " ^ code_txt
    36 ML%linenumbers{*fun ml_val code_txt = "val _ = " ^ code_txt
    34 
    37 
    42 
    45 
    43 text {*
    46 text {*
    44 
    47 
    45   Note that the parser @{ML "(Scan.lift Args.name)"} in line 9 parses a string, 
    48   Note that the parser @{ML "(Scan.lift Args.name)"} in line 9 parses a string, 
    46   in this case the code given as argument. As mentioned before, this argument 
    49   in this case the code given as argument. As mentioned before, this argument 
    47   is send to the ML-compiler in the line 4 using the function @{ML ml_val},
    50   is sent to the ML-compiler in the line 4 using the function @{ML ml_val},
    48   which constructs an apropriate ML-expression.
    51   which constructs the appropriate ML-expression.
    49   If the code is ``approved'' by the compiler, then the output function @{ML
    52   If the code is ``approved'' by the compiler, then the output function @{ML
    50   "ThyOutput.output_list (fn _ => fn s => Pretty.str s)"} in the next line pretty prints the
    53   "ThyOutput.output_list (fn _ => fn s => Pretty.str s)"} in the next line pretty prints the
    51   code. This function expects that the code is a list of strings where each
    54   code. This function expects that the code is a list of strings where each
    52   string correspond to a line in the output. Therefore the use of 
    55   string correspond to a line in the output. Therefore the use of 
    53   @{ML "(space_explode \"\\n\" txt)" for txt} 
    56   @{ML "(space_explode \"\\n\" txt)" for txt} 
    54   which produces this list according to linebreaks.  There are a number of options for antiquotations
    57   which produces this list according to linebreaks.  There are a number of options for antiquotations
    55   that are observed by @{ML ThyOutput.output_list} when printing the code (for
    58   that are observed by @{ML ThyOutput.output_list} when printing the code (including
    56   example @{text "[display]"}, @{text "[quotes]"} and @{text "[source]"}).
    59   @{text "[display]"}, @{text "[quotes]"} and @{text "[source]"}).
    57 
    60 
    58   \begin{readmore}
    61   \begin{readmore}
    59   For more information about options of antiquotations see \rsccite{sec:antiq}).
    62   For more information about options of antiquotations see \rsccite{sec:antiq}).
    60   \end{readmore}
    63   \end{readmore}
    61 
    64 
    62   Since we used the argument @{ML "Position.none"}, the compiler cannot give specific 
    65   Since we used the argument @{ML "Position.none"}, the compiler cannot give specific 
    63   information about the line number, in case an error is detected. We 
    66   information about the line number, in case an error is detected. We 
    64   can improve the code above slightly by writing 
    67   can improve the code above slightly by writing 
    65 
       
    66 *}
    68 *}
    67 
    69 
    68 ML%linenumbers{*fun output_ml src ctxt (code_txt,pos) =
    70 ML%linenumbers{*fun output_ml src ctxt (code_txt,pos) =
    69   (ML_Context.eval_in (SOME ctxt) false pos (ml_val code_txt);
    71   (ML_Context.eval_in (SOME ctxt) false pos (ml_val code_txt);
    70   ThyOutput.output_list (fn _ => fn s => Pretty.str s) src ctxt 
    72   ThyOutput.output_list (fn _ => fn s => Pretty.str s) src ctxt 
    82   We can now write in a document @{text "@{ML_checked \"2 + 3\"}"} in order to
    84   We can now write in a document @{text "@{ML_checked \"2 + 3\"}"} in order to
    83   obtain @{ML_checked "2 + 3"} and be sure that this code compiles until
    85   obtain @{ML_checked "2 + 3"} and be sure that this code compiles until
    84   somebody changes the definition of \mbox{@{ML "(op +)"}}.
    86   somebody changes the definition of \mbox{@{ML "(op +)"}}.
    85 
    87 
    86 
    88 
    87   The second antiquotation extends the first by allowing also to give
    89   The second antiquotation we describe extends the first by allowing also to give
    88   hints what the result of the ML-code should be and to check the consistency of 
    90   a pattern that specifies what the result of the ML-code should be and to check 
    89   the actual result with these hints. For this we are going to implement the 
    91   the consistency of the actual result with the given pattern. For this we are going 
    90   antiquotation  
    92   to implement the antiquotation  
    91   @{text "@{ML_resp \"expr\" \"pat\"}"}
    93   
    92   whose first argument is the ML-code and the second is a pattern specifying
    94   @{text [display] "@{ML_resp \"a_piece_of_code\" \"pattern\"}"}
    93   the result. To add some convenience and allow dealing with large outputs,
    95   
    94   the user can give a partial specification including abbreviations 
    96   To add some convenience and also to deal with large outputs,
    95   @{text [quotes] "\<dots>"}.
    97   the user can give a partial specification by giving the abbreviation 
       
    98   @{text [quotes] "\<dots>"}. For example @{text "(\<dots>,\<dots>)"} for a pair.
    96 
    99 
    97   Whereas in the antiquotation @{text "@{ML_checked \"expr\"}"} above, we have
   100   Whereas in the antiquotation @{text "@{ML_checked \"piece_of_code\"}"} above, 
    98   sent the expression 
   101   we have sent the expression 
    99   @{text [quotes] "val _ = expr"} to the compiler, in this antiquotation the wildcard
   102   @{text [quotes] "val _ = piece_of_code"} to the compiler, in the second the 
   100   @{text "_"} we will be replaced by a proper pattern constructed the hints. To
   103   wildcard @{text "_"} we will be replaced by a proper pattern. To do this we 
   101   do this we need to replace the @{text [quotes] "\<dots>"} by @{text [quotes] "_"} 
   104   need to replace the @{text [quotes] "\<dots>"} by 
   102   before sending the code to the compiler. The following function will do this:
   105   @{text [quotes] "_"}  before sending the code to the compiler. The following 
       
   106   function will do this:
   103 
   107 
   104 *}
   108 *}
   105 
   109 
   106 ML{*fun ml_pat (code_txt, pat) =
   110 ML{*fun ml_pat (code_txt, pat) =
   107    let val pat' = 
   111    let val pat' = 
   152   to obtain
   156   to obtain
   153 
   157 
   154   @{ML_resp [display] "let val i = 3 in (i * i,\"foo\") end" "(9,\<dots>)"} 
   158   @{ML_resp [display] "let val i = 3 in (i * i,\"foo\") end" "(9,\<dots>)"} 
   155 
   159 
   156   In both cases, the check by the compiler ensures that code and result match. A limitation
   160   In both cases, the check by the compiler ensures that code and result match. A limitation
   157   of this antiquotation is that the hints can only be given for results that can actually
   161   of this antiquotation, however, is that the hints can only be given in case
   158   be constructed as a pattern. This excludes values that are abstract datatypes, like 
   162   they can be constructed as a pattern. This excludes values that are abstract datatypes, like 
   159   theorems or cterms.
   163   theorems or cterms.
   160 
   164 
   161 *}
   165 *}
   162 
   166 
   163 
   167