ProgTutorial/Recipes/Antiquotes.thy
changeset 575 c3dbc04471a9
parent 574 034150db9d91
equal deleted inserted replaced
574:034150db9d91 575:c3dbc04471a9
    32   We first describe the antiquotation \<open>ML_checked\<close> with the syntax:
    32   We first describe the antiquotation \<open>ML_checked\<close> with the syntax:
    33  
    33  
    34   @{text [display] \<open>@{ML_checked "a_piece_of_code"}\<close>}
    34   @{text [display] \<open>@{ML_checked "a_piece_of_code"}\<close>}
    35 
    35 
    36   The code is checked by sending the ML-expression @{text [quotes] "val _ =
    36   The code is checked by sending the ML-expression @{text [quotes] "val _ =
    37   a_piece_of_code"} to the ML-compiler (i.e.~the function @{ML \<open>ML_Context.eval_source_in\<close>} in Line 7 below). The complete code of the
    37   a_piece_of_code"} to the ML-compiler (i.e.~the function @{ML \<open>ML_Context.eval_in\<close>} in Line 8 below). The complete code of the
    38   document antiquotation is as follows:
    38   document antiquotation is as follows:
    39 
    39 
    40 \<close>
    40 \<close>
    41 ML \<open>Input.pos_of\<close>
    41 ML \<open>Input.pos_of\<close>
    42 ML%linenosgray\<open>fun ml_enclose bg en source =
    42 ML%linenosgray\<open>fun ml_enclose bg en source =
    57 
    57 
    58 setup \<open>ml_checked_setup\<close>
    58 setup \<open>ml_checked_setup\<close>
    59 
    59 
    60 
    60 
    61 text \<open>
    61 text \<open>
    62   The parser @{ML \<open>(Scan.lift Args.name)\<close>} in Line 7 parses a string, in this
    62   The parser @{ML \<open>(Scan.lift Args.text_input)\<close>} in Line 15 parses a string, in this
    63   case the code, and then calls the function @{ML output_ml}. As mentioned
    63   case the code, and then we call the function @{ML output_ml}. As mentioned
    64   before, the parsed code is sent to the ML-compiler in Line 4 using the
    64   before, the parsed code is sent to the ML-compiler in Line 8 using the
    65   function @{ML ml_val}, which constructs the appropriate ML-expression, and
    65   function @{ML ml_val}, which constructs the appropriate ML-expression, and
    66   using @{ML \<open>eval_in\<close> in ML_Context}, which calls the compiler.  If the code is
    66   using @{ML \<open>eval_in\<close> in ML_Context}, which calls the compiler.  If the code is
    67   ``approved'' by the compiler, then the output function @{ML \<open>output\<close> in
    67   ``approved'' by the compiler, then the output given to @{ML \<open>antiquotation_pretty_source\<close> in
    68   Document_Antiquotation} in the next line pretty prints the code. This function expects
    68   Thy_Output} in the Line 15 pretty prints the code. This function expects
    69   that the code is a list of (pretty)strings where each string correspond to a
    69   that the code is (pretty) string. There are a
    70   line in the output. Therefore the use of @{ML \<open>(space_explode "\\n" txt)\<close>
       
    71   for txt} which produces such a list according to linebreaks.  There are a
       
    72   number of options for antiquotations that are observed by the function 
    70   number of options for antiquotations that are observed by the function 
    73   @{ML \<open>output\<close> in Document_Antiquotation} when printing the code (including \<open>[display]\<close> 
    71   @{ML \<open>output\<close> in Document_Antiquotation} when printing the code (including \<open>[display]\<close> 
    74   and \<open>[quotes]\<close>). The function @{ML \<open>antiquotation_raw\<close> in Thy_Output} in 
    72   and \<open>[quotes]\<close>). 
    75   Line 7 sets up the new document antiquotation.
       
    76 
    73 
    77   \begin{readmore}
    74   \begin{readmore}
    78   For more information about options of document antiquotations see \rsccite{sec:antiq}).
    75   For more information about options of document antiquotations see \rsccite{sec:antiq}).
    79   \end{readmore}
    76   \end{readmore}
    80 
    77 
    81   Since we used the argument @{ML \<open>Position.none\<close>}, the compiler cannot give specific 
       
    82   information about the line number, in case an error is detected. We 
       
    83   can improve the code above slightly by writing 
       
    84 \<close>
    78 \<close>
    85 
    79 
    86 text \<open>
    80 text \<open>
    87   where in Lines 1 and 2 the positional information is properly treated. The
       
    88   parser @{ML Parse.position} encodes the positional information in the 
       
    89   result.
       
    90 
       
    91   We can now write \<open>@{ML_checked2 "2 + 3"}\<close> in a document in order to
       
    92   obtain @{ML_checked "2 + 3"} and be sure that this code compiles until
       
    93   somebody changes the definition of addition.
       
    94 
       
    95 
       
    96   The second document antiquotation we describe extends the first by a pattern
    81   The second document antiquotation we describe extends the first by a pattern
    97   that specifies what the result of the ML-code should be and checks the
    82   that specifies what the result of the ML-code should be and checks the
    98   consistency of the actual result with the given pattern. For this we are
    83   consistency of the actual result with the given pattern. For this we are
    99   going to implement the document antiquotation:
    84   going to implement the document antiquotation:
   100 
    85 
   103   
    88   
   104   To add some convenience and also to deal with large outputs, the user can
    89   To add some convenience and also to deal with large outputs, the user can
   105   give a partial specification by using ellipses. For example \<open>(\<dots>, \<dots>)\<close>
    90   give a partial specification by using ellipses. For example \<open>(\<dots>, \<dots>)\<close>
   106   for specifying a pair.  In order to check consistency between the pattern
    91   for specifying a pair.  In order to check consistency between the pattern
   107   and the output of the code, we have to change the ML-expression that is sent 
    92   and the output of the code, we have to change the ML-expression that is sent 
   108   to the compiler: in \<open>ML_checked2\<close> we sent the expression @{text [quotes]
    93   to the compiler: 
   109   "val _ = a_piece_of_code"} to the compiler; now the wildcard \<open>_\<close>
       
   110   must be be replaced by the given pattern. However, we have to remove all
       
   111   ellipses from it and replace them by @{text [quotes] "_"}. The following 
       
   112   function will do this:
       
   113 \<close>
    94 \<close>
   114 
    95 
   115 ML%linenosgray\<open>fun ml_pat pat code =
    96 ML%linenosgray\<open>fun ml_pat pat code =
   116   ML_Lex.read "val" @ 
    97   ML_Lex.read "val" @ 
   117   ML_Lex.read_source pat @ 
    98   ML_Lex.read_source pat @ 
   119   ML_Lex.read_source code\<close>
   100   ML_Lex.read_source code\<close>
   120 
   101 
   121 text \<open>
   102 text \<open>
   122   Next we add a response indicator to the result using:
   103   Next we add a response indicator to the result using:
   123 \<close>
   104 \<close>
   124 
       
   125 
   105 
   126 ML %grayML\<open>fun add_resp pat = map (fn s => "> " ^ s) pat\<close>
   106 ML %grayML\<open>fun add_resp pat = map (fn s => "> " ^ s) pat\<close>
   127 
   107 
   128 text \<open>
   108 text \<open>
   129   The rest of the code of \<open>ML_resp\<close> is: 
   109   The rest of the code of \<open>ML_resp\<close> is: 
   149 
   129 
   150 setup \<open>ml_response_setup\<close>
   130 setup \<open>ml_response_setup\<close>
   151 
   131 
   152 (* FIXME *)
   132 (* FIXME *)
   153 text \<open>
   133 text \<open>
   154   In comparison with \<open>ML_checked\<close>, we only changed the line about 
   134   In comparison with \<open>ML_checked\<close>, we changed the line about 
   155   the compiler (Line~2), the lines about
   135   the compiler (Lines 4 to 5), the lines about
   156   the output (Lines 4 to 7) and the parser in the setup (Line 11). Now 
   136   the output (Lines 6 to 7 and 9) and the parser setup (Line 14). Now 
   157   you can write
   137   you can write
   158  
   138  
   159   @{text [display] \<open>@{ML_resp [display] "true andalso false" "false"}\<close>}
   139   @{text [display] \<open>@{ML_resp [display] "true andalso false" "false"}\<close>}
   160 
   140 
   161   to obtain
   141   to obtain