CookBook/Recipes/Antiquotes.thy
changeset 59 b5914f3c643c
parent 58 f3794c231898
child 65 c8e9a4f97916
equal deleted inserted replaced
58:f3794c231898 59:b5914f3c643c
    20   ML-code and also to check whether the given code actually compiles. This
    20   ML-code and also to check whether the given code actually compiles. This
    21   provides a sanity check for the code and also allows one to keep documents
    21   provides a sanity check for the code and also allows one to keep documents
    22   in sync with other code, for example Isabelle.
    22   in sync with other code, for example Isabelle.
    23 
    23 
    24   We first describe the antiquotation @{text "@{ML_checked \"expr\"}"}. This
    24   We first describe the antiquotation @{text "@{ML_checked \"expr\"}"}. This
    25   antiquotation takes a piece of code as argument. The argument is checked
    25   antiquotation takes a piece of code as argument. The code is checked
    26   by sending the ML-expression @{text [quotes] "val _ = \<dots>"} containing the
    26   by sending the ML-expression @{text [quotes] "val _ = \<dots>"} containing the
    27   given argument to the ML-compiler (i.e.~the function @{ML "ML_Context.eval_in"}
    27   given argument to the ML-compiler (i.e.~the function @{ML "ML_Context.eval_in"}
    28   in Line 4 in the code below). The code of @{text "@{ML_checked \"expr\"}"} 
    28   in Line 4 below). The complete code of the antiquotation is as follows:
    29   is as follows:
       
    30 
    29 
    31 *}
    30 *}
    32 
    31 
    33 ML %linenumbers {*fun ml_val txt = "val _ = " ^ txt
    32 ML %linenumbers {*fun ml_val code_txt = "val _ = " ^ code_txt
    34 
    33 
    35 fun output_ml ml src ctxt txt =
    34 fun output_ml src ctxt code_txt =
    36   (ML_Context.eval_in (SOME ctxt) false Position.none (ml txt); 
    35   (ML_Context.eval_in (SOME ctxt) false Position.none (ml_val code_txt); 
    37   ThyOutput.output_list (fn _ => fn s => Pretty.str s) src ctxt 
    36   ThyOutput.output_list (fn _ => fn s => Pretty.str s) src ctxt 
    38                                             (space_explode "\n" txt))
    37                                             (space_explode "\n" code_txt))
    39 
    38 
    40 val _ = ThyOutput.add_commands
    39 val _ = ThyOutput.add_commands
    41  [("ML_checked", ThyOutput.args (Scan.lift Args.name) (output_ml ml_val))]
    40  [("ML_checked", ThyOutput.args (Scan.lift Args.name) output_ml)]
    42 *}
    41 *}
    43 
    42 
    44 text {*
    43 text {*
    45 
    44 
    46   Note that the parser @{ML "(Scan.lift Args.name)"} in line 9 parses a string, 
    45   Note that the parser @{ML "(Scan.lift Args.name)"} in line 9 parses a string, 
    47   in this case the code given as argument. This argument is send to the ML-compiler in the line 4.  
    46   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},
       
    48   which constructs an apropriate ML-expression.
    48   If the code is ``approved'' by the compiler, then the output function @{ML
    49   If the code is ``approved'' by the compiler, then the output function @{ML
    49   "ThyOutput.output_list (fn _ => fn s => Pretty.str s)"} in the next line pretty prints the
    50   "ThyOutput.output_list (fn _ => fn s => Pretty.str s)"} in the next line pretty prints the
    50   code. This function expects that the code is a list of strings where each
    51   code. This function expects that the code is a list of strings where each
    51   string correspond to a line. Therefore the @{ML "(space_explode \"\\n\" txt)" for txt} 
    52   string correspond to a line in the output. Therefore the use of 
    52   which produces this list.  There are a number of options for antiquotations
    53   @{ML "(space_explode \"\\n\" txt)" for txt} 
       
    54   which produces this list according to linebreaks.  There are a number of options for antiquotations
    53   that are observed by @{ML ThyOutput.output_list} when printing the code (for
    55   that are observed by @{ML ThyOutput.output_list} when printing the code (for
    54   example @{text "[display]"}, @{text "[quotes]"} and @{text "[source]"}).
    56   example @{text "[display]"}, @{text "[quotes]"} and @{text "[source]"}).
    55 
    57 
    56   \begin{readmore}
    58   \begin{readmore}
    57   For more information about options of antiquotations see \rsccite{sec:antiq}).
    59   For more information about options of antiquotations see \rsccite{sec:antiq}).
    58   \end{readmore}
    60   \end{readmore}
    59 
    61 
    60   Since we used the argument @{ML "Position.none"}, the compiler cannot give specific 
    62   Since we used the argument @{ML "Position.none"}, the compiler cannot give specific 
    61   information about the line number where an error might have occurred. We 
    63   information about the line number, in case an error is detected. We 
    62   can improve this code slightly by writing 
    64   can improve the code above slightly by writing 
    63 
    65 
    64 *}
    66 *}
    65 
    67 
    66 ML %linenumbers {* fun output_ml ml src ctxt (txt,pos) =
    68 ML %linenumbers {* fun output_ml src ctxt (code_txt,pos) =
    67   (ML_Context.eval_in (SOME ctxt) false pos (ml txt);
    69   (ML_Context.eval_in (SOME ctxt) false pos (ml_val code_txt);
    68   ThyOutput.output_list (fn _ => fn s => Pretty.str s) src ctxt 
    70   ThyOutput.output_list (fn _ => fn s => Pretty.str s) src ctxt 
    69                                             (space_explode "\n" txt))
    71                                             (space_explode "\n" code_txt))
    70 
    72 
    71 val _ = ThyOutput.add_commands
    73 val _ = ThyOutput.add_commands
    72  [("ML_checked", ThyOutput.args 
    74  [("ML_checked", ThyOutput.args 
    73        (Scan.lift (OuterParse.position Args.name)) (output_ml ml_val))]
    75        (Scan.lift (OuterParse.position Args.name)) output_ml)]
    74 *}
    76 *}
    75 
    77 
    76 text {*
    78 text {*
    77   where in Lines 1 and 2 the positional information is properly treated.
    79   where in Lines 1 and 2 the positional information is properly treated.
    78 
    80 
    82   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
    83   somebody changes the definition of \mbox{@{ML "(op +)"}}.
    85   somebody changes the definition of \mbox{@{ML "(op +)"}}.
    84 
    86 
    85 
    87 
    86   The second antiquotation extends the first by allowing also to give
    88   The second antiquotation extends the first by allowing also to give
    87   hints what the result of the ML-code is and check the consistency of 
    89   hints what the result of the ML-code should be and to check the consistency of 
    88   the actual result with these hints. For this we use the antiquotation  
    90   the actual result with these hints. For this we are going to implement the 
       
    91   antiquotation  
    89   @{text "@{ML_response \"expr\" \"pat\"}"}
    92   @{text "@{ML_response \"expr\" \"pat\"}"}
    90   whose first argument is the ML-code and the second is a pattern specifying
    93   whose first argument is the ML-code and the second is a pattern specifying
    91   the result. To add some convenience we allow the user to give a partial
    94   the result. To add some convenience and allow dealing with large outputs,
    92   specification using @{text "\<dots>"}.
    95   the user can give a partial specification including abbreviations 
       
    96   @{text [quotes] "\<dots>"}.
    93 
    97 
    94   In the antiquotation @{text "@{ML_checked \"expr\"}"} we send the expression 
    98   Whereas in the antiquotation @{text "@{ML_checked \"expr\"}"} above, we have
    95   @{text [quotes] "val _ = expr"} to the compiler. Instead of the wildcard
    99   sent the expression 
    96   @{text "_"}, we will here use the hints to construct a proper pattern. To
   100   @{text [quotes] "val _ = expr"} to the compiler, in this antiquotation the wildcard
    97   do this we need to replace the @{text "\<dots>"} by @{text "_"} before sending the 
   101   @{text "_"} we will be replaced by a proper pattern constructed the hints. To
    98   code to the compiler. The function 
   102   do this we need to replace the @{text [quotes] "\<dots>"} by @{text [quotes] "_"} 
       
   103   before sending the code to the compiler. The following function will do this:
    99 
   104 
   100 *}
   105 *}
   101 
   106 
   102 ML {* 
   107 ML {* 
   103 fun ml_pat (rhs, pat) =
   108 fun ml_pat (code_txt, pat) =
   104 let val pat' = implode (map (fn "\<dots>" => "_" | s => s) (Symbol.explode pat))
   109    let val pat' = 
   105 in "val " ^ pat' ^ " = " ^ rhs end;
   110          implode (map (fn "\<dots>" => "_" | s => s) (Symbol.explode pat))
       
   111    in 
       
   112      "val " ^ pat' ^ " = " ^ code_txt 
       
   113    end
   106 *}
   114 *}
   107 
   115 
   108 text {* 
   116 text {* 
   109   will construct the pattern that the compiler can use. Next we like to add 
   117   Next we like to add a response indicator to the result using:
   110   a response indicator to the result using:
       
   111 *}
   118 *}
   112 
   119 
   113 
   120 
   114 ML {*
   121 ML {*
   115 fun add_response_indicator txt =
   122 fun add_response_indicator pat =
   116   map (fn s => "> " ^ s) (space_explode "\n" txt)
   123   map (fn s => "> " ^ s) (space_explode "\n" pat)
   117 *}
   124 *}
   118 
   125 
   119 text {* 
   126 text {* 
   120   The rest of the code of the antiquotation is
   127   The rest of the code of the antiquotation is
   121   *}
   128   *}
   122 
   129 
   123 ML {*
   130 ML {*
   124 fun output_ml_response ml src ctxt ((lhs,pat),pos) = 
   131 fun output_ml_response src ctxt ((code_txt,pat),pos) = 
   125   (ML_Context.eval_in (SOME ctxt) false pos (ml (lhs,pat));
   132   (ML_Context.eval_in (SOME ctxt) false pos (ml_pat (code_txt,pat));
   126   let val txt = (space_explode "\n" lhs) @ (add_response_indicator pat)
   133    let 
   127   in ThyOutput.output_list (fn _ => fn s => Pretty.str s) src ctxt txt end)
   134      val output = (space_explode "\n" code_txt) @ (add_response_indicator pat)
       
   135    in 
       
   136      ThyOutput.output_list (fn _ => fn s => Pretty.str s) src ctxt output 
       
   137    end)
   128 
   138 
   129 val _ = ThyOutput.add_commands
   139 val _ = ThyOutput.add_commands
   130  [("ML_response", 
   140  [("ML_response", 
   131      ThyOutput.args 
   141      ThyOutput.args 
   132       (Scan.lift (OuterParse.position (Args.name -- Args.name))) 
   142       (Scan.lift (OuterParse.position (Args.name -- Args.name))) 
   133       (output_ml_response ml_pat))]
   143         output_ml_response)]
   134 *}
   144 *}
   135 
   145 
   136 text {*
   146 text {*
   137   This extended antiquotation allows us to write 
   147   This extended antiquotation allows us to write 
   138   @{text [display] "@{ML_response [display] \"true andalso false\" \"false\"}"}
   148   @{text [display] "@{ML_response [display] \"true andalso false\" \"false\"}"}