CookBook/Recipes/Antiquotes.thy
changeset 175 7c09bd3227c5
parent 171 18f90044c777
child 188 8939b8fd8603
equal deleted inserted replaced
174:a29b81d4fa88 175:7c09bd3227c5
    18   obtain a list of all currently available document antiquotations and their options.
    18   obtain a list of all currently available document antiquotations and their options.
    19   You obtain the same list on the ML-level by typing
    19   You obtain the same list on the ML-level by typing
    20 
    20 
    21   @{ML [display,gray] "ThyOutput.print_antiquotations ()"}
    21   @{ML [display,gray] "ThyOutput.print_antiquotations ()"}
    22 
    22 
    23   Below we give the code for two additional antiquotations that can be used to typeset
    23   Below we give the code for two additional document antiquotations that can
    24   ML-code and also to check whether the given code actually compiles. This
    24   be used to typeset ML-code and also to check whether the given code actually
    25   provides a sanity check for the code and also allows one to keep documents
    25   compiles. This provides a sanity check for the code and also allows one to
    26   in sync with other code, for example Isabelle.
    26   keep documents in sync with other code, for example Isabelle.
    27 
    27 
    28   We first describe the antiquotation @{text "ML_checked"} with the syntax:
    28   We first describe the antiquotation @{text "ML_checked"} with the syntax:
    29  
    29  
    30   @{text [display] "@{ML_checked \"a_piece_of_code\"}"}
    30   @{text [display] "@{ML_checked \"a_piece_of_code\"}"}
    31 
    31 
    32   The code is checked by sending the ML-expression @{text [quotes] "val _ =
    32   The code is checked by sending the ML-expression @{text [quotes] "val _ =
    33   a_piece_of_code"} to the ML-compiler (i.e.~the function @{ML
    33   a_piece_of_code"} to the ML-compiler (i.e.~the function @{ML
    34   "ML_Context.eval_in"} in Line 4 below). The complete code of the
    34   "ML_Context.eval_in"} in Line 4 below). The complete code of the
    35   antiquotation is as follows:
    35   document antiquotation is as follows:
    36 
    36 
    37 *}
    37 *}
    38 
    38 
    39 ML%linenosgray{*fun ml_val code_txt = "val _ = " ^ code_txt
    39 ML%linenosgray{*fun ml_val code_txt = "val _ = " ^ code_txt
    40 
    40 
    43    ThyOutput.output (map Pretty.str (space_explode "\n" code_txt)))
    43    ThyOutput.output (map Pretty.str (space_explode "\n" code_txt)))
    44 
    44 
    45 val _ = ThyOutput.antiquotation "ML_checked" (Scan.lift Args.name) output_ml*}
    45 val _ = ThyOutput.antiquotation "ML_checked" (Scan.lift Args.name) output_ml*}
    46 
    46 
    47 text {*
    47 text {*
       
    48   The parser @{ML "(Scan.lift Args.name)"} in line 9 parses a string, in this
       
    49   case the code. As mentioned before, the code is sent to the ML-compiler in
       
    50   the line 4 using the function @{ML ml_val}, which constructs the appropriate
       
    51   ML-expression.  If the code is ``approved'' by the compiler, then the output
       
    52   function @{ML "ThyOutput.output"} in the next line pretty prints the
       
    53   code. This function expects that the code is a list of (pretty)strings where
       
    54   each string correspond to a line in the output. Therefore the use of @{ML
       
    55   "(space_explode \"\\n\" txt)" for txt} which produces this list according to
       
    56   linebreaks.  There are a number of options for antiquotations that are
       
    57   observed by @{ML ThyOutput.output} when printing the code (including @{text
       
    58   "[display]"} and @{text "[quotes]"}). Line 7 sets up the new document
       
    59   antiquotation.
    48 
    60 
    49   The parser @{ML "(Scan.lift Args.name)"} in line 9 parses a string, 
       
    50   in this case the code. As mentioned before, the code
       
    51   is sent to the ML-compiler in the line 4 using the function @{ML ml_val},
       
    52   which constructs the appropriate ML-expression.
       
    53   If the code is ``approved'' by the compiler, then the output function @{ML
       
    54   "ThyOutput.output"} in the next line pretty prints the
       
    55   code. This function expects that the code is a list of (pretty)strings where each
       
    56   string correspond to a line in the output. Therefore the use of 
       
    57   @{ML "(space_explode \"\\n\" txt)" for txt} 
       
    58   which produces this list according to linebreaks.  There are a number of options 
       
    59   for antiquotations that are observed by @{ML ThyOutput.output} when printing the 
       
    60   code (including @{text "[display]"} and @{text "[quotes]"}). Line 7 sets 
       
    61   up the new antiquotation.
       
    62 
    61 
    63   \begin{readmore}
    62   \begin{readmore}
    64   For more information about options of antiquotations see \rsccite{sec:antiq}).
    63   For more information about options of document antiquotations see \rsccite{sec:antiq}).
    65   \end{readmore}
    64   \end{readmore}
    66 
    65 
    67   Since we used the argument @{ML "Position.none"}, the compiler cannot give specific 
    66   Since we used the argument @{ML "Position.none"}, the compiler cannot give specific 
    68   information about the line number, in case an error is detected. We 
    67   information about the line number, in case an error is detected. We 
    69   can improve the code above slightly by writing 
    68   can improve the code above slightly by writing 
    84   We can now write in a document @{text "@{ML_checked \"2 + 3\"}"} in order to
    83   We can now write in a document @{text "@{ML_checked \"2 + 3\"}"} in order to
    85   obtain @{ML_checked "2 + 3"} and be sure that this code compiles until
    84   obtain @{ML_checked "2 + 3"} and be sure that this code compiles until
    86   somebody changes the definition of \mbox{@{ML "(op +)"}}.
    85   somebody changes the definition of \mbox{@{ML "(op +)"}}.
    87 
    86 
    88 
    87 
    89   The second antiquotation we describe extends the first by a pattern that 
    88   The second document antiquotation we describe extends the first by a pattern
    90   specifies what the result of the ML-code should be and check 
    89   that specifies what the result of the ML-code should be and check the
    91   the consistency of the actual result with the given pattern. For this we are going 
    90   consistency of the actual result with the given pattern. For this we are
    92   to implement the antiquotation  
    91   going to implement the document antiquotation
       
    92 
    93   
    93   
    94   @{text [display] "@{ML_resp \"a_piece_of_code\" \"a_pattern\"}"}
    94   @{text [display] "@{ML_resp \"a_piece_of_code\" \"a_pattern\"}"}
    95   
    95   
    96   To add some convenience and also to deal with large outputs, the user can
    96   To add some convenience and also to deal with large outputs, the user can
    97   give a partial specification inside the pattern by giving abbreviations of
    97   give a partial specification inside the pattern by giving abbreviations of
    98   the form @{text [quotes] "\<dots>"}. For example @{text "(\<dots>, \<dots>)"} to specify a
    98   the form @{text [quotes] "\<dots>"}. For example @{text "(\<dots>, \<dots>)"} for specifying a
    99   pair.
    99   pair.
   100 
   100 
   101   Whereas in the antiquotation @{text "@{ML_checked \"piece_of_code\"}"}
   101   In the document antiquotation @{text "@{ML_checked \"piece_of_code\"}"}
   102   above, we have sent the expression @{text [quotes] "val _ = piece_of_code"}
   102   above we have sent the expression @{text [quotes] "val _ = piece_of_code"}
   103   to the compiler, in the second the wildcard @{text "_"} we will be replaced
   103   to the compiler, now instead the wildcard @{text "_"} we will be replaced by
   104   by the given pattern. To do this we need to replace the @{text [quotes] "\<dots>"}
   104   the given pattern. To do this we need to replace in the input the @{text
   105   by @{text [quotes] "_"} before sending the code to the compiler. The
   105   [quotes] "\<dots>"} by @{text [quotes] "_"} before sending the code to the
   106   following function will do this:
   106   compiler. The following function will do this:
   107 *}
   107 *}
   108 
   108 
   109 ML{*fun ml_pat (code_txt, pat) =
   109 ML{*fun ml_pat (code_txt, pat) =
   110 let val pat' = 
   110 let val pat' = 
   111          implode (map (fn "\<dots>" => "_" | s => s) (Symbol.explode pat))
   111          implode (map (fn "\<dots>" => "_" | s => s) (Symbol.explode pat))
   116 text {* 
   116 text {* 
   117   Next we like to add a response indicator to the result using:
   117   Next we like to add a response indicator to the result using:
   118 *}
   118 *}
   119 
   119 
   120 
   120 
   121 ML{*fun add_resp_indicator pat =
   121 ML{*fun add_resp pat = map (fn s => "> " ^ s) pat*}
   122   map (fn s => "> " ^ s) (space_explode "\n" pat) *}
       
   123 
   122 
   124 text {* 
   123 text {* 
   125   The rest of the code of the antiquotation is
   124   The rest of the code of the document antiquotation is
   126 *}
   125 *}
   127 
   126 
   128 ML{*fun output_ml_resp {context = ctxt, ...} ((code_txt, pat), pos) = 
   127 ML{*fun output_ml_resp {context = ctxt, ...} ((code_txt, pat), pos) = 
   129   (ML_Context.eval_in (SOME ctxt) false pos (ml_pat (code_txt, pat));
   128   (ML_Context.eval_in (SOME ctxt) false pos (ml_pat (code_txt, pat));
   130    let 
   129    let 
   131      val output = (space_explode "\n" code_txt) @ (add_resp_indicator pat)
   130      val output1 = space_explode "\n" code_txt 
       
   131      val output2 = add_resp (space_explode "\n" pat)
   132    in 
   132    in 
   133      ThyOutput.output (map Pretty.str output) 
   133      ThyOutput.output (map Pretty.str (output1 @ output2)) 
   134    end)
   134    end)
   135 
   135 
   136 val _ = ThyOutput.antiquotation "ML_resp" 
   136 val _ = ThyOutput.antiquotation "ML_resp" 
   137          (Scan.lift (OuterParse.position (Args.name -- Args.name))) 
   137          (Scan.lift (OuterParse.position (Args.name -- Args.name))) 
   138            output_ml_resp*}
   138            output_ml_resp*}
   139 
   139 
   140 text {*
   140 text {*
   141   This extended antiquotation allows us to write
   141   This extended document antiquotation allows us to write
   142  
   142  
   143   @{text [display] "@{ML_resp [display] \"true andalso false\" \"false\"}"}
   143   @{text [display] "@{ML_resp [display] \"true andalso false\" \"false\"}"}
   144 
   144 
   145   to obtain
   145   to obtain
   146 
   146 
   153   to obtain
   153   to obtain
   154 
   154 
   155   @{ML_resp [display] "let val i = 3 in (i * i, \"foo\") end" "(9, \<dots>)"} 
   155   @{ML_resp [display] "let val i = 3 in (i * i, \"foo\") end" "(9, \<dots>)"} 
   156 
   156 
   157   In both cases, the check by the compiler ensures that code and result
   157   In both cases, the check by the compiler ensures that code and result
   158   match. A limitation of this antiquotation, however, is that the pattern can
   158   match. A limitation of this document antiquotation, however, is that the
   159   only be given for values that can be constructed. This excludes
   159   pattern can only be given for values that can be constructed. This excludes
   160   values that are abstract datatypes, like theorems or cterms.
   160   values that are abstract datatypes, like @{ML_type thm}s and @{ML_type cterm}s.
       
   161 
   161 *}
   162 *}
   162 end
   163 end