ProgTutorial/Recipes/Antiquotes.thy
changeset 191 0150cf5982ae
parent 189 069d525f8f1d
child 292 41a802bbb7df
equal deleted inserted replaced
190:ca0ac2e75f6d 191:0150cf5982ae
     1 
     1 
     2 theory Antiquotes
     2 theory Antiquotes
     3 imports "../Base"
     3 imports "../Base"
     4 begin
     4 begin
     5 
     5 
     6 
       
     7 section {* Useful Document Antiquotations *}
     6 section {* Useful Document Antiquotations *}
     8 
     7 
     9 text {*
     8 text {*
    10   (FIXME: update to to new antiquotation setup)
       
    11 
       
    12   {\bf Problem:} 
     9   {\bf Problem:} 
    13   How to keep your ML-code inside a document synchronised with the actual code?\smallskip
    10   How to keep your ML-code inside a document synchronised with the actual code?\smallskip
    14 
    11 
    15   {\bf Solution:} This can be achieved using document antiquotations.\smallskip
    12   {\bf Solution:} This can be achieved with document antiquotations.\smallskip
    16 
    13 
    17   Document antiquotations can be used for ensuring consistent type-setting of
    14   Document antiquotations can be used for ensuring consistent type-setting of
    18   various entities in a document. They can also be used for sophisticated
    15   various entities in a document. They can also be used for sophisticated
    19   \LaTeX-hacking. If you type @{text "Ctrl-c Ctrl-a h A"} inside ProofGeneral, you
    16   \LaTeX-hacking. If you type on the Isabelle level
    20   obtain a list of all currently available document antiquotations and their options.
    17 *}
    21   You obtain the same list on the ML-level by typing
       
    22 
    18 
    23   @{ML [display,gray] "ThyOutput.print_antiquotations ()"}
    19 print_antiquotations
    24 
    20 
    25   Below we give the code for two additional document antiquotations that can
    21 text {*
    26   be used to typeset ML-code and also to check whether the given code actually
    22   you obtain a list of all currently available document antiquotations and
    27   compiles. This provides a sanity check for the code and also allows one to
    23   their options.  
    28   keep documents in sync with other code, for example Isabelle.
    24 
       
    25   Below we will give the code for two additional document
       
    26   antiquotations both of which are intended to typeset ML-code. The crucial point
       
    27   of these document antiquotations is that they not just print the ML-code, but also 
       
    28   check whether it compiles. This will provide a sanity check for the code
       
    29   and also allows you to keep documents in sync with other code, for example
       
    30   Isabelle.
    29 
    31 
    30   We first describe the antiquotation @{text "ML_checked"} with the syntax:
    32   We first describe the antiquotation @{text "ML_checked"} with the syntax:
    31  
    33  
    32   @{text [display] "@{ML_checked \"a_piece_of_code\"}"}
    34   @{text [display] "@{ML_checked \"a_piece_of_code\"}"}
    33 
    35 
    45    ThyOutput.output (map Pretty.str (space_explode "\n" code_txt)))
    47    ThyOutput.output (map Pretty.str (space_explode "\n" code_txt)))
    46 
    48 
    47 val _ = ThyOutput.antiquotation "ML_checked" (Scan.lift Args.name) output_ml*}
    49 val _ = ThyOutput.antiquotation "ML_checked" (Scan.lift Args.name) output_ml*}
    48 
    50 
    49 text {*
    51 text {*
    50   The parser @{ML "(Scan.lift Args.name)"} in line 9 parses a string, in this
    52   The parser @{ML "(Scan.lift Args.name)"} in Line 7 parses a string, in this
    51   case the code. As mentioned before, the code is sent to the ML-compiler in
    53   case the code, and then calls the function @{ML output_ml}. As mentioned
    52   the line 4 using the function @{ML ml_val}, which constructs the appropriate
    54   before, the parsed code is sent to the ML-compiler in Line 4 using the
    53   ML-expression.  If the code is ``approved'' by the compiler, then the output
    55   function @{ML ml_val}, which constructs the appropriate ML-expression, and
    54   function @{ML "ThyOutput.output"} in the next line pretty prints the
    56   using @{ML "eval_in" in ML_Context}, which calls the compiler.  If the code is
    55   code. This function expects that the code is a list of (pretty)strings where
    57   ``approved'' by the compiler, then the output function @{ML "output" in
    56   each string correspond to a line in the output. Therefore the use of @{ML
    58   ThyOutput} in the next line pretty prints the code. This function expects
    57   "(space_explode \"\\n\" txt)" for txt} which produces this list according to
    59   that the code is a list of (pretty)strings where each string correspond to a
    58   linebreaks.  There are a number of options for antiquotations that are
    60   line in the output. Therefore the use of @{ML "(space_explode \"\\n\" txt)"
    59   observed by @{ML ThyOutput.output} when printing the code (including @{text
    61   for txt} which produces such a list according to linebreaks.  There are a
    60   "[display]"} and @{text "[quotes]"}). Line 7 sets up the new document
    62   number of options for antiquotations that are observed by the function 
    61   antiquotation.
    63   @{ML "output" in ThyOutput} when printing the code (including @{text "[display]"} 
    62 
    64   and @{text "[quotes]"}). The function @{ML "antiquotation" in ThyOutput} in 
       
    65   Line 7 sets up the new document antiquotation.
    63 
    66 
    64   \begin{readmore}
    67   \begin{readmore}
    65   For more information about options of document antiquotations see \rsccite{sec:antiq}).
    68   For more information about options of document antiquotations see \rsccite{sec:antiq}).
    66   \end{readmore}
    69   \end{readmore}
    67 
    70 
    80 text {*
    83 text {*
    81   where in Lines 1 and 2 the positional information is properly treated. The
    84   where in Lines 1 and 2 the positional information is properly treated. The
    82   parser @{ML OuterParse.position} encodes the positional information in the 
    85   parser @{ML OuterParse.position} encodes the positional information in the 
    83   result.
    86   result.
    84 
    87 
    85   We can now write in a document @{text "@{ML_checked \"2 + 3\"}"} in order to
    88   We can now write @{text "@{ML_checked \"2 + 3\"}"} in a document in order to
    86   obtain @{ML_checked "2 + 3"} and be sure that this code compiles until
    89   obtain @{ML_checked "2 + 3"} and be sure that this code compiles until
    87   somebody changes the definition of \mbox{@{ML "(op +)"}}.
    90   somebody changes the definition of addition.
    88 
    91 
    89 
    92 
    90   The second document antiquotation we describe extends the first by a pattern
    93   The second document antiquotation we describe extends the first by a pattern
    91   that specifies what the result of the ML-code should be and check the
    94   that specifies what the result of the ML-code should be and checks the
    92   consistency of the actual result with the given pattern. For this we are
    95   consistency of the actual result with the given pattern. For this we are
    93   going to implement the document antiquotation
    96   going to implement the document antiquotation:
    94 
    97 
    95   
    98   
    96   @{text [display] "@{ML_resp \"a_piece_of_code\" \"a_pattern\"}"}
    99   @{text [display] "@{ML_resp \"a_piece_of_code\" \"a_pattern\"}"}
    97   
   100   
    98   To add some convenience and also to deal with large outputs, the user can
   101   To add some convenience and also to deal with large outputs, the user can
    99   give a partial specification inside the pattern by giving abbreviations of
   102   give a partial specification by using ellipses. For example @{text "(\<dots>, \<dots>)"}
   100   the form @{text [quotes] "\<dots>"}. For example @{text "(\<dots>, \<dots>)"} for specifying a
   103   for specifying a pair.  In order to check consistency between the pattern
   101   pair.
   104   and the output of the code, we have to change the ML-expression that is sent 
   102 
   105   to the compiler: in @{text "ML_checked"} we sent the expression @{text [quotes]
   103   In the document antiquotation @{text "@{ML_checked \"piece_of_code\"}"}
   106   "val _ = a_piece_of_code"} to the compiler; now the wildcard @{text "_"}
   104   above we have sent the expression @{text [quotes] "val _ = piece_of_code"}
   107   must be be replaced by the given pattern. However, we have to remove all
   105   to the compiler, now instead the wildcard @{text "_"} we will be replaced by
   108   ellipses from it and replace them by @{text [quotes] "_"}. The following 
   106   the given pattern. To do this we need to replace in the input the @{text
   109   function will do this:
   107   [quotes] "\<dots>"} by @{text [quotes] "_"} before sending the code to the
       
   108   compiler. The following function will do this:
       
   109 *}
   110 *}
   110 
   111 
   111 ML{*fun ml_pat (code_txt, pat) =
   112 ML{*fun ml_pat (code_txt, pat) =
   112 let val pat' = 
   113 let val pat' = 
   113          implode (map (fn "\<dots>" => "_" | s => s) (Symbol.explode pat))
   114          implode (map (fn "\<dots>" => "_" | s => s) (Symbol.explode pat))
   114 in 
   115 in 
   115   "val " ^ pat' ^ " = " ^ code_txt 
   116   "val " ^ pat' ^ " = " ^ code_txt 
   116 end*}
   117 end*}
   117 
   118 
   118 text {* 
   119 text {* 
   119   Next we like to add a response indicator to the result using:
   120   Next we add a response indicator to the result using:
   120 *}
   121 *}
   121 
   122 
   122 
   123 
   123 ML{*fun add_resp pat = map (fn s => "> " ^ s) pat*}
   124 ML{*fun add_resp pat = map (fn s => "> " ^ s) pat*}
   124 
   125 
   125 text {* 
   126 text {* 
   126   The rest of the code of the document antiquotation is
   127   The rest of the code of @{text "ML_resp"} is: 
   127 *}
   128 *}
   128 
   129 
   129 ML{*fun output_ml_resp {context = ctxt, ...} ((code_txt, pat), pos) = 
   130 ML %linenosgray{*fun output_ml_resp {context = ctxt, ...} ((code_txt, pat), pos) = 
   130   (ML_Context.eval_in (SOME ctxt) false pos (ml_pat (code_txt, pat));
   131   (ML_Context.eval_in (SOME ctxt) false pos (ml_pat (code_txt, pat));
   131    let 
   132    let 
   132      val output1 = space_explode "\n" code_txt 
   133      val code_output = space_explode "\n" code_txt 
   133      val output2 = add_resp (space_explode "\n" pat)
   134      val resp_output = add_resp (space_explode "\n" pat)
   134    in 
   135    in 
   135      ThyOutput.output (map Pretty.str (output1 @ output2)) 
   136      ThyOutput.output (map Pretty.str (code_output @ resp_output)) 
   136    end)
   137    end)
   137 
   138 
   138 val _ = ThyOutput.antiquotation "ML_resp" 
   139 val _ = ThyOutput.antiquotation "ML_resp" 
   139          (Scan.lift (OuterParse.position (Args.name -- Args.name))) 
   140           (Scan.lift (OuterParse.position (Args.name -- Args.name))) 
   140            output_ml_resp*}
   141              output_ml_resp*}
   141 
   142 
   142 text {*
   143 text {*
   143   This extended document antiquotation allows us to write
   144   In comparison with @{text "ML_checked"}, we only changed the line about 
       
   145   the compiler (Line~2), the lines about
       
   146   the output (Lines 4 to 7) and the parser in the setup (Line 11). Now 
       
   147   you can write
   144  
   148  
   145   @{text [display] "@{ML_resp [display] \"true andalso false\" \"false\"}"}
   149   @{text [display] "@{ML_resp [display] \"true andalso false\" \"false\"}"}
   146 
   150 
   147   to obtain
   151   to obtain
   148 
   152